diff --git a/src/HOL/Analysis/Gamma_Function.thy b/src/HOL/Analysis/Gamma_Function.thy --- a/src/HOL/Analysis/Gamma_Function.thy +++ b/src/HOL/Analysis/Gamma_Function.thy @@ -1,3505 +1,3506 @@ (* Title: HOL/Analysis/Gamma_Function.thy Author: Manuel Eberl, TU München *) section \The Gamma Function\ theory Gamma_Function imports Equivalence_Lebesgue_Henstock_Integration Summation_Tests Harmonic_Numbers "HOL-Library.Nonpos_Ints" "HOL-Library.Periodic_Fun" begin text \ Several equivalent definitions of the Gamma function and its most important properties. Also contains the definition and some properties of the log-Gamma function and the Digamma function and the other Polygamma functions. Based on the Gamma function, we also prove the Weierstra{\ss} product form of the sin function and, based on this, the solution of the Basel problem (the sum over all \<^term>\1 / (n::nat)^2\. \ lemma pochhammer_eq_0_imp_nonpos_Int: "pochhammer (x::'a::field_char_0) n = 0 \ x \ \\<^sub>\\<^sub>0" by (auto simp: pochhammer_eq_0_iff) lemma closed_nonpos_Ints [simp]: "closed (\\<^sub>\\<^sub>0 :: 'a :: real_normed_algebra_1 set)" proof - have "\\<^sub>\\<^sub>0 = (of_int ` {n. n \ 0} :: 'a set)" by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) also have "closed \" by (rule closed_of_int_image) finally show ?thesis . qed lemma plus_one_in_nonpos_Ints_imp: "z + 1 \ \\<^sub>\\<^sub>0 \ z \ \\<^sub>\\<^sub>0" using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all lemma of_int_in_nonpos_Ints_iff: "(of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ 0" by (auto simp: nonpos_Ints_def) lemma one_plus_of_int_in_nonpos_Ints_iff: "(1 + of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ -1" proof - have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp also have "\ \ \\<^sub>\\<^sub>0 \ n + 1 \ 0" by (subst of_int_in_nonpos_Ints_iff) simp_all also have "\ \ n \ -1" by presburger finally show ?thesis . qed lemma one_minus_of_nat_in_nonpos_Ints_iff: "(1 - of_nat n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n > 0" proof - have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp also have "\ \ \\<^sub>\\<^sub>0 \ n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger finally show ?thesis . qed lemma fraction_not_in_ints: assumes "\(n dvd m)" "n \ 0" shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" proof assume "of_int m / (of_int n :: 'a) \ \" then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases) with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: field_split_simps) hence "m = k * n" by (subst (asm) of_int_eq_iff) hence "n dvd m" by simp with assms(1) show False by contradiction qed lemma fraction_not_in_nats: assumes "\n dvd m" "n \ 0" shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" proof assume "of_int m / of_int n \ (\ :: 'a set)" also note Nats_subset_Ints finally have "of_int m / of_int n \ (\ :: 'a set)" . moreover have "of_int m / of_int n \ (\ :: 'a set)" using assms by (intro fraction_not_in_ints) ultimately show False by contradiction qed lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \ \ \ z \ \\<^sub>\\<^sub>0" by (auto simp: Ints_def nonpos_Ints_def) lemma double_in_nonpos_Ints_imp: assumes "2 * (z :: 'a :: field_char_0) \ \\<^sub>\\<^sub>0" shows "z \ \\<^sub>\\<^sub>0 \ z + 1/2 \ \\<^sub>\\<^sub>0" proof- from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases') thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps) qed lemma sin_series: "(\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" proof - from sin_converges[of z] have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z" . also have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z \ (\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" by (subst sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: sin_coeff_def strict_mono_def ac_simps elim!: oddE) finally show ?thesis . qed lemma cos_series: "(\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" proof - from cos_converges[of z] have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z" . also have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z \ (\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" by (subst sums_mono_reindex[of "\n. 2*n", symmetric]) (auto simp: cos_coeff_def strict_mono_def ac_simps elim!: evenE) finally show ?thesis . qed lemma sin_z_over_z_series: fixes z :: "'a :: {real_normed_field,banach}" assumes "z \ 0" shows "(\n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)" proof - from sin_series[of z] have "(\n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z" by (simp add: field_simps scaleR_conv_of_real) from sums_mult[OF this, of "inverse z"] and assms show ?thesis by (simp add: field_simps) qed lemma sin_z_over_z_series': fixes z :: "'a :: {real_normed_field,banach}" assumes "z \ 0" shows "(\n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)" proof - from sums_split_initial_segment[OF sin_converges[of z], of 1] have "(\n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps) qed lemma has_field_derivative_sin_z_over_z: fixes A :: "'a :: {real_normed_field,banach} set" shows "((\z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)" (is "(?f has_field_derivative ?f') _") proof (rule has_field_derivative_at_within) have "((\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) has_field_derivative (\n. diffs (\n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)" proof (rule termdiffs_strong) from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1] show "summable (\n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def) qed simp also have "(\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) = ?f" proof fix z show "(\n. of_real (sin_coeff (n+1)) * z^n) = ?f z" by (cases "z = 0") (insert sin_z_over_z_series'[of z], simp_all add: scaleR_conv_of_real sums_iff sin_coeff_def) qed also have "(\n. diffs (\n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = diffs (\n. of_real (sin_coeff (Suc n))) 0" by simp also have "\ = 0" by (simp add: sin_coeff_def diffs_def) finally show "((\z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" . qed lemma round_Re_minimises_norm: "norm ((z::complex) - of_int m) \ norm (z - of_int (round (Re z)))" proof - let ?n = "round (Re z)" have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)" by (simp add: cmod_def) also have "\Re z - of_int ?n\ \ \Re z - of_int m\" by (rule round_diff_minimal) hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \ sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)" by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff) also have "\ = norm (z - of_int m)" by (simp add: cmod_def) finally show ?thesis . qed lemma Re_pos_in_ball: assumes "Re z > 0" "t \ ball z (Re z/2)" shows "Re t > 0" proof - have "Re (z - t) \ norm (z - t)" by (rule complex_Re_le_cmod) also from assms have "\ < Re z / 2" by (simp add: dist_complex_def) finally show "Re t > 0" using assms by simp qed lemma no_nonpos_Int_in_ball_complex: assumes "Re z > 0" "t \ ball z (Re z/2)" shows "t \ \\<^sub>\\<^sub>0" using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases) lemma no_nonpos_Int_in_ball: assumes "t \ ball z (dist z (round (Re z)))" shows "t \ \\<^sub>\\<^sub>0" proof assume "t \ \\<^sub>\\<^sub>0" then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases) have "dist z (of_int n) \ dist z t + dist t (of_int n)" by (rule dist_triangle) also from assms have "dist z t < dist z (round (Re z))" by simp also have "\ \ dist z (of_int n)" using round_Re_minimises_norm[of z] by (simp add: dist_complex_def) finally have "dist t (of_int n) > 0" by simp with \t = of_int n\ show False by simp qed lemma no_nonpos_Int_in_ball': assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \ \\<^sub>\\<^sub>0" obtains d where "d > 0" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" proof (rule that) from assms show "setdist {z} \\<^sub>\\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto next fix t assume "t \ ball z (setdist {z} \\<^sub>\\<^sub>0)" thus "t \ \\<^sub>\\<^sub>0" using setdist_le_dist[of z "{z}" t "\\<^sub>\\<^sub>0"] by force qed lemma no_nonpos_Real_in_ball: assumes z: "z \ \\<^sub>\\<^sub>0" and t: "t \ ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" shows "t \ \\<^sub>\\<^sub>0" using z proof (cases "Im z = 0") assume A: "Im z = 0" with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff) with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff) next assume A: "Im z \ 0" have "abs (Im z) - abs (Im t) \ abs (Im z - Im t)" by linarith also have "\ = abs (Im (z - t))" by simp also have "\ \ norm (z - t)" by (rule abs_Im_le_cmod) also from A t have "\ \ abs (Im z) / 2" by (simp add: dist_complex_def) finally have "abs (Im t) > 0" using A by simp thus ?thesis by (force simp add: complex_nonpos_Reals_iff) qed subsection \The Euler form and the logarithmic Gamma function\ text \ We define the Gamma function by first defining its multiplicative inverse \rGamma\. This is more convenient because \rGamma\ is entire, which makes proofs of its properties more convenient because one does not have to watch out for discontinuities. (e.g. \rGamma\ fulfils \rGamma z = z * rGamma (z + 1)\ everywhere, whereas the \\\ function does not fulfil the analogous equation on the non-positive integers) We define the \\\ function (resp.\ its reciprocale) in the Euler form. This form has the advantage that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 (due to division by 0). The functional equation \Gamma (z + 1) = z * Gamma z\ follows immediately from the definition. \ definition\<^marker>\tag important\ Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)" definition Gamma_series' :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n" definition rGamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))" lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)" and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)" unfolding Gamma_series_def rGamma_series_def by simp_all lemma rGamma_series_minus_of_nat: "eventually (\n. rGamma_series (- of_nat k) n = 0) sequentially" using eventually_ge_at_top[of k] by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff) lemma Gamma_series_minus_of_nat: "eventually (\n. Gamma_series (- of_nat k) n = 0) sequentially" using eventually_ge_at_top[of k] by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff) lemma Gamma_series'_minus_of_nat: "eventually (\n. Gamma_series' (- of_nat k) n = 0) sequentially" using eventually_gt_at_top[of k] by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff) lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ rGamma_series z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp) lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp) lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series' z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp) lemma Gamma_series_Gamma_series': assumes z: "z \ \\<^sub>\\<^sub>0" shows "(\n. Gamma_series' z n / Gamma_series z n) \ 1" proof (rule Lim_transform_eventually) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n" by (cases n, simp) (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec' dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp) also from n have "\ = z / of_nat n + 1" by (simp add: field_split_simps) finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" .. qed have "(\x. z / of_nat x) \ 0" by (rule tendsto_norm_zero_cancel) (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], simp add: norm_divide inverse_eq_divide) from tendsto_add[OF this tendsto_const[of 1]] show "(\n. z / of_nat n + 1) \ 1" by simp qed text \ We now show that the series that defines the \\\ function in the Euler form converges and that the function defined by it is continuous on the complex halfspace with positive real part. We do this by showing that the logarithm of the Euler series is continuous and converges locally uniformly, which means that the log-Gamma function defined by its limit is also continuous. This will later allow us to lift holomorphicity and continuity from the log-Gamma function to the inverse of the Gamma function, and from that to the Gamma function itself. \ definition\<^marker>\tag important\ ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\k=1..n. ln (z / of_nat k + 1))" definition\<^marker>\tag unimportant\ ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series' z n = - euler_mascheroni*z - ln z + (\k=1..n. z / of_nat n - ln (z / of_nat k + 1))" definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \ 'a" where "ln_Gamma z = lim (ln_Gamma_series z)" text \ We now show that the log-Gamma series converges locally uniformly for all complex numbers except the non-positive integers. We do this by proving that the series is locally Cauchy. \ context begin private lemma ln_Gamma_series_complex_converges_aux: fixes z :: complex and k :: nat assumes z: "z \ 0" and k: "of_nat k \ 2*norm z" "k \ 2" shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \ 2*(norm z + norm z^2) / of_nat k^2" proof - let ?k = "of_nat k :: complex" and ?z = "norm z" have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)" by (simp add: algebra_simps) also have "norm ... \ ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)" by (subst norm_mult [symmetric], rule norm_triangle_ineq) also have "norm (Ln (1 + -1/?k) - (-1/?k)) \ (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))" using k by (intro Ln_approx_linear) (simp add: norm_divide) hence "?z * norm (ln (1-1/?k) + 1/?k) \ ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" by (intro mult_left_mono) simp_all also have "... \ (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k by (simp add: field_simps power2_eq_square norm_divide) also have "... \ (?z * 2) / of_nat k^2" using k by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) also have "norm (ln (1+z/?k) - z/?k) \ norm (z/?k)^2 / (1 - norm (z/?k))" using k by (intro Ln_approx_linear) (simp add: norm_divide) hence "norm (ln (1+z/?k) - z/?k) \ ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" by (simp add: field_simps norm_divide) also have "... \ (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k by (simp add: field_simps power2_eq_square) also have "... \ (?z^2 * 2) / of_nat k^2" using k by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) also note add_divide_distrib [symmetric] finally show ?thesis by (simp only: distrib_left mult.commute) qed lemma ln_Gamma_series_complex_converges: assumes z: "z \ \\<^sub>\\<^sub>0" assumes d: "d > 0" "\n. n \ \\<^sub>\\<^sub>0 \ norm (z - of_int n) > d" shows "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n :: complex)" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') fix e :: real assume e: "e > 0" define e'' where "e'' = (SUP t\ball z d. norm t + norm t^2)" define e' where "e' = e / (2*e'')" have "bounded ((\t. norm t + norm t^2) ` cball z d)" by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) hence "bounded ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto hence bdd: "bdd_above ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above) with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z]) have e'': "norm t + norm t^2 \ e''" if "t \ ball z d" for t unfolding e''_def using that by (rule cSUP_upper[OF _ bdd]) from e z e''_pos have e': "e' > 0" unfolding e'_def by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps) have "summable (\k. inverse ((real_of_nat k)^2))" by (rule inverse_power_summable) simp from summable_partial_sum_bound[OF this e'] guess M . note M = this define N where "N = max 2 (max (nat \2 * (norm z + d)\) M)" { from d have "\2 * (cmod z + d)\ \ \0::real\" by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all hence "2 * (norm z + d) \ of_nat (nat \2 * (norm z + d)\)" unfolding N_def by (simp_all) also have "... \ of_nat N" unfolding N_def by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1) finally have "of_nat N \ 2 * (norm z + d)" . moreover have "N \ 2" "N \ M" unfolding N_def by simp_all moreover have "(\k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \ N" for m n using M[OF order.trans[OF \N \ M\ that]] unfolding real_norm_def by (subst (asm) abs_of_nonneg) (auto intro: sum_nonneg simp: field_split_simps) moreover note calculation } note N = this show "\M. \t\ball z d. \m\M. \n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e" unfolding dist_complex_def proof (intro exI[of _ N] ballI allI impI) fix t m n assume t: "t \ ball z d" and mn: "m \ N" "n > m" from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def) also have "dist z 0 - dist z t \ dist 0 t" using dist_triangle[of 0 z t] by (simp add: dist_commute) finally have t_nz: "t \ 0" by auto have "norm t \ norm z + norm (t - z)" by (rule norm_triangle_sub) also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute) also have "2 * (norm z + d) \ of_nat N" by (rule N) also have "N \ m" by (rule mn) finally have norm_t: "2 * norm t < of_nat m" by simp have "ln_Gamma_series t m - ln_Gamma_series t n = (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) + ((\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)))" by (simp add: ln_Gamma_series_def algebra_simps) also have "(\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)) = (\k\{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn by (simp add: sum_diff) also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = (\k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn by (subst sum_telescope'' [symmetric]) simp_all also have "... = (\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N by (intro sum_cong_Suc) (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat) also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \ {Suc m..n}" for k using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: field_split_simps) hence "(\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) = (\k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N by (intro sum.cong) simp_all also note sum.distrib [symmetric] also have "norm (\k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \ (\k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t by (intro order.trans[OF norm_sum sum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all also have "... \ 2 * (norm t + norm t^2) * (\k=Suc m..n. 1 / (of_nat k)\<^sup>2)" by (simp add: sum_distrib_left) also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" by (simp add: e'_def field_simps power2_eq_square) also from e''[OF t] e''_pos e have "\ \ e * 1" by (intro mult_left_mono) (simp_all add: field_simps) finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp qed qed end lemma ln_Gamma_series_complex_converges': assumes z: "(z :: complex) \ \\<^sub>\\<^sub>0" shows "\d>0. uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" proof - define d' where "d' = Re z" define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" have "of_int (round d') \ \\<^sub>\\<^sub>0" if "d' \ 0" using that by (intro nonpos_Ints_of_int) (simp_all add: round_def) with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) have "d < cmod (z - of_int n)" if "n \ \\<^sub>\\<^sub>0" for n proof (cases "Re z > 0") case True from nonpos_Ints_nonpos[OF that] have n: "n \ 0" by simp from True have "d = Re z/2" by (simp add: d_def d'_def) also from n True have "\ < Re (z - of_int n)" by simp also have "\ \ norm (z - of_int n)" by (rule complex_Re_le_cmod) finally show ?thesis . next case False with assms nonpos_Ints_of_int[of "round (Re z)"] have "z \ of_int (round d')" by (auto simp: not_less) with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def) also have "\ \ norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm) finally show ?thesis . qed hence conv: "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" by (intro ln_Gamma_series_complex_converges d_pos z) simp_all from d_pos conv show ?thesis by blast qed lemma ln_Gamma_series_complex_converges'': "(z :: complex) \ \\<^sub>\\<^sub>0 \ convergent (ln_Gamma_series z)" by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent) theorem ln_Gamma_complex_LIMSEQ: "(z :: complex) \ \\<^sub>\\<^sub>0 \ ln_Gamma_series z \ ln_Gamma z" using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def) lemma exp_ln_Gamma_series_complex: assumes "n > 0" "z \ \\<^sub>\\<^sub>0" shows "exp (ln_Gamma_series z n :: complex) = Gamma_series z n" proof - from assms obtain m where m: "n = Suc m" by (cases n) blast from assms have "z \ 0" by (intro notI) auto with assms have "exp (ln_Gamma_series z n) = (of_nat n) powr z / (z * (\k=1..n. exp (Ln (z / of_nat k + 1))))" unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_sum) also from assms have "(\k=1..n. exp (Ln (z / of_nat k + 1))) = (\k=1..n. z / of_nat k + 1)" by (intro prod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp) also have "... = (\k=1..n. z + k) / fact n" by (simp add: fact_prod) (subst prod_dividef [symmetric], simp_all add: field_simps) also from m have "z * ... = (\k=0..n. z + k) / fact n" by (simp add: prod.atLeast0_atMost_Suc_shift prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) also have "(\k=0..n. z + k) = pochhammer z (Suc n)" unfolding pochhammer_prod by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n" unfolding Gamma_series_def using assms by (simp add: field_split_simps powr_def) finally show ?thesis . qed lemma ln_Gamma_series'_aux: assumes "(z::complex) \ \\<^sub>\\<^sub>0" shows "(\k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s") unfolding sums_def proof (rule Lim_transform) show "(\n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \ ?s" (is "?g \ _") by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms) have A: "eventually (\n. (\k 0" have "(\kk=1..n. z / of_nat k - ln (1 + z / of_nat k))" by (subst atLeast0LessThan [symmetric], subst sum.shift_bounds_Suc_ivl [symmetric], subst atLeastLessThanSuc_atLeastAtMost) simp_all also have "\ = z * of_real (harm n) - (\k=1..n. ln (1 + z / of_nat k))" by (simp add: harm_def sum_subtractf sum_distrib_left divide_inverse) also from n have "\ - ?g n = 0" by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps) finally show "(\kn. (\k 0" by (subst tendsto_cong[OF A]) simp_all qed lemma uniformly_summable_deriv_ln_Gamma: assumes z: "(z :: 'a :: {real_normed_field,banach}) \ 0" and d: "d > 0" "d \ norm z/2" shows "uniformly_convergent_on (ball z d) (\k z. \ik z. \i ball z d" have "norm z = norm (t + (z - t))" by simp have "norm (t + (z - t)) \ norm t + norm (z - t)" by (rule norm_triangle_ineq) also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm) finally have A: "norm t > norm z / 2" using z by (simp add: field_simps) have "norm t = norm (z + (t - z))" by simp also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) also from t d have "norm (t - z) \ norm z / 2" by (simp add: dist_norm norm_minus_commute) also from z have "\ < norm z" by simp finally have B: "norm t < 2 * norm z" by simp note A B } note ball = this show "eventually (\n. \t\ball z d. norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially" using eventually_gt_at_top apply eventually_elim proof safe fix t :: 'a assume t: "t \ ball z d" from z ball[OF t] have t_nz: "t \ 0" by auto fix n :: nat assume n: "n > nat \4 * norm z\" from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp also from n have "\ < of_nat n" by linarith finally have n: "of_nat n > 2 * norm t" . hence "of_nat n > norm t" by simp hence t': "t \ -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc) with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))" by (simp add: field_split_simps eq_neg_iff_add_eq_0 del: of_nat_Suc) also have "norm \ = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))" by (simp add: norm_divide norm_mult field_split_simps del: of_nat_Suc) also { from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \ of_nat (Suc n) / (2 * norm t)" by (intro divide_left_mono mult_pos_pos) simp_all also have "\ < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc) also have "\ \ norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq) finally have "inverse (norm (of_nat (Suc n)/t + 1)) \ 4 * norm z / of_nat (Suc n)" using z by (simp add: field_split_simps norm_divide mult_ac del: of_nat_Suc) } also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) = 4 * norm z * inverse (of_nat (Suc n)^2)" by (simp add: field_split_simps power2_eq_square del: of_nat_Suc) finally show "norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)" by (simp del: of_nat_Suc) qed next show "summable (\n. 4 * norm z * inverse ((of_nat (Suc n))^2))" by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable) qed subsection \The Polygamma functions\ lemma summable_deriv_ln_Gamma: "z \ (0 :: 'a :: {real_normed_field,banach}) \ summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))" unfolding summable_iff_convergent by (rule uniformly_convergent_imp_convergent, rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all definition\<^marker>\tag important\ Polygamma :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" where "Polygamma n z = (if n = 0 then (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else (-1)^Suc n * fact n * (\k. inverse ((z + of_nat k)^Suc n)))" abbreviation\<^marker>\tag important\ Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where "Digamma \ Polygamma 0" lemma Digamma_def: "Digamma z = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni" by (simp add: Polygamma_def) lemma summable_Digamma: assumes "(z :: 'a :: {real_normed_field,banach}) \ 0" shows "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" proof - have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums (0 - inverse (z + of_nat 0))" by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]] show "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp qed lemma summable_offset: assumes "summable (\n. f (n + k) :: 'a :: real_normed_vector)" shows "summable f" proof - - from assms have "convergent (\m. \nm. \nm. (\nnm. (\nnm. \n {k..n\\. f n) = (\nn=k..n=k..n=0..nnna. sum f {.. lim (\m. sum f {.. 0" and n: "n \ 2" shows "uniformly_convergent_on (ball z d) (\k z. \inorm z * e\" { fix t assume t: "t \ ball z d" have "norm t = norm (z + (t - z))" by simp also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute) finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def) } note ball = this show "eventually (\k. \t\ball z d. norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)) sequentially" using eventually_gt_at_top[of m] apply eventually_elim proof (intro ballI) fix k :: nat and t :: 'a assume k: "k > m" and t: "t \ ball z d" from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff) also have "\ \ norm (of_nat k :: 'a) - norm z * e" unfolding m_def by (subst norm_of_nat) linarith also from ball[OF t] have "\ \ norm (of_nat k :: 'a) - norm t" by simp also have "\ \ norm (of_nat k + t)" by (rule norm_diff_ineq) finally have "inverse ((norm (t + of_nat k))^n) \ inverse (real_of_nat (k - m)^n)" using k n by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc) thus "norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)" by (simp add: norm_inverse norm_power power_inverse) qed have "summable (\k. inverse ((real_of_nat k)^n))" using inverse_power_summable[of n] n by simp hence "summable (\k. inverse ((real_of_nat (k + m - m))^n))" by simp thus "summable (\k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset) qed lemma Polygamma_converges': fixes z :: "'a :: {real_normed_field,banach}" assumes z: "z \ 0" and n: "n \ 2" shows "summable (\k. inverse ((z + of_nat k)^n))" using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z] by (simp add: summable_iff_convergent) theorem Digamma_LIMSEQ: fixes z :: "'a :: {banach,real_normed_field}" assumes z: "z \ 0" shows "(\m. of_real (ln (real m)) - (\n Digamma z" proof - have "(\n. of_real (ln (real n / (real (Suc n))))) \ (of_real (ln 1) :: 'a)" by (intro tendsto_intros LIMSEQ_n_over_Suc_n) simp_all hence "(\n. of_real (ln (real n / (real n + 1)))) \ (0 :: 'a)" by (simp add: add_ac) hence lim: "(\n. of_real (ln (real n)) - of_real (ln (real n + 1))) \ (0::'a)" proof (rule Lim_transform_eventually) show "eventually (\n. of_real (ln (real n / (real n + 1))) = of_real (ln (real n)) - (of_real (ln (real n + 1)) :: 'a)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_div) qed from summable_Digamma[OF z] have "(\n. inverse (of_nat (n+1)) - inverse (z + of_nat n)) sums (Digamma z + euler_mascheroni)" by (simp add: Digamma_def summable_sums) from sums_diff[OF this euler_mascheroni_sum] have "(\n. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)) - inverse (z + of_nat n)) sums Digamma z" by (simp add: add_ac) hence "(\m. (\nn Digamma z" by (simp add: sums_def sum_subtractf) also have "(\m. (\nm. of_real (ln (m + 1)) :: 'a)" by (subst sum_lessThan_telescope) simp_all finally show ?thesis by (rule Lim_transform) (insert lim, simp) qed theorem Polygamma_LIMSEQ: fixes z :: "'a :: {banach,real_normed_field}" assumes "z \ 0" and "n > 0" shows "(\k. inverse ((z + of_nat k)^Suc n)) sums ((-1) ^ Suc n * Polygamma n z / fact n)" using Polygamma_converges'[OF assms(1), of "Suc n"] assms(2) by (simp add: sums_iff Polygamma_def) theorem has_field_derivative_ln_Gamma_complex [derivative_intros]: fixes z :: complex assumes z: "z \ \\<^sub>\\<^sub>0" shows "(ln_Gamma has_field_derivative Digamma z) (at z)" proof - have not_nonpos_Int [simp]: "t \ \\<^sub>\\<^sub>0" if "Re t > 0" for t using that by (auto elim!: nonpos_Ints_cases') from z have z': "z \ \\<^sub>\\<^sub>0" and z'': "z \ 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I by blast+ let ?f' = "\z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" let ?f = "\z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\z. \n. ?f' z n" define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" from z have d: "d > 0" "norm z/2 \ d" by (auto simp add: complex_nonpos_Reals_iff d_def) have ball: "Im t = 0 \ Re t > 0" if "dist z t < d" for t using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums (0 - inverse (z + of_nat 0))" by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) have "((\z. \n. ?f z n) has_field_derivative ?F' z) (at z)" using d z ln_Gamma_series'_aux[OF z'] apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma) apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff simp del: of_nat_Suc) apply (auto simp add: complex_nonpos_Reals_iff) done with z have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative ?F' z - euler_mascheroni - inverse z) (at z)" by (force intro!: derivative_eq_intros simp: Digamma_def) also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp also from sums have "-inverse z = (\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))" by (simp add: sums_iff) also from sums summable_deriv_ln_Gamma[OF z''] have "?F' z + \ = (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by (subst suminf_add) (simp_all add: add_ac sums_iff) also have "\ - euler_mascheroni = Digamma z" by (simp add: Digamma_def) finally have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative Digamma z) (at z)" . moreover from eventually_nhds_ball[OF d(1), of z] have "eventually (\z. ln_Gamma z = (\k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)" proof eventually_elim fix t assume "t \ ball z d" hence "t \ \\<^sub>\\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases) from ln_Gamma_series'_aux[OF this] show "ln_Gamma t = (\k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff) qed ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) qed declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros] lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni" by (simp add: Digamma_def) lemma Digamma_plus1: assumes "z \ 0" shows "Digamma (z+1) = Digamma z + 1/z" proof - have sums: "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) sums (inverse (z + of_nat 0) - 0)" by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) have "Digamma (z+1) = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac) also have "suminf ?f = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + (\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))" using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff) also have "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z" using sums by (simp add: sums_iff inverse_eq_divide) finally show ?thesis by (simp add: Digamma_def[of z]) qed theorem Polygamma_plus1: assumes "z \ 0" shows "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" proof (cases "n = 0") assume n: "n \ 0" let ?f = "\k. inverse ((z + of_nat k) ^ Suc n)" have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\k. ?f (k+1))" using n by (simp add: Polygamma_def add_ac) also have "(\k. ?f (k+1)) + (\k<1. ?f k) = (\k. ?f k)" using Polygamma_converges'[OF assms, of "Suc n"] n by (subst suminf_split_initial_segment [symmetric]) simp_all hence "(\k. ?f (k+1)) = (\k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps) also have "(-1) ^ Suc n * fact n * ((\k. ?f k) - inverse (z ^ Suc n)) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n by (simp add: inverse_eq_divide algebra_simps Polygamma_def) finally show ?thesis . qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide) theorem Digamma_of_nat: "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni" proof (induction n) case (Suc n) have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp also have "\ = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc) also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc) also have "\ + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni" by (simp add: harm_Suc) finally show ?case . qed (simp add: harm_def) lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni" by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl) lemma Polygamma_of_real: "x \ 0 \ Polygamma n (of_real x) = of_real (Polygamma n x)" unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"] by (simp_all add: suminf_of_real) lemma Polygamma_Real: "z \ \ \ z \ 0 \ Polygamma n z \ \" by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all lemma Digamma_half_integer: "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = (\kk. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" by (simp add: Digamma_def add_ac) also have "(\k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) = (\k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))" by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide) also have "\ = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums'] by (subst suminf_mult) (simp_all add: algebra_simps sums_iff) finally show ?case by simp next case (Suc n) have nz: "2 * of_nat n + (1:: 'a) \ 0" using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac) hence nz': "of_nat n + (1/2::'a) \ 0" by (simp add: field_simps) have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp also from nz' have "\ = Digamma (of_nat n + 1/2) + 1 / (of_nat n + 1/2)" by (rule Digamma_plus1) also from nz nz' have "1 / (of_nat n + 1/2 :: 'a) = 2 / (2 * of_nat n + 1)" by (subst divide_eq_eq) simp_all also note Suc finally show ?case by (simp add: add_ac) qed lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)" using Digamma_half_integer[of 0] by simp lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0" proof - have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp also have "\ = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp also note euler_mascheroni_less_13_over_22 also note ln2_le_25_over_36 finally show ?thesis by simp qed theorem has_field_derivative_Polygamma [derivative_intros]: fixes z :: "'a :: {real_normed_field,euclidean_space}" assumes z: "z \ \\<^sub>\\<^sub>0" shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)" proof (rule has_field_derivative_at_within, cases "n = 0") assume n: "n = 0" let ?f = "\k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)" let ?F = "\z. \k. ?f k z" and ?f' = "\k z. inverse ((z + of_nat k)\<^sup>2)" from no_nonpos_Int_in_ball'[OF z] guess d . note d = this from z have summable: "summable (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))" by (intro summable_Digamma) force from z have conv: "uniformly_convergent_on (ball z d) (\k z. \i2))" by (intro Polygamma_converges) auto with d have "summable (\k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent ) have "(?F has_field_derivative (\k. ?f' k z)) (at z)" proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) fix k :: nat and t :: 'a assume t: "t \ ball z d" from t d(2)[of t] show "((\z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases) qed (insert d(1) summable conv, (assumption|simp)+) with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n by (force simp: power2_eq_square intro!: derivative_eq_intros) next assume n: "n \ 0" from z have z': "z \ 0" by auto from no_nonpos_Int_in_ball'[OF z] guess d . note d = this define n' where "n' = Suc n" from n have n': "n' \ 2" by (simp add: n'_def) have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative (\k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) fix k :: nat and t :: 'a assume t: "t \ ball z d" with d have t': "t \ \\<^sub>\\<^sub>0" "t \ 0" by auto show "((\a. inverse ((a + of_nat k) ^ n')) has_field_derivative - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t' by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp) next have "uniformly_convergent_on (ball z d) (\k z. (- of_nat n' :: 'a) * (\ik z. \ik. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) = (- of_nat n') * (\k. inverse ((z + of_nat k) ^ (n' + 1)))" using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all finally have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative - of_nat n' * (\k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" . from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"] show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps) qed declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros] lemma isCont_Polygamma [continuous_intros]: fixes f :: "_ \ 'a :: {real_normed_field,euclidean_space}" shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Polygamma n (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Polygamma]]) lemma continuous_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A (Polygamma n :: _ \ 'a :: {real_normed_field,euclidean_space})" by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast lemma isCont_ln_Gamma_complex [continuous_intros]: fixes f :: "'a::t2_space \ complex" shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\z. ln_Gamma (f z)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]]) lemma continuous_on_ln_Gamma_complex [continuous_intros]: fixes A :: "complex set" shows "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A ln_Gamma" by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) fastforce lemma deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "deriv (Polygamma m) z = Polygamma (Suc m) (z :: 'a :: {real_normed_field,euclidean_space})" by (intro DERIV_imp_deriv has_field_derivative_Polygamma assms) thm has_field_derivative_Polygamma lemma higher_deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "(deriv ^^ n) (Polygamma m) z = Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})" proof - have "eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)" proof (induction n) case (Suc n) from Suc.IH have "eventually (\z. eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)" by (simp add: eventually_eventually) hence "eventually (\z. deriv ((deriv ^^ n) (Polygamma m)) z = deriv (Polygamma (m + n)) z) (nhds z)" by eventually_elim (intro deriv_cong_ev refl) moreover have "eventually (\z. z \ UNIV - \\<^sub>\\<^sub>0) (nhds z)" using assms by (intro eventually_nhds_in_open open_Diff open_UNIV) auto ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma) qed simp_all thus ?thesis by (rule eventually_nhds_x_imp_x) qed lemma deriv_ln_Gamma_complex: assumes "z \ \\<^sub>\\<^sub>0" shows "deriv ln_Gamma z = Digamma (z :: complex)" by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_complex assms) text \ We define a type class that captures all the fundamental properties of the inverse of the Gamma function and defines the Gamma function upon that. This allows us to instantiate the type class both for the reals and for the complex numbers with a minimal amount of proof duplication. \ class\<^marker>\tag unimportant\ Gamma = real_normed_field + complete_space + fixes rGamma :: "'a \ 'a" assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \ (\n. z = - of_nat n)" assumes differentiable_rGamma_aux1: "(\n. z \ - of_nat n) \ let d = (THE d. (\n. \k d) - scaleR euler_mascheroni 1 in filterlim (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R norm (y - z)) (nhds 0) (at z)" assumes differentiable_rGamma_aux2: "let z = - of_nat n in filterlim (\y. (rGamma y - rGamma z - (-1)^n * (prod of_nat {1..n}) * (y - z)) /\<^sub>R norm (y - z)) (nhds 0) (at z)" assumes rGamma_series_aux: "(\n. z \ - of_nat n) \ let fact' = (\n. prod of_nat {1..n}); exp = (\x. THE e. (\n. \kR fact k) \ e); pochhammer' = (\a n. (\n = 0..n. a + of_nat n)) in filterlim (\n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) (nhds (rGamma z)) sequentially" begin subclass banach .. end definition "Gamma z = inverse (rGamma z)" subsection \Basic properties\ lemma Gamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ Gamma z = 0" and rGamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ rGamma z = 0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma Gamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ Gamma z \ 0" and rGamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ rGamma z \ 0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma Gamma_eq_zero_iff: "Gamma z = 0 \ z \ \\<^sub>\\<^sub>0" and rGamma_eq_zero_iff: "rGamma z = 0 \ z \ \\<^sub>\\<^sub>0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)" unfolding Gamma_def by simp lemma rGamma_series_LIMSEQ [tendsto_intros]: "rGamma_series z \ rGamma z" proof (cases "z \ \\<^sub>\\<^sub>0") case False hence "z \ - of_nat n" for n by auto from rGamma_series_aux[OF this] show ?thesis by (simp add: rGamma_series_def[abs_def] fact_prod pochhammer_Suc_prod exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost) qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ) theorem Gamma_series_LIMSEQ [tendsto_intros]: "Gamma_series z \ Gamma z" proof (cases "z \ \\<^sub>\\<^sub>0") case False hence "(\n. inverse (rGamma_series z n)) \ inverse (rGamma z)" by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff) also have "(\n. inverse (rGamma_series z n)) = Gamma_series z" by (simp add: rGamma_series_def Gamma_series_def[abs_def]) finally show ?thesis by (simp add: Gamma_def) qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ) lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)" using Gamma_series_LIMSEQ[of z] by (simp add: limI) lemma rGamma_1 [simp]: "rGamma 1 = 1" proof - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact field_split_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) moreover have "rGamma_series 1 \ rGamma 1" by (rule tendsto_intros) ultimately show ?thesis by (intro LIMSEQ_unique) qed lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z" proof - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) also from n have "\ = ?f n * rGamma_series z n" by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def) finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. qed moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" by (intro tendsto_intros lim_inverse_n) hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" by (blast intro: Lim_transform_eventually) moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" by (intro tendsto_intros) ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast qed lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)" proof (induction n arbitrary: z) case (Suc n z) have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH) also note rGamma_plus1 [symmetric] finally show ?case by (simp add: add_ac pochhammer_rec') qed simp_all theorem Gamma_plus1: "z \ \\<^sub>\\<^sub>0 \ Gamma (z + 1) = z * Gamma z" using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff) theorem pochhammer_Gamma: "z \ \\<^sub>\\<^sub>0 \ pochhammer z n = Gamma (z + of_nat n) / Gamma z" using pochhammer_rGamma[of z] by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps) lemma Gamma_0 [simp]: "Gamma 0 = 0" and rGamma_0 [simp]: "rGamma 0 = 0" and Gamma_neg_1 [simp]: "Gamma (- 1) = 0" and rGamma_neg_1 [simp]: "rGamma (- 1) = 0" and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0" and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0" and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0" and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0" by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff) lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp theorem Gamma_fact: "Gamma (1 + of_nat n) = fact n" by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff flip: of_nat_Suc) lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)" by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst of_nat_Suc, subst Gamma_fact) (rule refl) lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)" proof (cases "n > 0") case True hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int) lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)" by (simp add: Gamma_of_int rGamma_inverse_Gamma) lemma Gamma_seriesI: assumes "(\n. g n / Gamma_series z n) \ 1" shows "g \ Gamma z" proof (rule Lim_transform_eventually) have "1/2 > (0::real)" by simp from tendstoD[OF assms, OF this] show "eventually (\n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially" by (force elim!: eventually_mono simp: dist_real_def) from assms have "(\n. g n / Gamma_series z n * Gamma_series z n) \ 1 * Gamma z" by (intro tendsto_intros) thus "(\n. g n / Gamma_series z n * Gamma_series z n) \ Gamma z" by simp qed lemma Gamma_seriesI': assumes "f \ rGamma z" assumes "(\n. g n * f n) \ 1" assumes "z \ \\<^sub>\\<^sub>0" shows "g \ Gamma z" proof (rule Lim_transform_eventually) have "1/2 > (0::real)" by simp from tendstoD[OF assms(2), OF this] show "eventually (\n. g n * f n / f n = g n) sequentially" by (force elim!: eventually_mono simp: dist_real_def) from assms have "(\n. g n * f n / f n) \ 1 / rGamma z" by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff) thus "(\n. g n * f n / f n) \ Gamma z" by (simp add: Gamma_def divide_inverse) qed lemma Gamma_series'_LIMSEQ: "Gamma_series' z \ Gamma z" by (cases "z \ \\<^sub>\\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] Gamma_series'_nonpos_Ints_LIMSEQ[of z]) subsection \Differentiability\ lemma has_field_derivative_rGamma_no_nonpos_int: assumes "z \ \\<^sub>\\<^sub>0" shows "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)" proof (rule has_field_derivative_at_within) from assms have "z \ - of_nat n" for n by auto from differentiable_rGamma_aux1[OF this] show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)" unfolding Digamma_def suminf_def sums_def[abs_def] has_field_derivative_def has_derivative_def netlimit_at by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric]) qed lemma has_field_derivative_rGamma_nonpos_int: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" apply (rule has_field_derivative_at_within) using differentiable_rGamma_aux2[of n] unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_prod) simp lemma has_field_derivative_rGamma [derivative_intros]: "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \norm z\) * fact (nat \norm z\) else -rGamma z * Digamma z)) (at z within A)" using has_field_derivative_rGamma_no_nonpos_int[of z A] has_field_derivative_rGamma_nonpos_int[of "nat \norm z\" A] by (auto elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros] declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros] declare has_field_derivative_rGamma_nonpos_int [derivative_intros] declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros] declare has_field_derivative_rGamma [derivative_intros] theorem has_field_derivative_Gamma [derivative_intros]: "z \ \\<^sub>\\<^sub>0 \ (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)" unfolding Gamma_def [abs_def] by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff) declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros] (* TODO: Hide ugly facts properly *) hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2 differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma" by (rule DERIV_continuous_on has_field_derivative_rGamma)+ lemma continuous_on_Gamma [continuous_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A Gamma" by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast lemma isCont_rGamma [continuous_intros]: "isCont f z \ isCont (\x. rGamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_rGamma]]) lemma isCont_Gamma [continuous_intros]: "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Gamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) subsection\<^marker>\tag unimportant\ \The complex Gamma function\ instantiation\<^marker>\tag unimportant\ complex :: Gamma begin definition\<^marker>\tag unimportant\ rGamma_complex :: "complex \ complex" where "rGamma_complex z = lim (rGamma_series z)" lemma rGamma_series_complex_converges: "convergent (rGamma_series (z :: complex))" (is "?thesis1") and rGamma_complex_altdef: "rGamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2") proof - have "?thesis1 \ ?thesis2" proof (cases "z \ \\<^sub>\\<^sub>0") case False have "rGamma_series z \ exp (- ln_Gamma z)" proof (rule Lim_transform_eventually) from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE) from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] have "ln_Gamma_series z \ lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff) thus "(\n. exp (-ln_Gamma_series z n)) \ exp (- ln_Gamma z)" unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus) from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False show "eventually (\n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially" by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def) qed with False show ?thesis by (auto simp: convergent_def rGamma_complex_def intro!: limI) next case True then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases') also have "rGamma_series \ \ 0" by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const) finally show ?thesis using True by (auto simp: rGamma_complex_def convergent_def intro!: limI) qed thus "?thesis1" "?thesis2" by blast+ qed context\<^marker>\tag unimportant\ begin (* TODO: duplication *) private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)" proof - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) also from n have "\ = ?f n * rGamma_series z n" by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def add_ac) finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. qed moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" using rGamma_series_complex_converges by (intro tendsto_intros lim_inverse_n) (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def) hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" by (blast intro: Lim_transform_eventually) moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" using rGamma_series_complex_converges by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff) ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast qed private lemma has_field_derivative_rGamma_complex_no_nonpos_Int: assumes "(z :: complex) \ \\<^sub>\\<^sub>0" shows "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" proof - have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z proof (subst DERIV_cong_ev[OF refl _ refl]) from that have "eventually (\t. t \ ball z (Re z/2)) (nhds z)" by (intro eventually_nhds_in_nhd) simp_all thus "eventually (\t. rGamma t = exp (- ln_Gamma t)) (nhds z)" using no_nonpos_Int_in_ball_complex[OF that] by (auto elim!: eventually_mono simp: rGamma_complex_altdef) next have "z \ \\<^sub>\\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff) with that show "((\t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)" by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef) qed from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" proof (induction "nat \1 - Re z\" arbitrary: z) case (Suc n z) from Suc.prems have z: "z \ 0" by auto from Suc.hyps have "n = nat \- Re z\" by linarith hence A: "n = nat \1 - Re (z + 1)\" by simp from Suc.prems have B: "z + 1 \ \\<^sub>\\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp) have "((\z. z * (rGamma \ (\z. z + 1)) z) has_field_derivative -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)" by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps) also have "(\z. z * (rGamma \ (\z. z + 1 :: complex)) z) = rGamma" by (simp add: rGamma_complex_plus1) also from z have "Digamma (z + 1) * z - 1 = z * Digamma z" by (subst Digamma_plus1) (simp_all add: field_simps) also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z" by (simp add: rGamma_complex_plus1[of z, symmetric]) finally show ?case . qed (intro diff, simp) qed private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1" proof - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact field_split_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI) qed private lemma has_field_derivative_rGamma_complex_nonpos_Int: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))" proof (induction n) case 0 have A: "(0::complex) + 1 \ \\<^sub>\\<^sub>0" by simp have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)" by (rule derivative_eq_intros DERIV_chain refl has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1) thus ?case by (simp add: rGamma_complex_plus1) next case (Suc n) hence A: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat (Suc n) + 1 :: complex))" by simp have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))" by (rule derivative_eq_intros refl A DERIV_chain)+ (simp add: algebra_simps rGamma_complex_altdef) thus ?case by (simp add: rGamma_complex_plus1) qed instance proof fix z :: complex show "(rGamma z = 0) \ (\n. z = - of_nat n)" by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases') next fix z :: complex assume "\n. z \ - of_nat n" hence "z \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases') from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this] show "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] of_real_def[symmetric] suminf_def) next fix n :: nat from has_field_derivative_rGamma_complex_nonpos_Int[of n] show "let z = - of_nat n in (\y. (rGamma y - rGamma z - (- 1) ^ n * prod of_nat {1..n} * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" by (simp add: has_field_derivative_def has_derivative_def fact_prod Let_def) next fix z :: complex from rGamma_series_complex_converges[of z] have "rGamma_series z \ rGamma z" by (simp add: convergent_LIMSEQ_iff rGamma_complex_def) thus "let fact' = \n. prod of_nat {1..n}; exp = \x. THE e. (\n. \kR fact k) \ e; pochhammer' = \a n. \n = 0..n. a + of_nat n in (\n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma z" by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) qed end end lemma Gamma_complex_altdef: "Gamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))" unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus) lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)" proof - have "rGamma_series (cnj z) = (\n. cnj (rGamma_series z n))" by (intro ext) (simp_all add: rGamma_series_def exp_cnj) also have "... \ cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros) finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI]) qed lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)" unfolding Gamma_def by (simp add: cnj_rGamma) lemma Gamma_complex_real: "z \ \ \ Gamma z \ (\ :: complex set)" and rGamma_complex_real: "z \ \ \ rGamma z \ \" by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma) lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)" using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast lemma holomorphic_rGamma [holomorphic_intros]: "rGamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) lemma holomorphic_rGamma' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. rGamma (f x)) holomorphic_on A" proof - have "rGamma \ f holomorphic_on A" using assms by (intro holomorphic_on_compose assms holomorphic_rGamma) thus ?thesis by (simp only: o_def) qed lemma analytic_rGamma: "rGamma analytic_on A" unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_rGamma) lemma field_differentiable_Gamma: "z \ \\<^sub>\\<^sub>0 \ Gamma field_differentiable (at z within A)" using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto lemma holomorphic_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) lemma holomorphic_Gamma' [holomorphic_intros]: assumes "f holomorphic_on A" and "\x. x \ A \ f x \ \\<^sub>\\<^sub>0" shows "(\x. Gamma (f x)) holomorphic_on A" proof - have "Gamma \ f holomorphic_on A" using assms by (intro holomorphic_on_compose assms holomorphic_Gamma) auto thus ?thesis by (simp only: o_def) qed lemma analytic_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_Gamma) lemma field_differentiable_ln_Gamma_complex: "z \ \\<^sub>\\<^sub>0 \ ln_Gamma field_differentiable (at (z::complex) within A)" by (rule field_differentiable_within_subset[of _ _ UNIV]) (force simp: field_differentiable_def intro!: derivative_intros)+ lemma holomorphic_ln_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ ln_Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_ln_Gamma_complex) lemma analytic_ln_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ ln_Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_ln_Gamma) lemma has_field_derivative_rGamma_complex' [derivative_intros]: "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \-Re z\) * fact (nat \-Re z\) else -rGamma z * Digamma z)) (at z within A)" using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros] lemma field_differentiable_Polygamma: fixes z :: complex shows "z \ \\<^sub>\\<^sub>0 \ Polygamma n field_differentiable (at z within A)" using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto lemma holomorphic_on_Polygamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma) lemma analytic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_on_Polygamma) subsection\<^marker>\tag unimportant\ \The real Gamma function\ lemma rGamma_series_real: "eventually (\n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially" using eventually_gt_at_top[of "0 :: nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have "Re (rGamma_series (of_real x) n) = Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))" using n by (simp add: rGamma_series_def powr_def pochhammer_of_real) also from n have "\ = Re (of_real ((pochhammer x (Suc n)) / (fact n * (exp (x * ln (real_of_nat n))))))" by (subst exp_of_real) simp also from n have "\ = rGamma_series x n" by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def) finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" .. qed instantiation\<^marker>\tag unimportant\ real :: Gamma begin definition "rGamma_real x = Re (rGamma (of_real x :: complex))" instance proof fix x :: real have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def) also have "of_real \ = rGamma (of_real x :: complex)" by (intro of_real_Re rGamma_complex_real) simp_all also have "\ = 0 \ x \ \\<^sub>\\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff) also have "\ \ (\n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases') finally show "(rGamma x) = 0 \ (\n. x = - real_of_nat n)" by simp next fix x :: real assume "\n. x \ - of_nat n" hence x: "complex_of_real x \ \\<^sub>\\<^sub>0" by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases') then have "x \ 0" by auto with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)" by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real rGamma_real_def [abs_def]) thus "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma x + rGamma x * d * (y - x)) /\<^sub>R norm (y - x)) \x\ 0" by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] of_real_def[symmetric] suminf_def) next fix n :: nat have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))" by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real rGamma_real_def [abs_def]) thus "let x = - of_nat n in (\y. (rGamma y - rGamma x - (- 1) ^ n * prod of_nat {1..n} * (y - x)) /\<^sub>R norm (y - x)) \x::real\ 0" by (simp add: has_field_derivative_def has_derivative_def fact_prod Let_def) next fix x :: real have "rGamma_series x \ rGamma x" proof (rule Lim_transform_eventually) show "(\n. Re (rGamma_series (of_real x) n)) \ rGamma x" unfolding rGamma_real_def by (intro tendsto_intros) qed (insert rGamma_series_real, simp add: eq_commute) thus "let fact' = \n. prod of_nat {1..n}; exp = \x. THE e. (\n. \kR fact k) \ e; pochhammer' = \a n. \n = 0..n. a + of_nat n in (\n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma x" by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) qed end lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)" unfolding rGamma_real_def using rGamma_complex_real by simp lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)" unfolding Gamma_def by (simp add: rGamma_complex_of_real) lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))" by (rule sym, rule limI, rule tendsto_intros) lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))" by (rule sym, rule limI, rule tendsto_intros) lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))" using rGamma_complex_real[OF Reals_of_real[of x]] by (elim Reals_cases) (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real) lemma ln_Gamma_series_complex_of_real: "x > 0 \ n > 0 \ ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)" proof - assume xn: "x > 0" "n > 0" have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \ 1" for k using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps) with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_real) qed lemma ln_Gamma_real_converges: assumes "(x::real) > 0" shows "convergent (ln_Gamma_series x)" proof - have "(\n. ln_Gamma_series (complex_of_real x) n) \ ln_Gamma (of_real x)" using assms by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff) moreover from eventually_gt_at_top[of "0::nat"] have "eventually (\n. complex_of_real (ln_Gamma_series x n) = ln_Gamma_series (complex_of_real x) n) sequentially" by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms) ultimately have "(\n. complex_of_real (ln_Gamma_series x n)) \ ln_Gamma (of_real x)" by (subst tendsto_cong) assumption+ from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def) qed lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \ ln_Gamma_series x \ ln_Gamma x" using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff) lemma ln_Gamma_complex_of_real: "x > 0 \ ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)" proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually) assume x: "x > 0" show "eventually (\n. of_real (ln_Gamma_series x n) = ln_Gamma_series (complex_of_real x) n) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_Gamma_series_complex_of_real x) qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def) lemma Gamma_real_pos_exp: "x > (0 :: real) \ Gamma x = exp (ln_Gamma x)" by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff ln_Gamma_complex_of_real exp_of_real) lemma ln_Gamma_real_pos: "x > 0 \ ln_Gamma x = ln (Gamma x :: real)" unfolding Gamma_real_pos_exp by simp lemma ln_Gamma_complex_conv_fact: "n > 0 \ ln_Gamma (of_nat n :: complex) = ln (fact (n - 1))" using ln_Gamma_complex_of_real[of "real n"] Gamma_fact[of "n - 1", where 'a = real] by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric]) lemma ln_Gamma_real_conv_fact: "n > 0 \ ln_Gamma (real n) = ln (fact (n - 1))" using Gamma_fact[of "n - 1", where 'a = real] by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric]) lemma Gamma_real_pos [simp, intro]: "x > (0::real) \ Gamma x > 0" by (simp add: Gamma_real_pos_exp) lemma Gamma_real_nonneg [simp, intro]: "x > (0::real) \ Gamma x \ 0" by (simp add: Gamma_real_pos_exp) lemma has_field_derivative_ln_Gamma_real [derivative_intros]: assumes x: "x > (0::real)" shows "(ln_Gamma has_field_derivative Digamma x) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms show "((Re \ ln_Gamma \ complex_of_real) has_field_derivative Digamma x) (at x)" by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real o_def) from eventually_nhds_in_nhd[of x "{0<..}"] assms show "eventually (\y. ln_Gamma y = (Re \ ln_Gamma \ of_real) y) (nhds x)" by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open) qed lemma field_differentiable_ln_Gamma_real: "x > 0 \ ln_Gamma field_differentiable (at (x::real) within A)" by (rule field_differentiable_within_subset[of _ _ UNIV]) (auto simp: field_differentiable_def intro!: derivative_intros)+ declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros] lemma deriv_ln_Gamma_real: assumes "z > 0" shows "deriv ln_Gamma z = Digamma (z :: real)" by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_real assms) lemma has_field_derivative_rGamma_real' [derivative_intros]: "(rGamma has_field_derivative (if x \ \\<^sub>\\<^sub>0 then (-1)^(nat \-x\) * fact (nat \-x\) else -rGamma x * Digamma x)) (at x within A)" using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros] lemma Polygamma_real_odd_pos: assumes "(x::real) \ \\<^sub>\\<^sub>0" "odd n" shows "Polygamma n x > 0" proof - from assms have "x \ 0" by auto with assms show ?thesis unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] by (auto simp: zero_less_power_eq simp del: power_Suc dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos) qed lemma Polygamma_real_even_neg: assumes "(x::real) > 0" "n > 0" "even n" shows "Polygamma n x < 0" using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] by (auto intro!: mult_pos_pos suminf_pos) lemma Polygamma_real_strict_mono: assumes "x > 0" "x < (y::real)" "even n" shows "Polygamma n x < Polygamma n y" proof - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) then guess \ by (elim exE conjE) note \ = this note \(3) also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ > 0" by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases) finally show ?thesis by simp qed lemma Polygamma_real_strict_antimono: assumes "x > 0" "x < (y::real)" "odd n" shows "Polygamma n x > Polygamma n y" proof - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) then guess \ by (elim exE conjE) note \ = this note \(3) also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ < 0" by (intro mult_pos_neg Polygamma_real_even_neg) simp_all finally show ?thesis by simp qed lemma Polygamma_real_mono: assumes "x > 0" "x \ (y::real)" "even n" shows "Polygamma n x \ Polygamma n y" using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) by (cases "x = y") simp_all lemma Digamma_real_strict_mono: "(0::real) < x \ x < y \ Digamma x < Digamma y" by (rule Polygamma_real_strict_mono) simp_all lemma Digamma_real_mono: "(0::real) < x \ x \ y \ Digamma x \ Digamma y" by (rule Polygamma_real_mono) simp_all lemma Digamma_real_ge_three_halves_pos: assumes "x \ 3/2" shows "Digamma (x :: real) > 0" proof - have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos) also from assms have "\ \ Digamma x" by (intro Polygamma_real_mono) simp_all finally show ?thesis . qed lemma ln_Gamma_real_strict_mono: assumes "x \ 3/2" "x < y" shows "ln_Gamma (x :: real) < ln_Gamma y" proof - have "\\. x < \ \ \ < y \ ln_Gamma y - ln_Gamma x = (y - x) * Digamma \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) then guess \ by (elim exE conjE) note \ = this note \(3) also from \(1,2) assms have "(y - x) * Digamma \ > 0" by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all finally show ?thesis by simp qed lemma Gamma_real_strict_mono: assumes "x \ 3/2" "x < y" shows "Gamma (x :: real) < Gamma y" proof - from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp also have "\ < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms) also from Gamma_real_pos_exp[of y] assms have "\ = Gamma y" by simp finally show ?thesis . qed theorem log_convex_Gamma_real: "convex_on {0<..} (ln \ Gamma :: real \ real)" by (rule convex_on_realI[of _ _ Digamma]) (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases') subsection \The uniqueness of the real Gamma function\ text \ The following is a proof of the Bohr--Mollerup theorem, which states that any log-convex function \G\ on the positive reals that fulfils \G(1) = 1\ and satisfies the functional equation \G(x + 1) = x G(x)\ must be equal to the Gamma function. In principle, if \G\ is a holomorphic complex function, one could then extend this from the positive reals to the entire complex plane (minus the non-positive integers, where the Gamma function is not defined). \ context\<^marker>\tag unimportant\ fixes G :: "real \ real" assumes G_1: "G 1 = 1" assumes G_plus1: "x > 0 \ G (x + 1) = x * G x" assumes G_pos: "x > 0 \ G x > 0" assumes log_convex_G: "convex_on {0<..} (ln \ G)" begin private lemma G_fact: "G (of_nat n + 1) = fact n" using G_plus1[of "real n + 1" for n] by (induction n) (simp_all add: G_1 G_plus1) private definition S :: "real \ real \ real" where "S x y = (ln (G y) - ln (G x)) / (y - x)" private lemma S_eq: "n \ 2 \ S (of_nat n) (of_nat n + x) = (ln (G (real n + x)) - ln (fact (n - 1))) / x" by (subst G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff) private lemma G_lower: assumes x: "x > 0" and n: "n \ 1" shows "Gamma_series x n \ G x" proof - have "(ln \ G) (real (Suc n)) \ ((ln \ G) (real (Suc n) + x) - (ln \ G) (real (Suc n) - 1)) / (real (Suc n) + x - (real (Suc n) - 1)) * (real (Suc n) - (real (Suc n) - 1)) + (ln \ G) (real (Suc n) - 1)" using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto hence "S (of_nat n) (of_nat (Suc n)) \ S (of_nat (Suc n)) (of_nat (Suc n) + x)" unfolding S_def using x by (simp add: field_simps) also have "S (of_nat n) (of_nat (Suc n)) = ln (fact n) - ln (fact (n-1))" unfolding S_def using n by (subst (1 2) G_fact [symmetric]) (simp_all add: add_ac of_nat_diff) also have "\ = ln (fact n / fact (n-1))" by (subst ln_div) simp_all also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all finally have "x * ln (real n) + ln (fact n) \ ln (G (real (Suc n) + x))" using x n by (subst (asm) S_eq) (simp_all add: field_simps) also have "x * ln (real n) + ln (fact n) = ln (exp (x * ln (real n)) * fact n)" using x by (simp add: ln_mult) finally have "exp (x * ln (real n)) * fact n \ G (real (Suc n) + x)" using x by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos) also have "G (real (Suc n) + x) = pochhammer x (Suc n) * G x" using G_plus1[of "real (Suc n) + x" for n] G_plus1[of x] x by (induction n) (simp_all add: pochhammer_Suc add_ac) finally show "Gamma_series x n \ G x" using x by (simp add: field_simps pochhammer_pos Gamma_series_def) qed private lemma G_upper: assumes x: "x > 0" "x \ 1" and n: "n \ 2" shows "G x \ Gamma_series x n * (1 + x / real n)" proof - have "(ln \ G) (real n + x) \ ((ln \ G) (real n + 1) - (ln \ G) (real n)) / (real n + 1 - (real n)) * ((real n + x) - real n) + (ln \ G) (real n)" using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto hence "S (of_nat n) (of_nat n + x) \ S (of_nat n) (of_nat n + 1)" unfolding S_def using x by (simp add: field_simps) also from n have "S (of_nat n) (of_nat n + 1) = ln (fact n) - ln (fact (n-1))" by (subst (1 2) G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff) also have "\ = ln (fact n / (fact (n-1)))" using n by (subst ln_div) simp_all also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all finally have "ln (G (real n + x)) \ x * ln (real n) + ln (fact (n - 1))" using x n by (subst (asm) S_eq) (simp_all add: field_simps) also have "\ = ln (exp (x * ln (real n)) * fact (n - 1))" using x by (simp add: ln_mult) finally have "G (real n + x) \ exp (x * ln (real n)) * fact (n - 1)" using x by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos) also have "G (real n + x) = pochhammer x n * G x" using G_plus1[of "real n + x" for n] x by (induction n) (simp_all add: pochhammer_Suc add_ac) finally have "G x \ exp (x * ln (real n)) * fact (n- 1) / pochhammer x n" using x by (simp add: field_simps pochhammer_pos) also from n have "fact (n - 1) = fact n / n" by (cases n) simp_all also have "exp (x * ln (real n)) * \ / pochhammer x n = Gamma_series x n * (1 + x / real n)" using n x by (simp add: Gamma_series_def divide_simps pochhammer_Suc) finally show ?thesis . qed private lemma G_eq_Gamma_aux: assumes x: "x > 0" "x \ 1" shows "G x = Gamma x" proof (rule antisym) show "G x \ Gamma x" proof (rule tendsto_upperbound) from G_lower[of x] show "eventually (\n. Gamma_series x n \ G x) sequentially" using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "1::nat"]]) qed (simp_all add: Gamma_series_LIMSEQ) next show "G x \ Gamma x" proof (rule tendsto_lowerbound) have "(\n. Gamma_series x n * (1 + x / real n)) \ Gamma x * (1 + 0)" by (rule tendsto_intros real_tendsto_divide_at_top Gamma_series_LIMSEQ filterlim_real_sequentially)+ thus "(\n. Gamma_series x n * (1 + x / real n)) \ Gamma x" by simp next from G_upper[of x] show "eventually (\n. Gamma_series x n * (1 + x / real n) \ G x) sequentially" using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "2::nat"]]) qed simp_all qed theorem Gamma_pos_real_unique: assumes x: "x > 0" shows "G x = Gamma x" proof - have G_eq: "G (real n + x) = Gamma (real n + x)" if "x \ {0<..1}" for n x using that proof (induction n) case (Suc n) from Suc have "x + real n > 0" by simp hence "x + real n \ \\<^sub>\\<^sub>0" by auto with Suc show ?case using G_plus1[of "real n + x"] Gamma_plus1[of "real n + x"] by (auto simp: add_ac) qed (simp_all add: G_eq_Gamma_aux) show ?thesis proof (cases "frac x = 0") case True hence "x = of_int (floor x)" by (simp add: frac_def) with x have x_eq: "x = of_nat (nat (floor x) - 1) + 1" by simp show ?thesis by (subst (1 2) x_eq, rule G_eq) simp_all next case False from assms have x_eq: "x = of_nat (nat (floor x)) + frac x" by (simp add: frac_def) have frac_le_1: "frac x \ 1" unfolding frac_def by linarith show ?thesis by (subst (1 2) x_eq, rule G_eq, insert False frac_le_1) simp_all qed qed end subsection \The Beta function\ definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)" lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)" by (simp add: inverse_eq_divide Beta_def Gamma_def) lemma Beta_commute: "Beta a b = Beta b a" unfolding Beta_def by (simp add: ac_simps) lemma has_field_derivative_Beta1 [derivative_intros]: assumes "x \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" shows "((\x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) (at x within A)" unfolding Beta_altdef by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps) lemma Beta_pole1: "x \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma Beta_pole2: "y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma Beta_zero: "x + y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma has_field_derivative_Beta2 [derivative_intros]: assumes "y \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" shows "((\y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) (at y within A)" using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac) theorem Beta_plus1_plus1: assumes "x \ \\<^sub>\\<^sub>0" "y \ \\<^sub>\\<^sub>0" shows "Beta (x + 1) y + Beta x (y + 1) = Beta x y" proof - have "Beta (x + 1) y + Beta x (y + 1) = (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)" by (simp add: Beta_altdef add_divide_distrib algebra_simps) also have "\ = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))" by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps) also from assms have "\ = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp finally show ?thesis . qed theorem Beta_plus1_left: assumes "x \ \\<^sub>\\<^sub>0" shows "(x + y) * Beta (x + 1) y = x * Beta x y" proof - have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))" unfolding Beta_altdef by (simp only: ac_simps) also have "\ = x * Beta x y" unfolding Beta_altdef by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps) finally show ?thesis . qed theorem Beta_plus1_right: assumes "y \ \\<^sub>\\<^sub>0" shows "(x + y) * Beta x (y + 1) = y * Beta x y" using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute) lemma Gamma_Gamma_Beta: assumes "x + y \ \\<^sub>\\<^sub>0" shows "Gamma x * Gamma y = Beta x y * Gamma (x + y)" unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"] by (simp add: rGamma_inverse_Gamma) subsection \Legendre duplication theorem\ context begin private lemma Gamma_legendre_duplication_aux: fixes z :: "'a :: Gamma" assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)" proof - let ?powr = "\b a. exp (a * of_real (ln (of_nat b)))" let ?h = "\n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * exp (1/2 * of_real (ln (real_of_nat n)))" { fix z :: 'a assume z: "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" let ?g = "\n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / Gamma_series' (2*z) (2*n)" have "eventually (\n. ?g n = ?h n) sequentially" using eventually_gt_at_top proof eventually_elim fix n :: nat assume n: "n > 0" let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a" have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) / (pochhammer z n * pochhammer (z + 1/2) n)" by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac) have B: "Gamma_series' (2*z) (2*n) = ?f' * ?powr 2 (2*z) * ?powr n (2*z) / (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double) from z have "pochhammer z n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) moreover from z have "pochhammer (z + 1/2) n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))" using n unfolding A B by (simp add: field_split_simps exp_minus) also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)" by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib) finally show "?g n = ?h n" by (simp only: mult_ac) qed moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \ \\<^sub>\\<^sub>0" by auto hence "?g \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "(*)2" "2*z"] by (intro tendsto_intros Gamma_series'_LIMSEQ) (simp_all add: o_def strict_mono_def Gamma_eq_zero_iff) ultimately have "?h \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" by (blast intro: Lim_transform_eventually) } note lim = this from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \ \\<^sub>\\<^sub>0" by auto from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \ \\<^sub>\\<^sub>0" by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all with lim[of "1/2 :: 'a"] have "?h \ 2 * Gamma (1/2 :: 'a)" by (simp add: exp_of_real) from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis by (simp add: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real) qed text \ The following lemma is somewhat annoying. With a little bit of complex analysis (Cauchy's integral theorem, to be exact), this would be completely trivial. However, we want to avoid depending on the complex analysis session at this point, so we prove it the hard way. \ private lemma Gamma_reflection_aux: defines "h \ \z::complex. if z \ \ then 0 else (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" defines "a \ complex_of_real pi" obtains h' where "continuous_on UNIV h'" "\z. (h has_field_derivative (h' z)) (at z)" proof - define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z define g where "g n = complex_of_real (sin_coeff (n+1))" for n define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z have a_nz: "a \ 0" unfolding a_def by simp have "(\n. f n * (a*z)^n) sums (F z) \ (\n. g n * (a*z)^n) sums (G z)" if "abs (Re z) < 1" for z proof (cases "z = 0"; rule conjI) assume "z \ 0" note z = this that from z have sin_nz: "sin (a*z) \ 0" unfolding a_def by (auto simp: sin_eq_0) have "(\n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"] by (simp add: scaleR_conv_of_real) from sums_split_initial_segment[OF this, of 1] have "(\n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac) from sums_mult[OF this, of "inverse (a*z)"] z a_nz have A: "(\n. g n * (a*z)^n) sums (sin (a*z)/(a*z))" by (simp add: field_simps g_def) with z show "(\n. g n * (a*z)^n) sums (G z)" by (simp add: G_def) from A z a_nz sin_nz have g_nz: "(\n. g n * (a*z)^n) \ 0" by (simp add: sums_iff g_def) have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def) from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1] have "(\n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))" by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def) from sums_mult[OF this, of "inverse z"] z assms show "(\n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def) next assume z: "z = 0" have "(\n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp with z show "(\n. f n * (a * z) ^ n) sums (F z)" by (simp add: f_def F_def sin_coeff_def cos_coeff_def) have "(\n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp with z show "(\n. g n * (a * z) ^ n) sums (G z)" by (simp add: g_def G_def sin_coeff_def cos_coeff_def) qed note sums = conjunct1[OF this] conjunct2[OF this] define h2 where [abs_def]: "h2 z = (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z define POWSER where [abs_def]: "POWSER f z = (\n. f n * (z^n :: complex))" for f z define POWSER' where [abs_def]: "POWSER' f z = (\n. diffs f n * (z^n))" for f and z :: complex define h2' where [abs_def]: "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t proof - from that have t: "t \ \ \ t = 0" by (auto elim!: Ints_cases) hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def) also have "a*cot (a*t) - 1/t = (F t) / (G t)" using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def) also have "\ = (\n. f n * (a*t)^n) / (\n. g n * (a*t)^n)" using sums[of t] that by (simp add: sums_iff) finally show "h t = h2 t" by (simp only: h2_def) qed let ?A = "{z. abs (Re z) < 1}" have "open ({z. Re z < 1} \ {z. Re z > -1})" using open_halfspace_Re_gt open_halfspace_Re_lt by auto also have "({z. Re z < 1} \ {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto finally have open_A: "open ?A" . hence [simp]: "interior ?A = ?A" by (simp add: interior_open) have summable_f: "summable (\n. f n * z^n)" for z by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) (simp_all add: norm_mult a_def del: of_real_add) have summable_g: "summable (\n. g n * z^n)" for z by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) (simp_all add: norm_mult a_def del: of_real_add) have summable_fg': "summable (\n. diffs f n * z^n)" "summable (\n. diffs g n * z^n)" for z by (intro termdiff_converges_all summable_f summable_g)+ have "(POWSER f has_field_derivative (POWSER' f z)) (at z)" "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z unfolding POWSER_def POWSER'_def by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+ note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def] have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" for z unfolding POWSER_def POWSER'_def by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+ note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def] { fix z :: complex assume z: "abs (Re z) < 1" define d where "d = \ * of_real (norm z + 1)" have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) have "eventually (\z. h z = h2 z) (nhds z)" using eventually_nhds_in_nhd[of z ?A] using h_eq z by (auto elim!: eventually_mono) moreover from sums(2)[OF z] z have nz: "(\n. g n * (a * z) ^ n) \ 0" unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def) have A: "z \ \ \ z = 0" using z by (auto elim!: Ints_cases) have no_int: "1 + z \ \ \ z = 0" using z Ints_diff[of "1+z" 1] A by (auto elim!: nonpos_Ints_cases) have no_int': "1 - z \ \ \ z = 0" using z Ints_diff[of 1 "1-z"] A by (auto elim!: nonpos_Ints_cases) from no_int no_int' have no_int: "1 - z \ \\<^sub>\\<^sub>0" "1 + z \ \\<^sub>\\<^sub>0" by auto have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+) (auto simp: h2'_def POWSER_def field_simps power2_eq_square) ultimately have deriv: "(h has_field_derivative h2' z) (at z)" by (subst DERIV_cong_ev[OF refl _ refl]) from sums(2)[OF z] z have "(\n. g n * (a * z) ^ n) \ 0" unfolding G_def by (auto simp: sums_iff a_def sin_eq_0) hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def by (intro continuous_intros cont continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto note deriv and this } note A = this interpret h: periodic_fun_simple' h proof fix z :: complex show "h (z + 1) = h z" proof (cases "z \ \") assume z: "z \ \" hence A: "z + 1 \ \" "z \ 0" using Ints_diff[of "z+1" 1] by auto hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)" by (subst (1 2) Digamma_plus1) simp_all with A z show "h (z + 1) = h z" by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def) qed (simp add: h_def) qed have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z proof - have "((\z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)" by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) (insert z, auto intro!: derivative_eq_intros) hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1) moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) qed define h2'' where "h2'' z = h2' (z - of_int \Re z\)" for z have deriv: "(h has_field_derivative h2'' z) (at z)" for z proof - fix z :: complex have B: "\Re z - real_of_int \Re z\\ < 1" by linarith have "((\t. h (t - of_int \Re z\)) has_field_derivative h2'' z) (at z)" unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) (insert B, auto intro!: derivative_intros) thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int) qed have cont: "continuous_on UNIV h2''" proof (intro continuous_at_imp_continuous_on ballI) fix z :: complex define r where "r = \Re z\" define A where "A = {t. of_int r - 1 < Re t \ Re t < of_int r + 1}" have "continuous_on A (\t. h2' (t - of_int r))" unfolding A_def by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) (simp_all add: abs_real_def) moreover have "h2'' t = h2' (t - of_int r)" if t: "t \ A" for t proof (cases "Re t \ of_int r") case True from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) with True have "\Re t\ = \Re z\" unfolding r_def by linarith thus ?thesis by (auto simp: r_def h2''_def) next case False from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) with False have t': "\Re t\ = \Re z\ - 1" unfolding r_def by linarith moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)" by (intro h2'_eq) simp_all ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t') qed ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl]) moreover { have "open ({t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t})" by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt) also have "{t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t} = A" unfolding A_def by blast finally have "open A" . } ultimately have C: "isCont h2'' t" if "t \ A" for t using that by (subst (asm) continuous_on_eq_continuous_at) auto have "of_int r - 1 < Re z" "Re z < of_int r + 1" unfolding r_def by linarith+ thus "isCont h2'' z" by (intro C) (simp_all add: A_def) qed from that[OF cont deriv] show ?thesis . qed lemma Gamma_reflection_complex: fixes z :: complex shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" proof - let ?g = "\z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" define g where [abs_def]: "g z = (if z \ \ then of_real pi else ?g z)" for z :: complex let ?h = "\z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" define h where [abs_def]: "h z = (if z \ \ then 0 else ?h z)" for z :: complex \ \@{term g} is periodic with period 1.\ interpret g: periodic_fun_simple' g proof fix z :: complex show "g (z + 1) = g z" proof (cases "z \ \") case False hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def) also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints by (subst Beta_plus1_left [symmetric]) auto also have "\ * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))" using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi) also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def) finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto qed (simp add: g_def) qed \ \@{term g} is entire.\ have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex proof (cases "z \ \") let ?h' = "\z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) + of_real pi * cos (z * of_real pi))" case False from False have "eventually (\t. t \ UNIV - \) (nhds z)" by (intro eventually_nhds_in_open) (auto simp: open_Diff) hence "eventually (\t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def) moreover { from False Ints_diff[of 1 "1-z"] have "1 - z \ \" by auto hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def) also from False have "sin (of_real pi * z) \ 0" by (subst sin_eq_0) auto hence "?h' z = h z * g z" using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def) finally have "(?g has_field_derivative (h z * g z)) (at z)" . } ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) next case True then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases) let ?t = "(\z::complex. if z = 0 then 1 else sin z / z) \ (\z. of_real pi * z)" have deriv_0: "(g has_field_derivative 0) (at 0)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "eventually (\z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)" using eventually_nhds_ball[OF zero_less_one, of "0::complex"] proof eventually_elim fix z :: complex assume z: "z \ ball 0 1" show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z" proof (cases "z = 0") assume z': "z \ 0" with z have z'': "z \ \\<^sub>\\<^sub>0" "z \ \" by (auto elim!: Ints_cases) from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp with z'' z' show ?thesis by (simp add: g_def ac_simps) qed (simp add: g_def) qed have "(?t has_field_derivative (0 * of_real pi)) (at 0)" using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] by (intro DERIV_chain) simp_all thus "((\z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)" by (auto intro!: derivative_eq_intros simp: o_def) qed have "((g \ (\x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))" using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros) also have "g \ (\x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int) finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def) qed have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z proof (cases "z \ \") case True with that have "z = 0 \ z = 1" by (force elim!: Ints_cases) moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square) moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square Beta_def algebra_simps) ultimately show ?thesis by force next case False hence z: "z/2 \ \" "(z+1)/2 \ \" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases) hence z': "z/2 \ \\<^sub>\\<^sub>0" "(z+1)/2 \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) from z have "1-z/2 \ \" "1-((z+1)/2) \ \" using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto hence z'': "1-z/2 \ \\<^sub>\\<^sub>0" "1-((z+1)/2) \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) from z have "g (z/2) * g ((z+1)/2) = (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) * (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))" by (simp add: g_def) also from z' Gamma_legendre_duplication_aux[of "z/2"] have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z" by (simp add: add_divide_distrib) also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"] have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))" by (simp add: add_divide_distrib ac_simps) finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) * (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))" by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real) also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)" using cos_sin_eq[of "- of_real pi * z/2", symmetric] by (simp add: ring_distribs add_divide_distrib ac_simps) also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)" by (subst sin_times_cos) (simp add: field_simps) also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z" using \z \ \\ by (simp add: g_def) finally show ?thesis . qed have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z proof - define r where "r = \Re z / 2\" have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) also have "of_int (2*r) = 2 * of_int r" by simp also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" unfolding r_def by (intro g_eq[symmetric]) simp_all also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp also have "g \ = g (z/2)" by (rule g.minus_of_int) also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp also have "g \ = g ((z+1)/2)" by (rule g.minus_of_int) finally show ?thesis .. qed have g_nz [simp]: "g z \ 0" for z :: complex unfolding g_def using Ints_diff[of 1 "1 - z"] by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int) have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z proof - have "((\t. g (t/2) * g ((t+1)/2)) has_field_derivative (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps) hence "((\t. Gamma (1/2)^2 * g t) has_field_derivative Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" by (subst (1 2) g_eq[symmetric]) simp from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"] have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints) moreover have "(g has_field_derivative (g z * h z)) (at z)" using g_g'[of z] by (simp add: ac_simps) ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)" by (intro DERIV_unique) thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp qed obtain h' where h'_cont: "continuous_on UNIV h'" and h_h': "\z. (h has_field_derivative h' z) (at z)" unfolding h_def by (erule Gamma_reflection_aux) have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z proof - have "((\t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2]) hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)" by (subst (asm) h_eq[symmetric]) from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique) qed have h'_zero: "h' z = 0" for z proof - define m where "m = max 1 \Re z\" define B where "B = {t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" have "closed ({t. Re t \ -m} \ {t. Re t \ m} \ {t. Im t \ -\Im z\} \ {t. Im t \ \Im z\})" (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le closed_halfspace_Im_ge closed_halfspace_Im_le) also have "?B = B" unfolding B_def by fastforce finally have "closed B" . moreover have "bounded B" unfolding bounded_iff proof (intro ballI exI) fix t assume t: "t \ B" have "norm t \ \Re t\ + \Im t\" by (rule cmod_le) also from t have "\Re t\ \ m" unfolding B_def by blast also from t have "\Im t\ \ \Im z\" unfolding B_def by blast finally show "norm t \ m + \Im z\" by - simp qed ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast define M where "M = (SUP z\B. norm (h' z))" have "compact (h' ` B)" by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ hence bdd: "bdd_above ((\z. norm (h' z)) ` B)" using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded) have "norm (h' z) \ M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def) also have "M \ M/2" proof (subst M_def, subst cSUP_le_iff) have "z \ B" unfolding B_def m_def by simp thus "B \ {}" by auto next show "\z\B. norm (h' z) \ M/2" proof fix t :: complex assume t: "t \ B" from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp) also have "norm \ = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp also have "norm (h' (t/2) + h' ((t+1)/2)) \ norm (h' (t/2)) + norm (h' ((t+1)/2))" by (rule norm_triangle_ineq) also from t have "abs (Re ((t + 1)/2)) \ m" unfolding m_def B_def by auto with t have "t/2 \ B" "(t+1)/2 \ B" unfolding B_def by auto hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \ M + M" unfolding M_def by (intro add_mono cSUP_upper bdd) (auto simp: B_def) also have "(M + M) / 4 = M / 2" by simp finally show "norm (h' t) \ M/2" by - simp_all qed qed (insert bdd, auto) hence "M \ 0" by simp finally show "h' z = 0" by simp qed have h_h'_2: "(h has_field_derivative 0) (at z)" for z using h_h'[of z] h'_zero[of z] by simp have g_real: "g z \ \" if "z \ \" for z unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real) have h_real: "h z \ \" if "z \ \" for z unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real) have g_nz: "g z \ 0" for z unfolding g_def using Ints_diff[of 1 "1-z"] by (auto simp: Gamma_eq_zero_iff sin_eq_0) from h'_zero h_h'_2 have "\c. \z\UNIV. h z = c" by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm) then obtain c where c: "\z. h z = c" by auto have "\u. u \ closed_segment 0 1 \ Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))" by (intro complex_mvt_line g_g') then obtain u where u: "u \ closed_segment 0 1" "Re (g 1) - Re (g 0) = Re (h u * g u)" by auto from u(1) have u': "u \ \" unfolding closed_segment_def by (auto simp: scaleR_conv_of_real) from u' g_real[of u] g_nz[of u] have "Re (g u) \ 0" by (auto elim!: Reals_cases) with u(2) c[of u] g_real[of u] g_nz[of u] u' have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1) with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases) with c have A: "h z * g z = 0" for z by simp hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp hence "\c'. \z\UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all then obtain c' where c: "\z. g z = c'" by (force) from this[of 0] have "c' = pi" unfolding g_def by simp with c have "g z = pi" by simp show ?thesis proof (cases "z \ \") case False with \g z = pi\ show ?thesis by (auto simp: g_def divide_simps) next case True then obtain n where n: "z = of_int n" by (elim Ints_cases) with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force moreover have "of_int (1 - n) \ \\<^sub>\\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp ultimately show ?thesis using n by (cases "n \ 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int) qed qed lemma rGamma_reflection_complex: "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi" using Gamma_reflection_complex[of z] by (simp add: Gamma_def field_split_simps split: if_split_asm) lemma rGamma_reflection_complex': "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi" proof - have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))" using rGamma_plus1[of "-z", symmetric] by simp also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi" by (rule rGamma_reflection_complex) finally show ?thesis by simp qed lemma Gamma_reflection_complex': "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))" using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def field_split_simps) lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi" proof - from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1] have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square) hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp also have "\ = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all moreover have "Gamma (1/2 :: real) \ 0" using Gamma_real_pos[of "1/2"] by simp ultimately show ?thesis by (rule real_sqrt_unique [symmetric]) qed lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)" proof - have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp also have "\ = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real) finally show ?thesis . qed theorem Gamma_legendre_duplication: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)" using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex) end subsection\<^marker>\tag unimportant\ \Limits and residues\ text \ The inverse of the Gamma function has simple zeros: \ lemma rGamma_zeros: "(\z. rGamma z / (z + of_nat n)) \ (- of_nat n) \ ((-1)^n * fact n :: 'a :: Gamma)" proof (subst tendsto_cong) let ?f = "\z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a" from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] show "eventually (\z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" by (subst pochhammer_rGamma[of _ "Suc n"]) (auto elim!: eventually_mono simp: field_split_simps pochhammer_rec' eq_neg_iff_add_eq_0) have "isCont ?f (- of_nat n)" by (intro continuous_intros) thus "?f \ (- of_nat n) \ (- 1) ^ n * fact n" unfolding isCont_def by (simp add: pochhammer_same) qed text \ The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function, and their residues can easily be computed from the limit we have just proven: \ lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))" proof - from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] have "eventually (\z. rGamma z \ (0 :: 'a)) (at (- of_nat n))" by (auto elim!: eventually_mono nonpos_Ints_cases' simp: rGamma_eq_zero_iff dist_of_nat dist_minus) with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] have "filterlim (\z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))" unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) (simp_all add: filterlim_at) moreover have "(\z. inverse (rGamma z) :: 'a) = Gamma" by (intro ext) (simp add: rGamma_inverse_Gamma) ultimately show ?thesis by (simp only: ) qed lemma Gamma_residues: "(\z. Gamma z * (z + of_nat n)) \ (- of_nat n) \ ((-1)^n / fact n :: 'a :: Gamma)" proof (subst tendsto_cong) let ?c = "(- 1) ^ n / fact n :: 'a" from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] show "eventually (\z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) (at (- of_nat n))" by (auto elim!: eventually_mono simp: field_split_simps rGamma_inverse_Gamma) have "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ inverse ((- 1) ^ n * fact n :: 'a)" by (intro tendsto_intros rGamma_zeros) simp_all also have "inverse ((- 1) ^ n * fact n) = ?c" by (simp_all add: field_simps flip: power_mult_distrib) finally show "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ ?c" . qed subsection \Alternative definitions\ subsubsection \Variant of the Euler form\ definition Gamma_series_euler' where "Gamma_series_euler' z n = inverse z * (\k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))" context begin private lemma Gamma_euler'_aux1: fixes z :: "'a :: {real_normed_field,banach}" assumes n: "n > 0" shows "exp (z * of_real (ln (of_nat n + 1))) = (\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))" proof - have "(\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = exp (z * of_real (\k = 1..n. ln (1 + 1 / real_of_nat k)))" by (subst exp_sum [symmetric]) (simp_all add: sum_distrib_left) also have "(\k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\k=1..n. 1 + 1 / real_of_nat k)" by (subst ln_prod [symmetric]) (auto intro!: add_pos_nonneg) also have "(\k=1..n. 1 + 1 / of_nat k :: real) = (\k=1..n. (of_nat k + 1) / of_nat k)" by (intro prod.cong) (simp_all add: field_split_simps) also have "(\k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1" by (induction n) (simp_all add: prod.nat_ivl_Suc' field_split_simps) finally show ?thesis .. qed theorem Gamma_series_euler': assumes z: "(z :: 'a :: Gamma) \ \\<^sub>\\<^sub>0" shows "(\n. Gamma_series_euler' z n) \ Gamma z" proof (rule Gamma_seriesI, rule Lim_transform_eventually) let ?f = "\n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)" let ?r = "\n. ?f n / Gamma_series z n" let ?r' = "\n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))" from z have z': "z \ 0" by auto have "eventually (\n. ?r' n = ?r n) sequentially" using z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div intro: eventually_mono eventually_gt_at_top[of "0::nat"] dest: pochhammer_eq_0_imp_nonpos_Int) moreover have "?r' \ exp (z * of_real (ln 1))" by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all ultimately show "?r \ 1" by (force intro: Lim_transform_eventually) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" from n z' have "Gamma_series_euler' z n = exp (z * of_real (ln (of_nat n + 1))) / (z * (\k=1..n. (1 + z / of_nat k)))" by (subst Gamma_euler'_aux1) (simp_all add: Gamma_series_euler'_def prod.distrib prod_inversef[symmetric] divide_inverse) also have "(\k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n" proof (cases n) case (Suc n') then show ?thesis unfolding pochhammer_prod fact_prod by (simp add: atLeastLessThanSuc_atLeastAtMost field_simps prod_dividef prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) qed auto also have "z * \ = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec) finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp qed qed end subsubsection \Weierstrass form\ definition Gamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "Gamma_series_Weierstrass z n = exp (-euler_mascheroni * z) / z * (\k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" definition\<^marker>\tag unimportant\ rGamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "rGamma_series_Weierstrass z n = exp (euler_mascheroni * z) * z * (\k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" lemma Gamma_series_Weierstrass_nonpos_Ints: "eventually (\k. Gamma_series_Weierstrass (- of_nat n) k = 0) sequentially" using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_Weierstrass_def) lemma rGamma_series_Weierstrass_nonpos_Ints: "eventually (\k. rGamma_series_Weierstrass (- of_nat n) k = 0) sequentially" using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_Weierstrass_def) theorem Gamma_Weierstrass_complex: "Gamma_series_Weierstrass z \ Gamma (z :: complex)" proof (cases "z \ \\<^sub>\\<^sub>0") case True then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') also from True have "Gamma_series_Weierstrass \ \ Gamma z" by (simp add: tendsto_cong[OF Gamma_series_Weierstrass_nonpos_Ints] Gamma_nonpos_Int) finally show ?thesis . next case False hence z: "z \ 0" by auto let ?f = "(\x. \x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))" have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \ 1" for n :: nat using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp) have "(\n. \k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \ ln_Gamma z + euler_mascheroni * z + ln z" using ln_Gamma_series'_aux[OF False] by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def sum.shift_bounds_Suc_ivl sums_def atLeast0LessThan) from tendsto_exp[OF this] False z have "?f \ z * exp (euler_mascheroni * z) * Gamma z" by (simp add: exp_add exp_sum exp_diff mult_ac Gamma_complex_altdef A) from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z show "Gamma_series_Weierstrass z \ Gamma z" by (simp add: exp_minus field_split_simps Gamma_series_Weierstrass_def [abs_def]) qed lemma tendsto_complex_of_real_iff: "((\x. complex_of_real (f x)) \ of_real c) F = (f \ c) F" by (rule tendsto_of_real_iff) lemma Gamma_Weierstrass_real: "Gamma_series_Weierstrass x \ Gamma (x :: real)" using Gamma_Weierstrass_complex[of "of_real x"] unfolding Gamma_series_Weierstrass_def[abs_def] by (subst tendsto_complex_of_real_iff [symmetric]) (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real) lemma rGamma_Weierstrass_complex: "rGamma_series_Weierstrass z \ rGamma (z :: complex)" proof (cases "z \ \\<^sub>\\<^sub>0") case True then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') also from True have "rGamma_series_Weierstrass \ \ rGamma z" by (simp add: tendsto_cong[OF rGamma_series_Weierstrass_nonpos_Ints] rGamma_nonpos_Int) finally show ?thesis . next case False have "rGamma_series_Weierstrass z = (\n. inverse (Gamma_series_Weierstrass z n))" by (simp add: rGamma_series_Weierstrass_def[abs_def] Gamma_series_Weierstrass_def exp_minus divide_inverse prod_inversef[symmetric] mult_ac) also from False have "\ \ inverse (Gamma z)" by (intro tendsto_intros Gamma_Weierstrass_complex) (simp add: Gamma_eq_zero_iff) finally show ?thesis by (simp add: Gamma_def) qed subsubsection \Binomial coefficient form\ lemma Gamma_gbinomial: "(\n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \ rGamma (z+1)" proof (cases "z = 0") case False show ?thesis proof (rule Lim_transform_eventually) let ?powr = "\a b. exp (b * of_real (ln (of_nat a)))" show "eventually (\n. rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially" proof (intro always_eventually allI) fix n :: nat from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n" by (simp add: gbinomial_pochhammer' pochhammer_rec) also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z" by (simp add: rGamma_series_def field_split_simps exp_minus) finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" .. qed from False have "(\n. rGamma_series z n / z) \ rGamma z / z" by (intro tendsto_intros) also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] by (simp add: field_simps) finally show "(\n. rGamma_series z n / z) \ rGamma (z+1)" . qed qed (simp_all add: binomial_gbinomial [symmetric]) lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)" by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric]) lemma gbinomial_asymptotic: fixes z :: "'a :: Gamma" shows "(\n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \ inverse (Gamma (- z))" unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] by (subst (asm) gbinomial_minus') (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric]) lemma fact_binomial_limit: "(\n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \ 1 / fact k" proof (rule Lim_transform_eventually) have "(\n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n)))) \ 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \ _") using Gamma_gbinomial[of "of_nat k :: 'a"] by (simp add: binomial_gbinomial Gamma_def field_split_simps exp_of_real [symmetric] exp_minus) also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact) finally show "?f \ 1 / fact k" . show "eventually (\n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)" by (simp add: exp_of_nat_mult) thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp qed qed lemma binomial_asymptotic': "(\n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \ 1" using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp lemma gbinomial_Beta: assumes "z + 1 \ \\<^sub>\\<^sub>0" shows "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))" using assms proof (induction n arbitrary: z) case 0 hence "z + 2 \ \\<^sub>\\<^sub>0" using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute) with 0 show ?case by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute) next case (Suc n z) show ?case proof (cases "z \ \\<^sub>\\<^sub>0") case True with Suc.prems have "z = 0" by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff) show ?thesis proof (cases "n = 0") case True with \z = 0\ show ?thesis by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric]) next case False with \z = 0\ show ?thesis by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff) qed next case False have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp also have "\ = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)" by (subst gbinomial_factors) (simp add: field_simps) also from False have "\ = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1) also have "of_nat (Suc n) \ (\\<^sub>\\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)" by (subst Beta_plus1_right [symmetric]) simp_all finally show ?thesis . qed qed theorem gbinomial_Gamma: assumes "z + 1 \ \\<^sub>\\<^sub>0" shows "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))" proof - have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))" by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac) also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)" using Gamma_plus1[of "z+1"] by (auto simp add: field_split_simps) finally show ?thesis . qed subsubsection \Integral form\ lemma integrable_on_powr_from_0': assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0<..c}" proof - from c have *: "{0<..c} - {0..c} = {}" "{0..c} - {0<..c} = {0}" by auto show ?thesis by (rule integrable_spike_set [OF integrable_on_powr_from_0[OF a c]]) (simp_all add: *) qed lemma absolutely_integrable_Gamma_integral: assumes "Re z > 0" "a > 0" shows "(\t. complex_of_real t powr (z - 1) / of_real (exp (a * t))) absolutely_integrable_on {0<..}" (is "?f absolutely_integrable_on _") proof - have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" by (intro tendsto_intros ln_x_over_x_tendsto_0) hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp from order_tendstoD(2)[OF this, of "a/2"] and \a > 0\ have "eventually (\x. (Re z - 1) * ln x / x < a/2) at_top" by simp from eventually_conj[OF this eventually_gt_at_top[of 0]] obtain x0 where "\x\x0. (Re z - 1) * ln x / x < a/2 \ x > 0" by (auto simp: eventually_at_top_linorder) hence "x0 > 0" by simp have "x powr (Re z - 1) / exp (a * x) < exp (-(a/2) * x)" if "x \ x0" for x proof - from that and \\x\x0. _\ have x: "(Re z - 1) * ln x / x < a / 2" "x > 0" by auto have "x powr (Re z - 1) = exp ((Re z - 1) * ln x)" using \x > 0\ by (simp add: powr_def) also from x have "(Re z - 1) * ln x < (a * x) / 2" by (simp add: field_simps) finally show ?thesis by (simp add: field_simps exp_add [symmetric]) qed note x0 = \x0 > 0\ this have "?f absolutely_integrable_on ({0<..x0} \ {x0..})" proof (rule set_integrable_Un) show "?f absolutely_integrable_on {0<..x0}" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2]) show "integrable lebesgue (\x. indicat_real {0<..x0} x *\<^sub>R x powr (Re z - 1))" using x0(1) assms by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_powr_from_0') auto show "(\x. indicat_real {0<..x0} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \ borel_measurable lebesgue" by (intro measurable_completion) (auto intro!: borel_measurable_continuous_on_indicator continuous_intros) fix x :: real have "x powr (Re z - 1) / exp (a * x) \ x powr (Re z - 1) / 1" if "x \ 0" using that assms by (intro divide_left_mono) auto thus "norm (indicator {0<..x0} x *\<^sub>R ?f x) \ norm (indicator {0<..x0} x *\<^sub>R x powr (Re z - 1))" by (simp_all add: norm_divide norm_powr_real_powr indicator_def) qed next show "?f absolutely_integrable_on {x0..}" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2]) show "integrable lebesgue (\x. indicat_real {x0..} x *\<^sub>R exp (- (a / 2) * x))" using assms by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_exp_minus_to_infinity) auto show "(\x. indicat_real {x0..} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \ borel_measurable lebesgue" using x0(1) by (intro measurable_completion) (auto intro!: borel_measurable_continuous_on_indicator continuous_intros) fix x :: real show "norm (indicator {x0..} x *\<^sub>R ?f x) \ norm (indicator {x0..} x *\<^sub>R exp (-(a/2) * x))" using x0 by (auto simp: norm_divide norm_powr_real_powr indicator_def less_imp_le) qed qed auto also have "{0<..x0} \ {x0..} = {0<..}" using x0(1) by auto finally show ?thesis . qed lemma integrable_Gamma_integral_bound: fixes a c :: real assumes a: "a > -1" and c: "c \ 0" defines "f \ \x. if x \ {0..c} then x powr a else exp (-x/2)" shows "f integrable_on {0..}" proof - have "f integrable_on {0..c}" by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]]) (insert a c, simp_all add: f_def) moreover have A: "(\x. exp (-x/2)) integrable_on {c..}" using integrable_on_exp_minus_to_infinity[of "1/2"] by simp have "f integrable_on {c..}" by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def) ultimately show "f integrable_on {0..}" by (rule integrable_Un') (insert c, auto simp: max_def) qed theorem Gamma_integral_complex: assumes z: "Re z > 0" shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" proof - have A: "((\t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n)) has_integral (fact n / pochhammer z (n+1))) {0..1}" if "Re z > 0" for n z using that proof (induction n arbitrary: z) case 0 have "((\t. complex_of_real t powr (z - 1)) has_integral (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0 by (intro fundamental_theorem_of_calculus_interior) (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_field) thus ?case by simp next case (Suc n) let ?f = "\t. complex_of_real t powr z / z" let ?f' = "\t. complex_of_real t powr (z - 1)" let ?g = "\t. (1 - complex_of_real t) ^ Suc n" let ?g' = "\t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)" have "((\t. ?f' t * ?g t) has_integral (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}" (is "(_ has_integral ?I) _") proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g]) from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g" by (auto intro!: continuous_intros) next fix t :: real assume t: "t \ {0<..<1}" show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems by (auto intro!: derivative_eq_intros has_vector_derivative_real_field) show "(?g has_vector_derivative ?g' t) (at t)" by (rule has_vector_derivative_real_field derivative_eq_intros refl)+ simp_all next from Suc.prems have [simp]: "z \ 0" by auto from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp have [simp]: "z + of_nat n \ 0" "z + 1 + of_nat n \ 0" for n using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel) have "((\x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}" (is "(?A has_integral ?B) _") using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec) also have "?A = (\t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps) also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))" by (simp add: field_split_simps pochhammer_rec prod.shift_bounds_cl_Suc_ivl del: of_nat_Suc) finally show "((\t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}" by simp qed (simp_all add: bounded_bilinear_mult) thus ?case by simp qed have B: "((\t. if t \ {0..of_nat n} then of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0) has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n proof (cases "n > 0") case [simp]: True hence [simp]: "n \ 0" by auto with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0] have "((\x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n) has_integral fact n * of_nat n / pochhammer z (n+1)) ((\x. real n * x)`{0..1})" (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real) also from True have "((\x. real n*x)`{0..1}) = {0..real n}" by (subst image_mult_atLeastAtMost) simp_all also have "?f = (\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)" using True by (intro ext) (simp add: field_simps) finally have "((\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" (is ?P) . also have "?P \ ((\x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric]) also have "\ \ ((\x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div) finally have \ . note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"] have "((\x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n) has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P) by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric]) (simp add: algebra_simps) also have "?P \ ((\x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n) has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real) also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) = (of_nat n powr z * fact n / pochhammer z (n+1))" by (auto simp add: powr_def algebra_simps exp_diff exp_of_real) finally show ?thesis by (subst has_integral_restrict) simp_all next case False thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl) qed have "eventually (\n. Gamma_series z n = of_nat n powr z * fact n / pochhammer z (n+1)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: powr_def algebra_simps Gamma_series_def) from this and Gamma_series_LIMSEQ[of z] have C: "(\k. of_nat k powr z * fact k / pochhammer z (k+1)) \ Gamma z" by (blast intro: Lim_transform_eventually) { fix x :: real assume x: "x \ 0" have lim_exp: "(\k. (1 - x / real k) ^ k) \ exp (-x)" using tendsto_exp_limit_sequentially[of "-x"] by simp have "(\k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k)) \ of_real x powr (z - 1) * of_real (exp (-x))" (is ?P) by (intro tendsto_intros lim_exp) also from eventually_gt_at_top[of "nat \x\"] have "eventually (\k. of_nat k > x) sequentially" by eventually_elim linarith hence "?P \ (\k. if x \ of_nat k then of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0) \ of_real x powr (z - 1) * of_real (exp (-x))" by (intro tendsto_cong) (auto elim!: eventually_mono) finally have \ . } hence D: "\x\{0..}. (\k. if x \ {0..real k} then of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0) \ of_real x powr (z - 1) / of_real (exp x)" by (simp add: exp_minus field_simps cong: if_cong) have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" by (intro tendsto_intros ln_x_over_x_tendsto_0) hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp from order_tendstoD(2)[OF this, of "1/2"] have "eventually (\x. (Re z - 1) * ln x / x < 1/2) at_top" by simp from eventually_conj[OF this eventually_gt_at_top[of 0]] obtain x0 where "\x\x0. (Re z - 1) * ln x / x < 1/2 \ x > 0" by (auto simp: eventually_at_top_linorder) hence x0: "x0 > 0" "\x. x \ x0 \ (Re z - 1) * ln x < x / 2" by auto define h where "h = (\x. if x \ {0..x0} then x powr (Re z - 1) else exp (-x/2))" have le_h: "x powr (Re z - 1) * exp (-x) \ h x" if x: "x \ 0" for x proof (cases "x > x0") case True from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)" by (simp add: powr_def exp_diff exp_minus field_simps exp_add) also from x0(2)[of x] True have "\ < exp (-x/2)" by (simp add: field_simps) finally show ?thesis using True by (auto simp add: h_def) next case False from x have "x powr (Re z - 1) * exp (- x) \ x powr (Re z - 1) * 1" by (intro mult_left_mono) simp_all with False show ?thesis by (auto simp add: h_def) qed have E: "\x\{0..}. cmod (if x \ {0..real k} then of_real x powr (z - 1) * (1 - complex_of_real x / of_nat k) ^ k else 0) \ h x" (is "\x\_. ?f x \ _") for k proof safe fix x :: real assume x: "x \ 0" { fix x :: real and n :: nat assume x: "x \ of_nat n" have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp also have "norm \ = \(1 - x / real n)\" by (subst norm_of_real) (rule refl) also from x have "\ = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: field_split_simps) finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" . } note D = this from D[of x k] x have "?f x \ (if of_nat k \ x \ k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)" by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg) also have "\ \ x powr (Re z - 1) * exp (-x)" by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n) also from x have "\ \ h x" by (rule le_h) finally show "?f x \ h x" . qed have F: "h integrable_on {0..}" unfolding h_def by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all) show ?thesis by (rule has_integral_dominated_convergence[OF B F E D C]) qed lemma Gamma_integral_real: assumes x: "x > (0 :: real)" shows "((\t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}" proof - have A: "((\t. complex_of_real t powr (complex_of_real x - 1) / complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}" using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real) have "((\t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}" by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric]) from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def) qed lemma absolutely_integrable_Gamma_integral': assumes "Re z > 0" shows "(\t. complex_of_real t powr (z - 1) / of_real (exp t)) absolutely_integrable_on {0<..}" using absolutely_integrable_Gamma_integral [OF assms zero_less_one] by simp lemma Gamma_integral_complex': assumes z: "Re z > 0" shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}" proof - have "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" by (rule Gamma_integral_complex) fact+ hence "((\t. if t \ {0<..} then of_real t powr (z - 1) / of_real (exp t) else 0) has_integral Gamma z) {0..}" by (rule has_integral_spike [of "{0}", rotated 2]) auto also have "?this = ?thesis" by (subst has_integral_restrict) auto finally show ?thesis . qed lemma Gamma_conv_nn_integral_real: assumes "s > (0::real)" shows "Gamma s = nn_integral lborel (\t. ennreal (indicator {0..} t * t powr (s - 1) / exp t))" using nn_integral_has_integral_lebesgue[OF _ Gamma_integral_real[OF assms]] by simp lemma integrable_Beta: assumes "a > 0" "b > (0::real)" shows "set_integrable lborel {0..1} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" proof - define C where "C = max 1 ((1/2) powr (b - 1))" define D where "D = max 1 ((1/2) powr (a - 1))" have C: "(1 - x) powr (b - 1) \ C" if "x \ {0..1/2}" for x proof (cases "b < 1") case False with that have "(1 - x) powr (b - 1) \ (1 powr (b - 1))" by (intro powr_mono2) auto thus ?thesis by (auto simp: C_def) qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 C_def) have D: "x powr (a - 1) \ D" if "x \ {1/2..1}" for x proof (cases "a < 1") case False with that have "x powr (a - 1) \ (1 powr (a - 1))" by (intro powr_mono2) auto thus ?thesis by (auto simp: D_def) next case True qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 D_def) have [simp]: "C \ 0" "D \ 0" by (simp_all add: C_def D_def) have I1: "set_integrable lborel {0..1/2} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "(\t. t powr (a - 1)) integrable_on {0..1/2}" by (rule integrable_on_powr_from_0) (use assms in auto) hence "(\t. t powr (a - 1)) absolutely_integrable_on {0..1/2}" by (subst absolutely_integrable_on_iff_nonneg) auto from integrable_mult_right[OF this [unfolded set_integrable_def], of C] show "integrable lborel (\x. indicat_real {0..1/2} x *\<^sub>R (C * x powr (a - 1)))" by (subst (asm) integrable_completion) (auto simp: mult_ac) next fix x :: real have "x powr (a - 1) * (1 - x) powr (b - 1) \ x powr (a - 1) * C" if "x \ {0..1/2}" using that by (intro mult_left_mono powr_mono2 C) auto thus "norm (indicator {0..1/2} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \ norm (indicator {0..1/2} x *\<^sub>R (C * x powr (a - 1)))" by (auto simp: indicator_def abs_mult mult_ac) qed (auto intro!: AE_I2 simp: indicator_def) have I2: "set_integrable lborel {1/2..1} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "(\t. t powr (b - 1)) integrable_on {0..1/2}" by (rule integrable_on_powr_from_0) (use assms in auto) hence "(\t. t powr (b - 1)) integrable_on (cbox 0 (1/2))" by simp from integrable_affinity[OF this, of "-1" 1] have "(\t. (1 - t) powr (b - 1)) integrable_on {1/2..1}" by simp hence "(\t. (1 - t) powr (b - 1)) absolutely_integrable_on {1/2..1}" by (subst absolutely_integrable_on_iff_nonneg) auto from integrable_mult_right[OF this [unfolded set_integrable_def], of D] show "integrable lborel (\x. indicat_real {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))" by (subst (asm) integrable_completion) (auto simp: mult_ac) next fix x :: real have "x powr (a - 1) * (1 - x) powr (b - 1) \ D * (1 - x) powr (b - 1)" if "x \ {1/2..1}" using that by (intro mult_right_mono powr_mono2 D) auto thus "norm (indicator {1/2..1} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \ norm (indicator {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))" by (auto simp: indicator_def abs_mult mult_ac) qed (auto intro!: AE_I2 simp: indicator_def) have "set_integrable lborel ({0..1/2} \ {1/2..1}) (\t. t powr (a - 1) * (1 - t) powr (b - 1))" by (intro set_integrable_Un I1 I2) auto also have "{0..1/2} \ {1/2..1} = {0..(1::real)}" by auto finally show ?thesis . qed lemma integrable_Beta': assumes "a > 0" "b > (0::real)" shows "(\t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}" using integrable_Beta[OF assms] by (rule set_borel_integral_eq_integral) theorem has_integral_Beta_real: assumes a: "a > 0" and b: "b > (0 :: real)" shows "((\t. t powr (a - 1) * (1 - t) powr (b - 1)) has_integral Beta a b) {0..1}" proof - define B where "B = integral {0..1} (\x. x powr (a - 1) * (1 - x) powr (b - 1))" have [simp]: "B \ 0" unfolding B_def using a b by (intro integral_nonneg integrable_Beta') auto from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator ({0..}\{0..}) (t,u) * t powr (a - 1) * u powr (b - 1) / exp (t + u)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel \lborel)" proof (rule nn_integral_cong, goal_cases) case (1 t) have "(\\<^sup>+u. ennreal (indicator ({0..}\{0..}) (t,u) * t powr (a - 1) * u powr (b - 1) / exp (t + u)) \distr lborel borel ((+) (-t))) = (\\<^sup>+u. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel)" by (subst nn_integral_distr) (auto intro!: nn_integral_cong simp: indicator_def) thus ?case by (subst (asm) lborel_distr_plus) qed also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel \lborel)" by (subst lborel_pair.Fubini') (auto simp: case_prod_unfold indicator_def cong: measurable_cong_sets) also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) * ennreal (indicator {0..} u / exp u) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def ennreal_mult' [symmetric]) also have "\ = (\\<^sup>+u. (\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) * ennreal (indicator {0..} u / exp u) \lborel)" by (subst nn_integral_multc [symmetric]) auto also have "\ = (\\<^sup>+u. (\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) * ennreal (indicator {0<..} u / exp u) \lborel)" by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]]) (auto simp: indicator_def) also have "\ = (\\<^sup>+u. ennreal B * ennreal (indicator {0..} u / exp u * u powr (a + b - 1)) \lborel)" proof (intro nn_integral_cong, goal_cases) case (1 u) show ?case proof (cases "u > 0") case True have "(\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) = (\\<^sup>+t. ennreal (indicator {0..1} t * (u * t) powr (a - 1) * (u - u * t) powr (b - 1)) \distr lborel borel ((*) (1 / u)))" (is "_ = nn_integral _ ?f") using True by (subst nn_integral_distr) (auto simp: indicator_def field_simps intro!: nn_integral_cong) also have "distr lborel borel ((*) (1 / u)) = density lborel (\_. u)" using \u > 0\ by (subst lborel_distr_mult) auto also have "nn_integral \ ?f = (\\<^sup>+x. ennreal (indicator {0..1} x * (u * (u * x) powr (a - 1) * (u * (1 - x)) powr (b - 1))) \lborel)" using \u > 0\ by (subst nn_integral_density) (auto simp: ennreal_mult' [symmetric] algebra_simps) also have "\ = (\\<^sup>+x. ennreal (u powr (a + b - 1)) * ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel)" using \u > 0\ a b by (intro nn_integral_cong) (auto simp: indicator_def powr_mult powr_add powr_diff mult_ac ennreal_mult' [symmetric]) also have "\ = ennreal (u powr (a + b - 1)) * (\\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel)" by (subst nn_integral_cmult) auto also have "((\x. x powr (a - 1) * (1 - x) powr (b - 1)) has_integral integral {0..1} (\x. x powr (a - 1) * (1 - x) powr (b - 1))) {0..1}" using a b by (intro integrable_integral integrable_Beta') from nn_integral_has_integral_lebesgue[OF _ this] a b have "(\\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel) = B" by (simp add: mult_ac B_def) finally show ?thesis using \u > 0\ by (simp add: ennreal_mult' [symmetric] mult_ac) qed auto qed also have "\ = ennreal B * ennreal (Gamma (a + b))" using a b by (subst nn_integral_cmult) (auto simp: Gamma_conv_nn_integral_real) also have "\ = ennreal (B * Gamma (a + b))" by (subst (1 2) mult.commute, intro ennreal_mult' [symmetric]) (use a b in auto) finally have "B = Beta a b" using a b Gamma_real_pos[of "a + b"] by (subst (asm) ennreal_inj) (auto simp: field_simps Beta_def Gamma_eq_zero_iff) moreover have "(\t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}" by (intro integrable_Beta' a b) ultimately show ?thesis by (simp add: has_integral_iff B_def) qed subsection \The Weierstra{\ss} product formula for the sine\ theorem sin_product_formula_complex: fixes z :: complex shows "(\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k^2)) \ sin (of_real pi * z)" proof - let ?f = "rGamma_series_Weierstrass" have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n)) \ (- of_real pi * inverse z) * (rGamma z * rGamma (- z))" by (intro tendsto_intros rGamma_Weierstrass_complex) also have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) = (\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k ^ 2))" proof fix n :: nat have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * (\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)" by (simp add: rGamma_series_Weierstrass_def mult_ac exp_minus divide_simps prod.distrib[symmetric] power2_eq_square) also have "(\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) = (\k=1..n. 1 - z^2 / of_nat k ^ 2)" by (intro prod.cong) (simp_all add: power2_eq_square field_simps) finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \" by (simp add: field_split_simps) qed also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)" by (subst rGamma_reflection_complex') (simp add: field_split_simps) finally show ?thesis . qed lemma sin_product_formula_real: "(\n. pi * (x::real) * (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x)" proof - from sin_product_formula_complex[of "of_real x"] have "(\n. of_real pi * of_real x * (\k=1..n. 1 - (of_real x)^2 / (of_nat k)^2)) \ sin (of_real pi * of_real x :: complex)" (is "?f \ ?y") . also have "?f = (\n. of_real (pi * x * (\k=1..n. 1 - x^2 / (of_nat k^2))))" by simp also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult) finally show ?thesis by (subst (asm) tendsto_of_real_iff) qed lemma sin_product_formula_real': assumes "x \ (0::real)" shows "(\n. (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x) / (pi * x)" using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms by simp theorem wallis: "(\n. \k=1..n. (4*real k^2) / (4*real k^2 - 1)) \ pi / 2" proof - from tendsto_inverse[OF tendsto_mult[OF sin_product_formula_real[of "1/2"] tendsto_const[of "2/pi"]]] have "(\n. (\k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) \ pi/2" by (simp add: prod_inversef [symmetric]) also have "(\n. (\k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) = (\n. (\k=1..n. (4*real k^2)/(4*real k^2 - 1)))" by (intro ext prod.cong refl) (simp add: field_split_simps) finally show ?thesis . qed subsection \The Solution to the Basel problem\ theorem inverse_squares_sums: "(\n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" proof - define P where "P x n = (\k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n define K where "K = (\n. inverse (real_of_nat (Suc n))^2)" define f where [abs_def]: "f x = (\n. P x n / of_nat (Suc n)^2)" for x define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x have sums: "(\n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x proof (cases "x = 0") assume x: "x = 0" have "summable (\n. inverse ((real_of_nat (Suc n))\<^sup>2))" using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums) next assume x: "x \ 0" have "(\n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))" unfolding P_def using x by (intro telescope_sums' sin_product_formula_real') also have "(\n. P x n - P x (Suc n)) = (\n. (x^2 / of_nat (Suc n)^2) * P x n)" unfolding P_def by (simp add: prod.nat_ivl_Suc' algebra_simps) also have "P x 0 = 1" by (simp add: P_def) finally have "(\n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" . from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp qed have "continuous_on (ball 0 1) f" proof (rule uniform_limit_theorem; (intro always_eventually allI)?) show "uniform_limit (ball 0 1) (\n x. \k ball 0 1" { fix k :: nat assume k: "k \ 1" from x have "x^2 < 1" by (auto simp: abs_square_less_1) also from k have "\ \ of_nat k^2" by simp finally have "(1 - x^2 / of_nat k^2) \ {0..1}" using k by (simp_all add: field_simps del: of_nat_Suc) } hence "(\k=1..n. abs (1 - x^2 / of_nat k^2)) \ (\k=1..n. 1)" by (intro prod_mono) simp thus "norm (P x n / (of_nat (Suc n)^2)) \ 1 / of_nat (Suc n)^2" unfolding P_def by (simp add: field_simps abs_prod del: of_nat_Suc) qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide) qed (auto simp: P_def intro!: continuous_intros) hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all hence "(f \ 0 \ f 0)" by (simp add: isCont_def) also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide) finally have "f \ 0 \ K" . moreover have "f \ 0 \ pi^2 / 6" proof (rule Lim_transform_eventually) define f' where [abs_def]: "f' x = (\n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x have "eventually (\x. x \ (0::real)) (at 0)" by (auto simp add: eventually_at intro!: exI[of _ 1]) thus "eventually (\x. f' x = f x) (at 0)" proof eventually_elim fix x :: real assume x: "x \ 0" have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def) with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"] have "(\n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))" by (simp add: eval_nat_numeral) from sums_divide[OF this, of "x^3 * pi"] x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)" by (simp add: field_split_simps eval_nat_numeral) with x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)" by (simp add: g_def) hence "f' x = g x / x^2" by (simp add: sums_iff f'_def) also have "\ = f x" using sums[of x] x by (simp add: sums_iff g_def f_def) finally show "f' x = f x" . qed have "isCont f' 0" unfolding f'_def proof (intro isCont_powser_converges_everywhere) fix x :: real show "summable (\n. -sin_coeff (n+3) * pi^(n+2) * x^n)" proof (cases "x = 0") assume x: "x \ 0" from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x show ?thesis by (simp add: field_split_simps eval_nat_numeral) qed (simp only: summable_0_powser) qed hence "f' \ 0 \ f' 0" by (simp add: isCont_def) also have "f' 0 = pi * pi / fact 3" unfolding f'_def by (subst powser_zero) (simp add: sin_coeff_def) finally show "f' \ 0 \ pi^2 / 6" by (simp add: eval_nat_numeral) qed ultimately have "K = pi^2 / 6" by (rule LIM_unique) moreover from inverse_power_summable[of 2] have "summable (\n. (inverse (real_of_nat (Suc n)))\<^sup>2)" by (subst summable_Suc_iff) (simp add: power_inverse) ultimately show ?thesis unfolding K_def by (auto simp add: sums_iff power_divide inverse_eq_divide) qed end diff --git a/src/HOL/Analysis/Infinite_Products.thy b/src/HOL/Analysis/Infinite_Products.thy --- a/src/HOL/Analysis/Infinite_Products.thy +++ b/src/HOL/Analysis/Infinite_Products.thy @@ -1,1811 +1,1811 @@ (*File: HOL/Analysis/Infinite_Product.thy Author: Manuel Eberl & LC Paulson Basic results about convergence and absolute convergence of infinite products and their connection to summability. *) section \Infinite Products\ theory Infinite_Products imports Topology_Euclidean_Space Complex_Transcendental begin subsection\<^marker>\tag unimportant\ \Preliminaries\ lemma sum_le_prod: fixes f :: "'a \ 'b :: linordered_semidom" assumes "\x. x \ A \ f x \ 0" shows "sum f A \ (\x\A. 1 + f x)" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) from insert.hyps have "sum f A + f x * (\x\A. 1) \ (\x\A. 1 + f x) + f x * (\x\A. 1 + f x)" by (intro add_mono insert mult_left_mono prod_mono) (auto intro: insert.prems) with insert.hyps show ?case by (simp add: algebra_simps) qed simp_all lemma prod_le_exp_sum: fixes f :: "'a \ real" assumes "\x. x \ A \ f x \ 0" shows "prod (\x. 1 + f x) A \ exp (sum f A)" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) have "(1 + f x) * (\x\A. 1 + f x) \ exp (f x) * exp (sum f A)" using insert.prems by (intro mult_mono insert prod_nonneg exp_ge_add_one_self) auto with insert.hyps show ?case by (simp add: algebra_simps exp_add) qed simp_all lemma lim_ln_1_plus_x_over_x_at_0: "(\x::real. ln (1 + x) / x) \0\ 1" proof (rule lhopital) show "(\x::real. ln (1 + x)) \0\ 0" by (rule tendsto_eq_intros refl | simp)+ have "eventually (\x::real. x \ {-1/2<..<1/2}) (nhds 0)" by (rule eventually_nhds_in_open) auto hence *: "eventually (\x::real. x \ {-1/2<..<1/2}) (at 0)" by (rule filter_leD [rotated]) (simp_all add: at_within_def) show "eventually (\x::real. ((\x. ln (1 + x)) has_field_derivative inverse (1 + x)) (at x)) (at 0)" using * by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) show "eventually (\x::real. ((\x. x) has_field_derivative 1) (at x)) (at 0)" using * by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) show "\\<^sub>F x in at 0. x \ 0" by (auto simp: at_within_def eventually_inf_principal) show "(\x::real. inverse (1 + x) / 1) \0\ 1" by (rule tendsto_eq_intros refl | simp)+ qed auto subsection\Definitions and basic properties\ definition\<^marker>\tag important\ raw_has_prod :: "[nat \ 'a::{t2_space, comm_semiring_1}, nat, 'a] \ bool" where "raw_has_prod f M p \ (\n. \i\n. f (i+M)) \ p \ p \ 0" text\The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\ text\<^marker>\tag important\ \%whitespace\ definition\<^marker>\tag important\ has_prod :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a \ bool" (infixr "has'_prod" 80) where "f has_prod p \ raw_has_prod f 0 p \ (\i q. p = 0 \ f i = 0 \ raw_has_prod f (Suc i) q)" definition\<^marker>\tag important\ convergent_prod :: "(nat \ 'a :: {t2_space,comm_semiring_1}) \ bool" where "convergent_prod f \ \M p. raw_has_prod f M p" definition\<^marker>\tag important\ prodinf :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a" (binder "\" 10) where "prodinf f = (THE p. f has_prod p)" lemmas prod_defs = raw_has_prod_def has_prod_def convergent_prod_def prodinf_def lemma has_prod_subst[trans]: "f = g \ g has_prod z \ f has_prod z" by simp lemma has_prod_cong: "(\n. f n = g n) \ f has_prod c \ g has_prod c" by presburger lemma raw_has_prod_nonzero [simp]: "\ raw_has_prod f M 0" by (simp add: raw_has_prod_def) lemma raw_has_prod_eq_0: fixes f :: "nat \ 'a::{semidom,t2_space}" assumes p: "raw_has_prod f m p" and i: "f i = 0" "i \ m" shows "p = 0" proof - have eq0: "(\k\n. f (k+m)) = 0" if "i - m \ n" for n proof - have "\k\n. f (k + m) = 0" using i that by auto then show ?thesis by auto qed have "(\n. \i\n. f (i + m)) \ 0" by (rule LIMSEQ_offset [where k = "i-m"]) (simp add: eq0) with p show ?thesis unfolding raw_has_prod_def using LIMSEQ_unique by blast qed lemma raw_has_prod_Suc: "raw_has_prod f (Suc M) a \ raw_has_prod (\n. f (Suc n)) M a" unfolding raw_has_prod_def by auto lemma has_prod_0_iff: "f has_prod 0 \ (\i. f i = 0 \ (\p. raw_has_prod f (Suc i) p))" by (simp add: has_prod_def) lemma has_prod_unique2: fixes f :: "nat \ 'a::{semidom,t2_space}" assumes "f has_prod a" "f has_prod b" shows "a = b" using assms by (auto simp: has_prod_def raw_has_prod_eq_0) (meson raw_has_prod_def sequentially_bot tendsto_unique) lemma has_prod_unique: fixes f :: "nat \ 'a :: {semidom,t2_space}" shows "f has_prod s \ s = prodinf f" by (simp add: has_prod_unique2 prodinf_def the_equality) lemma convergent_prod_altdef: fixes f :: "nat \ 'a :: {t2_space,comm_semiring_1}" shows "convergent_prod f \ (\M L. (\n\M. f n \ 0) \ (\n. \i\n. f (i+M)) \ L \ L \ 0)" proof assume "convergent_prod f" then obtain M L where *: "(\n. \i\n. f (i+M)) \ L" "L \ 0" by (auto simp: prod_defs) have "f i \ 0" if "i \ M" for i proof assume "f i = 0" have **: "eventually (\n. (\i\n. f (i+M)) = 0) sequentially" using eventually_ge_at_top[of "i - M"] proof eventually_elim case (elim n) with \f i = 0\ and \i \ M\ show ?case by (auto intro!: bexI[of _ "i - M"] prod_zero) qed have "(\n. (\i\n. f (i+M))) \ 0" unfolding filterlim_iff by (auto dest!: eventually_nhds_x_imp_x intro!: eventually_mono[OF **]) from tendsto_unique[OF _ this *(1)] and *(2) show False by simp qed with * show "(\M L. (\n\M. f n \ 0) \ (\n. \i\n. f (i+M)) \ L \ L \ 0)" by blast qed (auto simp: prod_defs) subsection\Absolutely convergent products\ definition\<^marker>\tag important\ abs_convergent_prod :: "(nat \ _) \ bool" where "abs_convergent_prod f \ convergent_prod (\i. 1 + norm (f i - 1))" lemma abs_convergent_prodI: assumes "convergent (\n. \i\n. 1 + norm (f i - 1))" shows "abs_convergent_prod f" proof - from assms obtain L where L: "(\n. \i\n. 1 + norm (f i - 1)) \ L" by (auto simp: convergent_def) have "L \ 1" proof (rule tendsto_le) show "eventually (\n. (\i\n. 1 + norm (f i - 1)) \ 1) sequentially" proof (intro always_eventually allI) fix n have "(\i\n. 1 + norm (f i - 1)) \ (\i\n. 1)" by (intro prod_mono) auto thus "(\i\n. 1 + norm (f i - 1)) \ 1" by simp qed qed (use L in simp_all) hence "L \ 0" by auto with L show ?thesis unfolding abs_convergent_prod_def prod_defs by (intro exI[of _ "0::nat"] exI[of _ L]) auto qed lemma fixes f :: "nat \ 'a :: {topological_semigroup_mult,t2_space,idom}" assumes "convergent_prod f" - shows convergent_prod_imp_convergent: "convergent (\n. \i\n. f i)" - and convergent_prod_to_zero_iff: "(\n. \i\n. f i) \ 0 \ (\i. f i = 0)" + shows convergent_prod_imp_convergent: "convergent (\n. \i\n. f i)" + and convergent_prod_to_zero_iff [simp]: "(\n. \i\n. f i) \ 0 \ (\i. f i = 0)" proof - from assms obtain M L where M: "\n. n \ M \ f n \ 0" and "(\n. \i\n. f (i + M)) \ L" and "L \ 0" by (auto simp: convergent_prod_altdef) note this(2) also have "(\n. \i\n. f (i + M)) = (\n. \i=M..M+n. f i)" by (intro ext prod.reindex_bij_witness[of _ "\n. n - M" "\n. n + M"]) auto finally have "(\n. (\ii=M..M+n. f i)) \ (\in. (\ii=M..M+n. f i)) = (\n. (\i\{..{M..M+n}. f i))" by (subst prod.union_disjoint) auto also have "(\n. {.. {M..M+n}) = (\n. {..n+M})" by auto finally have lim: "(\n. prod f {..n}) \ prod f {..n. \i\n. f i)" by (auto simp: convergent_def) show "(\n. \i\n. f i) \ 0 \ (\i. f i = 0)" proof assume "\i. f i = 0" then obtain i where "f i = 0" by auto moreover with M have "i < M" by (cases "i < M") auto ultimately have "(\in. \i\n. f i) \ 0" by simp next assume "(\n. \i\n. f i) \ 0" from tendsto_unique[OF _ this lim] and \L \ 0\ show "\i. f i = 0" by auto qed qed lemma convergent_prod_iff_nz_lim: fixes f :: "nat \ 'a :: {topological_semigroup_mult,t2_space,idom}" assumes "\i. f i \ 0" shows "convergent_prod f \ (\L. (\n. \i\n. f i) \ L \ L \ 0)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs using assms convergentD convergent_prod_imp_convergent convergent_prod_to_zero_iff by blast next assume ?rhs then show ?lhs unfolding prod_defs by (rule_tac x=0 in exI) auto qed lemma\<^marker>\tag important\ convergent_prod_iff_convergent: fixes f :: "nat \ 'a :: {topological_semigroup_mult,t2_space,idom}" assumes "\i. f i \ 0" shows "convergent_prod f \ convergent (\n. \i\n. f i) \ lim (\n. \i\n. f i) \ 0" by (force simp: convergent_prod_iff_nz_lim assms convergent_def limI) lemma bounded_imp_convergent_prod: fixes a :: "nat \ real" assumes 1: "\n. a n \ 1" and bounded: "\n. (\i\n. a i) \ B" shows "convergent_prod a" proof - have "bdd_above (range(\n. \i\n. a i))" by (meson bdd_aboveI2 bounded) moreover have "incseq (\n. \i\n. a i)" unfolding mono_def by (metis 1 prod_mono2 atMost_subset_iff dual_order.trans finite_atMost zero_le_one) ultimately obtain p where p: "(\n. \i\n. a i) \ p" using LIMSEQ_incseq_SUP by blast then have "p \ 0" by (metis "1" not_one_le_zero prod_ge_1 LIMSEQ_le_const) with 1 p show ?thesis by (metis convergent_prod_iff_nz_lim not_one_le_zero) qed lemma abs_convergent_prod_altdef: fixes f :: "nat \ 'a :: {one,real_normed_vector}" shows "abs_convergent_prod f \ convergent (\n. \i\n. 1 + norm (f i - 1))" proof assume "abs_convergent_prod f" thus "convergent (\n. \i\n. 1 + norm (f i - 1))" by (auto simp: abs_convergent_prod_def intro!: convergent_prod_imp_convergent) qed (auto intro: abs_convergent_prodI) lemma Weierstrass_prod_ineq: fixes f :: "'a \ real" assumes "\x. x \ A \ f x \ {0..1}" shows "1 - sum f A \ (\x\A. 1 - f x)" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) from insert.hyps and insert.prems have "1 - sum f A + f x * (\x\A. 1 - f x) \ (\x\A. 1 - f x) + f x * (\x\A. 1)" by (intro insert.IH add_mono mult_left_mono prod_mono) auto with insert.hyps show ?case by (simp add: algebra_simps) qed simp_all lemma norm_prod_minus1_le_prod_minus1: fixes f :: "nat \ 'a :: {real_normed_div_algebra,comm_ring_1}" shows "norm (prod (\n. 1 + f n) A - 1) \ prod (\n. 1 + norm (f n)) A - 1" proof (induction A rule: infinite_finite_induct) case (insert x A) from insert.hyps have "norm ((\n\insert x A. 1 + f n) - 1) = norm ((\n\A. 1 + f n) - 1 + f x * (\n\A. 1 + f n))" by (simp add: algebra_simps) also have "\ \ norm ((\n\A. 1 + f n) - 1) + norm (f x * (\n\A. 1 + f n))" by (rule norm_triangle_ineq) also have "norm (f x * (\n\A. 1 + f n)) = norm (f x) * (\x\A. norm (1 + f x))" by (simp add: prod_norm norm_mult) also have "(\x\A. norm (1 + f x)) \ (\x\A. norm (1::'a) + norm (f x))" by (intro prod_mono norm_triangle_ineq ballI conjI) auto also have "norm (1::'a) = 1" by simp also note insert.IH also have "(\n\A. 1 + norm (f n)) - 1 + norm (f x) * (\x\A. 1 + norm (f x)) = (\n\insert x A. 1 + norm (f n)) - 1" using insert.hyps by (simp add: algebra_simps) finally show ?case by - (simp_all add: mult_left_mono) qed simp_all lemma convergent_prod_imp_ev_nonzero: fixes f :: "nat \ 'a :: {t2_space,comm_semiring_1}" assumes "convergent_prod f" shows "eventually (\n. f n \ 0) sequentially" using assms by (auto simp: eventually_at_top_linorder convergent_prod_altdef) lemma convergent_prod_imp_LIMSEQ: fixes f :: "nat \ 'a :: {real_normed_field}" assumes "convergent_prod f" shows "f \ 1" proof - from assms obtain M L where L: "(\n. \i\n. f (i+M)) \ L" "\n. n \ M \ f n \ 0" "L \ 0" by (auto simp: convergent_prod_altdef) hence L': "(\n. \i\Suc n. f (i+M)) \ L" by (subst filterlim_sequentially_Suc) have "(\n. (\i\Suc n. f (i+M)) / (\i\n. f (i+M))) \ L / L" using L L' by (intro tendsto_divide) simp_all also from L have "L / L = 1" by simp also have "(\n. (\i\Suc n. f (i+M)) / (\i\n. f (i+M))) = (\n. f (n + Suc M))" using assms L by (auto simp: fun_eq_iff atMost_Suc) finally show ?thesis by (rule LIMSEQ_offset) qed lemma abs_convergent_prod_imp_summable: fixes f :: "nat \ 'a :: real_normed_div_algebra" assumes "abs_convergent_prod f" shows "summable (\i. norm (f i - 1))" proof - from assms have "convergent (\n. \i\n. 1 + norm (f i - 1))" unfolding abs_convergent_prod_def by (rule convergent_prod_imp_convergent) then obtain L where L: "(\n. \i\n. 1 + norm (f i - 1)) \ L" unfolding convergent_def by blast have "convergent (\n. \i\n. norm (f i - 1))" proof (rule Bseq_monoseq_convergent) have "eventually (\n. (\i\n. 1 + norm (f i - 1)) < L + 1) sequentially" using L(1) by (rule order_tendstoD) simp_all hence "\\<^sub>F x in sequentially. norm (\i\x. norm (f i - 1)) \ L + 1" proof eventually_elim case (elim n) have "norm (\i\n. norm (f i - 1)) = (\i\n. norm (f i - 1))" unfolding real_norm_def by (intro abs_of_nonneg sum_nonneg) simp_all also have "\ \ (\i\n. 1 + norm (f i - 1))" by (rule sum_le_prod) auto also have "\ < L + 1" by (rule elim) finally show ?case by simp qed thus "Bseq (\n. \i\n. norm (f i - 1))" by (rule BfunI) next show "monoseq (\n. \i\n. norm (f i - 1))" by (rule mono_SucI1) auto qed thus "summable (\i. norm (f i - 1))" by (simp add: summable_iff_convergent') qed lemma summable_imp_abs_convergent_prod: fixes f :: "nat \ 'a :: real_normed_div_algebra" assumes "summable (\i. norm (f i - 1))" shows "abs_convergent_prod f" proof (intro abs_convergent_prodI Bseq_monoseq_convergent) show "monoseq (\n. \i\n. 1 + norm (f i - 1))" by (intro mono_SucI1) (auto simp: atMost_Suc algebra_simps intro!: mult_nonneg_nonneg prod_nonneg) next show "Bseq (\n. \i\n. 1 + norm (f i - 1))" proof (rule Bseq_eventually_mono) show "eventually (\n. norm (\i\n. 1 + norm (f i - 1)) \ norm (exp (\i\n. norm (f i - 1)))) sequentially" by (intro always_eventually allI) (auto simp: abs_prod exp_sum intro!: prod_mono) next from assms have "(\n. \i\n. norm (f i - 1)) \ (\i. norm (f i - 1))" using sums_def_le by blast hence "(\n. exp (\i\n. norm (f i - 1))) \ exp (\i. norm (f i - 1))" by (rule tendsto_exp) hence "convergent (\n. exp (\i\n. norm (f i - 1)))" by (rule convergentI) thus "Bseq (\n. exp (\i\n. norm (f i - 1)))" by (rule convergent_imp_Bseq) qed qed theorem abs_convergent_prod_conv_summable: fixes f :: "nat \ 'a :: real_normed_div_algebra" shows "abs_convergent_prod f \ summable (\i. norm (f i - 1))" by (blast intro: abs_convergent_prod_imp_summable summable_imp_abs_convergent_prod) lemma abs_convergent_prod_imp_LIMSEQ: fixes f :: "nat \ 'a :: {comm_ring_1,real_normed_div_algebra}" assumes "abs_convergent_prod f" shows "f \ 1" proof - from assms have "summable (\n. norm (f n - 1))" by (rule abs_convergent_prod_imp_summable) from summable_LIMSEQ_zero[OF this] have "(\n. f n - 1) \ 0" by (simp add: tendsto_norm_zero_iff) from tendsto_add[OF this tendsto_const[of 1]] show ?thesis by simp qed lemma abs_convergent_prod_imp_ev_nonzero: fixes f :: "nat \ 'a :: {comm_ring_1,real_normed_div_algebra}" assumes "abs_convergent_prod f" shows "eventually (\n. f n \ 0) sequentially" proof - from assms have "f \ 1" by (rule abs_convergent_prod_imp_LIMSEQ) hence "eventually (\n. dist (f n) 1 < 1) at_top" by (auto simp: tendsto_iff) thus ?thesis by eventually_elim auto qed subsection\<^marker>\tag unimportant\ \Ignoring initial segments\ lemma convergent_prod_offset: assumes "convergent_prod (\n. f (n + m))" shows "convergent_prod f" proof - from assms obtain M L where "(\n. \k\n. f (k + (M + m))) \ L" "L \ 0" by (auto simp: prod_defs add.assoc) thus "convergent_prod f" unfolding prod_defs by blast qed lemma abs_convergent_prod_offset: assumes "abs_convergent_prod (\n. f (n + m))" shows "abs_convergent_prod f" using assms unfolding abs_convergent_prod_def by (rule convergent_prod_offset) lemma raw_has_prod_ignore_initial_segment: fixes f :: "nat \ 'a :: real_normed_field" assumes "raw_has_prod f M p" "N \ M" obtains q where "raw_has_prod f N q" proof - have p: "(\n. \k\n. f (k + M)) \ p" and "p \ 0" using assms by (auto simp: raw_has_prod_def) then have nz: "\n. n \ M \ f n \ 0" using assms by (auto simp: raw_has_prod_eq_0) define C where "C = (\k 0" by (auto simp: C_def) from p have "(\i. \k\i + (N-M). f (k + M)) \ p" by (rule LIMSEQ_ignore_initial_segment) also have "(\i. \k\i + (N-M). f (k + M)) = (\n. C * (\k\n. f (k + N)))" proof (rule ext, goal_cases) case (1 n) have "{..n+(N-M)} = {..<(N-M)} \ {(N-M)..n+(N-M)}" by auto also have "(\k\\. f (k + M)) = C * (\k=(N-M)..n+(N-M). f (k + M))" unfolding C_def by (rule prod.union_disjoint) auto also have "(\k=(N-M)..n+(N-M). f (k + M)) = (\k\n. f (k + (N-M) + M))" by (intro ext prod.reindex_bij_witness[of _ "\k. k + (N-M)" "\k. k - (N-M)"]) auto finally show ?case using \N \ M\ by (simp add: add_ac) qed finally have "(\n. C * (\k\n. f (k + N)) / C) \ p / C" by (intro tendsto_divide tendsto_const) auto hence "(\n. \k\n. f (k + N)) \ p / C" by simp moreover from \p \ 0\ have "p / C \ 0" by simp ultimately show ?thesis using raw_has_prod_def that by blast qed corollary\<^marker>\tag unimportant\ convergent_prod_ignore_initial_segment: fixes f :: "nat \ 'a :: real_normed_field" assumes "convergent_prod f" shows "convergent_prod (\n. f (n + m))" using assms unfolding convergent_prod_def apply clarify apply (erule_tac N="M+m" in raw_has_prod_ignore_initial_segment) apply (auto simp add: raw_has_prod_def add_ac) done corollary\<^marker>\tag unimportant\ convergent_prod_ignore_nonzero_segment: fixes f :: "nat \ 'a :: real_normed_field" assumes f: "convergent_prod f" and nz: "\i. i \ M \ f i \ 0" shows "\p. raw_has_prod f M p" using convergent_prod_ignore_initial_segment [OF f] by (metis convergent_LIMSEQ_iff convergent_prod_iff_convergent le_add_same_cancel2 nz prod_defs(1) zero_order(1)) corollary\<^marker>\tag unimportant\ abs_convergent_prod_ignore_initial_segment: assumes "abs_convergent_prod f" shows "abs_convergent_prod (\n. f (n + m))" using assms unfolding abs_convergent_prod_def by (rule convergent_prod_ignore_initial_segment) subsection\More elementary properties\ theorem abs_convergent_prod_imp_convergent_prod: fixes f :: "nat \ 'a :: {real_normed_div_algebra,complete_space,comm_ring_1}" assumes "abs_convergent_prod f" shows "convergent_prod f" proof - from assms have "eventually (\n. f n \ 0) sequentially" by (rule abs_convergent_prod_imp_ev_nonzero) then obtain N where N: "f n \ 0" if "n \ N" for n by (auto simp: eventually_at_top_linorder) let ?P = "\n. \i\n. f (i + N)" and ?Q = "\n. \i\n. 1 + norm (f (i + N) - 1)" have "Cauchy ?P" proof (rule CauchyI', goal_cases) case (1 \) from assms have "abs_convergent_prod (\n. f (n + N))" by (rule abs_convergent_prod_ignore_initial_segment) hence "Cauchy ?Q" unfolding abs_convergent_prod_def by (intro convergent_Cauchy convergent_prod_imp_convergent) from CauchyD[OF this 1] obtain M where M: "norm (?Q m - ?Q n) < \" if "m \ M" "n \ M" for m n by blast show ?case proof (rule exI[of _ M], safe, goal_cases) case (1 m n) have "dist (?P m) (?P n) = norm (?P n - ?P m)" by (simp add: dist_norm norm_minus_commute) also from 1 have "{..n} = {..m} \ {m<..n}" by auto hence "norm (?P n - ?P m) = norm (?P m * (\k\{m<..n}. f (k + N)) - ?P m)" by (subst prod.union_disjoint [symmetric]) (auto simp: algebra_simps) also have "\ = norm (?P m * ((\k\{m<..n}. f (k + N)) - 1))" by (simp add: algebra_simps) also have "\ = (\k\m. norm (f (k + N))) * norm ((\k\{m<..n}. f (k + N)) - 1)" by (simp add: norm_mult prod_norm) also have "\ \ ?Q m * ((\k\{m<..n}. 1 + norm (f (k + N) - 1)) - 1)" using norm_prod_minus1_le_prod_minus1[of "\k. f (k + N) - 1" "{m<..n}"] norm_triangle_ineq[of 1 "f k - 1" for k] by (intro mult_mono prod_mono ballI conjI norm_prod_minus1_le_prod_minus1 prod_nonneg) auto also have "\ = ?Q m * (\k\{m<..n}. 1 + norm (f (k + N) - 1)) - ?Q m" by (simp add: algebra_simps) also have "?Q m * (\k\{m<..n}. 1 + norm (f (k + N) - 1)) = (\k\{..m}\{m<..n}. 1 + norm (f (k + N) - 1))" by (rule prod.union_disjoint [symmetric]) auto also from 1 have "{..m}\{m<..n} = {..n}" by auto also have "?Q n - ?Q m \ norm (?Q n - ?Q m)" by simp also from 1 have "\ < \" by (intro M) auto finally show ?case . qed qed hence conv: "convergent ?P" by (rule Cauchy_convergent) then obtain L where L: "?P \ L" by (auto simp: convergent_def) have "L \ 0" proof assume [simp]: "L = 0" from tendsto_norm[OF L] have limit: "(\n. \k\n. norm (f (k + N))) \ 0" by (simp add: prod_norm) from assms have "(\n. f (n + N)) \ 1" by (intro abs_convergent_prod_imp_LIMSEQ abs_convergent_prod_ignore_initial_segment) hence "eventually (\n. norm (f (n + N) - 1) < 1) sequentially" by (auto simp: tendsto_iff dist_norm) then obtain M0 where M0: "norm (f (n + N) - 1) < 1" if "n \ M0" for n by (auto simp: eventually_at_top_linorder) { fix M assume M: "M \ M0" with M0 have M: "norm (f (n + N) - 1) < 1" if "n \ M" for n using that by simp have "(\n. \k\n. 1 - norm (f (k+M+N) - 1)) \ 0" proof (rule tendsto_sandwich) show "eventually (\n. (\k\n. 1 - norm (f (k+M+N) - 1)) \ 0) sequentially" using M by (intro always_eventually prod_nonneg allI ballI) (auto intro: less_imp_le) have "norm (1::'a) - norm (f (i + M + N) - 1) \ norm (f (i + M + N))" for i using norm_triangle_ineq3[of "f (i + M + N)" 1] by simp thus "eventually (\n. (\k\n. 1 - norm (f (k+M+N) - 1)) \ (\k\n. norm (f (k+M+N)))) at_top" using M by (intro always_eventually allI prod_mono ballI conjI) (auto intro: less_imp_le) define C where "C = (\k 0" by (auto simp: C_def) from L have "(\n. norm (\k\n+M. f (k + N))) \ 0" by (intro LIMSEQ_ignore_initial_segment) (simp add: tendsto_norm_zero_iff) also have "(\n. norm (\k\n+M. f (k + N))) = (\n. C * (\k\n. norm (f (k + M + N))))" proof (rule ext, goal_cases) case (1 n) have "{..n+M} = {.. {M..n+M}" by auto also have "norm (\k\\. f (k + N)) = C * norm (\k=M..n+M. f (k + N))" unfolding C_def by (subst prod.union_disjoint) (auto simp: norm_mult prod_norm) also have "(\k=M..n+M. f (k + N)) = (\k\n. f (k + N + M))" by (intro prod.reindex_bij_witness[of _ "\i. i + M" "\i. i - M"]) auto finally show ?case by (simp add: add_ac prod_norm) qed finally have "(\n. C * (\k\n. norm (f (k + M + N))) / C) \ 0 / C" by (intro tendsto_divide tendsto_const) auto thus "(\n. \k\n. norm (f (k + M + N))) \ 0" by simp qed simp_all have "1 - (\i. norm (f (i + M + N) - 1)) \ 0" proof (rule tendsto_le) show "eventually (\n. 1 - (\k\n. norm (f (k+M+N) - 1)) \ (\k\n. 1 - norm (f (k+M+N) - 1))) at_top" using M by (intro always_eventually allI Weierstrass_prod_ineq) (auto intro: less_imp_le) show "(\n. \k\n. 1 - norm (f (k+M+N) - 1)) \ 0" by fact show "(\n. 1 - (\k\n. norm (f (k + M + N) - 1))) \ 1 - (\i. norm (f (i + M + N) - 1))" by (intro tendsto_intros summable_LIMSEQ' summable_ignore_initial_segment abs_convergent_prod_imp_summable assms) qed simp_all hence "(\i. norm (f (i + M + N) - 1)) \ 1" by simp also have "\ + (\ii. norm (f (i + N) - 1))" by (intro suminf_split_initial_segment [symmetric] summable_ignore_initial_segment abs_convergent_prod_imp_summable assms) finally have "1 + (\i (\i. norm (f (i + N) - 1))" by simp } note * = this have "1 + (\i. norm (f (i + N) - 1)) \ (\i. norm (f (i + N) - 1))" proof (rule tendsto_le) show "(\M. 1 + (\i 1 + (\i. norm (f (i + N) - 1))" by (intro tendsto_intros summable_LIMSEQ summable_ignore_initial_segment abs_convergent_prod_imp_summable assms) show "eventually (\M. 1 + (\i (\i. norm (f (i + N) - 1))) at_top" using eventually_ge_at_top[of M0] by eventually_elim (use * in auto) qed simp_all thus False by simp qed with L show ?thesis by (auto simp: prod_defs) qed lemma raw_has_prod_cases: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" assumes "raw_has_prod f M p" obtains i where "in. \i\n. f (i + M)) \ p" "p \ 0" using assms unfolding raw_has_prod_def by blast+ then have "(\n. prod f {..i\n. f (i + M))) \ prod f {..i\n. f (i + M)) = prod f {..n+M}" for n proof - have "{..n+M} = {.. {M..n+M}" by auto then have "prod f {..n+M} = prod f {.. = prod f {..i\n. f (i + M))" by (metis (mono_tags, lifting) add.left_neutral atMost_atLeast0 prod.shift_bounds_cl_nat_ivl) finally show ?thesis by metis qed ultimately have "(\n. prod f {..n}) \ prod f {..i 0" using \p \ 0\ assms that by (auto simp: raw_has_prod_def) then show thesis using that by blast qed corollary convergent_prod_offset_0: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" assumes "convergent_prod f" "\i. f i \ 0" shows "\p. raw_has_prod f 0 p" using assms convergent_prod_def raw_has_prod_cases by blast lemma prodinf_eq_lim: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" assumes "convergent_prod f" "\i. f i \ 0" shows "prodinf f = lim (\n. \i\n. f i)" using assms convergent_prod_offset_0 [OF assms] by (simp add: prod_defs lim_def) (metis (no_types) assms(1) convergent_prod_to_zero_iff) lemma has_prod_one[simp, intro]: "(\n. 1) has_prod 1" unfolding prod_defs by auto lemma convergent_prod_one[simp, intro]: "convergent_prod (\n. 1)" unfolding prod_defs by auto lemma prodinf_cong: "(\n. f n = g n) \ prodinf f = prodinf g" by presburger lemma convergent_prod_cong: fixes f g :: "nat \ 'a::{field,topological_semigroup_mult,t2_space}" assumes ev: "eventually (\x. f x = g x) sequentially" and f: "\i. f i \ 0" and g: "\i. g i \ 0" shows "convergent_prod f = convergent_prod g" proof - from assms obtain N where N: "\n\N. f n = g n" by (auto simp: eventually_at_top_linorder) define C where "C = (\k 0" by (simp add: f) have *: "eventually (\n. prod f {..n} = C * prod g {..n}) sequentially" using eventually_ge_at_top[of N] proof eventually_elim case (elim n) then have "{..n} = {.. {N..n}" by auto also have "prod f \ = prod f {.. {N..n})" by (intro prod.union_disjoint [symmetric]) auto also from elim have "{.. {N..n} = {..n}" by auto finally show "prod f {..n} = C * prod g {..n}" . qed then have cong: "convergent (\n. prod f {..n}) = convergent (\n. C * prod g {..n})" by (rule convergent_cong) show ?thesis proof assume cf: "convergent_prod f" then have "\ (\n. prod g {..n}) \ 0" using tendsto_mult_left * convergent_prod_to_zero_iff f filterlim_cong by fastforce then show "convergent_prod g" by (metis convergent_mult_const_iff \C \ 0\ cong cf convergent_LIMSEQ_iff convergent_prod_iff_convergent convergent_prod_imp_convergent g) next assume cg: "convergent_prod g" have "\a. C * a \ 0 \ (\n. prod g {..n}) \ a" by (metis (no_types) \C \ 0\ cg convergent_prod_iff_nz_lim divide_eq_0_iff g nonzero_mult_div_cancel_right) then show "convergent_prod f" using "*" tendsto_mult_left filterlim_cong by (fastforce simp add: convergent_prod_iff_nz_lim f) qed qed lemma has_prod_finite: fixes f :: "nat \ 'a::{semidom,t2_space}" assumes [simp]: "finite N" and f: "\n. n \ N \ f n = 1" shows "f has_prod (\n\N. f n)" proof - have eq: "prod f {..n + Suc (Max N)} = prod f N" for n proof (rule prod.mono_neutral_right) show "N \ {..n + Suc (Max N)}" by (auto simp: le_Suc_eq trans_le_add2) show "\i\{..n + Suc (Max N)} - N. f i = 1" using f by blast qed auto show ?thesis proof (cases "\n\N. f n \ 0") case True then have "prod f N \ 0" by simp moreover have "(\n. prod f {..n}) \ prod f N" by (rule LIMSEQ_offset[of _ "Suc (Max N)"]) (simp add: eq atLeast0LessThan del: add_Suc_right) ultimately show ?thesis by (simp add: raw_has_prod_def has_prod_def) next case False then obtain k where "k \ N" "f k = 0" by auto let ?Z = "{n \ N. f n = 0}" have maxge: "Max ?Z \ n" if "f n = 0" for n using Max_ge [of ?Z] \finite N\ \f n = 0\ by (metis (mono_tags) Collect_mem_eq f finite_Collect_conjI mem_Collect_eq zero_neq_one) let ?q = "prod f {Suc (Max ?Z)..Max N}" have [simp]: "?q \ 0" using maxge Suc_n_not_le_n le_trans by force have eq: "(\i\n + Max N. f (Suc (i + Max ?Z))) = ?q" for n proof - have "(\i\n + Max N. f (Suc (i + Max ?Z))) = prod f {Suc (Max ?Z)..n + Max N + Suc (Max ?Z)}" proof (rule prod.reindex_cong [where l = "\i. i + Suc (Max ?Z)", THEN sym]) show "{Suc (Max ?Z)..n + Max N + Suc (Max ?Z)} = (\i. i + Suc (Max ?Z)) ` {..n + Max N}" using le_Suc_ex by fastforce qed (auto simp: inj_on_def) also have "\ = ?q" by (rule prod.mono_neutral_right) (use Max.coboundedI [OF \finite N\] f in \force+\) finally show ?thesis . qed have q: "raw_has_prod f (Suc (Max ?Z)) ?q" proof (simp add: raw_has_prod_def) show "(\n. \i\n. f (Suc (i + Max ?Z))) \ ?q" by (rule LIMSEQ_offset[of _ "(Max N)"]) (simp add: eq) qed show ?thesis unfolding has_prod_def proof (intro disjI2 exI conjI) show "prod f N = 0" using \f k = 0\ \k \ N\ \finite N\ prod_zero by blast show "f (Max ?Z) = 0" using Max_in [of ?Z] \finite N\ \f k = 0\ \k \ N\ by auto qed (use q in auto) qed qed corollary\<^marker>\tag unimportant\ has_prod_0: fixes f :: "nat \ 'a::{semidom,t2_space}" assumes "\n. f n = 1" shows "f has_prod 1" by (simp add: assms has_prod_cong) lemma prodinf_zero[simp]: "prodinf (\n. 1::'a::real_normed_field) = 1" using has_prod_unique by force lemma convergent_prod_finite: fixes f :: "nat \ 'a::{idom,t2_space}" assumes "finite N" "\n. n \ N \ f n = 1" shows "convergent_prod f" proof - have "\n p. raw_has_prod f n p" using assms has_prod_def has_prod_finite by blast then show ?thesis by (simp add: convergent_prod_def) qed lemma has_prod_If_finite_set: fixes f :: "nat \ 'a::{idom,t2_space}" shows "finite A \ (\r. if r \ A then f r else 1) has_prod (\r\A. f r)" using has_prod_finite[of A "(\r. if r \ A then f r else 1)"] by simp lemma has_prod_If_finite: fixes f :: "nat \ 'a::{idom,t2_space}" shows "finite {r. P r} \ (\r. if P r then f r else 1) has_prod (\r | P r. f r)" using has_prod_If_finite_set[of "{r. P r}"] by simp lemma convergent_prod_If_finite_set[simp, intro]: fixes f :: "nat \ 'a::{idom,t2_space}" shows "finite A \ convergent_prod (\r. if r \ A then f r else 1)" by (simp add: convergent_prod_finite) lemma convergent_prod_If_finite[simp, intro]: fixes f :: "nat \ 'a::{idom,t2_space}" shows "finite {r. P r} \ convergent_prod (\r. if P r then f r else 1)" using convergent_prod_def has_prod_If_finite has_prod_def by fastforce lemma has_prod_single: fixes f :: "nat \ 'a::{idom,t2_space}" shows "(\r. if r = i then f r else 1) has_prod f i" using has_prod_If_finite[of "\r. r = i"] by simp context fixes f :: "nat \ 'a :: real_normed_field" begin lemma convergent_prod_imp_has_prod: assumes "convergent_prod f" shows "\p. f has_prod p" proof - obtain M p where p: "raw_has_prod f M p" using assms convergent_prod_def by blast then have "p \ 0" using raw_has_prod_nonzero by blast with p have fnz: "f i \ 0" if "i \ M" for i using raw_has_prod_eq_0 that by blast define C where "C = (\nn\M. f n \ 0") case True then have "C \ 0" by (simp add: C_def) then show ?thesis by (meson True assms convergent_prod_offset_0 fnz has_prod_def nat_le_linear) next case False let ?N = "GREATEST n. f n = 0" have 0: "f ?N = 0" using fnz False by (metis (mono_tags, lifting) GreatestI_ex_nat nat_le_linear) have "f i \ 0" if "i > ?N" for i by (metis (mono_tags, lifting) Greatest_le_nat fnz leD linear that) then have "\p. raw_has_prod f (Suc ?N) p" using assms by (auto simp: intro!: convergent_prod_ignore_nonzero_segment) then show ?thesis unfolding has_prod_def using 0 by blast qed qed lemma convergent_prod_has_prod [intro]: shows "convergent_prod f \ f has_prod (prodinf f)" unfolding prodinf_def by (metis convergent_prod_imp_has_prod has_prod_unique theI') lemma convergent_prod_LIMSEQ: shows "convergent_prod f \ (\n. \i\n. f i) \ prodinf f" by (metis convergent_LIMSEQ_iff convergent_prod_has_prod convergent_prod_imp_convergent convergent_prod_to_zero_iff raw_has_prod_eq_0 has_prod_def prodinf_eq_lim zero_le) theorem has_prod_iff: "f has_prod x \ convergent_prod f \ prodinf f = x" proof assume "f has_prod x" then show "convergent_prod f \ prodinf f = x" apply safe using convergent_prod_def has_prod_def apply blast using has_prod_unique by blast qed auto lemma convergent_prod_has_prod_iff: "convergent_prod f \ f has_prod prodinf f" by (auto simp: has_prod_iff convergent_prod_has_prod) lemma prodinf_finite: assumes N: "finite N" and f: "\n. n \ N \ f n = 1" shows "prodinf f = (\n\N. f n)" using has_prod_finite[OF assms, THEN has_prod_unique] by simp end subsection\<^marker>\tag unimportant\ \Infinite products on ordered topological monoids\ lemma LIMSEQ_prod_0: fixes f :: "nat \ 'a::{semidom,topological_space}" assumes "f i = 0" shows "(\n. prod f {..n}) \ 0" proof (subst tendsto_cong) show "\\<^sub>F n in sequentially. prod f {..n} = 0" proof show "prod f {..n} = 0" if "n \ i" for n using that assms by auto qed qed auto lemma LIMSEQ_prod_nonneg: fixes f :: "nat \ 'a::{linordered_semidom,linorder_topology}" assumes 0: "\n. 0 \ f n" and a: "(\n. prod f {..n}) \ a" shows "a \ 0" by (simp add: "0" prod_nonneg LIMSEQ_le_const [OF a]) context fixes f :: "nat \ 'a::{linordered_semidom,linorder_topology}" begin lemma has_prod_le: assumes f: "f has_prod a" and g: "g has_prod b" and le: "\n. 0 \ f n \ f n \ g n" shows "a \ b" proof (cases "a=0 \ b=0") case True then show ?thesis proof assume [simp]: "a=0" have "b \ 0" proof (rule LIMSEQ_prod_nonneg) show "(\n. prod g {..n}) \ b" using g by (auto simp: has_prod_def raw_has_prod_def LIMSEQ_prod_0) qed (use le order_trans in auto) then show ?thesis by auto next assume [simp]: "b=0" then obtain i where "g i = 0" using g by (auto simp: prod_defs) then have "f i = 0" using antisym le by force then have "a=0" using f by (auto simp: prod_defs LIMSEQ_prod_0 LIMSEQ_unique) then show ?thesis by auto qed next case False then show ?thesis using assms unfolding has_prod_def raw_has_prod_def by (force simp: LIMSEQ_prod_0 intro!: LIMSEQ_le prod_mono) qed lemma prodinf_le: assumes f: "f has_prod a" and g: "g has_prod b" and le: "\n. 0 \ f n \ f n \ g n" shows "prodinf f \ prodinf g" using has_prod_le [OF assms] has_prod_unique f g by blast end lemma prod_le_prodinf: fixes f :: "nat \ 'a::{linordered_idom,linorder_topology}" assumes "f has_prod a" "\i. 0 \ f i" "\i. i\n \ 1 \ f i" shows "prod f {.. prodinf f" by(rule has_prod_le[OF has_prod_If_finite_set]) (use assms has_prod_unique in auto) lemma prodinf_nonneg: fixes f :: "nat \ 'a::{linordered_idom,linorder_topology}" assumes "f has_prod a" "\i. 1 \ f i" shows "1 \ prodinf f" using prod_le_prodinf[of f a 0] assms by (metis order_trans prod_ge_1 zero_le_one) lemma prodinf_le_const: fixes f :: "nat \ real" assumes "convergent_prod f" "\n. prod f {.. x" shows "prodinf f \ x" by (metis lessThan_Suc_atMost assms convergent_prod_LIMSEQ LIMSEQ_le_const2) -lemma prodinf_eq_one_iff: +lemma prodinf_eq_one_iff [simp]: fixes f :: "nat \ real" assumes f: "convergent_prod f" and ge1: "\n. 1 \ f n" shows "prodinf f = 1 \ (\n. f n = 1)" proof assume "prodinf f = 1" then have "(\n. \i 1" using convergent_prod_LIMSEQ[of f] assms by (simp add: LIMSEQ_lessThan_iff_atMost) then have "\i. (\n\{i}. f n) \ 1" proof (rule LIMSEQ_le_const) have "1 \ prod f n" for n by (simp add: ge1 prod_ge_1) have "prod f {..\n. 1 \ prod f n\ \prodinf f = 1\ antisym f convergent_prod_has_prod ge1 order_trans prod_le_prodinf zero_le_one) then have "(\n\{i}. f n) \ prod f {.. Suc i" for i n by (metis mult.left_neutral order_refl prod.cong prod.neutral_const prod.lessThan_Suc) then show "\N. \n\N. (\n\{i}. f n) \ prod f {..n. f n = 1" by (auto intro!: antisym) qed (metis prodinf_zero fun_eq_iff) lemma prodinf_pos_iff: fixes f :: "nat \ real" assumes "convergent_prod f" "\n. 1 \ f n" shows "1 < prodinf f \ (\i. 1 < f i)" using prod_le_prodinf[of f 1] prodinf_eq_one_iff by (metis convergent_prod_has_prod assms less_le prodinf_nonneg) lemma less_1_prodinf2: fixes f :: "nat \ real" assumes "convergent_prod f" "\n. 1 \ f n" "1 < f i" shows "1 < prodinf f" proof - have "1 < (\n \ prodinf f" by (intro prod_le_prodinf) (use assms order_trans zero_le_one in \blast+\) finally show ?thesis . qed lemma less_1_prodinf: fixes f :: "nat \ real" shows "\convergent_prod f; \n. 1 < f n\ \ 1 < prodinf f" by (intro less_1_prodinf2[where i=1]) (auto intro: less_imp_le) lemma prodinf_nonzero: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" assumes "convergent_prod f" "\i. f i \ 0" shows "prodinf f \ 0" by (metis assms convergent_prod_offset_0 has_prod_unique raw_has_prod_def has_prod_def) lemma less_0_prodinf: fixes f :: "nat \ real" assumes f: "convergent_prod f" and 0: "\i. f i > 0" shows "0 < prodinf f" proof - have "prodinf f \ 0" by (metis assms less_irrefl prodinf_nonzero) moreover have "0 < (\n 0" using convergent_prod_LIMSEQ [OF f] LIMSEQ_prod_nonneg 0 less_le by blast ultimately show ?thesis by auto qed lemma prod_less_prodinf2: fixes f :: "nat \ real" assumes f: "convergent_prod f" and 1: "\m. m\n \ 1 \ f m" and 0: "\m. 0 < f m" and i: "n \ i" "1 < f i" shows "prod f {.. prod f {.. prodinf f" using prod_le_prodinf[of f _ "Suc i"] by (meson "0" "1" Suc_leD convergent_prod_has_prod f \n \ i\ le_trans less_eq_real_def) ultimately show ?thesis by (metis le_less_trans mult.commute not_le prod.lessThan_Suc) qed lemma prod_less_prodinf: fixes f :: "nat \ real" assumes f: "convergent_prod f" and 1: "\m. m\n \ 1 < f m" and 0: "\m. 0 < f m" shows "prod f {.. real" assumes pos: "\n. 1 \ f n" and le: "\n. (\i x" shows "\p. raw_has_prod f 0 p" unfolding raw_has_prod_def add_0_right proof (rule exI LIMSEQ_incseq_SUP conjI)+ show "bdd_above (range (\n. prod f {..n}))" by (metis bdd_aboveI2 le lessThan_Suc_atMost) then have "(SUP i. prod f {..i}) > 0" by (metis UNIV_I cSUP_upper less_le_trans pos prod_pos zero_less_one) then show "(SUP i. prod f {..i}) \ 0" by auto show "incseq (\n. prod f {..n})" using pos order_trans [OF zero_le_one] by (auto simp: mono_def intro!: prod_mono2) qed lemma convergent_prodI_nonneg_bounded: fixes f :: "nat \ real" assumes "\n. 1 \ f n" "\n. (\i x" shows "convergent_prod f" using convergent_prod_def raw_has_prodI_bounded [OF assms] by blast subsection\<^marker>\tag unimportant\ \Infinite products on topological spaces\ context fixes f g :: "nat \ 'a::{t2_space,topological_semigroup_mult,idom}" begin lemma raw_has_prod_mult: "\raw_has_prod f M a; raw_has_prod g M b\ \ raw_has_prod (\n. f n * g n) M (a * b)" by (force simp add: prod.distrib tendsto_mult raw_has_prod_def) lemma has_prod_mult_nz: "\f has_prod a; g has_prod b; a \ 0; b \ 0\ \ (\n. f n * g n) has_prod (a * b)" by (simp add: raw_has_prod_mult has_prod_def) end context fixes f g :: "nat \ 'a::real_normed_field" begin lemma has_prod_mult: assumes f: "f has_prod a" and g: "g has_prod b" shows "(\n. f n * g n) has_prod (a * b)" using f [unfolded has_prod_def] proof (elim disjE exE conjE) assume f0: "raw_has_prod f 0 a" show ?thesis using g [unfolded has_prod_def] proof (elim disjE exE conjE) assume g0: "raw_has_prod g 0 b" with f0 show ?thesis by (force simp add: has_prod_def prod.distrib tendsto_mult raw_has_prod_def) next fix j q assume "b = 0" and "g j = 0" and q: "raw_has_prod g (Suc j) q" obtain p where p: "raw_has_prod f (Suc j) p" using f0 raw_has_prod_ignore_initial_segment by blast then have "Ex (raw_has_prod (\n. f n * g n) (Suc j))" using q raw_has_prod_mult by blast then show ?thesis using \b = 0\ \g j = 0\ has_prod_0_iff by fastforce qed next fix i p assume "a = 0" and "f i = 0" and p: "raw_has_prod f (Suc i) p" show ?thesis using g [unfolded has_prod_def] proof (elim disjE exE conjE) assume g0: "raw_has_prod g 0 b" obtain q where q: "raw_has_prod g (Suc i) q" using g0 raw_has_prod_ignore_initial_segment by blast then have "Ex (raw_has_prod (\n. f n * g n) (Suc i))" using raw_has_prod_mult p by blast then show ?thesis using \a = 0\ \f i = 0\ has_prod_0_iff by fastforce next fix j q assume "b = 0" and "g j = 0" and q: "raw_has_prod g (Suc j) q" obtain p' where p': "raw_has_prod f (Suc (max i j)) p'" by (metis raw_has_prod_ignore_initial_segment max_Suc_Suc max_def p) moreover obtain q' where q': "raw_has_prod g (Suc (max i j)) q'" by (metis raw_has_prod_ignore_initial_segment max.cobounded2 max_Suc_Suc q) ultimately show ?thesis using \b = 0\ by (simp add: has_prod_def) (metis \f i = 0\ \g j = 0\ raw_has_prod_mult max_def) qed qed lemma convergent_prod_mult: assumes f: "convergent_prod f" and g: "convergent_prod g" shows "convergent_prod (\n. f n * g n)" unfolding convergent_prod_def proof - obtain M p N q where p: "raw_has_prod f M p" and q: "raw_has_prod g N q" using convergent_prod_def f g by blast+ then obtain p' q' where p': "raw_has_prod f (max M N) p'" and q': "raw_has_prod g (max M N) q'" by (meson raw_has_prod_ignore_initial_segment max.cobounded1 max.cobounded2) then show "\M p. raw_has_prod (\n. f n * g n) M p" using raw_has_prod_mult by blast qed lemma prodinf_mult: "convergent_prod f \ convergent_prod g \ prodinf f * prodinf g = (\n. f n * g n)" by (intro has_prod_unique has_prod_mult convergent_prod_has_prod) end context fixes f :: "'i \ nat \ 'a::real_normed_field" and I :: "'i set" begin lemma has_prod_prod: "(\i. i \ I \ (f i) has_prod (x i)) \ (\n. \i\I. f i n) has_prod (\i\I. x i)" by (induct I rule: infinite_finite_induct) (auto intro!: has_prod_mult) lemma prodinf_prod: "(\i. i \ I \ convergent_prod (f i)) \ (\n. \i\I. f i n) = (\i\I. \n. f i n)" using has_prod_unique[OF has_prod_prod, OF convergent_prod_has_prod] by simp lemma convergent_prod_prod: "(\i. i \ I \ convergent_prod (f i)) \ convergent_prod (\n. \i\I. f i n)" using convergent_prod_has_prod_iff has_prod_prod prodinf_prod by force end subsection\<^marker>\tag unimportant\ \Infinite summability on real normed fields\ context fixes f :: "nat \ 'a::real_normed_field" begin lemma raw_has_prod_Suc_iff: "raw_has_prod f M (a * f M) \ raw_has_prod (\n. f (Suc n)) M a \ f M \ 0" proof - have "raw_has_prod f M (a * f M) \ (\i. \j\Suc i. f (j+M)) \ a * f M \ a * f M \ 0" by (subst filterlim_sequentially_Suc) (simp add: raw_has_prod_def) also have "\ \ (\i. (\j\i. f (Suc j + M)) * f M) \ a * f M \ a * f M \ 0" by (simp add: ac_simps atMost_Suc_eq_insert_0 image_Suc_atMost prod.atLeast1_atMost_eq lessThan_Suc_atMost del: prod.cl_ivl_Suc) also have "\ \ raw_has_prod (\n. f (Suc n)) M a \ f M \ 0" proof safe assume tends: "(\i. (\j\i. f (Suc j + M)) * f M) \ a * f M" and 0: "a * f M \ 0" with tendsto_divide[OF tends tendsto_const, of "f M"] show "raw_has_prod (\n. f (Suc n)) M a" by (simp add: raw_has_prod_def) qed (auto intro: tendsto_mult_right simp: raw_has_prod_def) finally show ?thesis . qed lemma has_prod_Suc_iff: assumes "f 0 \ 0" shows "(\n. f (Suc n)) has_prod a \ f has_prod (a * f 0)" proof (cases "a = 0") case True then show ?thesis proof (simp add: has_prod_def, safe) fix i x assume "f (Suc i) = 0" and "raw_has_prod (\n. f (Suc n)) (Suc i) x" then obtain y where "raw_has_prod f (Suc (Suc i)) y" by (metis (no_types) raw_has_prod_eq_0 Suc_n_not_le_n raw_has_prod_Suc_iff raw_has_prod_ignore_initial_segment raw_has_prod_nonzero linear) then show "\i. f i = 0 \ Ex (raw_has_prod f (Suc i))" using \f (Suc i) = 0\ by blast next fix i x assume "f i = 0" and x: "raw_has_prod f (Suc i) x" then obtain j where j: "i = Suc j" by (metis assms not0_implies_Suc) moreover have "\ y. raw_has_prod (\n. f (Suc n)) i y" using x by (auto simp: raw_has_prod_def) then show "\i. f (Suc i) = 0 \ Ex (raw_has_prod (\n. f (Suc n)) (Suc i))" using \f i = 0\ j by blast qed next case False then show ?thesis by (auto simp: has_prod_def raw_has_prod_Suc_iff assms) qed -lemma convergent_prod_Suc_iff: +lemma convergent_prod_Suc_iff [simp]: shows "convergent_prod (\n. f (Suc n)) = convergent_prod f" proof assume "convergent_prod f" then obtain M L where M_nz:"\n\M. f n \ 0" and M_L:"(\n. \i\n. f (i + M)) \ L" and "L \ 0" unfolding convergent_prod_altdef by auto have "(\n. \i\n. f (Suc (i + M))) \ L / f M" proof - have "(\n. \i\{0..Suc n}. f (i + M)) \ L" using M_L apply (subst (asm) filterlim_sequentially_Suc[symmetric]) using atLeast0AtMost by auto then have "(\n. f M * (\i\{0..n}. f (Suc (i + M)))) \ L" apply (subst (asm) prod.atLeast0_atMost_Suc_shift) by simp then have "(\n. (\i\{0..n}. f (Suc (i + M)))) \ L/f M" apply (drule_tac tendsto_divide) using M_nz[rule_format,of M,simplified] by auto then show ?thesis unfolding atLeast0AtMost . qed then show "convergent_prod (\n. f (Suc n))" unfolding convergent_prod_altdef apply (rule_tac exI[where x=M]) apply (rule_tac exI[where x="L/f M"]) using M_nz \L\0\ by auto next assume "convergent_prod (\n. f (Suc n))" then obtain M where "\L. (\n\M. f (Suc n) \ 0) \ (\n. \i\n. f (Suc (i + M))) \ L \ L \ 0" unfolding convergent_prod_altdef by auto then show "convergent_prod f" unfolding convergent_prod_altdef apply (rule_tac exI[where x="Suc M"]) using Suc_le_D by auto qed lemma raw_has_prod_inverse: assumes "raw_has_prod f M a" shows "raw_has_prod (\n. inverse (f n)) M (inverse a)" using assms unfolding raw_has_prod_def by (auto dest: tendsto_inverse simp: prod_inversef [symmetric]) lemma has_prod_inverse: assumes "f has_prod a" shows "(\n. inverse (f n)) has_prod (inverse a)" using assms raw_has_prod_inverse unfolding has_prod_def by auto lemma convergent_prod_inverse: assumes "convergent_prod f" shows "convergent_prod (\n. inverse (f n))" using assms unfolding convergent_prod_def by (blast intro: raw_has_prod_inverse elim: ) end context fixes f :: "nat \ 'a::real_normed_field" begin lemma raw_has_prod_Suc_iff': "raw_has_prod f M a \ raw_has_prod (\n. f (Suc n)) M (a / f M) \ f M \ 0" by (metis raw_has_prod_eq_0 add.commute add.left_neutral raw_has_prod_Suc_iff raw_has_prod_nonzero le_add1 nonzero_mult_div_cancel_right times_divide_eq_left) lemma has_prod_divide: "f has_prod a \ g has_prod b \ (\n. f n / g n) has_prod (a / b)" unfolding divide_inverse by (intro has_prod_inverse has_prod_mult) lemma convergent_prod_divide: assumes f: "convergent_prod f" and g: "convergent_prod g" shows "convergent_prod (\n. f n / g n)" using f g has_prod_divide has_prod_iff by blast lemma prodinf_divide: "convergent_prod f \ convergent_prod g \ prodinf f / prodinf g = (\n. f n / g n)" by (intro has_prod_unique has_prod_divide convergent_prod_has_prod) lemma prodinf_inverse: "convergent_prod f \ (\n. inverse (f n)) = inverse (\n. f n)" by (intro has_prod_unique [symmetric] has_prod_inverse convergent_prod_has_prod) lemma has_prod_Suc_imp: assumes "(\n. f (Suc n)) has_prod a" shows "f has_prod (a * f 0)" proof - have "f has_prod (a * f 0)" when "raw_has_prod (\n. f (Suc n)) 0 a" apply (cases "f 0=0") using that unfolding has_prod_def raw_has_prod_Suc by (auto simp add: raw_has_prod_Suc_iff) moreover have "f has_prod (a * f 0)" when "(\i q. a = 0 \ f (Suc i) = 0 \ raw_has_prod (\n. f (Suc n)) (Suc i) q)" proof - from that obtain i q where "a = 0" "f (Suc i) = 0" "raw_has_prod (\n. f (Suc n)) (Suc i) q" by auto then show ?thesis unfolding has_prod_def by (auto intro!:exI[where x="Suc i"] simp:raw_has_prod_Suc) qed ultimately show "f has_prod (a * f 0)" using assms unfolding has_prod_def by auto qed lemma has_prod_iff_shift: assumes "\i. i < n \ f i \ 0" shows "(\i. f (i + n)) has_prod a \ f has_prod (a * (\ii. f (Suc i + n)) has_prod a \ (\i. f (i + n)) has_prod (a * f n)" by (subst has_prod_Suc_iff) auto with Suc show ?case by (simp add: ac_simps) qed corollary\<^marker>\tag unimportant\ has_prod_iff_shift': assumes "\i. i < n \ f i \ 0" shows "(\i. f (i + n)) has_prod (a / (\i f has_prod a" by (simp add: assms has_prod_iff_shift) lemma has_prod_one_iff_shift: assumes "\i. i < n \ f i = 1" shows "(\i. f (i+n)) has_prod a \ (\i. f i) has_prod a" by (simp add: assms has_prod_iff_shift) lemma convergent_prod_iff_shift [simp]: shows "convergent_prod (\i. f (i + n)) \ convergent_prod f" apply safe using convergent_prod_offset apply blast using convergent_prod_ignore_initial_segment convergent_prod_def by blast lemma has_prod_split_initial_segment: assumes "f has_prod a" "\i. i < n \ f i \ 0" shows "(\i. f (i + n)) has_prod (a / (\ii. i < n \ f i \ 0" shows "(\i. f (i + n)) = (\i. f i) / (\ii. i < n \ f i \ 0" shows "prodinf f = (\i. f (i + n)) * (\i 0" shows "(\n. f (Suc n)) = prodinf f / f 0" using prodinf_split_initial_segment[of 1] assms by simp end context fixes f :: "nat \ 'a::real_normed_field" begin -lemma convergent_prod_inverse_iff: "convergent_prod (\n. inverse (f n)) \ convergent_prod f" +lemma convergent_prod_inverse_iff [simp]: "convergent_prod (\n. inverse (f n)) \ convergent_prod f" by (auto dest: convergent_prod_inverse) -lemma convergent_prod_const_iff: +lemma convergent_prod_const_iff [simp]: fixes c :: "'a :: {real_normed_field}" shows "convergent_prod (\_. c) \ c = 1" proof assume "convergent_prod (\_. c)" then show "c = 1" using convergent_prod_imp_LIMSEQ LIMSEQ_unique by blast next assume "c = 1" then show "convergent_prod (\_. c)" by auto qed lemma has_prod_power: "f has_prod a \ (\i. f i ^ n) has_prod (a ^ n)" by (induction n) (auto simp: has_prod_mult) lemma convergent_prod_power: "convergent_prod f \ convergent_prod (\i. f i ^ n)" by (induction n) (auto simp: convergent_prod_mult) lemma prodinf_power: "convergent_prod f \ prodinf (\i. f i ^ n) = prodinf f ^ n" by (metis has_prod_unique convergent_prod_imp_has_prod has_prod_power) end subsection\Exponentials and logarithms\ context fixes f :: "nat \ 'a::{real_normed_field,banach}" begin lemma sums_imp_has_prod_exp: assumes "f sums s" shows "raw_has_prod (\i. exp (f i)) 0 (exp s)" using assms continuous_on_exp [of UNIV "\x::'a. x"] using continuous_on_tendsto_compose [of UNIV exp "(\n. sum f {..n})" s] by (simp add: prod_defs sums_def_le exp_sum) lemma convergent_prod_exp: assumes "summable f" shows "convergent_prod (\i. exp (f i))" using sums_imp_has_prod_exp assms unfolding summable_def convergent_prod_def by blast lemma prodinf_exp: assumes "summable f" shows "prodinf (\i. exp (f i)) = exp (suminf f)" proof - have "f sums suminf f" using assms by blast then have "(\i. exp (f i)) has_prod exp (suminf f)" by (simp add: has_prod_def sums_imp_has_prod_exp) then show ?thesis by (rule has_prod_unique [symmetric]) qed end theorem convergent_prod_iff_summable_real: fixes a :: "nat \ real" assumes "\n. a n > 0" shows "convergent_prod (\k. 1 + a k) \ summable a" (is "?lhs = ?rhs") proof assume ?lhs then obtain p where "raw_has_prod (\k. 1 + a k) 0 p" by (metis assms add_less_same_cancel2 convergent_prod_offset_0 not_one_less_zero) then have to_p: "(\n. \k\n. 1 + a k) \ p" by (auto simp: raw_has_prod_def) moreover have le: "(\k\n. a k) \ (\k\n. 1 + a k)" for n by (rule sum_le_prod) (use assms less_le in force) have "(\k\n. 1 + a k) \ p" for n proof (rule incseq_le [OF _ to_p]) show "incseq (\n. \k\n. 1 + a k)" using assms by (auto simp: mono_def order.strict_implies_order intro!: prod_mono2) qed with le have "(\k\n. a k) \ p" for n by (metis order_trans) with assms bounded_imp_summable show ?rhs by (metis not_less order.asym) next assume R: ?rhs have "(\k\n. 1 + a k) \ exp (suminf a)" for n proof - have "(\k\n. 1 + a k) \ exp (\k\n. a k)" for n by (rule prod_le_exp_sum) (use assms less_le in force) moreover have "exp (\k\n. a k) \ exp (suminf a)" for n unfolding exp_le_cancel_iff by (meson sum_le_suminf R assms finite_atMost less_eq_real_def) ultimately show ?thesis by (meson order_trans) qed then obtain L where L: "(\n. \k\n. 1 + a k) \ L" by (metis assms bounded_imp_convergent_prod convergent_prod_iff_nz_lim le_add_same_cancel1 le_add_same_cancel2 less_le not_le zero_le_one) moreover have "L \ 0" proof assume "L = 0" with L have "(\n. \k\n. 1 + a k) \ 0" by simp moreover have "(\k\n. 1 + a k) > 1" for n by (simp add: assms less_1_prod) ultimately show False by (meson Lim_bounded2 not_one_le_zero less_imp_le) qed ultimately show ?lhs using assms convergent_prod_iff_nz_lim by (metis add_less_same_cancel1 less_le not_le zero_less_one) qed lemma exp_suminf_prodinf_real: fixes f :: "nat \ real" assumes ge0:"\n. f n \ 0" and ac: "abs_convergent_prod (\n. exp (f n))" shows "prodinf (\i. exp (f i)) = exp (suminf f)" proof - have "summable f" using ac unfolding abs_convergent_prod_conv_summable proof (elim summable_comparison_test') fix n have "\f n\ = f n" by (simp add: ge0) also have "\ \ exp (f n) - 1" by (metis diff_diff_add exp_ge_add_one_self ge_iff_diff_ge_0) finally show "norm (f n) \ norm (exp (f n) - 1)" by simp qed then show ?thesis by (simp add: prodinf_exp) qed lemma has_prod_imp_sums_ln_real: fixes f :: "nat \ real" assumes "raw_has_prod f 0 p" and 0: "\x. f x > 0" shows "(\i. ln (f i)) sums (ln p)" proof - have "p > 0" using assms unfolding prod_defs by (metis LIMSEQ_prod_nonneg less_eq_real_def) then show ?thesis using assms continuous_on_ln [of "{0<..}" "\x. x"] using continuous_on_tendsto_compose [of "{0<..}" ln "(\n. prod f {..n})" p] by (auto simp: prod_defs sums_def_le ln_prod order_tendstoD) qed lemma summable_ln_real: fixes f :: "nat \ real" assumes f: "convergent_prod f" and 0: "\x. f x > 0" shows "summable (\i. ln (f i))" proof - obtain M p where "raw_has_prod f M p" using f convergent_prod_def by blast then consider i where "i real" assumes f: "convergent_prod f" and 0: "\x. f x > 0" shows "suminf (\i. ln (f i)) = ln (prodinf f)" proof - have "f has_prod prodinf f" by (simp add: f has_prod_iff) then have "raw_has_prod f 0 (prodinf f)" by (metis "0" has_prod_def less_irrefl) then have "(\i. ln (f i)) sums ln (prodinf f)" using "0" has_prod_imp_sums_ln_real by blast then show ?thesis by (rule sums_unique [symmetric]) qed lemma prodinf_exp_real: fixes f :: "nat \ real" assumes f: "convergent_prod f" and 0: "\x. f x > 0" shows "prodinf f = exp (suminf (\i. ln (f i)))" by (simp add: "0" f less_0_prodinf suminf_ln_real) theorem Ln_prodinf_complex: fixes z :: "nat \ complex" assumes z: "\j. z j \ 0" and \: "\ \ 0" shows "((\n. \j\n. z j) \ \) \ (\k. (\n. (\j\n. Ln (z j))) \ Ln \ + of_int k * (of_real(2*pi) * \))" (is "?lhs = ?rhs") proof assume L: ?lhs have pnz: "(\j\n. z j) \ 0" for n using z by auto define \ where "\ \ Arg \ + 2*pi" then have "\ > pi" using Arg_def mpi_less_Im_Ln by fastforce have \_eq: "\ = cmod \ * exp (\ * \)" using Arg_def Arg_eq \ unfolding \_def by (simp add: algebra_simps exp_add) define \ where "\ \ \n. THE t. is_Arg (\j\n. z j) t \ t \ {\-pi<..\+pi}" have uniq: "\!s. is_Arg (\j\n. z j) s \ s \ {\-pi<..\+pi}" for n using Argument_exists_unique [OF pnz] by metis have \: "is_Arg (\j\n. z j) (\ n)" and \_interval: "\ n \ {\-pi<..\+pi}" for n unfolding \_def using theI' [OF uniq] by metis+ have \_pos: "\j. \ j > 0" using \_interval \\ > pi\ by simp (meson diff_gt_0_iff_gt less_trans) have "(\j\n. z j) = cmod (\j\n. z j) * exp (\ * \ n)" for n using \ by (auto simp: is_Arg_def) then have eq: "(\n. \j\n. z j) = (\n. cmod (\j\n. z j) * exp (\ * \ n))" by simp then have "(\n. (cmod (\j\n. z j)) * exp (\ * (\ n))) \ \" using L by force then obtain k where k: "(\j. \ j - of_int (k j) * (2 * pi)) \ \" using L by (subst (asm) \_eq) (auto simp add: eq z \ polar_convergence) moreover have "\\<^sub>F n in sequentially. k n = 0" proof - have *: "kj = 0" if "dist (vj - real_of_int kj * 2) V < 1" "vj \ {V - 1<..V + 1}" for kj vj V using that by (auto simp: dist_norm) have "\\<^sub>F j in sequentially. dist (\ j - of_int (k j) * (2 * pi)) \ < pi" using tendstoD [OF k] pi_gt_zero by blast then show ?thesis proof (rule eventually_mono) fix j assume d: "dist (\ j - real_of_int (k j) * (2 * pi)) \ < pi" show "k j = 0" by (rule * [of "\ j/pi" _ "\/pi"]) (use \_interval [of j] d in \simp_all add: divide_simps dist_norm\) qed qed ultimately have \to\: "\ \ \" apply (simp only: tendsto_def) apply (erule all_forward imp_forward asm_rl)+ apply (drule (1) eventually_conj) apply (auto elim: eventually_mono) done then have to0: "(\n. \\ (Suc n) - \ n\) \ 0" by (metis (full_types) diff_self filterlim_sequentially_Suc tendsto_diff tendsto_rabs_zero) have "\k. Im (\j\n. Ln (z j)) - of_int k * (2*pi) = \ n" for n proof (rule is_Arg_exp_diff_2pi) show "is_Arg (exp (\j\n. Ln (z j))) (\ n)" using pnz \ by (simp add: is_Arg_def exp_sum prod_norm) qed then have "\k. (\j\n. Im (Ln (z j))) = \ n + of_int k * (2*pi)" for n by (simp add: algebra_simps) then obtain k where k: "\n. (\j\n. Im (Ln (z j))) = \ n + of_int (k n) * (2*pi)" by metis obtain K where "\\<^sub>F n in sequentially. k n = K" proof - have k_le: "(2*pi) * \k (Suc n) - k n\ \ \\ (Suc n) - \ n\ + \Im (Ln (z (Suc n)))\" for n proof - have "(\j\Suc n. Im (Ln (z j))) - (\j\n. Im (Ln (z j))) = Im (Ln (z (Suc n)))" by simp then show ?thesis using k [of "Suc n"] k [of n] by (auto simp: abs_if algebra_simps) qed have "z \ 1" using L \ convergent_prod_iff_nz_lim z by (blast intro: convergent_prod_imp_LIMSEQ) with z have "(\n. Ln (z n)) \ Ln 1" using isCont_tendsto_compose [OF continuous_at_Ln] nonpos_Reals_one_I by blast then have "(\n. Ln (z n)) \ 0" by simp then have "(\n. \Im (Ln (z (Suc n)))\) \ 0" by (metis LIMSEQ_unique \z \ 1\ continuous_at_Ln filterlim_sequentially_Suc isCont_tendsto_compose nonpos_Reals_one_I tendsto_Im tendsto_rabs_zero_iff zero_complex.simps(2)) then have "\\<^sub>F n in sequentially. \Im (Ln (z (Suc n)))\ < 1" by (simp add: order_tendsto_iff) moreover have "\\<^sub>F n in sequentially. \\ (Suc n) - \ n\ < 1" using to0 by (simp add: order_tendsto_iff) ultimately have "\\<^sub>F n in sequentially. (2*pi) * \k (Suc n) - k n\ < 1 + 1" proof (rule eventually_elim2) fix n assume "\Im (Ln (z (Suc n)))\ < 1" and "\\ (Suc n) - \ n\ < 1" with k_le [of n] show "2 * pi * real_of_int \k (Suc n) - k n\ < 1 + 1" by linarith qed then have "\\<^sub>F n in sequentially. real_of_int\k (Suc n) - k n\ < 1" proof (rule eventually_mono) fix n :: "nat" assume "2 * pi * \k (Suc n) - k n\ < 1 + 1" then have "\k (Suc n) - k n\ < 2 / (2*pi)" by (simp add: field_simps) also have "... < 1" using pi_ge_two by auto finally show "real_of_int \k (Suc n) - k n\ < 1" . qed then obtain N where N: "\n. n\N \ \k (Suc n) - k n\ = 0" using eventually_sequentially less_irrefl of_int_abs by fastforce have "k (N+i) = k N" for i proof (induction i) case (Suc i) with N [of "N+i"] show ?case by auto qed simp then have "\n. n\N \ k n = k N" using le_Suc_ex by auto then show ?thesis by (force simp add: eventually_sequentially intro: that) qed with \to\ have "(\n. (\j\n. Im (Ln (z j)))) \ \ + of_int K * (2*pi)" by (simp add: k tendsto_add tendsto_mult tendsto_eventually) moreover have "(\n. (\k\n. Re (Ln (z k)))) \ Re (Ln \)" using assms continuous_imp_tendsto [OF isCont_ln tendsto_norm [OF L]] by (simp add: o_def flip: prod_norm ln_prod) ultimately show ?rhs by (rule_tac x="K+1" in exI) (auto simp: tendsto_complex_iff \_def Arg_def assms algebra_simps) next assume ?rhs then obtain r where r: "(\n. (\k\n. Ln (z k))) \ Ln \ + of_int r * (of_real(2*pi) * \)" .. have "(\n. exp (\k\n. Ln (z k))) \ \" using assms continuous_imp_tendsto [OF isCont_exp r] exp_integer_2pi [of r] by (simp add: o_def exp_add algebra_simps) moreover have "exp (\k\n. Ln (z k)) = (\k\n. z k)" for n by (simp add: exp_sum add_eq_0_iff assms) ultimately show ?lhs by auto qed text\Prop 17.2 of Bak and Newman, Complex Analysis, p.242\ proposition convergent_prod_iff_summable_complex: fixes z :: "nat \ complex" assumes "\k. z k \ 0" shows "convergent_prod (\k. z k) \ summable (\k. Ln (z k))" (is "?lhs = ?rhs") proof assume ?lhs then obtain p where p: "(\n. \k\n. z k) \ p" and "p \ 0" using convergent_prod_LIMSEQ prodinf_nonzero add_eq_0_iff assms by fastforce then show ?rhs using Ln_prodinf_complex assms by (auto simp: prodinf_nonzero summable_def sums_def_le) next assume R: ?rhs have "(\k\n. z k) = exp (\k\n. Ln (z k))" for n by (simp add: exp_sum add_eq_0_iff assms) then have "(\n. \k\n. z k) \ exp (suminf (\k. Ln (z k)))" using continuous_imp_tendsto [OF isCont_exp summable_LIMSEQ' [OF R]] by (simp add: o_def) then show ?lhs by (subst convergent_prod_iff_convergent) (auto simp: convergent_def tendsto_Lim assms add_eq_0_iff) qed text\Prop 17.3 of Bak and Newman, Complex Analysis\ proposition summable_imp_convergent_prod_complex: fixes z :: "nat \ complex" assumes z: "summable (\k. norm (z k))" and non0: "\k. z k \ -1" shows "convergent_prod (\k. 1 + z k)" proof - note if_cong [cong] power_Suc [simp del] obtain N where N: "\k. k\N \ norm (z k) < 1/2" using summable_LIMSEQ_zero [OF z] by (metis diff_zero dist_norm half_gt_zero_iff less_numeral_extra(1) lim_sequentially tendsto_norm_zero_iff) have "norm (Ln (1 + z k)) \ 2 * norm (z k)" if "k \ N" for k proof (cases "z k = 0") case False let ?f = "\i. cmod ((- 1) ^ i * z k ^ i / of_nat (Suc i))" have normf: "norm (?f n) \ (1 / 2) ^ n" for n proof - have "norm (?f n) = cmod (z k) ^ n / cmod (1 + of_nat n)" by (auto simp: norm_divide norm_mult norm_power) also have "\ \ cmod (z k) ^ n" by (auto simp: field_split_simps mult_le_cancel_left1 in_Reals_norm) also have "\ \ (1 / 2) ^ n" using N [OF that] by (simp add: power_mono) finally show "norm (?f n) \ (1 / 2) ^ n" . qed have summablef: "summable ?f" by (intro normf summable_comparison_test' [OF summable_geometric [of "1/2"]]) auto have "(\n. (- 1) ^ Suc n / of_nat n * z k ^ n) sums Ln (1 + z k)" using Ln_series [of "z k"] N that by fastforce then have *: "(\i. z k * (((- 1) ^ i * z k ^ i) / (Suc i))) sums Ln (1 + z k)" using sums_split_initial_segment [where n= 1] by (force simp: power_Suc mult_ac) then have "norm (Ln (1 + z k)) = norm (suminf (\i. z k * (((- 1) ^ i * z k ^ i) / (Suc i))))" using sums_unique by force also have "\ = norm (z k * suminf (\i. ((- 1) ^ i * z k ^ i) / (Suc i)))" apply (subst suminf_mult) using * False by (auto simp: sums_summable intro: summable_mult_D [of "z k"]) also have "\ = norm (z k) * norm (suminf (\i. ((- 1) ^ i * z k ^ i) / (Suc i)))" by (simp add: norm_mult) also have "\ \ norm (z k) * suminf (\i. norm (((- 1) ^ i * z k ^ i) / (Suc i)))" by (intro mult_left_mono summable_norm summablef) auto also have "\ \ norm (z k) * suminf (\i. (1/2) ^ i)" by (intro mult_left_mono suminf_le) (use summable_geometric [of "1/2"] summablef normf in auto) also have "\ \ norm (z k) * 2" using suminf_geometric [of "1/2::real"] by simp finally show ?thesis by (simp add: mult_ac) qed simp then have "summable (\k. Ln (1 + z k))" by (metis summable_comparison_test summable_mult z) with non0 show ?thesis by (simp add: add_eq_0_iff convergent_prod_iff_summable_complex) qed lemma summable_Ln_complex: fixes z :: "nat \ complex" assumes "convergent_prod z" "\k. z k \ 0" shows "summable (\k. Ln (z k))" using convergent_prod_def assms convergent_prod_iff_summable_complex by blast subsection\<^marker>\tag unimportant\ \Embeddings from the reals into some complete real normed field\ lemma tendsto_eq_of_real_lim: assumes "(\n. of_real (f n) :: 'a::{complete_space,real_normed_field}) \ q" shows "q = of_real (lim f)" proof - have "convergent (\n. of_real (f n) :: 'a)" using assms convergent_def by blast then have "convergent f" unfolding convergent_def by (simp add: convergent_eq_Cauchy Cauchy_def) then show ?thesis by (metis LIMSEQ_unique assms convergentD sequentially_bot tendsto_Lim tendsto_of_real) qed lemma tendsto_eq_of_real: assumes "(\n. of_real (f n) :: 'a::{complete_space,real_normed_field}) \ q" obtains r where "q = of_real r" using tendsto_eq_of_real_lim assms by blast -lemma has_prod_of_real_iff: +lemma has_prod_of_real_iff [simp]: "(\n. of_real (f n) :: 'a::{complete_space,real_normed_field}) has_prod of_real c \ f has_prod c" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (auto simp: prod_defs LIMSEQ_prod_0 tendsto_of_real_iff simp flip: of_real_prod) using tendsto_eq_of_real by (metis of_real_0 tendsto_of_real_iff) next assume ?rhs with tendsto_of_real_iff show ?lhs by (fastforce simp: prod_defs simp flip: of_real_prod) qed end diff --git a/src/HOL/Analysis/Lindelof_Spaces.thy b/src/HOL/Analysis/Lindelof_Spaces.thy --- a/src/HOL/Analysis/Lindelof_Spaces.thy +++ b/src/HOL/Analysis/Lindelof_Spaces.thy @@ -1,274 +1,274 @@ section\Lindel\"of spaces\ theory Lindelof_Spaces imports T1_Spaces begin definition Lindelof_space where "Lindelof_space X \ \\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. countable \ \ \ \ \ \ \\ = topspace X)" lemma Lindelof_spaceD: "\Lindelof_space X; \U. U \ \ \ openin X U; \\ = topspace X\ \ \\. countable \ \ \ \ \ \ \\ = topspace X" by (auto simp: Lindelof_space_def) lemma Lindelof_space_alt: "Lindelof_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. countable \ \ \ \ \ \ topspace X \ \\))" unfolding Lindelof_space_def using openin_subset by fastforce lemma compact_imp_Lindelof_space: "compact_space X \ Lindelof_space X" unfolding Lindelof_space_def compact_space by (meson uncountable_infinite) lemma Lindelof_space_topspace_empty: "topspace X = {} \ Lindelof_space X" using compact_imp_Lindelof_space compact_space_topspace_empty by blast lemma Lindelof_space_Union: assumes \: "countable \" and lin: "\U. U \ \ \ Lindelof_space (subtopology X U)" shows "Lindelof_space (subtopology X (\\))" proof - have "\\. countable \ \ \ \ \ \ \\ \ \\ = topspace X \ \\" if \: "\ \ Collect (openin X)" and UF: "\\ \ \\ = topspace X \ \\" for \ proof - have "\U. \U \ \; U \ \\ = topspace X \ U\ \ \\. countable \ \ \ \ \ \ U \ \\ = topspace X \ U" using lin \ unfolding Lindelof_space_def openin_subtopology_alt Ball_def subset_iff [symmetric] by (simp add: all_subset_image imp_conjL ex_countable_subset_image) then obtain g where g: "\U. \U \ \; U \ \\ = topspace X \ U\ \ countable (g U) \ (g U) \ \ \ U \ \(g U) = topspace X \ U" by metis show ?thesis proof (intro exI conjI) show "countable (\(g ` \))" using Int_commute UF g by (fastforce intro: countable_UN [OF \]) show "\(g ` \) \ \" using g UF by blast show "\\ \ \(\(g ` \)) = topspace X \ \\" proof show "\\ \ \(\(g ` \)) \ topspace X \ \\" using g UF by blast show "topspace X \ \\ \ \\ \ \(\(g ` \))" proof clarsimp show "\y\\. \W\g y. x \ W" if "x \ topspace X" "x \ V" "V \ \" for x V proof - have "V \ \\ = topspace X \ V" using UF \V \ \\ by blast with that g [OF \V \ \\] show ?thesis by blast qed qed qed qed qed then show ?thesis unfolding Lindelof_space_def openin_subtopology_alt Ball_def subset_iff [symmetric] by (simp add: all_subset_image imp_conjL ex_countable_subset_image) qed lemma countable_imp_Lindelof_space: assumes "countable(topspace X)" shows "Lindelof_space X" proof - have "Lindelof_space (subtopology X (\x \ topspace X. {x}))" proof (rule Lindelof_space_Union) show "countable ((\x. {x}) ` topspace X)" using assms by blast show "Lindelof_space (subtopology X U)" if "U \ (\x. {x}) ` topspace X" for U proof - have "compactin X U" using that by force then show ?thesis by (meson compact_imp_Lindelof_space compact_space_subtopology) qed qed then show ?thesis by simp qed lemma Lindelof_space_subtopology: "Lindelof_space(subtopology X S) \ (\\. (\U \ \. openin X U) \ topspace X \ S \ \\ \ (\V. countable V \ V \ \ \ topspace X \ S \ \V))" proof - have *: "(S \ \\ = topspace X \ S) = (topspace X \ S \ \\)" if "\x. x \ \ \ openin X x" for \ by (blast dest: openin_subset [OF that]) moreover have "(\ \ \ \ S \ \\ = topspace X \ S) = (\ \ \ \ topspace X \ S \ \\)" if "\x. x \ \ \ openin X x" "topspace X \ S \ \\" "countable \" for \ \ using that * by blast ultimately show ?thesis unfolding Lindelof_space_def openin_subtopology_alt Ball_def apply (simp add: all_subset_image imp_conjL ex_countable_subset_image flip: subset_iff) apply (intro all_cong1 imp_cong ex_cong, auto) done qed lemma Lindelof_space_subtopology_subset: "S \ topspace X \ (Lindelof_space(subtopology X S) \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\V. countable V \ V \ \ \ S \ \V)))" by (metis Lindelof_space_subtopology topspace_subtopology topspace_subtopology_subset) lemma Lindelof_space_closedin_subtopology: assumes X: "Lindelof_space X" and clo: "closedin X S" shows "Lindelof_space (subtopology X S)" proof - have "S \ topspace X" by (simp add: clo closedin_subset) then show ?thesis proof (clarsimp simp add: Lindelof_space_subtopology_subset) show "\V. countable V \ V \ \ \ S \ \V" if "\U\\. openin X U" and "S \ \\" for \ proof - have "\\. countable \ \ \ \ insert (topspace X - S) \ \ \\ = topspace X" proof (rule Lindelof_spaceD [OF X, of "insert (topspace X - S) \"]) show "openin X U" if "U \ insert (topspace X - S) \" for U using that \\U\\. openin X U\ clo by blast show "\(insert (topspace X - S) \) = topspace X" apply auto apply (meson in_mono openin_closedin_eq that(1)) using UnionE \S \ \\\ by auto qed then obtain \ where "countable \" "\ \ insert (topspace X - S) \" "\\ = topspace X" by metis with \S \ topspace X\ show ?thesis by (rule_tac x="(\ - {topspace X - S})" in exI) auto qed qed qed lemma Lindelof_space_continuous_map_image: assumes X: "Lindelof_space X" and f: "continuous_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "Lindelof_space Y" proof - have "\\. countable \ \ \ \ \ \ \\ = topspace Y" if \: "\U. U \ \ \ openin Y U" and UU: "\\ = topspace Y" for \ proof - define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "\V. V \ \ \ openin X V" unfolding \_def using \ continuous_map f by fastforce moreover have "\\ = topspace X" unfolding \_def using UU fim by fastforce ultimately have "\\. countable \ \ \ \ \ \ \\ = topspace X" using X by (simp add: Lindelof_space_def) then obtain \ where "countable \" "\ \ \" and \: "(\U\\. {x \ topspace X. f x \ U}) = topspace X" by (metis (no_types, lifting) \_def countable_subset_image) moreover have "\\ = topspace Y" proof show "\\ \ topspace Y" using UU \ \\ \ \\ by fastforce have "y \ \\" if "y \ topspace Y" for y proof - obtain x where "x \ topspace X" "y = f x" using that fim by (metis \y \ topspace Y\ imageE) with \ show ?thesis by auto qed then show "topspace Y \ \\" by blast qed ultimately show ?thesis by blast qed then show ?thesis unfolding Lindelof_space_def by auto qed lemma Lindelof_space_quotient_map_image: "\quotient_map X Y q; Lindelof_space X\ \ Lindelof_space Y" by (meson Lindelof_space_continuous_map_image quotient_imp_continuous_map quotient_imp_surjective_map) lemma Lindelof_space_retraction_map_image: "\retraction_map X Y r; Lindelof_space X\ \ Lindelof_space Y" using Abstract_Topology.retraction_imp_quotient_map Lindelof_space_quotient_map_image by blast lemma locally_finite_cover_of_Lindelof_space: assumes X: "Lindelof_space X" and UU: "topspace X \ \\" and fin: "locally_finite_in X \" shows "countable \" proof - have UU_eq: "\\ = topspace X" by (meson UU fin locally_finite_in_def subset_antisym) obtain T where T: "\x. x \ topspace X \ openin X (T x) \ x \ T x \ finite {U \ \. U \ T x \ {}}" - using fin unfolding locally_finite_in_def by meson + using fin unfolding locally_finite_in_def by metis then obtain I where "countable I" "I \ topspace X" and I: "topspace X \ \(T ` I)" using X unfolding Lindelof_space_alt by (drule_tac x="image T (topspace X)" in spec) (auto simp: ex_countable_subset_image) show ?thesis proof (rule countable_subset) have "\i. i \ I \ countable {U \ \. U \ T i \ {}}" using T by (meson \I \ topspace X\ in_mono uncountable_infinite) then show "countable (insert {} (\i\I. {U \ \. U \ T i \ {}}))" by (simp add: \countable I\) qed (use UU_eq I in auto) qed lemma Lindelof_space_proper_map_preimage: assumes f: "proper_map X Y f" and Y: "Lindelof_space Y" shows "Lindelof_space X" proof (clarsimp simp: Lindelof_space_alt) show "\\. countable \ \ \ \ \ \ topspace X \ \\" if \: "\U\\. openin X U" and sub_UU: "topspace X \ \\" for \ proof - have "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" if "y \ topspace Y" for y proof (rule compactinD) show "compactin X {x \ topspace X. f x = y}" using f proper_map_def that by fastforce qed (use sub_UU \ in auto) then obtain \ where \: "\y. y \ topspace Y \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by meson define \ where "\ \ (\y. topspace Y - image f (topspace X - \(\ y))) ` topspace Y" have "\U \ \. openin Y U" using f \ \ unfolding \_def proper_map_def closed_map_def by (simp add: closedin_diff openin_Union openin_diff subset_iff) moreover have "topspace Y \ \\" using \ unfolding \_def by clarsimp fastforce ultimately have "\\. countable \ \ \ \ \ \ topspace Y \ \\" using Y by (simp add: Lindelof_space_alt) then obtain I where "countable I" "I \ topspace Y" and I: "topspace Y \ (\i\I. topspace Y - f ` (topspace X - \(\ i)))" unfolding \_def ex_countable_subset_image by metis show ?thesis proof (intro exI conjI) have "\i. i \ I \ countable (\ i)" by (meson \ \I \ topspace Y\ in_mono uncountable_infinite) with \countable I\ show "countable (\(\ ` I))" by auto show "\(\ ` I) \ \" using \ \I \ topspace Y\ by fastforce show "topspace X \ \(\(\ ` I))" proof show "x \ \ (\ (\ ` I))" if "x \ topspace X" for x proof - have "f x \ topspace Y" by (meson f image_subset_iff proper_map_imp_subset_topspace that) then show ?thesis using that I by auto qed qed qed qed qed lemma Lindelof_space_perfect_map_image: "\Lindelof_space X; perfect_map X Y f\ \ Lindelof_space Y" using Lindelof_space_quotient_map_image perfect_imp_quotient_map by blast lemma Lindelof_space_perfect_map_image_eq: "perfect_map X Y f \ Lindelof_space X \ Lindelof_space Y" using Lindelof_space_perfect_map_image Lindelof_space_proper_map_preimage perfect_map_def by blast end diff --git a/src/HOL/Series.thy b/src/HOL/Series.thy --- a/src/HOL/Series.thy +++ b/src/HOL/Series.thy @@ -1,1333 +1,1333 @@ (* Title : Series.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Converted to Isar and polished by lcp Converted to sum and polished yet more by TNN Additional contributions by Jeremy Avigad *) section \Infinite Series\ theory Series imports Limits Inequalities begin subsection \Definition of infinite summability\ definition sums :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a \ bool" (infixr "sums" 80) where "f sums s \ (\n. \i s" definition summable :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ bool" where "summable f \ (\s. f sums s)" definition suminf :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a" (binder "\" 10) where "suminf f = (THE s. f sums s)" text\Variants of the definition\ lemma sums_def': "f sums s \ (\n. \i = 0..n. f i) \ s" unfolding sums_def apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp only: lessThan_Suc_atMost atLeast0AtMost) done lemma sums_def_le: "f sums s \ (\n. \i\n. f i) \ s" by (simp add: sums_def' atMost_atLeast0) lemma bounded_imp_summable: fixes a :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology,linordered_comm_semiring_strict}" assumes 0: "\n. a n \ 0" and bounded: "\n. (\k\n. a k) \ B" shows "summable a" proof - have "bdd_above (range(\n. \k\n. a k))" by (meson bdd_aboveI2 bounded) moreover have "incseq (\n. \k\n. a k)" by (simp add: mono_def "0" sum_mono2) ultimately obtain s where "(\n. \k\n. a k) \ s" using LIMSEQ_incseq_SUP by blast then show ?thesis by (auto simp: sums_def_le summable_def) qed subsection \Infinite summability on topological monoids\ lemma sums_subst[trans]: "f = g \ g sums z \ f sums z" by simp lemma sums_cong: "(\n. f n = g n) \ f sums c \ g sums c" by (drule ext) simp lemma sums_summable: "f sums l \ summable f" by (simp add: sums_def summable_def, blast) lemma summable_iff_convergent: "summable f \ convergent (\n. \i convergent (\n. sum f {..n})" by (simp_all only: summable_iff_convergent convergent_def lessThan_Suc_atMost [symmetric] filterlim_sequentially_Suc[of "\n. sum f {..n. \in. 0) sums 0" unfolding sums_def by simp lemma summable_zero[simp, intro]: "summable (\n. 0)" by (rule sums_zero [THEN sums_summable]) lemma sums_group: "f sums s \ 0 < k \ (\n. sum f {n * k ..< n * k + k}) sums s" apply (simp only: sums_def sum.nat_group tendsto_def eventually_sequentially) apply (erule all_forward imp_forward exE| assumption)+ apply (rule_tac x="N" in exI) by (metis le_square mult.commute mult.left_neutral mult_le_cancel2 mult_le_mono) lemma suminf_cong: "(\n. f n = g n) \ suminf f = suminf g" by (rule arg_cong[of f g], rule ext) simp lemma summable_cong: fixes f g :: "nat \ 'a::real_normed_vector" assumes "eventually (\x. f x = g x) sequentially" shows "summable f = summable g" proof - from assms obtain N where N: "\n\N. f n = g n" by (auto simp: eventually_at_top_linorder) define C where "C = (\kn. sum f {.. {N.. {N.. {N..n. n \ N \ f n = 0" shows "f sums (\n\N. f n)" proof - have eq: "sum f {..n. f n = 0) \ (f sums 0)" by (metis (no_types) finite.emptyI sum.empty sums_finite) lemma summable_finite: "finite N \ (\n. n \ N \ f n = 0) \ summable f" by (rule sums_summable) (rule sums_finite) lemma sums_If_finite_set: "finite A \ (\r. if r \ A then f r else 0) sums (\r\A. f r)" using sums_finite[of A "(\r. if r \ A then f r else 0)"] by simp lemma summable_If_finite_set[simp, intro]: "finite A \ summable (\r. if r \ A then f r else 0)" by (rule sums_summable) (rule sums_If_finite_set) lemma sums_If_finite: "finite {r. P r} \ (\r. if P r then f r else 0) sums (\r | P r. f r)" using sums_If_finite_set[of "{r. P r}"] by simp lemma summable_If_finite[simp, intro]: "finite {r. P r} \ summable (\r. if P r then f r else 0)" by (rule sums_summable) (rule sums_If_finite) lemma sums_single: "(\r. if r = i then f r else 0) sums f i" using sums_If_finite[of "\r. r = i"] by simp lemma summable_single[simp, intro]: "summable (\r. if r = i then f r else 0)" by (rule sums_summable) (rule sums_single) context fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" begin lemma summable_sums[intro]: "summable f \ f sums (suminf f)" by (simp add: summable_def sums_def suminf_def) (metis convergent_LIMSEQ_iff convergent_def lim_def) lemma summable_LIMSEQ: "summable f \ (\n. \i suminf f" by (rule summable_sums [unfolded sums_def]) lemma summable_LIMSEQ': "summable f \ (\n. \i\n. f i) \ suminf f" using sums_def_le by blast lemma sums_unique: "f sums s \ s = suminf f" by (metis limI suminf_eq_lim sums_def) lemma sums_iff: "f sums x \ summable f \ suminf f = x" by (metis summable_sums sums_summable sums_unique) lemma summable_sums_iff: "summable f \ f sums suminf f" by (auto simp: sums_iff summable_sums) lemma sums_unique2: "f sums a \ f sums b \ a = b" for a b :: 'a by (simp add: sums_iff) lemma sums_Uniq: "\\<^sub>\\<^sub>1a. f sums a" for a b :: 'a by (simp add: sums_unique2 Uniq_def) lemma suminf_finite: assumes N: "finite N" and f: "\n. n \ N \ f n = 0" shows "suminf f = (\n\N. f n)" using sums_finite[OF assms, THEN sums_unique] by simp end lemma suminf_zero[simp]: "suminf (\n. 0::'a::{t2_space, comm_monoid_add}) = 0" by (rule sums_zero [THEN sums_unique, symmetric]) subsection \Infinite summability on ordered, topological monoids\ lemma sums_le: "(\n. f n \ g n) \ f sums s \ g sums t \ s \ t" for f g :: "nat \ 'a::{ordered_comm_monoid_add,linorder_topology}" by (rule LIMSEQ_le) (auto intro: sum_mono simp: sums_def) context fixes f :: "nat \ 'a::{ordered_comm_monoid_add,linorder_topology}" begin lemma suminf_le: "(\n. f n \ g n) \ summable f \ summable g \ suminf f \ suminf g" using sums_le by blast lemma sum_le_suminf: shows "summable f \ finite I \ (\n. n \- I \ 0 \ f n) \ sum f I \ suminf f" by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto lemma suminf_nonneg: "summable f \ (\n. 0 \ f n) \ 0 \ suminf f" using sum_le_suminf by force lemma suminf_le_const: "summable f \ (\n. sum f {.. x) \ suminf f \ x" by (metis LIMSEQ_le_const2 summable_LIMSEQ) lemma suminf_eq_zero_iff: assumes "summable f" and pos: "\n. 0 \ f n" shows "suminf f = 0 \ (\n. f n = 0)" proof assume "suminf f = 0" then have f: "(\n. \i 0" using summable_LIMSEQ[of f] assms by simp then have "\i. (\n\{i}. f n) \ 0" proof (rule LIMSEQ_le_const) show "\N. \n\N. (\n\{i}. f n) \ sum f {..n. f n = 0" by (auto intro!: antisym) qed (metis suminf_zero fun_eq_iff) lemma suminf_pos_iff: "summable f \ (\n. 0 \ f n) \ 0 < suminf f \ (\i. 0 < f i)" using sum_le_suminf[of "{}"] suminf_eq_zero_iff by (simp add: less_le) lemma suminf_pos2: assumes "summable f" "\n. 0 \ f n" "0 < f i" shows "0 < suminf f" proof - have "0 < (\n \ suminf f" using assms by (intro sum_le_suminf) auto finally show ?thesis . qed lemma suminf_pos: "summable f \ (\n. 0 < f n) \ 0 < suminf f" by (intro suminf_pos2[where i=0]) (auto intro: less_imp_le) end context fixes f :: "nat \ 'a::{ordered_cancel_comm_monoid_add,linorder_topology}" begin lemma sum_less_suminf2: "summable f \ (\m. m\n \ 0 \ f m) \ n \ i \ 0 < f i \ sum f {.. (\m. m\n \ 0 < f m) \ sum f {.. 'a::{ordered_comm_monoid_add,linorder_topology,conditionally_complete_linorder}" assumes pos[simp]: "\n. 0 \ f n" and le: "\n. (\i x" shows "summable f" unfolding summable_def sums_def [abs_def] proof (rule exI LIMSEQ_incseq_SUP)+ show "bdd_above (range (\n. sum f {..n. sum f {.. 'a::{canonically_ordered_monoid_add,linorder_topology,complete_linorder}" by (intro summableI_nonneg_bounded[where x=top] zero_le top_greatest) lemma suminf_eq_SUP_real: assumes X: "summable X" "\i. 0 \ X i" shows "suminf X = (SUP i. \nInfinite summability on topological monoids\ context fixes f g :: "nat \ 'a::{t2_space,topological_comm_monoid_add}" begin lemma sums_Suc: assumes "(\n. f (Suc n)) sums l" shows "f sums (l + f 0)" proof - have "(\n. (\i l + f 0" using assms by (auto intro!: tendsto_add simp: sums_def) moreover have "(\ii g sums b \ (\n. f n + g n) sums (a + b)" unfolding sums_def by (simp add: sum.distrib tendsto_add) lemma summable_add: "summable f \ summable g \ summable (\n. f n + g n)" unfolding summable_def by (auto intro: sums_add) lemma suminf_add: "summable f \ summable g \ suminf f + suminf g = (\n. f n + g n)" by (intro sums_unique sums_add summable_sums) end context fixes f :: "'i \ nat \ 'a::{t2_space,topological_comm_monoid_add}" and I :: "'i set" begin lemma sums_sum: "(\i. i \ I \ (f i) sums (x i)) \ (\n. \i\I. f i n) sums (\i\I. x i)" by (induct I rule: infinite_finite_induct) (auto intro!: sums_add) lemma suminf_sum: "(\i. i \ I \ summable (f i)) \ (\n. \i\I. f i n) = (\i\I. \n. f i n)" using sums_unique[OF sums_sum, OF summable_sums] by simp lemma summable_sum: "(\i. i \ I \ summable (f i)) \ summable (\n. \i\I. f i n)" using sums_summable[OF sums_sum[OF summable_sums]] . end lemma sums_If_finite_set': fixes f g :: "nat \ 'a::{t2_space,topological_ab_group_add}" assumes "g sums S" and "finite A" and "S' = S + (\n\A. f n - g n)" shows "(\n. if n \ A then f n else g n) sums S'" proof - have "(\n. g n + (if n \ A then f n - g n else 0)) sums (S + (\n\A. f n - g n))" by (intro sums_add assms sums_If_finite_set) also have "(\n. g n + (if n \ A then f n - g n else 0)) = (\n. if n \ A then f n else g n)" by (simp add: fun_eq_iff) finally show ?thesis using assms by simp qed subsection \Infinite summability on real normed vector spaces\ context fixes f :: "nat \ 'a::real_normed_vector" begin lemma sums_Suc_iff: "(\n. f (Suc n)) sums s \ f sums (s + f 0)" proof - have "f sums (s + f 0) \ (\i. \j s + f 0" by (subst filterlim_sequentially_Suc) (simp add: sums_def) also have "\ \ (\i. (\j s + f 0" by (simp add: ac_simps lessThan_Suc_eq_insert_0 image_Suc_lessThan sum.atLeast1_atMost_eq) also have "\ \ (\n. f (Suc n)) sums s" proof assume "(\i. (\j s + f 0" with tendsto_add[OF this tendsto_const, of "- f 0"] show "(\i. f (Suc i)) sums s" by (simp add: sums_def) qed (auto intro: tendsto_add simp: sums_def) finally show ?thesis .. qed lemma summable_Suc_iff: "summable (\n. f (Suc n)) = summable f" proof assume "summable f" then have "f sums suminf f" by (rule summable_sums) then have "(\n. f (Suc n)) sums (suminf f - f 0)" by (simp add: sums_Suc_iff) then show "summable (\n. f (Suc n))" unfolding summable_def by blast qed (auto simp: sums_Suc_iff summable_def) lemma sums_Suc_imp: "f 0 = 0 \ (\n. f (Suc n)) sums s \ (\n. f n) sums s" using sums_Suc_iff by simp end context (* Separate contexts are necessary to allow general use of the results above, here. *) fixes f :: "nat \ 'a::real_normed_vector" begin lemma sums_diff: "f sums a \ g sums b \ (\n. f n - g n) sums (a - b)" unfolding sums_def by (simp add: sum_subtractf tendsto_diff) lemma summable_diff: "summable f \ summable g \ summable (\n. f n - g n)" unfolding summable_def by (auto intro: sums_diff) lemma suminf_diff: "summable f \ summable g \ suminf f - suminf g = (\n. f n - g n)" by (intro sums_unique sums_diff summable_sums) lemma sums_minus: "f sums a \ (\n. - f n) sums (- a)" unfolding sums_def by (simp add: sum_negf tendsto_minus) lemma summable_minus: "summable f \ summable (\n. - f n)" unfolding summable_def by (auto intro: sums_minus) lemma suminf_minus: "summable f \ (\n. - f n) = - (\n. f n)" by (intro sums_unique [symmetric] sums_minus summable_sums) lemma sums_iff_shift: "(\i. f (i + n)) sums s \ f sums (s + (\ii. f (Suc i + n)) sums s \ (\i. f (i + n)) sums (s + f n)" by (subst sums_Suc_iff) simp with Suc show ?case by (simp add: ac_simps) qed corollary sums_iff_shift': "(\i. f (i + n)) sums (s - (\i f sums s" by (simp add: sums_iff_shift) lemma sums_zero_iff_shift: assumes "\i. i < n \ f i = 0" shows "(\i. f (i+n)) sums s \ (\i. f i) sums s" by (simp add: assms sums_iff_shift) -lemma summable_iff_shift: "summable (\n. f (n + k)) \ summable f" +lemma summable_iff_shift [simp]: "summable (\n. f (n + k)) \ summable f" by (metis diff_add_cancel summable_def sums_iff_shift [abs_def]) lemma sums_split_initial_segment: "f sums s \ (\i. f (i + n)) sums (s - (\i summable (\n. f(n + k))" by (simp add: summable_iff_shift) lemma suminf_minus_initial_segment: "summable f \ (\n. f (n + k)) = (\n. f n) - (\i suminf f = (\n. f(n + k)) + (\i (\n. f (Suc n)) = suminf f - f 0" using suminf_split_initial_segment[of 1] by simp lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable f" shows "\N. \n\N. norm (\i. f (i + n)) < r" proof - from LIMSEQ_D[OF summable_LIMSEQ[OF \summable f\] \0 < r\] obtain N :: nat where "\ n \ N. norm (sum f {..summable f\]) qed lemma summable_LIMSEQ_zero: assumes "summable f" shows "f \ 0" proof - have "Cauchy (\n. sum f {.. convergent f" by (force dest!: summable_LIMSEQ_zero simp: convergent_def) lemma summable_imp_Bseq: "summable f \ Bseq f" by (simp add: convergent_imp_Bseq summable_imp_convergent) end lemma summable_minus_iff: "summable (\n. - f n) \ summable f" for f :: "nat \ 'a::real_normed_vector" by (auto dest: summable_minus) (* used two ways, hence must be outside the context above *) lemma (in bounded_linear) sums: "(\n. X n) sums a \ (\n. f (X n)) sums (f a)" unfolding sums_def by (drule tendsto) (simp only: sum) lemma (in bounded_linear) summable: "summable (\n. X n) \ summable (\n. f (X n))" unfolding summable_def by (auto intro: sums) lemma (in bounded_linear) suminf: "summable (\n. X n) \ f (\n. X n) = (\n. f (X n))" by (intro sums_unique sums summable_sums) lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real] lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real] lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real] lemmas sums_scaleR_left = bounded_linear.sums[OF bounded_linear_scaleR_left] lemmas summable_scaleR_left = bounded_linear.summable[OF bounded_linear_scaleR_left] lemmas suminf_scaleR_left = bounded_linear.suminf[OF bounded_linear_scaleR_left] lemmas sums_scaleR_right = bounded_linear.sums[OF bounded_linear_scaleR_right] lemmas summable_scaleR_right = bounded_linear.summable[OF bounded_linear_scaleR_right] lemmas suminf_scaleR_right = bounded_linear.suminf[OF bounded_linear_scaleR_right] lemma summable_const_iff: "summable (\_. c) \ c = 0" for c :: "'a::real_normed_vector" proof - have "\ summable (\_. c)" if "c \ 0" proof - from that have "filterlim (\n. of_nat n * norm c) at_top sequentially" by (subst mult.commute) (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_real_sequentially) then have "\ convergent (\n. norm (\kInfinite summability on real normed algebras\ context fixes f :: "nat \ 'a::real_normed_algebra" begin lemma sums_mult: "f sums a \ (\n. c * f n) sums (c * a)" by (rule bounded_linear.sums [OF bounded_linear_mult_right]) lemma summable_mult: "summable f \ summable (\n. c * f n)" by (rule bounded_linear.summable [OF bounded_linear_mult_right]) lemma suminf_mult: "summable f \ suminf (\n. c * f n) = c * suminf f" by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric]) lemma sums_mult2: "f sums a \ (\n. f n * c) sums (a * c)" by (rule bounded_linear.sums [OF bounded_linear_mult_left]) lemma summable_mult2: "summable f \ summable (\n. f n * c)" by (rule bounded_linear.summable [OF bounded_linear_mult_left]) lemma suminf_mult2: "summable f \ suminf f * c = (\n. f n * c)" by (rule bounded_linear.suminf [OF bounded_linear_mult_left]) end lemma sums_mult_iff: fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" shows "(\n. c * f n) sums (c * d) \ f sums d" using sums_mult[of f d c] sums_mult[of "\n. c * f n" "c * d" "inverse c"] by (force simp: field_simps assms) lemma sums_mult2_iff: fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" shows "(\n. f n * c) sums (d * c) \ f sums d" using sums_mult_iff[OF assms, of f d] by (simp add: mult.commute) lemma sums_of_real_iff: "(\n. of_real (f n) :: 'a::real_normed_div_algebra) sums of_real c \ f sums c" by (simp add: sums_def of_real_sum[symmetric] tendsto_of_real_iff del: of_real_sum) subsection \Infinite summability on real normed fields\ context fixes c :: "'a::real_normed_field" begin lemma sums_divide: "f sums a \ (\n. f n / c) sums (a / c)" by (rule bounded_linear.sums [OF bounded_linear_divide]) lemma summable_divide: "summable f \ summable (\n. f n / c)" by (rule bounded_linear.summable [OF bounded_linear_divide]) lemma suminf_divide: "summable f \ suminf (\n. f n / c) = suminf f / c" by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric]) lemma summable_inverse_divide: "summable (inverse \ f) \ summable (\n. c / f n)" by (auto dest: summable_mult [of _ c] simp: field_simps) lemma sums_mult_D: "(\n. c * f n) sums a \ c \ 0 \ f sums (a/c)" using sums_mult_iff by fastforce lemma summable_mult_D: "summable (\n. c * f n) \ c \ 0 \ summable f" by (auto dest: summable_divide) text \Sum of a geometric progression.\ lemma geometric_sums: assumes "norm c < 1" shows "(\n. c^n) sums (1 / (1 - c))" proof - have neq_0: "c - 1 \ 0" using assms by auto then have "(\n. c ^ n / (c - 1) - 1 / (c - 1)) \ 0 / (c - 1) - 1 / (c - 1)" by (intro tendsto_intros assms) then have "(\n. (c ^ n - 1) / (c - 1)) \ 1 / (1 - c)" by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) with neq_0 show "(\n. c ^ n) sums (1 / (1 - c))" by (simp add: sums_def geometric_sum) qed lemma summable_geometric: "norm c < 1 \ summable (\n. c^n)" by (rule geometric_sums [THEN sums_summable]) lemma suminf_geometric: "norm c < 1 \ suminf (\n. c^n) = 1 / (1 - c)" by (rule sums_unique[symmetric]) (rule geometric_sums) lemma summable_geometric_iff [simp]: "summable (\n. c ^ n) \ norm c < 1" proof assume "summable (\n. c ^ n :: 'a :: real_normed_field)" then have "(\n. norm c ^ n) \ 0" by (simp add: norm_power [symmetric] tendsto_norm_zero_iff summable_LIMSEQ_zero) from order_tendstoD(2)[OF this zero_less_one] obtain n where "norm c ^ n < 1" by (auto simp: eventually_at_top_linorder) then show "norm c < 1" using one_le_power[of "norm c" n] by (cases "norm c \ 1") (linarith, simp) qed (rule summable_geometric) end text \Biconditional versions for constant factors\ context fixes c :: "'a::real_normed_field" begin lemma summable_cmult_iff [simp]: "summable (\n. c * f n) \ c=0 \ summable f" proof - have "\summable (\n. c * f n); c \ 0\ \ summable f" using summable_mult_D by blast then show ?thesis by (auto simp: summable_mult) qed lemma summable_divide_iff [simp]: "summable (\n. f n / c) \ c=0 \ summable f" proof - have "\summable (\n. f n / c); c \ 0\ \ summable f" by (auto dest: summable_divide [where c = "1/c"]) then show ?thesis by (auto simp: summable_divide) qed end lemma power_half_series: "(\n. (1/2::real)^Suc n) sums 1" proof - have 2: "(\n. (1/2::real)^n) sums 2" using geometric_sums [of "1/2::real"] by auto have "(\n. (1/2::real)^Suc n) = (\n. (1 / 2) ^ n / 2)" by (simp add: mult.commute) then show ?thesis using sums_divide [OF 2, of 2] by simp qed subsection \Telescoping\ lemma telescope_sums: fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "(\n. f (Suc n) - f n) sums (c - f 0)" unfolding sums_def proof (subst filterlim_sequentially_Suc [symmetric]) have "(\n. \kn. f (Suc n) - f 0)" by (simp add: lessThan_Suc_atMost atLeast0AtMost [symmetric] sum_Suc_diff) also have "\ \ c - f 0" by (intro tendsto_diff LIMSEQ_Suc[OF assms] tendsto_const) finally show "(\n. \n c - f 0" . qed lemma telescope_sums': fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "(\n. f n - f (Suc n)) sums (f 0 - c)" using sums_minus[OF telescope_sums[OF assms]] by (simp add: algebra_simps) lemma telescope_summable: fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "summable (\n. f (Suc n) - f n)" using telescope_sums[OF assms] by (simp add: sums_iff) lemma telescope_summable': fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "summable (\n. f n - f (Suc n))" using summable_minus[OF telescope_summable[OF assms]] by (simp add: algebra_simps) subsection \Infinite summability on Banach spaces\ text \Cauchy-type criterion for convergence of series (c.f. Harrison).\ lemma summable_Cauchy: "summable f \ (\e>0. \N. \m\N. \n. norm (sum f {m.. 'a::banach" proof assume f: "summable f" show ?rhs proof clarify fix e :: real assume "0 < e" then obtain M where M: "\m n. \m\M; n\M\ \ norm (sum f {.. M" for m n proof (cases m n rule: linorder_class.le_cases) assume "m \ n" then show ?thesis by (metis (mono_tags, hide_lams) M atLeast0LessThan order_trans sum_diff_nat_ivl that zero_le) next assume "n \ m" then show ?thesis by (simp add: \0 < e\) qed then show "\N. \m\N. \n. norm (sum f {m..m n. m \ N \ norm (sum f {m..N" "n\N" for m n proof (cases m n rule: linorder_class.le_cases) assume "m \ n" then show ?thesis by (metis Groups_Big.sum_diff N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff norm_minus_commute \m\N\) next assume "n \ m" then show ?thesis by (metis Groups_Big.sum_diff N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff \n\N\) qed then show "\M. \m\M. \n\M. norm (sum f {.. 'a :: banach" assumes "eventually (\m. \n\m. norm (sum f {m.. g m) sequentially" assumes "filterlim g (nhds 0) sequentially" shows "summable f" proof (subst summable_Cauchy, intro allI impI, goal_cases) case (1 e) from order_tendstoD(2)[OF assms(2) this] and assms(1) have "eventually (\m. \n. norm (sum f {m.. m") auto qed qed thus ?case by (auto simp: eventually_at_top_linorder) qed context fixes f :: "nat \ 'a::banach" begin text \Absolute convergence imples normal convergence.\ lemma summable_norm_cancel: "summable (\n. norm (f n)) \ summable f" unfolding summable_Cauchy apply (erule all_forward imp_forward ex_forward | assumption)+ apply (fastforce simp add: order_le_less_trans [OF norm_sum] order_le_less_trans [OF abs_ge_self]) done lemma summable_norm: "summable (\n. norm (f n)) \ norm (suminf f) \ (\n. norm (f n))" by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_sum) text \Comparison tests.\ lemma summable_comparison_test: assumes fg: "\N. \n\N. norm (f n) \ g n" and g: "summable g" shows "summable f" proof - obtain N where N: "\n. n\N \ norm (f n) \ g n" using assms by blast show ?thesis proof (clarsimp simp add: summable_Cauchy) fix e :: real assume "0 < e" then obtain Ng where Ng: "\m n. m \ Ng \ norm (sum g {m..max N Ng" for m n proof - have "norm (sum f {m.. sum g {m.. norm (sum g {m..N. \m\N. \n. norm (sum f {m..n. norm (f n) \ g n) sequentially \ summable g \ summable f" by (rule summable_comparison_test) (auto simp: eventually_at_top_linorder) text \A better argument order.\ lemma summable_comparison_test': "summable g \ (\n. n \ N \ norm (f n) \ g n) \ summable f" by (rule summable_comparison_test) auto subsection \The Ratio Test\ lemma summable_ratio_test: assumes "c < 1" "\n. n \ N \ norm (f (Suc n)) \ c * norm (f n)" shows "summable f" proof (cases "0 < c") case True show "summable f" proof (rule summable_comparison_test) show "\N'. \n\N'. norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" proof (intro exI allI impI) fix n assume "N \ n" then show "norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" proof (induct rule: inc_induct) case base with True show ?case by simp next case (step m) have "norm (f (Suc m)) / c ^ Suc m * c ^ n \ norm (f m) / c ^ m * c ^ n" using \0 < c\ \c < 1\ assms(2)[OF \N \ m\] by (simp add: field_simps) with step show ?case by simp qed qed show "summable (\n. norm (f N) / c ^ N * c ^ n)" using \0 < c\ \c < 1\ by (intro summable_mult summable_geometric) simp qed next case False have "f (Suc n) = 0" if "n \ N" for n proof - from that have "norm (f (Suc n)) \ c * norm (f n)" by (rule assms(2)) also have "\ \ 0" using False by (simp add: not_less mult_nonpos_nonneg) finally show ?thesis by auto qed then show "summable f" by (intro sums_summable[OF sums_finite, of "{.. Suc N}"]) (auto simp: not_le Suc_less_eq2) qed end text \Relations among convergence and absolute convergence for power series.\ lemma Abel_lemma: fixes a :: "nat \ 'a::real_normed_vector" assumes r: "0 \ r" and r0: "r < r0" and M: "\n. norm (a n) * r0^n \ M" shows "summable (\n. norm (a n) * r^n)" proof (rule summable_comparison_test') show "summable (\n. M * (r / r0) ^ n)" using assms by (auto simp add: summable_mult summable_geometric) show "norm (norm (a n) * r ^ n) \ M * (r / r0) ^ n" for n using r r0 M [of n] dual_order.order_iff_strict by (fastforce simp add: abs_mult field_simps) qed text \Summability of geometric series for real algebras.\ lemma complete_algebra_summable_geometric: fixes x :: "'a::{real_normed_algebra_1,banach}" assumes "norm x < 1" shows "summable (\n. x ^ n)" proof (rule summable_comparison_test) show "\N. \n\N. norm (x ^ n) \ norm x ^ n" by (simp add: norm_power_ineq) from assms show "summable (\n. norm x ^ n)" by (simp add: summable_geometric) qed subsection \Cauchy Product Formula\ text \ Proof based on Analysis WebNotes: Chapter 07, Class 41 \<^url>\http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm\ \ lemma Cauchy_product_sums: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes a: "summable (\k. norm (a k))" and b: "summable (\k. norm (b k))" shows "(\k. \i\k. a i * b (k - i)) sums ((\k. a k) * (\k. b k))" proof - let ?S1 = "\n::nat. {.. {..m n. m \ n \ ?S1 m \ ?S1 n" by auto have S2_le_S1: "\n. ?S2 n \ ?S1 n" by auto have S1_le_S2: "\n. ?S1 (n div 2) \ ?S2 n" by auto have finite_S1: "\n. finite (?S1 n)" by simp with S2_le_S1 have finite_S2: "\n. finite (?S2 n)" by (rule finite_subset) let ?g = "\(i,j). a i * b j" let ?f = "\(i,j). norm (a i) * norm (b j)" have f_nonneg: "\x. 0 \ ?f x" by auto then have norm_sum_f: "\A. norm (sum ?f A) = sum ?f A" unfolding real_norm_def by (simp only: abs_of_nonneg sum_nonneg [rule_format]) have "(\n. (\kk (\k. a k) * (\k. b k)" by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b]) then have 1: "(\n. sum ?g (?S1 n)) \ (\k. a k) * (\k. b k)" by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan) have "(\n. (\kk (\k. norm (a k)) * (\k. norm (b k))" using a b by (intro tendsto_mult summable_LIMSEQ) then have "(\n. sum ?f (?S1 n)) \ (\k. norm (a k)) * (\k. norm (b k))" by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan) then have "convergent (\n. sum ?f (?S1 n))" by (rule convergentI) then have Cauchy: "Cauchy (\n. sum ?f (?S1 n))" by (rule convergent_Cauchy) have "Zfun (\n. sum ?f (?S1 n - ?S2 n)) sequentially" proof (rule ZfunI, simp only: eventually_sequentially norm_sum_f) fix r :: real assume r: "0 < r" from CauchyD [OF Cauchy r] obtain N where "\m\N. \n\N. norm (sum ?f (?S1 m) - sum ?f (?S1 n)) < r" .. then have "\m n. N \ n \ n \ m \ norm (sum ?f (?S1 m - ?S1 n)) < r" by (simp only: sum_diff finite_S1 S1_mono) then have N: "\m n. N \ n \ n \ m \ sum ?f (?S1 m - ?S1 n) < r" by (simp only: norm_sum_f) show "\N. \n\N. sum ?f (?S1 n - ?S2 n) < r" proof (intro exI allI impI) fix n assume "2 * N \ n" then have n: "N \ n div 2" by simp have "sum ?f (?S1 n - ?S2 n) \ sum ?f (?S1 n - ?S1 (n div 2))" by (intro sum_mono2 finite_Diff finite_S1 f_nonneg Diff_mono subset_refl S1_le_S2) also have "\ < r" using n div_le_dividend by (rule N) finally show "sum ?f (?S1 n - ?S2 n) < r" . qed qed then have "Zfun (\n. sum ?g (?S1 n - ?S2 n)) sequentially" apply (rule Zfun_le [rule_format]) apply (simp only: norm_sum_f) apply (rule order_trans [OF norm_sum sum_mono]) apply (auto simp add: norm_mult_ineq) done then have 2: "(\n. sum ?g (?S1 n) - sum ?g (?S2 n)) \ 0" unfolding tendsto_Zfun_iff diff_0_right by (simp only: sum_diff finite_S1 S2_le_S1) with 1 have "(\n. sum ?g (?S2 n)) \ (\k. a k) * (\k. b k)" by (rule Lim_transform2) then show ?thesis by (simp only: sums_def sum.triangle_reindex) qed lemma Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes "summable (\k. norm (a k))" and "summable (\k. norm (b k))" shows "(\k. a k) * (\k. b k) = (\k. \i\k. a i * b (k - i))" using assms by (rule Cauchy_product_sums [THEN sums_unique]) lemma summable_Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes "summable (\k. norm (a k))" and "summable (\k. norm (b k))" shows "summable (\k. \i\k. a i * b (k - i))" using Cauchy_product_sums[OF assms] by (simp add: sums_iff) subsection \Series on \<^typ>\real\s\ lemma summable_norm_comparison_test: "\N. \n\N. norm (f n) \ g n \ summable g \ summable (\n. norm (f n))" by (rule summable_comparison_test) auto lemma summable_rabs_comparison_test: "\N. \n\N. \f n\ \ g n \ summable g \ summable (\n. \f n\)" for f :: "nat \ real" by (rule summable_comparison_test) auto lemma summable_rabs_cancel: "summable (\n. \f n\) \ summable f" for f :: "nat \ real" by (rule summable_norm_cancel) simp lemma summable_rabs: "summable (\n. \f n\) \ \suminf f\ \ (\n. \f n\)" for f :: "nat \ real" by (fold real_norm_def) (rule summable_norm) lemma summable_zero_power [simp]: "summable (\n. 0 ^ n :: 'a::{comm_ring_1,topological_space})" proof - have "(\n. 0 ^ n :: 'a) = (\n. if n = 0 then 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp qed lemma summable_zero_power' [simp]: "summable (\n. f n * 0 ^ n :: 'a::{ring_1,topological_space})" proof - have "(\n. f n * 0 ^ n :: 'a) = (\n. if n = 0 then f 0 * 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp qed lemma summable_power_series: fixes z :: real assumes le_1: "\i. f i \ 1" and nonneg: "\i. 0 \ f i" and z: "0 \ z" "z < 1" shows "summable (\i. f i * z^i)" proof (rule summable_comparison_test[OF _ summable_geometric]) show "norm z < 1" using z by (auto simp: less_imp_le) show "\n. \N. \na\N. norm (f na * z ^ na) \ z ^ na" using z by (auto intro!: exI[of _ 0] mult_left_le_one_le simp: abs_mult nonneg power_abs less_imp_le le_1) qed lemma summable_0_powser: "summable (\n. f n * 0 ^ n :: 'a::real_normed_div_algebra)" proof - have A: "(\n. f n * 0 ^ n) = (\n. if n = 0 then f n else 0)" by (intro ext) auto then show ?thesis by (subst A) simp_all qed lemma summable_powser_split_head: "summable (\n. f (Suc n) * z ^ n :: 'a::real_normed_div_algebra) = summable (\n. f n * z ^ n)" proof - have "summable (\n. f (Suc n) * z ^ n) \ summable (\n. f (Suc n) * z ^ Suc n)" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs using summable_mult2[OF that, of z] by (simp add: power_commutes algebra_simps) show ?lhs if ?rhs using summable_mult2[OF that, of "inverse z"] by (cases "z \ 0", subst (asm) power_Suc2) (simp_all add: algebra_simps) qed also have "\ \ summable (\n. f n * z ^ n)" by (rule summable_Suc_iff) finally show ?thesis . qed lemma summable_powser_ignore_initial_segment: fixes f :: "nat \ 'a :: real_normed_div_algebra" shows "summable (\n. f (n + m) * z ^ n) \ summable (\n. f n * z ^ n)" proof (induction m) case (Suc m) have "summable (\n. f (n + Suc m) * z ^ n) = summable (\n. f (Suc n + m) * z ^ n)" by simp also have "\ = summable (\n. f (n + m) * z ^ n)" by (rule summable_powser_split_head) also have "\ = summable (\n. f n * z ^ n)" by (rule Suc.IH) finally show ?case . qed simp_all lemma powser_split_head: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" assumes "summable (\n. f n * z ^ n)" shows "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" and "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" and "summable (\n. f (Suc n) * z ^ n)" proof - from assms show "summable (\n. f (Suc n) * z ^ n)" by (subst summable_powser_split_head) from suminf_mult2[OF this, of z] have "(\n. f (Suc n) * z ^ n) * z = (\n. f (Suc n) * z ^ Suc n)" by (simp add: power_commutes algebra_simps) also from assms have "\ = suminf (\n. f n * z ^ n) - f 0" by (subst suminf_split_head) simp_all finally show "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" by simp then show "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" by simp qed lemma summable_partial_sum_bound: fixes f :: "nat \ 'a :: banach" and e :: real assumes summable: "summable f" and e: "e > 0" obtains N where "\m n. m \ N \ norm (\k=m..n. f k) < e" proof - from summable have "Cauchy (\n. \km n. m \ N \ n \ N \ norm ((\kkk=m..n. f k) < e" if m: "m \ N" for m n proof (cases "n \ m") case True with m have "norm ((\kkkkk=m..n. f k)" by (subst sum_diff [symmetric]) (simp_all add: sum.last_plus) finally show ?thesis . next case False with e show ?thesis by simp_all qed then show ?thesis by (rule that) qed lemma powser_sums_if: "(\n. (if n = m then (1 :: 'a::{ring_1,topological_space}) else 0) * z^n) sums z^m" proof - have "(\n. (if n = m then 1 else 0) * z^n) = (\n. if n = m then z^n else 0)" by (intro ext) auto then show ?thesis by (simp add: sums_single) qed lemma fixes f :: "nat \ real" assumes "summable f" and "inj g" and pos: "\x. 0 \ f x" shows summable_reindex: "summable (f \ g)" and suminf_reindex_mono: "suminf (f \ g) \ suminf f" and suminf_reindex: "(\x. x \ range g \ f x = 0) \ suminf (f \ g) = suminf f" proof - from \inj g\ have [simp]: "\A. inj_on g A" by (rule subset_inj_on) simp have smaller: "\n. (\i g) i) \ suminf f" proof fix n have "\ n' \ (g ` {..n'. n' < n \ g n' < m" by blast have "(\i \ (\i \ suminf f" using \summable f\ by (rule sum_le_suminf) (simp_all add: pos) finally show "(\i g) i) \ suminf f" by simp qed have "incseq (\n. \i g) i)" by (rule incseq_SucI) (auto simp add: pos) then obtain L where L: "(\ n. \i g) i) \ L" using smaller by(rule incseq_convergent) then have "(f \ g) sums L" by (simp add: sums_def) then show "summable (f \ g)" by (auto simp add: sums_iff) then have "(\n. \i g) i) \ suminf (f \ g)" by (rule summable_LIMSEQ) then show le: "suminf (f \ g) \ suminf f" by(rule LIMSEQ_le_const2)(blast intro: smaller[rule_format]) assume f: "\x. x \ range g \ f x = 0" from \summable f\ have "suminf f \ suminf (f \ g)" proof (rule suminf_le_const) fix n have "\ n' \ (g -` {..n'. g n' < n \ n' < m" by blast have "(\ii\{.. range g. f i)" using f by(auto intro: sum.mono_neutral_cong_right) also have "\ = (\i\g -` {.. g) i)" by (rule sum.reindex_cong[where l=g])(auto) also have "\ \ (\i g) i)" by (rule sum_mono2)(auto simp add: pos n) also have "\ \ suminf (f \ g)" using \summable (f \ g)\ by (rule sum_le_suminf) (simp_all add: pos) finally show "sum f {.. suminf (f \ g)" . qed with le show "suminf (f \ g) = suminf f" by (rule antisym) qed lemma sums_mono_reindex: assumes subseq: "strict_mono g" and zero: "\n. n \ range g \ f n = 0" shows "(\n. f (g n)) sums c \ f sums c" unfolding sums_def proof assume lim: "(\n. \k c" have "(\n. \kn. \kkk\g`{.. = (\kkk \ c" by (simp only: o_def) finally show "(\n. \k c" . next assume lim: "(\n. \k c" define g_inv where "g_inv n = (LEAST m. g m \ n)" for n from filterlim_subseq[OF subseq] have g_inv_ex: "\m. g m \ n" for n by (auto simp: filterlim_at_top eventually_at_top_linorder) then have g_inv: "g (g_inv n) \ n" for n unfolding g_inv_def by (rule LeastI_ex) have g_inv_least: "m \ g_inv n" if "g m \ n" for m n using that unfolding g_inv_def by (rule Least_le) have g_inv_least': "g m < n" if "m < g_inv n" for m n using that g_inv_least[of n m] by linarith have "(\n. \kn. \k {.. range g" proof (rule notI, elim imageE) fix l assume l: "k = g l" have "g l < g (g_inv n)" by (rule less_le_trans[OF _ g_inv]) (use k l in simp_all) with subseq have "l < g_inv n" by (simp add: strict_mono_less) with k l show False by simp qed then have "f k = 0" by (rule zero) } with g_inv_least' g_inv have "(\kk\g`{.. = (\kkk n" also have "n \ g (g_inv n)" by (rule g_inv) finally have "K \ g_inv n" using subseq by (simp add: strict_mono_less_eq) } then have "filterlim g_inv at_top sequentially" by (auto simp: filterlim_at_top eventually_at_top_linorder) with lim have "(\n. \k c" by (rule filterlim_compose) finally show "(\n. \k c" . qed lemma summable_mono_reindex: assumes subseq: "strict_mono g" and zero: "\n. n \ range g \ f n = 0" shows "summable (\n. f (g n)) \ summable f" using sums_mono_reindex[of g f, OF assms] by (simp add: summable_def) lemma suminf_mono_reindex: fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" assumes "strict_mono g" "\n. n \ range g \ f n = 0" shows "suminf (\n. f (g n)) = suminf f" proof (cases "summable f") case True with sums_mono_reindex [of g f, OF assms] and summable_mono_reindex [of g f, OF assms] show ?thesis by (simp add: sums_iff) next case False then have "\(\c. f sums c)" unfolding summable_def by blast then have "suminf f = The (\_. False)" by (simp add: suminf_def) moreover from False have "\ summable (\n. f (g n))" using summable_mono_reindex[of g f, OF assms] by simp then have "\(\c. (\n. f (g n)) sums c)" unfolding summable_def by blast then have "suminf (\n. f (g n)) = The (\_. False)" by (simp add: suminf_def) ultimately show ?thesis by simp qed lemma summable_bounded_partials: fixes f :: "nat \ 'a :: {real_normed_vector,complete_space}" assumes bound: "eventually (\x0. \a\x0. \b>a. norm (sum f {a<..b}) \ g a) sequentially" assumes g: "g \ 0" shows "summable f" unfolding summable_iff_convergent' proof (intro Cauchy_convergent CauchyI', goal_cases) case (1 \) with g have "eventually (\x. \g x\ < \) sequentially" by (auto simp: tendsto_iff) from eventually_conj[OF this bound] obtain x0 where x0: "\x. x \ x0 \ \g x\ < \" "\a b. x0 \ a \ a < b \ norm (sum f {a<..b}) \ g a" unfolding eventually_at_top_linorder by auto show ?case proof (intro exI[of _ x0] allI impI) fix m n assume mn: "x0 \ m" "m < n" have "dist (sum f {..m}) (sum f {..n}) = norm (sum f {..n} - sum f {..m})" by (simp add: dist_norm norm_minus_commute) also have "sum f {..n} - sum f {..m} = sum f ({..n} - {..m})" using mn by (intro Groups_Big.sum_diff [symmetric]) auto also have "{..n} - {..m} = {m<..n}" using mn by auto also have "norm (sum f {m<..n}) \ g m" using mn by (intro x0) auto also have "\ \ \g m\" by simp also have "\ < \" using mn by (intro x0) auto finally show "dist (sum f {..m}) (sum f {..n}) < \" . qed qed end