diff --git a/thys/Buffons_Needle/Buffons_Needle.thy b/thys/Buffons_Needle/Buffons_Needle.thy --- a/thys/Buffons_Needle/Buffons_Needle.thy +++ b/thys/Buffons_Needle/Buffons_Needle.thy @@ -1,495 +1,487 @@ (* File: Buffons_Needle.thy Author: Manuel Eberl A formal solution of Buffon's needle problem. *) section \Buffon's Needle Problem\ theory Buffons_Needle imports "HOL-Probability.Probability" begin subsection \Auxiliary material\ lemma sin_le_zero': "sin x \ 0" if "x \ -pi" "x \ 0" for x by (metis minus_le_iff neg_0_le_iff_le sin_ge_zero sin_minus that(1) that(2)) lemma emeasure_Un': assumes "A \ sets M" "B \ sets M" "A \ B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A + emeasure M B" proof - have "A \ B = A \ (B - A \ B)" by blast also have "emeasure M \ = emeasure M A + emeasure M (B - A \ B)" using assms by (subst plus_emeasure) auto also have "emeasure M (B - A \ B) = emeasure M B" using assms by (intro emeasure_Diff_null_set) auto finally show ?thesis . qed -lemma singleton_null_set_lborel [simp,intro]: "{x} \ null_sets lborel" - by (simp add: null_sets_def) - -lemma continuous_on_min [continuous_intros]: - fixes f g :: "'a::topological_space \ 'b::linorder_topology" - shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. min (f x) (g x))" - by (auto simp: continuous_on_def intro!: tendsto_min) - lemma integral_shift: fixes f :: "real \ 'a::euclidean_space" assumes cont: "continuous_on {a + c..b + c} f" shows "integral {a..b} (f \ (\x. x + c)) = integral {a + c..b + c} f" proof (cases "a \ b") case True have "((\x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}" using True cont by (intro has_integral_substitution[where c = "a + c" and d = "b + c"]) (auto intro!: derivative_eq_intros) thus ?thesis by (simp add: has_integral_iff o_def) qed auto lemma arcsin_le_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma le_arcsin_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed subsection \Problem definition\ text \ Consider a needle of length $l$ whose centre has the $x$-coordinate $x$. The following then defines the set of all $x$-coordinates that the needle covers (i.e. the projection of the needle onto the $x$-axis.) \ definition needle :: "real \ real \ real \ real set" where "needle l x \ = closed_segment (x - l / 2 * sin \) (x + l / 2 * sin \)" text \ Buffon's Needle problem is then this: Assuming the needle's $x$ position is chosen uniformly at random in a strip of width $d$ centred at the origin, what is the probability that the needle crosses at least one of the left/right boundaries of that strip (located at $x = \pm\frac{1}{2}d$)? \ definition buffon :: "real \ real \ bool measure" where "buffon l d = do { (x, \) \ uniform_measure lborel ({-d/2..d/2} \ {-pi..pi}); return (count_space UNIV) (needle l x \ \ {-d/2, d/2} \ {}) }" subsection \Derivation of the solution\ text \ The following form is a bit easier to handle. \ lemma buffon_altdef: "buffon l d = do { (x, \) \ uniform_measure lborel ({-d/2..d/2} \ {-pi..pi}); return (count_space UNIV) (let a = x - l / 2 * sin \; b = x + l / 2 * sin \ in min a b + d/2 \ 0 \ max a b + d/2 \ 0 \ min a b - d/2 \ 0 \ max a b - d/2 \ 0) }" proof - note buffon_def[of l d] also { have "(\(x,\). needle l x \ \ {-d/2, d/2} \ {}) = (\(x,\). let a = x - l / 2 * sin \; b = x + l / 2 * sin \ in -d/2 \ min a b \ -d/2 \ max a b \ min a b \ d/2 \ max a b \ d/2)" by (auto simp: needle_def Let_def closed_segment_eq_real_ivl min_def max_def) also have "\ = (\(x,\). let a = x - l / 2 * sin \; b = x + l / 2 * sin \ in min a b + d/2 \ 0 \ max a b + d/2 \ 0 \ min a b - d/2 \ 0 \ max a b - d/2 \ 0)" by (auto simp add: algebra_simps Let_def) finally have "(\(x, \). return (count_space UNIV) (needle l x \ \ {- d/2, d/2} \ {})) = (\(x,\). return (count_space UNIV) (let a = x - l / 2 * sin \; b = x + l / 2 * sin \ in min a b + d/2 \ 0 \ max a b + d/2 \ 0 \ min a b - d/2 \ 0 \ max a b - d/2 \ 0))" by (simp add: case_prod_unfold fun_eq_iff) } finally show ?thesis . qed text \ It is obvious that the problem boils down to determining the measure of the following set: \ definition buffon_set :: "real \ real \ (real \ real) set" where "buffon_set l d = {(x,\) \ {-d/2..d/2} \ {-pi..pi}. abs x \ d / 2 - abs (sin \) * l / 2}" text \ By using the symmetry inherent in the problem, we can reduce the problem to the following set, which corresponds to one quadrant of the original set: \ definition buffon_set' :: "real \ real \ (real \ real) set" where "buffon_set' l d = {(x,\) \ {0..d/2} \ {0..pi}. x \ d / 2 - sin \ * l / 2}" lemma closed_buffon_set [simp, intro, measurable]: "closed (buffon_set l d)" proof - have "buffon_set l d = ({-d/2..d/2} \ {-pi..pi}) \ (\z. abs (fst z) + abs (sin (snd z)) * l / 2 - d / 2) -` {0..}" (is "_ = ?A") unfolding buffon_set_def by auto also have "closed \" by (intro closed_Int closed_vimage closed_Times) (auto intro!: continuous_intros) finally show ?thesis by simp qed lemma closed_buffon_set' [simp, intro, measurable]: "closed (buffon_set' l d)" proof - have "buffon_set' l d = ({0..d/2} \ {0..pi}) \ (\z. fst z + sin (snd z) * l / 2 - d / 2) -` {0..}" (is "_ = ?A") unfolding buffon_set'_def by auto also have "closed \" by (intro closed_Int closed_vimage closed_Times) (auto intro!: continuous_intros) finally show ?thesis by simp qed lemma measurable_buffon_set [measurable]: "buffon_set l d \ sets borel" by measurable lemma measurable_buffon_set' [measurable]: "buffon_set' l d \ sets borel" by measurable context fixes d l :: real assumes d: "d > 0" and l: "l > 0" begin lemma buffon_altdef': "buffon l d = distr (uniform_measure lborel ({-d/2..d/2} \ {-pi..pi})) (count_space UNIV) (\z. z \ buffon_set l d)" proof - let ?P = "\(x,\). let a = x - l / 2 * sin \; b = x + l / 2 * sin \ in min a b + d/2 \ 0 \ max a b + d/2 \ 0 \ min a b - d/2 \ 0 \ max a b - d/2 \ 0" have "buffon l d = uniform_measure lborel ({- d / 2..d / 2} \ {-pi..pi}) \ (\z. return (count_space UNIV) (?P z))" unfolding buffon_altdef case_prod_unfold by simp also have "\ = uniform_measure lborel ({- d / 2..d / 2} \ {-pi..pi}) \ (\z. return (count_space UNIV) (z \ buffon_set l d))" proof (intro bind_cong_AE AE_uniform_measureI AE_I2 impI refl return_measurable, goal_cases) show "(\z. return (count_space UNIV) (?P z)) \ uniform_measure lborel ({- d / 2..d / 2} \ {- pi..pi}) \\<^sub>M subprob_algebra (count_space UNIV)" unfolding Let_def case_prod_unfold lborel_prod [symmetric] by measurable show "(\z. return (count_space UNIV) (z \ buffon_set l d)) \ uniform_measure lborel ({- d / 2..d / 2} \ {- pi..pi}) \\<^sub>M subprob_algebra (count_space UNIV)" by simp case (4 z) hence "?P z \ z \ buffon_set l d" proof (cases "snd z \ 0") case True with 4 have "fst z - l / 2 * sin (snd z) \ fst z + l / 2 * sin (snd z)" using l by (auto simp: sin_ge_zero) moreover from True and 4 have "sin (snd z) \ 0" by (auto simp: sin_ge_zero) ultimately show ?thesis using 4 True unfolding buffon_set_def by (force simp: field_simps Let_def min_def max_def case_prod_unfold abs_if) next case False with 4 have "fst z - l / 2 * sin (snd z) \ fst z + l / 2 * sin (snd z)" using l by (auto simp: sin_le_zero' mult_nonneg_nonpos) moreover from False and 4 have "sin (snd z) \ 0" by (auto simp: sin_le_zero') ultimately show ?thesis using 4 and False unfolding buffon_set_def using l d by (force simp: field_simps Let_def min_def max_def case_prod_unfold abs_if) qed thus ?case by (simp only: ) qed (simp_all add: borel_prod [symmetric]) also have "\ = distr (uniform_measure lborel ({-d/2..d/2} \ {-pi..pi})) (count_space UNIV) (\z. z \ buffon_set l d)" by (rule bind_return_distr') simp_all finally show ?thesis . qed lemma buffon_prob_aux: "emeasure (buffon l d) {True} = emeasure lborel (buffon_set l d) / ennreal (2 * d * pi)" proof - have [measurable]: "A \ B \ sets borel" if "A \ sets borel" "B \ sets borel" for A B :: "real set" using that unfolding borel_prod [symmetric] by simp have "emeasure (buffon l d) {True} = emeasure (uniform_measure lborel ({- (d / 2)..d / 2} \ {-pi..pi})) ((\z. z \ buffon_set l d) -` {True})" (is "_ = emeasure ?M _") by (simp add: buffon_altdef' emeasure_distr) also have "(\z. z \ buffon_set l d) -` {True} = buffon_set l d" by auto also have "buffon_set l d \ {-d/2..d/2} \ {-pi..pi}" using l d by (auto simp: buffon_set_def) hence "emeasure ?M (buffon_set l d) = emeasure lborel (buffon_set l d) / emeasure lborel ({- (d / 2)..d / 2} \ {-pi..pi})" by (subst emeasure_uniform_measure) (simp_all add: Int_absorb1) also have "emeasure lborel ({- (d / 2)..d / 2} \ {-pi..pi}) = ennreal (2 * pi * d)" using d by (simp add: lborel_prod [symmetric] lborel.emeasure_pair_measure_Times ennreal_mult algebra_simps) finally show ?thesis by (simp add: mult_ac) qed lemma emeasure_buffon_set_conv_buffon_set': "emeasure lborel (buffon_set l d) = 4 * emeasure lborel (buffon_set' l d)" proof - have distr_lborel [simp]: "distr M lborel f = distr M borel f" for M and f :: "real \ real" by (rule distr_cong) simp_all define A where "A = buffon_set' l d" define B C D where "B = (\x. (-fst x, snd x)) -` A" and "C = (\x. (fst x, -snd x)) -` A" and "D = (\x. (-fst x, -snd x)) -` A" have meas [measurable]: "(\x::real \ real. (-fst x, snd x)) \ borel_measurable borel" "(\x::real \ real. (fst x, -snd x)) \ borel_measurable borel" "(\x::real \ real. (-fst x, -snd x)) \ borel_measurable borel" unfolding borel_prod [symmetric] by measurable have meas' [measurable]: "A \ sets borel" "B \ sets borel" "C \ sets borel" "D \ sets borel" unfolding A_def B_def C_def D_def by (rule measurable_buffon_set' measurable_sets_borel meas)+ have *: "buffon_set l d = A \ B \ C \ D" proof (intro equalityI subsetI, goal_cases) case (1 z) show ?case proof (cases "fst z \ 0"; cases "snd z \ 0") assume "fst z \ 0" "snd z \ 0" with 1 have "z \ A" by (auto split: prod.splits simp: buffon_set_def buffon_set'_def sin_ge_zero A_def) thus ?thesis by blast next assume "\(fst z \ 0)" "snd z \ 0" with 1 have "z \ B" by (auto split: prod.splits simp: buffon_set_def buffon_set'_def sin_ge_zero A_def B_def) thus ?thesis by blast next assume "fst z \ 0" "\(snd z \ 0)" with 1 have "z \ C" by (auto split: prod.splits simp: buffon_set_def buffon_set'_def sin_le_zero' A_def C_def) thus ?thesis by blast next assume "\(fst z \ 0)" "\(snd z \ 0)" with 1 have "z \ D" by (auto split: prod.splits simp: buffon_set_def buffon_set'_def sin_le_zero' A_def D_def) thus ?thesis by blast qed qed (auto simp: buffon_set_def buffon_set'_def sin_ge_zero sin_le_zero' A_def B_def C_def D_def) have "A \ B = {0} \ ({0..pi} \ {\. sin \ * l - d \ 0})" using d l by (auto simp: buffon_set'_def A_def B_def C_def D_def) moreover have "emeasure lborel \ = 0" unfolding lborel_prod [symmetric] by (subst lborel.emeasure_pair_measure_Times) simp_all ultimately have AB: "(A \ B) \ null_sets lborel" unfolding lborel_prod [symmetric] by (simp add: null_sets_def) have "C \ D = {0} \ ({-pi..0} \ {\. -sin \ * l - d \ 0})" using d l by (auto simp: buffon_set'_def A_def B_def C_def D_def) moreover have "emeasure lborel \ = 0" unfolding lborel_prod [symmetric] by (subst lborel.emeasure_pair_measure_Times) simp_all ultimately have CD: "(C \ D) \ null_sets lborel" unfolding lborel_prod [symmetric] by (simp add: null_sets_def) have "A \ D = {}" "B \ C = {}" using d l by (auto simp: buffon_set'_def A_def D_def B_def C_def) moreover have "A \ C = {(d/2, 0)}" "B \ D = {(-d/2, 0)}" using d l by (auto simp: case_prod_unfold buffon_set'_def A_def B_def C_def D_def) ultimately have AD: "A \ D \ null_sets lborel" and BC: "B \ C \ null_sets lborel" and - AC: "A \ C \ null_sets lborel" and BD: "B \ D \ null_sets lborel" by simp_all + AC: "A \ C \ null_sets lborel" and BD: "B \ D \ null_sets lborel" by auto note * also have "emeasure lborel (A \ B \ C \ D) = emeasure lborel (A \ B \ C) + emeasure lborel D" using AB AC AD BC BD CD by (intro emeasure_Un') (auto simp: Int_Un_distrib2) also have "emeasure lborel (A \ B \ C) = emeasure lborel (A \ B) + emeasure lborel C" using AB AC BC using AB AC AD BC BD CD by (intro emeasure_Un') (auto simp: Int_Un_distrib2) also have "emeasure lborel (A \ B) = emeasure lborel A + emeasure lborel B" using AB using AB AC AD BC BD CD by (intro emeasure_Un') (auto simp: Int_Un_distrib2) also have "emeasure lborel B = emeasure (distr lborel lborel (\(x,y). (-x, y))) A" (is "_ = emeasure ?M _") unfolding B_def by (subst emeasure_distr) (simp_all add: case_prod_unfold) also have "?M = lborel" unfolding lborel_prod [symmetric] by (subst pair_measure_distr [symmetric]) (simp_all add: sigma_finite_lborel lborel_distr_uminus) also have "emeasure lborel C = emeasure (distr lborel lborel (\(x,y). (x, -y))) A" (is "_ = emeasure ?M _") unfolding C_def by (subst emeasure_distr) (simp_all add: case_prod_unfold) also have "?M = lborel" unfolding lborel_prod [symmetric] by (subst pair_measure_distr [symmetric]) (simp_all add: sigma_finite_lborel lborel_distr_uminus) also have "emeasure lborel D = emeasure (distr lborel lborel (\(x,y). (-x, -y))) A" (is "_ = emeasure ?M _") unfolding D_def by (subst emeasure_distr) (simp_all add: case_prod_unfold) also have "?M = lborel" unfolding lborel_prod [symmetric] by (subst pair_measure_distr [symmetric]) (simp_all add: sigma_finite_lborel lborel_distr_uminus) finally have "emeasure lborel (buffon_set l d) = of_nat (Suc (Suc (Suc (Suc 0)))) * emeasure lborel A" unfolding of_nat_Suc ring_distribs by simp also have "of_nat (Suc (Suc (Suc (Suc 0)))) = (4 :: ennreal)" by simp finally show ?thesis unfolding A_def . qed text \ It only remains now to compute the measure of @{const buffon_set'}. We first reduce this problem to a relatively simple integral: \ lemma emeasure_buffon_set': "emeasure lborel (buffon_set' l d) = ennreal (integral {0..pi} (\x. min (d / 2) (sin x * l / 2)))" (is "emeasure lborel ?A = _") proof - have "emeasure lborel ?A = nn_integral lborel (\x. indicator ?A x)" by (intro nn_integral_indicator [symmetric]) simp_all also have "(lborel :: (real \ real) measure) = lborel \\<^sub>M lborel" by (simp only: lborel_prod) also have "nn_integral \ (indicator ?A) = (\\<^sup>+\. \\<^sup>+x. indicator ?A (x, \) \lborel \lborel)" by (subst lborel_pair.nn_integral_snd [symmetric]) (simp_all add: lborel_prod borel_prod) also have "\ = (\\<^sup>+\. \\<^sup>+x. indicator {0..pi} \ * indicator {max 0 (d/2 - sin \ * l / 2) .. d/2} x \lborel \lborel)" using d l by (intro nn_integral_cong) (auto simp: indicator_def field_simps buffon_set'_def) also have "\ = \\<^sup>+ \. indicator {0..pi} \ * emeasure lborel {max 0 (d / 2 - sin \ * l / 2)..d / 2} \lborel" by (subst nn_integral_cmult) simp_all also have "\ = \\<^sup>+ \. ennreal (indicator {0..pi} \ * min (d / 2) (sin \ * l / 2)) \lborel" (is "_ = ?I") using d l by (intro nn_integral_cong) (auto simp: indicator_def sin_ge_zero max_def min_def) also have "integrable lborel (\\. (d / 2) * indicator {0..pi} \)" by simp hence int: "integrable lborel (\\. indicator {0..pi} \ * min (d / 2) (sin \ * l / 2))" by (rule Bochner_Integration.integrable_bound) (insert l d, auto intro!: AE_I2 simp: indicator_def min_def sin_ge_zero) hence "?I = set_lebesgue_integral lborel {0..pi} (\\. min (d / 2) (sin \ * l / 2))" by (subst nn_integral_eq_integral, assumption) (insert d l, auto intro!: AE_I2 simp: sin_ge_zero min_def indicator_def set_lebesgue_integral_def) also have "\ = ennreal (integral {0..pi} (\x. min (d / 2) (sin x * l / 2)))" (is "_ = ennreal ?I") using int by (subst set_borel_integral_eq_integral) (simp_all add: set_integrable_def) finally show ?thesis by (simp add: lborel_prod) qed text \ We now have to distinguish two cases: The first and easier one is that where the length of the needle, $l$, is less than or equal to the strip width, $d$: \ context assumes l_le_d: "l \ d" begin lemma emeasure_buffon_set'_short: "emeasure lborel (buffon_set' l d) = ennreal l" proof - have "emeasure lborel (buffon_set' l d) = ennreal (integral {0..pi} (\x. min (d / 2) (sin x * l / 2)))" (is "_ = ennreal ?I") by (rule emeasure_buffon_set') also have *: "sin \ * l \ d" if "\ \ 0" "\ \ pi" for \ using mult_mono[OF l_le_d sin_le_one _ sin_ge_zero] that d by (simp add: algebra_simps) have "?I = integral {0..pi} (\x. (l / 2) * sin x)" using l d l_le_d by (intro integral_cong) (auto dest: * simp: min_def sin_ge_zero) also have "\ = l / 2 * integral {0..pi} sin" by simp also have "(sin has_integral (-cos pi - (- cos 0))) {0..pi}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "integral {0..pi} sin = -cos pi - (-cos 0)" by (simp add: has_integral_iff) finally show ?thesis by (simp add: lborel_prod) qed lemma emeasure_buffon_set_short: "emeasure lborel (buffon_set l d) = 4 * ennreal l" by (simp add: emeasure_buffon_set_conv_buffon_set' emeasure_buffon_set'_short l_le_d) theorem buffon_short: "emeasure (buffon l d) {True} = ennreal (2 * l / (d * pi))" proof - have "emeasure (buffon l d) {True} = ennreal (4 * l) / ennreal (2 * d * pi)" using d l by (subst buffon_prob_aux) (simp add: emeasure_buffon_set_short ennreal_mult) also have "\ = ennreal (4 * l / (2 * d * pi))" using d l by (subst divide_ennreal) simp_all also have "4 * l / (2 * d * pi) = 2 * l / (d * pi)" by simp finally show ?thesis . qed end text \ The other case where the needle is at least as long as the strip width is more complicated: \ context assumes l_ge_d: "l \ d" begin lemma emeasure_buffon_set'_long: "emeasure lborel (buffon_set' l d) = ennreal (l * (1 - sqrt (1 - (d / l)\<^sup>2)) + arccos (d / l) * d)" proof - define \' where "\' = arcsin (d / l)" have \'_nonneg: "\' \ 0" unfolding \'_def using d l l_ge_d arcsin_le_mono[of 0 "d/l"] by (simp add: \'_def) have \'_le: "\' \ pi / 2" unfolding \'_def using arcsin_bounded[of "d/l"] d l l_ge_d by (simp add: field_simps) have ge_phi': "sin \ \ d / l" if "\ \ \'" "\ \ pi / 2" for \ using arcsin_le_iff[of "d / l" "\"] d l_ge_d that \'_nonneg by (auto simp: \'_def field_simps) have le_phi': "sin \ \ d / l" if "\ \ \'" "\ \ 0" for \ using le_arcsin_iff[of "d / l" "\"] d l_ge_d that \'_le by (auto simp: \'_def field_simps) let ?f = "(\x. min (d / 2) (sin x * l / 2))" have "emeasure lborel (buffon_set' l d) = ennreal (integral {0..pi} ?f)" (is "_ = ennreal ?I") by (rule emeasure_buffon_set') also have "?I = integral {0..pi/2} ?f + integral {pi/2..pi} ?f" - by (rule integral_combine [symmetric]) (auto intro!: integrable_continuous_real continuous_intros) + by (rule Henstock_Kurzweil_Integration.integral_combine [symmetric]) (auto intro!: integrable_continuous_real continuous_intros) also have "integral {pi/2..pi} ?f = integral {-pi/2..0} (?f \ (\\. \ + pi))" by (subst integral_shift) (auto intro!: continuous_intros) also have "\ = integral {-(pi/2)..-0} (\x. min (d / 2) (sin (-x) * l / 2))" by (simp add: o_def) - also have "\ = integral {0..pi/2} ?f" (is "_ = ?I") by (subst integral_reflect_real) simp_all + also have "\ = integral {0..pi/2} ?f" (is "_ = ?I") by (subst Henstock_Kurzweil_Integration.integral_reflect_real) simp_all also have "\ + \ = 2 * \" by simp also have "?I = integral {0..\'} ?f + integral {\'..pi/2} ?f" using l d l_ge_d \'_nonneg \'_le - by (intro integral_combine [symmetric]) (auto intro!: integrable_continuous_real continuous_intros) + by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric]) (auto intro!: integrable_continuous_real continuous_intros) also have "integral {0..\'} ?f = integral {0..\'} (\x. l / 2 * sin x)" using l by (intro integral_cong) (auto simp: min_def field_simps dest: le_phi') also have "((\x. l / 2 * sin x) has_integral (- (l / 2 * cos \') - (- (l / 2 * cos 0)))) {0..\'}" using \'_nonneg by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: derivative_eq_intros) hence "integral {0..\'} (\x. l / 2 * sin x) = (1 - cos \') * l / 2" by (simp add: has_integral_iff algebra_simps) also have "integral {\'..pi/2} ?f = integral {\'..pi/2} (\_. d / 2)" using l by (intro integral_cong) (auto simp: min_def field_simps dest: ge_phi') also have "\ = arccos (d / l) * d / 2" using \'_le d l l_ge_d by (subst arccos_arcsin_eq) (auto simp: field_simps \'_def) also have "cos \' = sqrt (1 - (d / l)^2)" unfolding \'_def by (rule cos_arcsin) (insert d l l_ge_d, auto simp: field_simps) also have "2 * ((1 - sqrt (1 - (d / l)\<^sup>2)) * l / 2 + arccos (d / l) * d / 2) = l * (1 - sqrt (1 - (d / l)\<^sup>2)) + arccos (d / l) * d" using d l by (simp add: field_simps) finally show ?thesis . qed lemma emeasure_buffon_set_long: "emeasure lborel (buffon_set l d) = 4 * ennreal (l * (1 - sqrt (1 - (d / l)\<^sup>2)) + arccos (d / l) * d)" by (simp add: emeasure_buffon_set_conv_buffon_set' emeasure_buffon_set'_long l_ge_d) theorem buffon_long: "emeasure (buffon l d) {True} = ennreal (2 / pi * ((l / d) - sqrt ((l / d)\<^sup>2 - 1) + arccos (d / l)))" proof - have *: "l * sqrt ((l\<^sup>2 - d\<^sup>2) / l\<^sup>2) + 0 \ l + d * arccos (d / l)" using d l_ge_d by (intro add_mono mult_nonneg_nonneg arccos_lbound) (auto simp: field_simps) have "emeasure (buffon l d) {True} = ennreal (4 * (l - l * sqrt (1 - (d / l)\<^sup>2) + arccos (d / l) * d)) / ennreal (2 * d * pi)" using d l l_ge_d * unfolding buffon_prob_aux emeasure_buffon_set_long ennreal_numeral [symmetric] by (subst ennreal_mult [symmetric]) (auto intro!: add_nonneg_nonneg mult_nonneg_nonneg simp: field_simps) also have "\ = ennreal ((4 * (l - l * sqrt (1 - (d / l)\<^sup>2) + arccos (d / l) * d)) / (2 * d * pi))" using d l * by (subst divide_ennreal) (auto simp: field_simps) also have "(4 * (l - l * sqrt (1 - (d / l)\<^sup>2) + arccos (d / l) * d)) / (2 * d * pi) = 2 / pi * (l / d - l / d * sqrt ((d / l)^2 * ((l / d)^2 - 1)) + arccos (d / l))" using d l by (simp add: field_simps) also have "l / d * sqrt ((d / l)^2 * ((l / d)^2 - 1)) = sqrt ((l / d) ^ 2 - 1)" using d l l_ge_d unfolding real_sqrt_mult real_sqrt_abs by simp finally show ?thesis . qed end end end diff --git a/thys/Euler_MacLaurin/Euler_MacLaurin.thy b/thys/Euler_MacLaurin/Euler_MacLaurin.thy --- a/thys/Euler_MacLaurin/Euler_MacLaurin.thy +++ b/thys/Euler_MacLaurin/Euler_MacLaurin.thy @@ -1,1618 +1,1618 @@ section \The Euler--MacLaurin summation formula\ theory Euler_MacLaurin imports "HOL-Analysis.Analysis" "HOL-Library.Multiset" Bernoulli.Periodic_Bernpoly Bernoulli.Bernoulli_FPS begin subsection \Auxiliary facts\ (* TODO Move? *) lemma pbernpoly_of_int [simp]: "pbernpoly n (of_int a) = bernoulli n" by (simp add: pbernpoly_def) lemma continuous_on_bernpoly' [continuous_intros]: assumes "continuous_on A f" shows "continuous_on A (\x. bernpoly n (f x) :: 'a :: real_normed_algebra_1)" using continuous_on_compose2[OF continuous_on_bernpoly assms, of UNIV n] by auto lemma sum_atLeastAtMost_int_last: assumes "a < (b :: int)" shows "sum f {a..b} = sum f {a.. = sum f {a.. = f a + sum f {a<..b}" by (subst sum.insert) auto finally show ?thesis . qed lemma not_in_nonpos_Reals_imp_add_nonzero: assumes "z \ \\<^sub>\\<^sub>0" "x \ 0" shows "z + of_real x \ 0" using assms by (auto simp: add_eq_0_iff2) (* END TODO *) lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" by (cases "b < a") auto lemma integrable_on_negligible: "negligible A \ (f :: 'n :: euclidean_space \ 'a :: banach) integrable_on A" by (subst integrable_spike_set_eq[of _ "{}"]) (simp_all add: integrable_on_empty) lemma Union_atLeastAtMost_real_of_int: assumes "a < b" shows "(\n\{a.. {real_of_int a..real_of_int b}" thus "x \ (\n\{a.. real_of_int a" "x < real_of_int b" by simp_all hence "x \ of_int \x\" "x \ of_int \x\ + 1" by linarith+ moreover from x have "\x\ \ a" "\x\ < b" by linarith+ ultimately have "\n\{a.. {of_int n..of_int (n + 1)}" by (intro bexI[of _ "\x\"]) simp_all thus ?thesis by blast qed qed auto subsection \The remainder terms\ text \ The following describes the remainder term in the classical version of the Euler--MacLaurin formula. \ definition EM_remainder' :: "nat \ (real \ 'a :: banach) \ real \ real \ 'a" where "EM_remainder' n f a b = ((-1) ^ Suc n / fact n) *\<^sub>R integral {a..b} (\t. pbernpoly n t *\<^sub>R f t)" text \ Next, we define the remainder term that occurs when one lets the right bound of summation in the Euler--MacLaurin formula tend to infinity. \ definition EM_remainder_converges :: "nat \ (real \ 'a :: banach) \ int \ bool" where "EM_remainder_converges n f a \ (\L. ((\x. EM_remainder' n f a (of_int x)) \ L) at_top)" definition EM_remainder :: "nat \ (real \ 'a :: banach) \ int \ 'a" where "EM_remainder n f a = (if EM_remainder_converges n f a then Lim at_top (\x. EM_remainder' n f a (of_int x)) else 0)" text \ The following lemmas are fairly obvious -- but tedious to prove -- properties of the remainder terms. \ lemma EM_remainder_eqI: fixes L assumes "((\x. EM_remainder' n f b (of_int x)) \ L) at_top" shows "EM_remainder n f b = L" using assms by (auto simp: EM_remainder_def EM_remainder_converges_def intro!: tendsto_Lim) lemma integrable_EM_remainder'_int: fixes a b :: int and f :: "real \ 'a :: banach" assumes "continuous_on {of_int a..of_int b} f" shows "(\t. pbernpoly n t *\<^sub>R f t) integrable_on {a..b}" proof - have [continuous_intros]: "continuous_on A f" if "A \ {of_int a..of_int b}" for A using continuous_on_subset[OF assms that] . consider "a > b" | "a = b" | "a < b" "n = 1" | "a < b" "n \ 1" by (cases a b rule: linorder_cases) auto thus ?thesis proof cases assume "a < b" and "n \ 1" thus ?thesis by (intro integrable_continuous_real continuous_intros) auto next assume ab: "a < b" and [simp]: "n = 1" let ?A = "(\n. {real_of_int n..real_of_int (n+1)}) ` {a.. ?A" then obtain i where i: "i \ {a..t. pbernpoly n t *\<^sub>R f t) integrable_on I" proof (rule integrable_spike) show "(\t. (t - of_int i - 1/2) *\<^sub>R f t) integrable_on I" using i by (auto intro!: integrable_continuous_real continuous_intros) next fix x assume "x \ I - {of_int (i + 1)}" with i have "of_int i \ x" "x < of_int i + 1" by simp_all hence "floor x = i" by linarith thus "pbernpoly n x *\<^sub>R f x = (x - of_int i - 1 / 2) *\<^sub>R f x" by (simp add: pbernpoly_def bernpoly_def frac_def) qed simp_all qed qed (simp_all add: integrable_on_negligible) qed lemma integrable_EM_remainder': fixes a b :: real and f :: "real \ 'a :: banach" assumes "continuous_on {a..b} f" shows "(\t. pbernpoly n t *\<^sub>R f t) integrable_on {a..b}" proof (cases "\a\ \ \b\") case True define a' b' where "a' = \a\" and "b' = \b\" from True have *: "a' \ b'" "a' \ a" "b' \ b" by (auto simp: a'_def b'_def) from * have A: "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a'..b'})" by (intro integrable_EM_remainder'_int continuous_on_subset[OF assms]) auto have B: "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a..a'})" proof (rule integrable_spike) show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - of_int (floor a)) *\<^sub>R f x" if "x \ {a..real_of_int a'} - {real_of_int a'}"for x proof - have "x \ a" "x t. pbernpoly n t *\<^sub>R f t) integrable_on ({b'..b})" proof (rule integrable_spike) show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - of_int b') *\<^sub>R f x" if "x \ {real_of_int b'..b} - {real_of_int b'}" for x proof - have "x \ b" "x > real_of_int b'" using that by auto with True have "floor x = b'" unfolding b'_def by (intro floor_unique; linarith) thus ?thesis by (simp add: pbernpoly_def frac_def) qed qed (insert *, auto intro!: continuous_intros integrable_continuous_real continuous_on_subset[OF assms]) have "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a..a'} \ {a'..b'} \ {b'..b})" using * A B C by (intro integrable_Un; (subst ivl_disj_un)?) (auto simp: ivl_disj_un max_def min_def) also have "{a..a'} \ {a'..b'} \ {b'..b} = {a..b}" using * by auto finally show ?thesis . next assume *: "\ceiling a \ floor b" show ?thesis proof (rule integrable_spike) show "(\t. bernpoly n (t - floor a) *\<^sub>R f t) integrable_on {a..b}" using * by (auto intro!: integrable_continuous_real continuous_intros assms) next show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - floor a) *\<^sub>R f x" if "x \ {a..b} - {}" for x proof - from * have **: "b < floor a + 1" unfolding ceiling_altdef by (auto split: if_splits simp: le_floor_iff) from that have x: "x \ a" "x \ b" by simp_all with * ** have "floor x = floor a" by linarith thus ?thesis by (simp add: pbernpoly_def frac_def) qed qed simp_all qed lemma EM_remainder'_bounded_linear: assumes "bounded_linear h" assumes "continuous_on {a..b} f" shows "EM_remainder' n (\x. h (f x)) a b = h (EM_remainder' n f a b)" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R h (f t)) = integral {a..b} (\t. h (pbernpoly n t *\<^sub>R f t))" using assms by (simp add: linear_simps) also have "\ = h (integral {a..b} (\t. pbernpoly n t *\<^sub>R f t))" by (subst integral_linear [OF _ assms(1), symmetric]) (auto intro!: integrable_EM_remainder' assms(2) simp: o_def) finally show ?thesis using assms(1) by (simp add: EM_remainder'_def linear_simps) qed lemma EM_remainder_converges_of_real: assumes "EM_remainder_converges n f a" "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. of_real (f x)) a" proof - from assms obtain L where L: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have "((\b. EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) \ of_real L) at_top" proof (rule Lim_transform_eventually) show "eventually (\b. of_real (EM_remainder' n f (of_int a) (of_int b)) = EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) at_top" using eventually_ge_at_top[of a] by eventually_elim (intro EM_remainder'_bounded_linear [OF bounded_linear_of_real, symmetric] continuous_on_subset[OF assms(2)], auto) qed (intro tendsto_intros L) thus ?thesis unfolding EM_remainder_converges_def .. qed lemma EM_remainder_converges_of_real_iff: fixes f :: "real \ real" assumes "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. of_real (f x) :: 'a :: {banach, real_normed_algebra_1, real_inner}) a \ EM_remainder_converges n f a" proof assume "EM_remainder_converges n (\x. of_real (f x) :: 'a) a" then obtain L :: 'a where L: "((\b. EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have "((\b. EM_remainder' n f (of_int a) (of_int b)) \ L \ 1) at_top" proof (rule Lim_transform_eventually) show "eventually (\b. EM_remainder' n (\x. of_real (f x) :: 'a) (of_int a) (of_int b) \ 1 = EM_remainder' n f (of_int a) (of_int b)) at_top" using eventually_ge_at_top[of a] by eventually_elim (subst EM_remainder'_bounded_linear [OF bounded_linear_of_real], auto intro!: continuous_on_subset[OF assms]) qed (intro tendsto_intros L) thus "EM_remainder_converges n f a" unfolding EM_remainder_converges_def .. qed (intro EM_remainder_converges_of_real assms) lemma EM_remainder_of_real: assumes "continuous_on {a..} f" shows "EM_remainder n (\x. of_real (f x) :: 'a :: {banach, real_normed_algebra_1, real_inner}) a = of_real (EM_remainder n f a)" proof - have eq: "EM_remainder' n (\x. of_real (f x) :: 'a) (real_of_int a) = (\x::int. of_real (EM_remainder' n f a x))" by (intro ext EM_remainder'_bounded_linear[OF bounded_linear_of_real] continuous_on_subset[OF assms]) auto show ?thesis proof (cases "EM_remainder_converges n f a") case False with EM_remainder_converges_of_real_iff[OF assms, of n] show ?thesis by (auto simp: EM_remainder_def) next case True then obtain L where L: "((\x. EM_remainder' n f a (real_of_int x)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have L': "((\x. EM_remainder' n (\x. of_real (f x) :: 'a) a (real_of_int x)) \ of_real L) at_top" unfolding eq by (intro tendsto_of_real L) from L L' tendsto_Lim[OF _ L] tendsto_Lim[OF _ L'] show ?thesis by (auto simp: EM_remainder_def EM_remainder_converges_def) qed qed lemma EM_remainder'_cong: assumes "\x. x \ {a..b} \ f x = g x" "n = n'" "a = a'" "b = b'" shows "EM_remainder' n f a b = EM_remainder' n' g a' b'" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = integral {a'..b'} (\t. pbernpoly n' t *\<^sub>R g t)" unfolding assms using assms by (intro integral_cong) auto with assms show ?thesis by (simp add: EM_remainder'_def) qed lemma EM_remainder_converges_cong: assumes "\x. x \ of_int a \ f x = g x" "n = n'" "a = a'" shows "EM_remainder_converges n f a = EM_remainder_converges n' g a'" unfolding EM_remainder_converges_def by (subst EM_remainder'_cong[OF _ refl refl refl, of _ _ f g]) (use assms in auto) lemma EM_remainder_cong: assumes "\x. x \ of_int a \ f x = g x" "n = n'" "a = a'" shows "EM_remainder n f a = EM_remainder n' g a'" proof - have *: "EM_remainder_converges n f a = EM_remainder_converges n' g a'" using assms by (intro EM_remainder_converges_cong) auto show ?thesis unfolding EM_remainder_def by (subst EM_remainder'_cong[OF _ refl refl refl, of _ _ f g]) (use assms * in auto) qed lemma EM_remainder_converges_cnj: assumes "continuous_on {a..} f" and "EM_remainder_converges n f a" shows "EM_remainder_converges n (\x. cnj (f x)) a" proof - interpret bounded_linear cnj by (rule bounded_linear_cnj) obtain L where L: "((\x. EM_remainder' n f (real_of_int a) (real_of_int x)) \ L) at_top" using assms unfolding EM_remainder_converges_def by blast note tendsto_cnj [OF this] also have "(\x. cnj (EM_remainder' n f (real_of_int a) (real_of_int x))) = (\x. EM_remainder' n (\x. cnj (f x)) (real_of_int a) (real_of_int x))" by (subst EM_remainder'_bounded_linear [OF bounded_linear_cnj]) (rule continuous_on_subset [OF assms(1)], auto) finally have L': "(\ \ cnj L) at_top" . thus "EM_remainder_converges n (\x. cnj (f x)) a" by (auto simp: EM_remainder_converges_def) qed lemma EM_remainder_converges_cnj_iff: assumes "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. cnj (f x)) a \ EM_remainder_converges n f a" proof assume "EM_remainder_converges n (\x. cnj (f x)) a" hence "EM_remainder_converges n (\x. cnj (cnj (f x))) a" by (rule EM_remainder_converges_cnj [rotated]) (auto intro: continuous_intros assms) thus "EM_remainder_converges n f a" by simp qed (intro EM_remainder_converges_cnj assms) lemma EM_remainder_cnj: assumes "continuous_on {a..} f" shows "EM_remainder n (\x. cnj (f x)) a = cnj (EM_remainder n f a)" proof (cases "EM_remainder_converges n f a") case False hence "\EM_remainder_converges n (\x. cnj (f x)) a" by (subst EM_remainder_converges_cnj_iff [OF assms]) with False show ?thesis by (simp add: EM_remainder_def) next case True then obtain L where L: "((\x. EM_remainder' n f (real_of_int a) (real_of_int x)) \ L) at_top" unfolding EM_remainder_converges_def by blast note tendsto_cnj [OF this] also have "(\x. cnj (EM_remainder' n f (real_of_int a) (real_of_int x))) = (\x. EM_remainder' n (\x. cnj (f x)) (real_of_int a) (real_of_int x))" by (subst EM_remainder'_bounded_linear [OF bounded_linear_cnj]) (rule continuous_on_subset [OF assms(1)], auto) finally have L': "(\ \ cnj L) at_top" . moreover from assms and L have "EM_remainder n f a = L" by (intro EM_remainder_eqI) ultimately show "EM_remainder n (\x. cnj (f x)) a = cnj (EM_remainder n f a)" using L' by (intro EM_remainder_eqI) simp_all qed lemma EM_remainder'_combine: fixes f :: "real \ 'a :: banach" assumes [continuous_intros]: "continuous_on {a..c} f" assumes "a \ b" "b \ c" shows "EM_remainder' n f a b + EM_remainder' n f b c = EM_remainder' n f a c" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) + integral {b..c} (\t. pbernpoly n t *\<^sub>R f t) = integral {a..c} (\t. pbernpoly n t *\<^sub>R f t)" - by (intro integral_combine assms integrable_EM_remainder') + by (intro Henstock_Kurzweil_Integration.integral_combine assms integrable_EM_remainder') from this [symmetric] show ?thesis by (simp add: EM_remainder'_def algebra_simps) qed lemma uniformly_convergent_EM_remainder': fixes f :: "'a \ real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b y. y \ A \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b s. EM_remainder' n (f s) a b)" proof - interpret bounded_linear "\x::'b. ((- 1) ^ Suc n / fact n) *\<^sub>R x" by (rule bounded_linear_scaleR_right) from bounded_pbernpoly[of n] guess C . note C = this from C[of 0] have [simp]: "C \ 0" by simp show ?thesis unfolding EM_remainder'_def proof (intro uniformly_convergent_on uniformly_convergent_improper_integral') fix x assume "x \ a" thus "((\x. C * G x) has_real_derivative C * g x) (at x within {a..})" by (intro DERIV_cmult deriv) next fix y a' b assume "y \ A" "a \ a'" "a' \ b" thus "(\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" by (rule integrable) next from conv obtain L where "(\y. G (real y)) \ L" by (auto simp: convergent_def) from tendsto_mult[OF tendsto_const[of C] this] show "convergent (\y. C * G (real y))" by (auto simp: convergent_def) next show "\\<^sub>F x in at_top. \y\A. norm (pbernpoly n x *\<^sub>R f y x) \ C * g x" using C unfolding norm_scaleR by (intro eventually_mono[OF bound] ballI mult_mono) auto qed qed lemma uniform_limit_EM_remainder: fixes f :: "'a \ real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b y. y \ A \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniform_limit A (\b s. EM_remainder' n (f s) a b) (\s. EM_remainder n (f s) a) sequentially" proof - have *: "uniformly_convergent_on A (\b s. EM_remainder' n (f s) a b)" by (rule uniformly_convergent_EM_remainder'[OF assms]) also have "?this \ ?thesis" unfolding uniformly_convergent_uniform_limit_iff proof (intro uniform_limit_cong refl always_eventually allI ballI) fix s assume "s \ A" with * have **: "convergent (\b. EM_remainder' n (f s) a b)" by (rule uniformly_convergent_imp_convergent) show "lim (\b. EM_remainder' n (f s) a b) = EM_remainder n (f s) a" proof (rule sym, rule EM_remainder_eqI) have "((\x. EM_remainder' n (f s) (real_of_int a) (real x)) \ lim (\x. EM_remainder' n (f s) (real_of_int a) (real x))) at_top" (is "(_ \ ?L) _") using ** unfolding convergent_LIMSEQ_iff by blast hence "((\x. EM_remainder' n (f s) (real_of_int a) (real (nat x))) \ ?L) at_top" by (rule filterlim_compose) (fact filterlim_nat_sequentially) thus "((\x. EM_remainder' n (f s) (real_of_int a) (real_of_int x)) \ ?L) at_top" by (rule Lim_transform_eventually) (auto intro: eventually_mono[OF eventually_ge_at_top[of 0]]) qed qed finally show \ . qed lemma tendsto_EM_remainder: fixes f :: "real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b . a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. norm (f x) \ g x) at_top" shows "filterlim (\b. EM_remainder' n f a b) (nhds (EM_remainder n f a)) sequentially" proof - have "uniform_limit {()} (\b s. EM_remainder' n f a b) (\s. EM_remainder n f a) sequentially" using assms by (intro uniform_limit_EM_remainder[where G = G and g = g]) auto moreover have "() \ {()}" by simp ultimately show ?thesis by (rule tendsto_uniform_limitI) qed lemma EM_remainder_0 [simp]: "EM_remainder n (\x. 0) a = 0" by (rule EM_remainder_eqI) (simp add: EM_remainder'_def) lemma holomorphic_EM_remainder': assumes deriv: "\z t. z \ U \ t \ {a..x} \ ((\z. f z t) has_field_derivative f' z t) (at z within U)" assumes int: "\b c z e. a \ b \ c \ x \ z \ U \ (\t. of_real (bernpoly n (t - e)) * f z t) integrable_on {b..c}" assumes cont: "continuous_on (U \ {a..x}) (\(z, t). f' z t)" assumes "convex U" shows "(\s. EM_remainder' n (f s) a x) holomorphic_on U" unfolding EM_remainder'_def scaleR_conv_of_real proof (intro holomorphic_intros) have holo: "(\z. integral (cbox b c) (\t. of_real (bernpoly n (t - e)) * f z t)) holomorphic_on U" if "b \ a" "c \ x" for b c e :: real proof (rule leibniz_rule_holomorphic) fix z t assume "z \ U" "t \ cbox b c" thus "((\z. complex_of_real (bernpoly n (t - e)) * f z t) has_field_derivative complex_of_real (bernpoly n (t - e)) * f' z t) (at z within U)" using that by (intro DERIV_cmult deriv) auto next fix z assume "z \ U" thus "(\t. complex_of_real (bernpoly n (t - e)) * f z t) integrable_on cbox b c" using that int[of b c z] by auto next have "continuous_on (U \ {b..c}) (\(z, t). f' z t)" using cont by (rule continuous_on_subset) (insert that, auto) thus "continuous_on (U \ cbox b c) (\(z, t). complex_of_real (bernpoly n (t - e)) * f' z t)" by (auto simp: case_prod_unfold intro!: continuous_intros) qed fact+ consider "a > x" | "a \ x" "floor x \ a" | "a \ x" "floor x > a" by force hence "(\z. integral (cbox a x) (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" (is "?f a x holomorphic_on _") proof cases case 2 have "(\z. integral (cbox a x) (\t. of_real (bernpoly n (t - of_int \x\)) * f z t)) holomorphic_on U" by (intro holo) auto also have "(\z. integral (cbox a x) (\t. of_real (bernpoly n (t - of_int \x\)) * f z t)) = ?f a x" proof (intro ext integral_cong, goal_cases) case (1 z t) hence "t \ a" "t \ x" by auto hence "floor t = floor x" using 2 by linarith thus ?case by (simp add: pbernpoly_def frac_def) qed finally show ?thesis . next case 3 define N :: "int set" where "N = {\a\..<\x\}" define A where "A = insert {a..of_int \a\} (insert {of_int \x\..x} ((\n. {of_int n..of_int n + 1}) ` N))" { fix X assume "X \ A" then consider "X = {a..of_int \a\}" | "X = {of_int \x\..x}" | n where "X = {of_int n..of_int n + 1}" "n \ N" by (auto simp: A_def) } note A_cases = this have division: "A division_of {a..x}" proof (rule division_ofI) show "finite A" by (auto simp: A_def N_def) fix K assume K: "K \ A" from 3 have "of_int \a\ \ x" using ceiling_le[of a "floor x"] by linarith moreover from 3 have "of_int \x\ \ a" by linarith ultimately show "K \ {a..x}" using K 3 by (auto simp: A_def N_def) linarith+ from K show "K \ {}" and "\a b. K = cbox a b" by (auto simp: A_def) next fix K1 K2 assume K: "K1 \ A" "K2 \ A" "K1 \ K2" have F1: "interior {a..\a\} \ interior {\x\..x} = {}" using 3 ceiling_le[of a "floor x"] by (auto simp: min_def max_def) hence F2: "interior {\x\..x} \ interior {a..\a\} = {}" by simp have F3: "interior {a..\a\} \ interior {of_int n..of_int n+1} = {}" "interior {\x\..x} \ interior {of_int n..of_int n+1} = {}" "interior {of_int n..of_int n+1} \ interior {a..\a\} = {}" "interior {of_int n..of_int n+1} \ interior {\x\..x} = {}"if "n \ N" for n using 3 ceiling_le[of a "floor x"] that by (auto simp: min_def max_def N_def) have F4: "interior {real_of_int n..of_int n+1} \ interior {of_int m..of_int m+1} = {}" if "{real_of_int n..of_int n+1} \ {of_int m..of_int m+1}" for m n proof - from that have "n \ m" by auto thus ?thesis by simp qed from F1 F2 F3 F4 K show "interior K1 \ interior K2 = {}" by (elim A_cases) (simp_all only: not_False_eq_True) next show "\A = {a..x}" proof (cases "\a\ = \x\") case True thus ?thesis using 3 by (auto simp: A_def N_def intro: order.trans) linarith+ next case False with 3 have *: "\a\ < \x\" by linarith have "\A = {a..of_int \a\} \ (\n\N. {of_int n..of_int (n + 1)}) \ {of_int \x\..x}" by (simp add: A_def Un_ac) also have "(\n\N. {of_int n..of_int (n + 1)}) = {of_int \a\..real_of_int \x\}" using * unfolding N_def by (intro Union_atLeastAtMost_real_of_int) also have "{a..of_int \a\} \ \ = {a..real_of_int \x\}" using 3 * by (intro ivl_disj_un) auto also have "\ \ {of_int \x\..x} = {a..x}" using 3 * by (intro ivl_disj_un) auto finally show ?thesis . qed qed have "(\z. \X\A. integral X (\t. of_real (bernpoly n (t - \Inf X\)) * f z t)) holomorphic_on U" proof (intro holomorphic_on_sum holo, goal_cases) case (1 X) from 1 and division have subset: "X \ {a..x}" by (auto simp: division_of_def) from 1 obtain b c where [simp]: "X = cbox b c" "b \ c" by (auto simp: A_def) from subset have "b \ a" "c \ x" by auto hence "(\x. integral (cbox b c) (\t. of_real (bernpoly n (t - \Inf {b..c}\)) * f x t)) holomorphic_on U" by (intro holo) auto thus ?case by simp qed also have "?this \ (\z. integral {a..x} (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" proof (intro holomorphic_cong refl, goal_cases) case (1 z) have "((\t. of_real (pbernpoly n t) * f z t) has_integral (\X\A. integral X (\t. of_real (bernpoly n (t - \Inf X\)) * f z t))) {a..x}" using division proof (rule has_integral_combine_division) fix X assume X: "X \ A" then obtain b c where X': "X = {b..c}" "b \ c" by (elim A_cases) auto from X and division have "X \ {a..x}" by (auto simp: division_of_def) with X' have bc: "b \ a" "c \ x" by auto have "((\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t) has_integral integral X (\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t)) X" unfolding X' using \z \ U\ bc by (intro integrable_integral int) also have "?this \ ((\t. of_real (pbernpoly n t) * f z t) has_integral integral X (\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t)) X" proof (rule has_integral_spike_eq[of "{Sup X}"], goal_cases) case (2 t) note t = this from \X \ A\ have "\t\ = \Inf X\" proof (cases rule: A_cases [consumes 1]) case 1 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef split: if_splits, (linarith+)?) next case 2 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef split: if_splits, (linarith+)?) next case 3 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef N_def split: if_splits) qed thus ?case by (simp add: pbernpoly_def frac_def) qed auto finally show \ . qed thus ?case by (simp add: has_integral_iff) qed finally show ?thesis by simp qed auto thus "(\z. integral {a..x} (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" by simp qed lemma assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes deriv': "\z t x. z \ U \ x \ a \ t \ {a..x} \ ((\z. f z t) has_field_derivative f' z t) (at z within U)" assumes cont: "continuous_on (U \ {of_int a..}) (\(z, t). f' z t)" assumes int: "\b c z e. a \ b \ z \ U \ (\t. of_real (bernpoly n (t - e)) * f z t) integrable_on {b..c}" assumes int': "\a' b y. y \ U \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\U. norm (f y x) \ g x) at_top" assumes "open U" shows analytic_EM_remainder: "(\s::complex. EM_remainder n (f s) a) analytic_on U" and holomorphic_EM_remainder: "(\s::complex. EM_remainder n (f s) a) holomorphic_on U" proof - show "(\s::complex. EM_remainder n (f s) a) analytic_on U" unfolding analytic_on_def proof fix z assume "z \ U" from \z \ U\ and \open U\ obtain \ where \: "\ > 0" "ball z \ \ U" by (auto simp: open_contains_ball) have "(\s. EM_remainder n (f s) a) holomorphic_on ball z \" proof (rule holomorphic_uniform_sequence) fix x :: nat show "(\s. EM_remainder' n (f s) a x) holomorphic_on ball z \" proof (rule holomorphic_EM_remainder', goal_cases) fix s t assume "s \ ball z \" "t \ {real_of_int a..real x}" thus "((\z. f z t) has_field_derivative f' s t) (at s within ball z \)" using \ by (intro DERIV_subset[OF deriv'[of _ x]]) auto next case (2 b c s e) with \ have "s \ U" by blast with 2 show ?case using \ int[of b s e c] by (cases "a \ x") auto next from cont show "continuous_on (ball z \ \ {real_of_int a..real x}) (\(z, t). f' z t)" by (rule continuous_on_subset) (insert \, auto) qed (auto) next fix s assume s: "s \ ball z \" have "open (ball z \)" by simp with s obtain \ where \: "\ > 0" "cball s \ \ ball z \" unfolding open_contains_cball by blast moreover have bound': "eventually (\x. \y\cball s \. norm (f y x) \ g x) at_top" by (intro eventually_mono [OF bound]) (insert \ \, auto) have "uniform_limit (cball s \) (\x s. EM_remainder' n (f s) (real_of_int a) (real x)) (\s. EM_remainder n (f s) a) sequentially" by (rule uniform_limit_EM_remainder[OF deriv int' conv bound']) (insert \ \ s, auto) ultimately show "\\>0. cball s \ \ ball z \ \ uniform_limit (cball s \) (\x s. EM_remainder' n (f s) (real_of_int a) (real x)) (\s. EM_remainder n (f s) a) sequentially" by blast qed auto with \ show "\\>0. (\s. EM_remainder n (f s) a) holomorphic_on ball z \" by blast qed thus "(\s::complex. EM_remainder n (f s) a) holomorphic_on U" by (rule analytic_imp_holomorphic) qed text \ The following lemma is the first step in the proof of the Euler--MacLaurin formula: We show the relationship between the first-order remainder term and the difference of the integral and the sum. \ context fixes f f' :: "real \ 'a :: banach" fixes a b :: int and I S :: 'a fixes Y :: "real set" assumes "a \ b" assumes fin: "finite Y" assumes cont: "continuous_on {real_of_int a..real_of_int b} f" assumes deriv [derivative_intros]: "\x::real. x \ {a..b} - Y \ (f has_vector_derivative f' x) (at x)" defines S_def: "S \ (\i\{a<..b}. f i)" and I_def: "I \ integral {a..b} f" begin lemma diff_sum_integral_has_integral_int: "((\t. (frac t - 1/2) *\<^sub>R f' t) has_integral (S - I - (f b - f a) /\<^sub>R 2)) {a..b}" proof (cases "a = b") case True thus ?thesis by (simp_all add: S_def I_def has_integral_refl) next case False with \a \ b\ have ab: "a < b" by simp let ?A = "(\n. {real_of_int n..real_of_int (n+1)}) ` {a.. {of_int a..of_int b}" for A using continuous_on_subset[OF cont that] . define d where "d = (\x. (f x + f (x + 1)) /\<^sub>R 2 - integral {x..x+1} f)" have "((\t. (frac t - 1/2) *\<^sub>R f' t) has_integral d i) {of_int i..of_int (i+1)}" if i: "i \ {a..R f' x = (x - of_int i - 1 / 2) *\<^sub>R f' x" if "x \ {of_int i..of_int (i + 1)} - {of_int (i + 1)}" for x proof - have "x \ of_int i" "x < of_int (i + 1)" using that by auto hence "floor x = of_int i" by (subst floor_unique) auto thus ?thesis by (simp add: frac_def) qed next define h where "h = (\x::real. (x - of_int i - 1 / 2) *\<^sub>R f' x)" define g where "g = (\x::real. (x - of_int i - 1/2) *\<^sub>R f x - integral {of_int i..x} f)" have *: "((\x. integral {real_of_int i..x} f) has_vector_derivative f x) (at x within {i..i+1})" if "x \ {of_int i<..x. integral {real_of_int i..x} f) has_vector_derivative f x) (at x)" if "x \ {of_int i<..I. I\?A \ ((\x. (frac x - 1 / 2) *\<^sub>R f' x) has_integral d (\Inf I\)) I" by (auto simp: add_ac) have "((\x::real. (frac x - 1 / 2) *\<^sub>R f' x) has_integral (\I\?A. d (\Inf I\))) (\?A)" by (intro has_integral_Union * finite_imageI) (force intro!: negligible_atLeastAtMostI pairwiseI)+ also have "\?A = {of_int a..of_int b}" by (intro Union_atLeastAtMost_real_of_int ab) also have "(\I\?A. d (\Inf I\)) = (\i=a.. = (1 / 2) *\<^sub>R ((\i = a..i = a..i = a..R (?S1 + ?S2) - ?S3") by (simp add: d_def algebra_simps sum.distrib sum_subtractf scaleR_sum_right) also have "?S1 = (\i = a..b. f (real_of_int i)) - f b" unfolding S_def using ab by (subst sum_atLeastAtMost_int_last) auto also have "(\i = a..b. f (real_of_int i)) = S + f a" unfolding S_def using ab by (subst sum_atLeastAtMost_int_head) auto also have "?S2 = S" unfolding S_def by (intro sum.reindex_bij_witness[of _ "\i. i-1" "\i. i+1"]) auto also have "(1 / 2) *\<^sub>R (S + f a - f b + S) = (1/2) *\<^sub>R S + (1/2) *\<^sub>R S - (f b - f a) /\<^sub>R 2" by (simp add: algebra_simps) also have "(1/2) *\<^sub>R S + (1/2) *\<^sub>R S = S" by (simp add: scaleR_add_right [symmetric]) also have "?S3 = (\I\?A. integral I f)" by (subst sum.reindex) (auto simp: inj_on_def add_ac) also have "\ = I" unfolding I_def by (intro integral_combine_division_topdown [symmetric] division integrable_continuous_real continuous_intros) simp_all finally show ?thesis by (simp add: algebra_simps) qed lemma diff_sum_integral_has_integral_int': "((\t. pbernpoly 1 t *\<^sub>R f' t) has_integral (S - I - (f b - f a) /\<^sub>R 2 )) {a..b}" using diff_sum_integral_has_integral_int by (simp add: pbernpoly_def bernpoly_def) lemma EM_remainder'_Suc_0: "EM_remainder' (Suc 0) f' a b = S - I - (f b - f a) /\<^sub>R 2" using diff_sum_integral_has_integral_int' by (simp add: has_integral_iff EM_remainder'_def) end text \ Next, we show that the $n$-th-order remainder can be expressed in terms of the $n+1$-th-order remainder term. Iterating this essentially yields the Euler--MacLaurin formula. \ context fixes f f' :: "real \ 'a :: banach" and a b :: int and n :: nat and A :: "real set" assumes ab: "a \ b" and n: "n > 0" assumes fin: "finite A" assumes cont: "continuous_on {of_int a..of_int b} f" assumes cont': "continuous_on {of_int a..of_int b} f'" assumes deriv: "\x. x \ {of_int a<.. (f has_vector_derivative f' x) (at x)" begin lemma EM_remainder'_integral_conv_Suc: shows "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = (bernoulli (Suc n) / real (Suc n)) *\<^sub>R (f b - f a) - integral {a..b} (\t. pbernpoly (Suc n) t *\<^sub>R f' t) /\<^sub>R real (Suc n)" unfolding EM_remainder'_def proof - let ?h = "\i. (pbernpoly (Suc n) (real_of_int i) / real (Suc n)) *\<^sub>R f (real_of_int i)" define T where "T = integral {a..b} (\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t)" note [derivative_intros] = has_field_derivative_pbernpoly_Suc' let ?A = "real_of_int ` {a..b} \ A" have "((\t. pbernpoly n t *\<^sub>R f t) has_integral (-T + (?h b - ?h a))) {a..b}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_scaleR]) from fin show "finite ?A" by simp from \n > 0\ show "continuous_on {of_int a..of_int b} (\t. pbernpoly (Suc n) t / real (Suc n))" by (intro continuous_intros) auto show "continuous_on {of_int a..of_int b} f" by fact show "(f has_vector_derivative f' t) (at t)" if "t \ {of_int a<..t. pbernpoly (Suc n) t *\<^sub>R f' t) integrable_on {a..b}" by (intro integrable_EM_remainder' cont') hence "(\t. (1 / real (Suc n)) *\<^sub>R pbernpoly (Suc n) t *\<^sub>R f' t) integrable_on {a..b}" by (rule integrable_cmul) also have "(\t. (1 / real (Suc n)) *\<^sub>R pbernpoly (Suc n) t *\<^sub>R f' t) = (\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t)" by (rule ext) (simp add: algebra_simps) finally show "((\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t) has_integral ?h b - ?h a - (- T + (?h b - ?h a))) {a..b}" using integrable_EM_remainder'[of a b f' "Suc n"] by (simp add: has_integral_iff T_def) qed (insert ab n, auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] not_le elim!: Ints_cases) also have "?h b - ?h a = (bernoulli (Suc n) / real (Suc n)) *\<^sub>R (f b - f a)" using n by (simp add: algebra_simps bernoulli'_def) finally have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = \ - T" by (simp add: has_integral_iff) also have "T = integral {a..b} (\t. (1 / real (Suc n)) *\<^sub>R (pbernpoly (Suc n) t) *\<^sub>R f' t)" by (simp add: T_def) also have "\ = integral {a..b} (\t. pbernpoly (Suc n) t *\<^sub>R f' t) /\<^sub>R real (Suc n)" by (subst integral_cmul) (simp_all add: divide_simps) finally show ?thesis . qed lemma EM_remainder'_conv_Suc: "EM_remainder' n f a b = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (f b - f a) + EM_remainder' (Suc n) f' a b" by (simp add: EM_remainder'_def EM_remainder'_integral_conv_Suc scaleR_diff_right scaleR_add_right field_simps del: of_nat_Suc) end context fixes f f' :: "real \ 'a :: banach" and a :: int and n :: nat and A :: "real set" and C assumes n: "n > 0" assumes fin: "finite A" assumes cont: "continuous_on {of_int a..} f" assumes cont': "continuous_on {of_int a..} f'" assumes lim: "(f \ C) at_top" assumes deriv: "\x. x \ {of_int a<..} - A \ (f has_vector_derivative f' x) (at x)" begin lemma shows EM_remainder_converges_iff_Suc_converges: "EM_remainder_converges n f a \ EM_remainder_converges (Suc n) f' a" and EM_remainder_conv_Suc: "EM_remainder_converges n f a \ EM_remainder n f a = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (C - f a) + EM_remainder (Suc n) f' a" proof (rule iffI) define g where "g = (\x. ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (f x - f a))" define G where "G = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (C - f a)" have limit_g: "(g \ G) at_top" unfolding g_def G_def by (intro tendsto_intros lim) have *: "eventually (\x. EM_remainder' n f (real_of_int a) (real_of_int x) = g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)) at_top" using eventually_ge_at_top[of a] proof eventually_elim case (elim b) thus ?case using EM_remainder'_conv_Suc[OF elim n fin continuous_on_subset[OF cont] continuous_on_subset[OF cont'] deriv] by (auto simp: g_def) qed { assume "EM_remainder_converges n f a" then obtain L where L: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have *: "((\b. EM_remainder' (Suc n) f' (real_of_int a) (real_of_int b)) \ L - G) at_top" proof (rule Lim_transform_eventually) show "\\<^sub>F x in at_top. EM_remainder' n f (real_of_int a) (real_of_int x) - g x = EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)" using * by (simp add: algebra_simps) show "((\x. EM_remainder' n f (real_of_int a) (real_of_int x) - g x) \ L - G) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g] L) qed from * show "EM_remainder_converges (Suc n) f' a" unfolding EM_remainder_converges_def .. from * have "EM_remainder (Suc n) f' a = L - G" by (rule EM_remainder_eqI) moreover from L have "EM_remainder n f a = L" by (rule EM_remainder_eqI) ultimately show "EM_remainder n f a = G + EM_remainder (Suc n) f' a" by (simp add: G_def) } { assume "EM_remainder_converges (Suc n) f' a" then obtain L where L: "((\b. EM_remainder' (Suc n) f' (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have *: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ G + L) at_top" proof (rule Lim_transform_eventually) show "\\<^sub>F x in at_top. g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x) = EM_remainder' n f (real_of_int a) (real_of_int x)" using * by (subst eq_commute) show "((\x. g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)) \ G + L) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g] L) qed thus "EM_remainder_converges n f a" unfolding EM_remainder_converges_def .. } qed end subsection \The conventional version of the Euler--MacLaurin formula\ text \ The following theorems are the classic Euler--MacLaurin formula that can be found, with slight variations, in many sources (e.\,g.\ \cite{AS_HMF, apostol99, GKP_CM}). \ context fixes f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a b :: int assumes ab: "a \ b" fixes N :: nat assumes N: "N > 0" fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real_of_int a..real_of_int b} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {a..b} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" begin theorem euler_maclaurin_raw_strong_int: defines "S \ (\i\{a<..b}. f (of_int i))" defines "I \ integral {of_int a..of_int b} f" defines "c' \ \k. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (fs k b - fs k a)" shows "S - I = (\k 'a" where "c = (\k. ((-1) ^ (Suc k) * bernoulli (Suc k) / fact (Suc k)) *\<^sub>R (fs k b - fs k a))" have "S - I = (\k 1" "m \ N" for m using that proof (induction m rule: dec_induct) case base with ab fin fs_cont[of 0] show ?case using fs_deriv[of 0] N unfolding One_nat_def by (subst EM_remainder'_Suc_0[of _ _ Y f]) (simp_all add: algebra_simps S_def I_def c_def) next case (step n) from step.prems have "S - I = (\kkkR (fs n b - fs n a))" (is "_ = _ + ?c") by (simp add: EM_remainder'_Suc_0 c_def) also have "\ + EM_remainder' n (fs n) a b = (\k b" "0 < N" "finite Y" "fs 0 = f" "(\k. k \ N \ continuous_on {real a..real b} (fs k))" "(\k x. k < N \ x \ {real a..real b} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x))" shows "(\i\{a<..b}. f (real i)) - integral {real a..real b} f = (\kR (fs k (real b) - fs k (real a))) + EM_remainder' N (fs N) (real a) (real b)" proof - have "(\i\{int a<..int b}. f (real_of_int i)) - integral {real_of_int (int a)..real_of_int (int b)} f = (\kR (fs k (real_of_int (int b)) - fs k (real_of_int (int a)))) + EM_remainder' N (fs N) (real_of_int (int a)) (real_of_int (int b))" using assms by (intro euler_maclaurin_raw_strong_int[where Y = Y] assms) simp_all also have "(\i\{int a<..int b}. f (real_of_int i)) = (\i\{a<..b}. f (real i))" by (intro sum.reindex_bij_witness[of _ int nat]) auto finally show ?thesis by simp qed subsection \The ``Concrete Mathematics'' version of the Euler--MacLaurin formula\ text \ As explained in \textit{Concrete Mathematics}~\cite{GKP_CM}, the above form of the formula has some drawbacks: When applying it to determine the asymptotics of some concrete function, one is usually left with several different unwieldy constant terms that are difficult to get rid of. There is no general way to determine what these constant terms are, but in concrete applications, they can often be determined or estimated by other means. We can therefore simply group all the constant terms into a single constant and have the user provide a proof of what it is. \ locale euler_maclaurin_int = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: int fixes N :: nat assumes N: "N > 0" fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real_of_int a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {of_int a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {of_int a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {of_int a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (of_int b) - (\iR fs i (of_int b))) \ C) at_top" begin context fixes C' T defines "C' \ -f a + F a + C + (\kR (fs k (of_int a)))" and "T \ (\x. \iR fs i x)" begin lemma euler_maclaurin_strong_int_aux: assumes ab: "a \ b" defines "S \ (\k=a..b. f (of_int k))" shows "S - F (of_int b) - T (of_int b) = EM_remainder' N (fs N) (of_int a) (of_int b) + (C - C')" proof (cases "a = b") case True thus ?thesis unfolding C'_def by (simp add: S_def EM_remainder'_def T_def) next case False with assms have ab: "a < b" by simp define T' where "T' = (\kR (fs k (of_int a)))" have "(\i\{a<..b}. f (of_int i)) - integral {of_int a..of_int b} f = (\kR (fs k (of_int b) - fs k (of_int a))) + EM_remainder' N (fs N) (of_int a) (of_int b)" using ab by (intro euler_maclaurin_raw_strong_int [where Y = Y] N fin fs_0 continuous_on_subset[OF fs_cont] fs_deriv) auto also have "(f has_integral (F b - F a)) {of_int a..of_int b}" using ab by (intro fundamental_theorem_of_calculus_strong[OF fin]) (auto intro!: continuous_on_subset[OF F_cont] derivative_intros) hence "integral {of_int a..of_int b} f = F (of_int b) - F (of_int a)" by (simp add: has_integral_iff) also have "(\kR (fs k (of_int b) - fs k (of_int a))) = T (of_int b) - T'" by (simp add: T_def T'_def algebra_simps sum_subtractf) also have "(\i\{a<..b}. f (of_int i)) = S - f (of_int a)" unfolding S_def using ab by (subst sum_atLeastAtMost_int_head) auto finally show ?thesis by (simp add: algebra_simps C'_def T'_def) qed lemma EM_remainder_limit: assumes ab: "a \ b" defines "D \ EM_remainder' N (fs N) (of_int a) (of_int b)" shows "EM_remainder N (fs N) b = C' - D" and EM_remainder_converges: "EM_remainder_converges N (fs N) b" proof - note limit also have "((\b. (\k = a..b. f (of_int k)) - F (of_int b) - (\iR fs i (of_int b))) \ C) at_top = ((\b. (\k = a..b. f (of_int k)) - F (of_int b) - T (of_int b)) \ C) at_top" unfolding T_def .. also have "eventually (\x. (\k=a..x. f k) - F (of_int x) - T (of_int x) = EM_remainder' N (fs N) (of_int a) (of_int x) + (C - C')) at_top" (is "eventually (\x. ?f x = ?g x) _") using eventually_gt_at_top[of b] by eventually_elim (rule euler_maclaurin_strong_int_aux, insert ab, simp_all) hence "(?f \ C) at_top \ (?g \ C) at_top" by (intro filterlim_cong refl) finally have "((\x. ?g x - (C - C')) \ (C - (C - C'))) at_top" by (rule tendsto_diff[OF _ tendsto_const]) hence *: "((\x. EM_remainder' N (fs N) (of_int a) (of_int x)) \ C') at_top" by simp have "((\x. EM_remainder' N (fs N) (of_int a) (of_int x) - D) \ C' - D) at_top" by (intro tendsto_intros *) also have "eventually (\x. EM_remainder' N (fs N) (of_int a) (of_int x) - D = EM_remainder' N (fs N) (of_int b) (of_int x)) at_top" (is "eventually (\x. ?f x = ?g x) _") using eventually_ge_at_top[of b] proof eventually_elim case (elim x) have "EM_remainder' N (fs N) (of_int a) (of_int x) = D + EM_remainder' N (fs N) (of_int b) (of_int x)" using elim ab unfolding D_def by (intro EM_remainder'_combine [symmetric] continuous_on_subset[OF fs_cont]) auto thus ?case by simp qed hence "(?f \ C' - D) at_top \ (?g \ C' - D) at_top" by (intro filterlim_cong refl) finally have *: \ . from * show "EM_remainder_converges N (fs N) b" unfolding EM_remainder_converges_def .. from * show "EM_remainder N (fs N) b = C' - D" by (rule EM_remainder_eqI) qed theorem euler_maclaurin_strong_int: assumes ab: "a \ b" defines "S \ (\k=a..b. f (of_int k))" shows "S = F (of_int b) + C + T (of_int b) - EM_remainder N (fs N) b" proof - have "S = F (of_int b) + T (of_int b) + - (C' - EM_remainder' N (fs N) (of_int a) (of_int b)) + C" using euler_maclaurin_strong_int_aux[OF ab] by (simp add: algebra_simps S_def) also have "C' - EM_remainder' N (fs N) (of_int a) (of_int b) = EM_remainder N (fs N) b" using ab by (rule EM_remainder_limit(1) [symmetric]) finally show ?thesis by simp qed end end text \ The following version of the formula removes all the terms where the associated Bernoulli numbers vanish. \ locale euler_maclaurin_int' = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: int fixes N :: nat fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ 2*N+1 \ continuous_on {real_of_int a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k \ 2*N \ x \ {of_int a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {of_int a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {of_int a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (of_int b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b))) \ C) at_top" begin sublocale euler_maclaurin_int F f fs a "2*N+1" C Y by standard (insert fin fs_0 fs_cont fs_deriv F_cont F_deriv limit, simp_all) theorem euler_maclaurin_strong_int': assumes "a \ b" shows "(\k=a..b. f (of_int k)) = F (of_int b) + C + (1 / 2) *\<^sub>R f (of_int b) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" proof - have "(\k=a..b. f (real_of_int k)) = F (of_int b) + C + (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" by (rule euler_maclaurin_strong_int) (simp_all only: lessThan_Suc_atMost Suc_eq_plus1 [symmetric] assms) also have "{..<2*N+1} = insert 0 {1..2*N}" by auto also have "(\i\\. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) = (1 / 2) *\<^sub>R f (of_int b) + (\i\{1..2*N}. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b))" by (subst sum.insert) (auto simp: assms bernoulli'_def) also have "(\i\{1..2*N}. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) = (\i\{1..N}. (bernoulli' (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b))" proof (rule sym, rule sum.reindex_bij_witness_not_neutral) fix i assume "i \ {1..2*N} - {i\{1..2*N}. even i}" thus "2 * ((i + 1) div 2) - 1 = i" "(i + 1) div 2 \ {1..N} - {}" by (auto elim!: oddE) qed (auto simp: bernoulli_odd_eq_0 bernoulli'_def algebra_simps) also have "\ = (\i\{1..N}. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b))" by (intro sum.cong refl) (auto simp: bernoulli'_def) finally show ?thesis by (simp only: add_ac) qed end text \ For convenience, we also offer a version where the sum ranges over natural numbers instead of integers. \ lemma sum_atLeastAtMost_of_int_nat_transfer: "(\k=int a..int b. f (of_int k)) = (\k=a..b. f (of_nat k))" by (intro sum.reindex_bij_witness[of _ int nat]) auto lemma euler_maclaurin_nat_int_transfer: fixes F and f :: "real \ 'a :: real_normed_vector" assumes "((\b. (\k=a..b. f (real k)) - F (real b) - T (real b)) \ C) at_top" shows "((\b. (\k=int a..b. f (of_int k)) - F (of_int b) - T (of_int b)) \ C) at_top" proof - have *: "((\b. (\k=int a..int b. f (of_int k)) - F (of_int (int b)) - T (of_int (int b))) \ C) at_top" using assms by (subst sum_atLeastAtMost_of_int_nat_transfer) simp thus ?thesis by (rule filterlim_int_of_nat_at_topD) qed locale euler_maclaurin_nat = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: nat fixes N :: nat assumes N: "N > 0" fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {real a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {real a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {real a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (real b) - (\iR fs i (real b))) \ C) at_top" begin sublocale euler_maclaurin_int F f fs "int a" N C Y by standard (insert N fin fs_cont fs_deriv F_cont F_deriv euler_maclaurin_nat_int_transfer[OF limit], simp_all) theorem euler_maclaurin_strong_nat: assumes ab: "a \ b" defines "S \ (\k=a..b. f (real k))" shows "S = F (real b) + C + (\iR fs i (real b)) - EM_remainder N (fs N) (int b)" using euler_maclaurin_strong_int[of "int b"] by (simp add: assms sum_atLeastAtMost_of_int_nat_transfer) end locale euler_maclaurin_nat' = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: nat fixes N :: nat fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ 2*N+1 \ continuous_on {real a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k \ 2*N \ x \ {real a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {real a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {real a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ C) at_top" begin sublocale euler_maclaurin_int' F f fs "int a" N C Y by standard (insert fin fs_cont fs_deriv F_cont F_deriv euler_maclaurin_nat_int_transfer[OF limit], simp_all) theorem euler_maclaurin_strong_nat': assumes "a \ b" shows "(\k=a..b. f (real k)) = F (real b) + C + (1 / 2) *\<^sub>R f (real b) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" using euler_maclaurin_strong_int'[of b] by (simp add: assms sum_atLeastAtMost_of_int_nat_transfer) end subsection \Bounds on the remainder term\ text \ The following theorems provide some simple means to bound the remainder terms. In practice, better bounds can often be obtained e.\,g. for the $n$-th remainder term by expanding it to the sum of the first discarded term in the expansion and the $n+1$-th remainder term. \ lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes pbernpoly_bound: "\x. \pbernpoly n x\ \ D" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ C) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative g' x) (at x)" shows norm_EM_remainder_le_strong_int: "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * (C - g x)" and norm_EM_remainder_le_strong_nat: "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * (C - g x)" proof - from pbernpoly_bound have D: "norm (pbernpoly n x) \ D" for x by auto from this[of 0] have D_nonneg: "D \ 0" by simp define D' where "D' = D / fact n" from D_nonneg have D'_nonneg: "D' \ 0" by (simp add: D'_def) have bound: "norm (EM_remainder' n f x y) \ D' * (g y - g x)" if xy: "x \ a" "x \ y" for x y :: real proof - have "norm (EM_remainder' n f x y) = norm (integral {x..y} (\t. pbernpoly n t *\<^sub>R f t)) / fact n" by (simp add: EM_remainder'_def) also have "(\t. D * g' t) integrable_on {x..y}" using xy by (intro integrable_continuous_real continuous_intros continuous_on_subset[OF cont_g']) auto hence "norm (integral {x..y} (\t. pbernpoly n t *\<^sub>R f t)) \ integral {x..y} (\t. D * g' t)" using D D_nonneg xy by (intro integral_norm_bound_integral integrable_EM_remainder' continuous_on_subset[OF cont_f]) (auto intro!: mult_mono f_bound) also have "\ = D * integral {x..y} g'" by simp also have "(g' has_integral (g y - g x)) {x..y}" using xy by (intro fundamental_theorem_of_calculus_strong[OF fin] continuous_on_subset[OF cont_g]) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: deriv) hence "integral {x..y} g' = g y - g x" by (simp add: has_integral_iff) finally show ?thesis by (simp add: D'_def divide_simps) qed have lim: "((\y. EM_remainder' n f x (of_int y)) \ EM_remainder n f x) at_top" if x: "x \ a" for x :: int proof - have "(\n. g (real n)) \ C" by (rule filterlim_compose[OF limit_g filterlim_real_sequentially]) hence Cauchy: "Cauchy (\n. g (real n))" using convergent_eq_Cauchy by blast have "Cauchy (\y. EM_remainder' n f x (int y))" proof (rule CauchyI', goal_cases) case (1 \) define \' where "\' = (if D' = 0 then 1 else \ / (2*D'))" from \\ > 0\ D'_nonneg have \': "\' > 0" by (simp add: \'_def divide_simps) from CauchyD[OF Cauchy this] obtain M where M: "\m n. m \ M \ n \ M \ norm (g (real m) - g (real n)) < \'" by blast show ?case proof (intro CauchyI' exI[of _ "max (max 0 M) (nat x)"] allI impI, goal_cases) case (1 k l) have "EM_remainder' n f x k + EM_remainder' n f k l = EM_remainder' n f x l" using 1 x by (intro EM_remainder'_combine continuous_on_subset[OF cont_f]) auto hence "EM_remainder' n f x l - EM_remainder' n f x k = EM_remainder' n f k l" by (simp add: algebra_simps) also from 1 x have "norm \ \ D' * (g l - g k)" by (intro bound) auto also have "g l - g k \ norm (g l - g k)" by simp also from 1 have "\ \ \'" using M[of l k] by auto also from \\ > 0\ have "D' * \' \ \ / 2" by (simp add: \'_def) also from \\ > 0\ have "\ < \" by simp finally show ?case by (simp add: D'_nonneg mult_left_mono dist_norm norm_minus_commute) qed qed then obtain L where "(\y. EM_remainder' n f x (int y)) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) from filterlim_int_of_nat_at_topD[OF this] have "((\y. EM_remainder' n f x (of_int y)) \ L) at_top" by simp moreover from this have "EM_remainder n f x = L" by (rule EM_remainder_eqI) ultimately show "((\y. EM_remainder' n f x (of_int y)) \ EM_remainder n f x) at_top" by simp qed have *: "norm (EM_remainder n f x) \ D' * (C - g x)" if x: "x \ a" for x :: int proof (rule tendsto_le) show "((\y. D' * (g (of_int y) - g (of_int x))) \ D' * (C - g (of_int x))) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g]) show "((\y. norm (EM_remainder' n f x (of_int y))) \ norm (EM_remainder n f x)) at_top" using x by (intro tendsto_norm lim) show "eventually (\y. norm (EM_remainder' n f (of_int x) (of_int y)) \ D' * (g (of_int y) - g (of_int x))) at_top" using eventually_ge_at_top[of x] by eventually_elim (rule bound, insert x, simp_all) qed simp_all thus "\x. of_int x \ a \ norm (EM_remainder n f x) \ D' * (C - g x)" by blast have "norm (EM_remainder n f x) \ D' * (C - g x)" if x: "x \ a" for x :: nat using x *[of "int x"] by simp thus "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D' * (C - g x)" by blast qed lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes pbernpoly_bound: "\x. \pbernpoly n x\ \ D" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ 0) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative -g' x) (at x)" shows norm_EM_remainder_le_strong_int': "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" and norm_EM_remainder_le_strong_nat': "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * g x" proof - have "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * (0 - (-g x))" using assms by (intro norm_EM_remainder_le_strong_int[OF fin pbernpoly_bound _ _ cont_g']) (auto intro!: continuous_intros tendsto_eq_intros derivative_eq_intros) thus "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" by auto next have "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * (0 - (-g x))" using assms by (intro norm_EM_remainder_le_strong_nat[OF fin pbernpoly_bound _ _ cont_g']) (auto intro!: continuous_intros tendsto_eq_intros derivative_eq_intros) thus "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * g x" by auto qed lemma norm_EM_remainder'_le: fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes cont_f: "continuous_on {a..} f" assumes cont_g': "continuous_on {a..} g'" assumes f_bigo: "eventually (\x. norm (f x) \ g' x) at_top" assumes deriv: "eventually (\x. (g has_field_derivative g' x) (at x)) at_top" obtains C D where "eventually (\x. norm (EM_remainder' n f a x) \ C + D * g x) at_top" proof - note cont = continuous_on_subset[OF cont_f] continuous_on_subset[OF cont_g'] from bounded_pbernpoly[of n] obtain D where D: "\x. norm (pbernpoly n x) \ D" by blast from this[of 0] have D_nonneg: "D \ 0" by simp from eventually_conj[OF f_bigo eventually_conj[OF deriv eventually_ge_at_top[of a]]] obtain x0 where x0: "x0 \ a" "\x. x \ x0 \ norm (f x) \ g' x" "\x. x \ x0 \ (g has_field_derivative g' x) (at x)" by (auto simp: eventually_at_top_linorder) define C where "C = (norm (integral {a..x0} (\t. pbernpoly n t *\<^sub>R f t)) - D * g x0) / fact n" have "eventually (\x. norm (EM_remainder' n f a x) \ C + D / fact n * g x) at_top" using eventually_ge_at_top[of x0] proof eventually_elim case (elim x) have "integral {a..x} (\t. pbernpoly n t *\<^sub>R f t) = integral {a..x0} (\t. pbernpoly n t *\<^sub>R f t) + integral {x0..x} (\t. pbernpoly n t *\<^sub>R f t)" (is "_ = ?I1 + ?I2") using elim x0(1) - by (intro integral_combine [symmetric] integrable_EM_remainder' cont) auto + by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] integrable_EM_remainder' cont) auto also have "norm \ \ norm ?I1 + norm ?I2" by (rule norm_triangle_ineq) also have "norm ?I2 \ integral {x0..x} (\t. D * g' t)" using x0 D D_nonneg by (intro integral_norm_bound_integral integrable_EM_remainder') (auto intro!: integrable_continuous_real continuous_intros cont mult_mono) also have "\ = D * integral {x0..x} g'" by simp also from elim have "(g' has_integral (g x - g x0)) {x0..x}" by (intro fundamental_theorem_of_calculus) (auto intro!: has_field_derivative_at_within[OF x0(3)] simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "integral {x0..x} g' = g x - g x0" by (simp add: has_integral_iff) finally have "norm (integral {a..x} (\t. pbernpoly n t *\<^sub>R f t)) \ norm ?I1 + D * (g x - g x0)" by simp_all thus "norm (EM_remainder' n f a x) \ C + D / fact n * g x" by (simp add: EM_remainder'_def field_simps C_def) qed thus ?thesis by (rule that) qed subsection \Application to harmonic numbers\ text \ As a first application, we can apply the machinery we have developed to the harmonic numbers. \ definition harm_remainder :: "nat \ nat \ real" where "harm_remainder N n = EM_remainder (2*N+1) (\x. -fact (2*N+1) / x ^ (2*N+2)) (int n)" lemma harm_expansion: assumes n: "n > 0" and N: "N > 0" shows "harm n = ln n + euler_mascheroni + 1 / (2*n) - (\i=1..N. bernoulli (2*i) / ((2*i) * n ^ (2*i))) - harm_remainder N n" proof - define fs where "fs = (\k x. (-1) ^ k * fact k / x ^ (Suc k) :: real)" interpret euler_maclaurin_nat' ln "\x. 1/x" fs 1 N euler_mascheroni "{}" proof fix k x assume "k \ 2*N" "x \ {real 1..} - {}" thus "(fs k has_vector_derivative fs (Suc k) x) (at x)" by (cases "k = 0") (auto intro!: derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] field_simps power_diff) next have "(\b. harm b - ln (real b) - (\i<2*N+1. bernoulli' (Suc i) * (- 1) ^ i / (real (Suc i) * (real b ^ Suc i)))) \ (euler_mascheroni - (\i<2*N+1. 0))" by (intro tendsto_diff euler_mascheroni_LIMSEQ tendsto_sum real_tendsto_divide_at_top[OF tendsto_const] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_real_sequentially) auto thus "(\b. (\k = 1..b. 1 / real k) - ln (real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ euler_mascheroni" by (simp add: harm_def divide_simps fs_def) qed (insert n N, auto intro!: continuous_intros derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric]) have "harm n = (\k=1..n. 1 / real k)" by (simp add: harm_def divide_simps) also have "\ = ln (real n) + euler_mascheroni + (1/2) *\<^sub>R (1 / real n) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real n)) - EM_remainder (2*N+1) (fs (2*N+1)) (int n)" using n N using n by (intro euler_maclaurin_strong_nat') simp_all also have "(\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R (fs (2*i-1) (real n))) = (\i=1..N. -(bernoulli (2*i) / (real (2*i) * real n ^ (2*i))))" by (intro sum.cong refl) (simp_all add: fs_def divide_simps fact_reduce del: of_nat_Suc power_Suc) also have "\ = -(\i=1..N. bernoulli (2*i) / (real (2*i) * real n ^ (2*i)))" by (simp add: sum_negf) finally show ?thesis unfolding fs_def by (simp add: harm_remainder_def) qed lemma of_nat_ge_1_iff: "of_nat x \ (1 :: 'a :: linordered_semidom) \ x \ 1" using of_nat_le_iff[of 1 x] by (simp del: of_nat_le_iff) lemma harm_remainder_bound: fixes N :: nat assumes N: "N > 0" shows "\C. \n\1. norm (harm_remainder N n) \ C / real n ^ (2*N+1)" proof - from bounded_pbernpoly[of "2*N+1"] obtain D where D: "\x. \pbernpoly (2*N+1) x\ \ D" by auto have "\x. 1 \ real x \ norm (harm_remainder N x) \ D / fact (2*N+1) * (fact (2*N) / x ^ (2*N+1))" unfolding harm_remainder_def of_int_of_nat_eq proof (rule norm_EM_remainder_le_strong_nat'[of "{}"]) fix x :: real assume x: "x \ 1" show "norm (-fact (2*N+1) / x ^ (2 * N + 2)) \ fact (2*N+1) / x ^ (2*N+2)" using x by simp next show "((\x::real. fact (2 * N) / x ^ (2 * N + 1)) \ 0) at_top" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_ident) simp_all qed (insert N D, auto intro!: derivative_eq_intros continuous_intros simp: field_simps power_diff) hence "\x. 1 \ x \ norm (harm_remainder N x) \ D / (2*N+1) / real x ^ (2*N+1)" by simp thus ?thesis by blast qed subsection \Application to sums of inverse squares\ text \ In the same vein, we can derive the asymptotics of the partial sum of inverse squares. \ lemma sum_inverse_squares_expansion: assumes n: "n > 0" and N: "N > 0" shows "(\k=1..n. 1 / real k ^ 2) = pi ^ 2 / 6 - 1 / real n + 1 / (2 * real n ^ 2) - (\i=1..N. bernoulli (2*i) / n ^ (2*i+1)) - EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n)" proof - have 3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) define fs where "fs = (\k x. (-1) ^ k * fact (Suc k) / x ^ (k+2) :: real)" interpret euler_maclaurin_nat' "\x. -1/x" "\x. 1/x^2" fs 1 N "pi^2/6" "{}" proof fix k x assume "k \ 2*N" "x \ {real 1..} - {}" thus "(fs k has_vector_derivative fs (Suc k) x) (at x)" by (cases "k = 0") (auto intro!: derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] field_simps power_diff) next from inverse_squares_sums have "(\n. \k pi\<^sup>2 / 6" by (simp add: sums_def) also have "(\n. \kn. \k=1..n. 1 / real k ^ 2)" by (intro ext sum.reindex_bij_witness[of _ "\n. n - 1" Suc]) auto finally have "(\b. (\k = 1..b. 1 / real k^2) + 1 / real b - (\i<2*N+1. bernoulli' (Suc i) * (- 1) ^ i / (real b ^ (i+2)))) \ (pi^2/6 + 0 - (\i<2*N+1. 0))" by (intro tendsto_diff tendsto_add real_tendsto_divide_at_top[OF tendsto_const] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_real_sequentially tendsto_sum) auto thus "(\b. (\k = 1..b. 1 / real k^2) - (- 1 / real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ pi^2/6" by (simp add: harm_def field_simps fs_def del: power_Suc of_nat_Suc) qed (insert n N, auto intro!: continuous_intros derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] power2_eq_square) have "(\k=1..n. 1 / real k ^ 2) = - 1 / real n + pi^2/6 + (1/2) *\<^sub>R (1 / real n^2) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real n)) - EM_remainder (2*N+1) (fs (2*N+1)) (int n)" using n N using n by (intro euler_maclaurin_strong_nat') simp_all also have "(\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R (fs (2*i-1) (real n))) = (\i=1..N. -(bernoulli (2*i) / (real n ^ (2*i+1))))" by (intro sum.cong refl) (simp_all add: fs_def divide_simps fact_reduce del: of_nat_Suc power_Suc) also have "\ = -(\i=1..N. bernoulli (2*i) / real n ^ (2*i+1))" by (simp add: sum_negf) finally show ?thesis unfolding fs_def by (simp add: fs_def 3) qed lemma sum_inverse_squares_remainder_bound: fixes N :: nat assumes N: "N > 0" defines "R \ (\n. EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n))" shows "\C. \n\1. norm (R n) \ C / real n ^ (2*N+2)" proof - have 3: "3 = Suc (Suc (Suc 0))" by simp from bounded_pbernpoly[of "2*N+1"] obtain D where D: "\x. \pbernpoly (2*N+1) x\ \ D" by auto have "\x. 1 \ real x \ norm (R x) \ D / fact (2*N+1) * (fact (2*N+1) / x ^ (2*N+2))" unfolding R_def of_int_of_nat_eq proof (rule norm_EM_remainder_le_strong_nat'[of "{}"]) fix x :: real assume x: "x \ 1" show "norm (-fact (2*N+2) / x ^ (2*N+3)) \ fact (2*N+2) / x ^ (2*N+3)" using x by simp next show "((\x::real. fact (2*N+1) / x ^ (2*N+2)) \ 0) at_top" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_ident) simp_all qed (insert N D, auto intro!: derivative_eq_intros continuous_intros simp: field_simps power_diff 3) hence "\x\1. norm (R x) \ D / real x ^ (2 * N + 2)" by simp thus ?thesis by blast qed end diff --git a/thys/Green/Integrals.thy b/thys/Green/Integrals.thy --- a/thys/Green/Integrals.thy +++ b/thys/Green/Integrals.thy @@ -1,948 +1,946 @@ theory Integrals imports "HOL-Analysis.Analysis" General_Utils begin lemma gauge_integral_Fubini_universe_x: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" assumes fun_lesbegue_integrable: "integrable lborel f" and x_axis_integral_measurable: "(\x. integral UNIV (\y. f(x, y))) \ borel_measurable lborel" shows "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(x,y)))" "(\x. integral UNIV (\y. f(x,y))) integrable_on UNIV" proof - have f_is_measurable: "f \ borel_measurable lborel" using fun_lesbegue_integrable and borel_measurable_integrable by auto have fun_lborel_prod_integrable: "integrable (lborel \\<^sub>M lborel) f" using fun_lesbegue_integrable by (simp add: lborel_prod) then have region_integral_is_one_twoD_integral: "LBINT x. LBINT y. f (x, y) = integral\<^sup>L (lborel \\<^sub>M lborel) f" using lborel_pair.integral_fst' by auto then have AE_one_D_integrals_eq: "AE x in lborel. LBINT y. f (x, y) = integral UNIV (\y. f(x,y))" proof - have "AE x in lborel. integrable lborel (\y. f(x,y))" using lborel_pair.AE_integrable_fst' and fun_lborel_prod_integrable by blast then show ?thesis using integral_lborel and always_eventually and AE_mp by fastforce qed have one_D_integral_measurable: "(\x. LBINT y. f (x, y)) \ borel_measurable lborel" using f_is_measurable and lborel.borel_measurable_lebesgue_integral by auto then have second_lesbegue_integral_eq: "LBINT x. LBINT y. f (x, y) = LBINT x. (integral UNIV (\y. f(x,y)))" using x_axis_integral_measurable and integral_cong_AE and AE_one_D_integrals_eq by blast have "integrable lborel (\x. LBINT y. f (x, y))" using fun_lborel_prod_integrable and lborel_pair.integrable_fst' by auto then have oneD_gauge_integral_lesbegue_integrable: "integrable lborel (\x. integral UNIV (\y. f(x,y)))" using x_axis_integral_measurable and AE_one_D_integrals_eq and integrable_cong_AE_imp by blast then show one_D_gauge_integral_integrable: "(\x. integral UNIV (\y. f(x,y))) integrable_on UNIV" using integrable_on_lborel by auto have "LBINT x. (integral UNIV (\y. f(x,y))) = integral UNIV (\x. integral UNIV (\y. f(x, y)))" using integral_lborel oneD_gauge_integral_lesbegue_integrable by fastforce then have twoD_lesbeuge_eq_twoD_gauge: "LBINT x. LBINT y. f (x, y) = integral UNIV (\x. integral UNIV (\y. f(x, y)))" using second_lesbegue_integral_eq by auto then show "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(x,y)))" using fun_lesbegue_integrable and integral_lborel and region_integral_is_one_twoD_integral by (metis lborel_prod) qed lemma gauge_integral_Fubini_universe_y: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" assumes fun_lesbegue_integrable: "integrable lborel f" and y_axis_integral_measurable: "(\x. integral UNIV (\y. f(y, x))) \ borel_measurable lborel" shows "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(y, x)))" "(\x. integral UNIV (\y. f(y, x))) integrable_on UNIV" proof - have f_is_measurable: "f \ borel_measurable lborel" using fun_lesbegue_integrable and borel_measurable_integrable by auto have fun_lborel_prod_integrable: "integrable (lborel \\<^sub>M lborel) f" using fun_lesbegue_integrable by (simp add: lborel_prod) then have region_integral_is_one_twoD_integral: "LBINT x. LBINT y. f (y, x) = integral\<^sup>L (lborel \\<^sub>M lborel) f" using lborel_pair.integral_fst' f_is_measurable lborel_pair.integrable_product_swap lborel_pair.integral_fst lborel_pair.integral_product_swap lborel_prod by force then have AE_one_D_integrals_eq: "AE x in lborel. LBINT y. f (y, x) = integral UNIV (\y. f(y,x))" proof - have "AE x in lborel. integrable lborel (\y. f(y,x))" using lborel_pair.AE_integrable_fst' and fun_lborel_prod_integrable lborel_pair.AE_integrable_fst lborel_pair.integrable_product_swap by blast then show ?thesis using integral_lborel and always_eventually and AE_mp by fastforce qed have one_D_integral_measurable: "(\x. LBINT y. f (y, x)) \ borel_measurable lborel" using f_is_measurable and lborel.borel_measurable_lebesgue_integral by auto then have second_lesbegue_integral_eq: "LBINT x. LBINT y. f (y, x) = LBINT x. (integral UNIV (\y. f(y, x)))" using y_axis_integral_measurable and integral_cong_AE and AE_one_D_integrals_eq by blast have "integrable lborel (\x. LBINT y. f (y, x))" using fun_lborel_prod_integrable and lborel_pair.integrable_fst' by (simp add: lborel_pair.integrable_fst lborel_pair.integrable_product_swap) then have oneD_gauge_integral_lesbegue_integrable: "integrable lborel (\x. integral UNIV (\y. f(y, x)))" using y_axis_integral_measurable and AE_one_D_integrals_eq and integrable_cong_AE_imp by blast then show one_D_gauge_integral_integrable: "(\x. integral UNIV (\y. f(y, x))) integrable_on UNIV" using integrable_on_lborel by auto have "LBINT x. (integral UNIV (\y. f(y, x))) = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using integral_lborel oneD_gauge_integral_lesbegue_integrable by fastforce then have twoD_lesbeuge_eq_twoD_gauge: "LBINT x. LBINT y. f (y, x) = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using second_lesbegue_integral_eq by auto then show "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using fun_lesbegue_integrable and integral_lborel and region_integral_is_one_twoD_integral by (metis lborel_prod) qed lemma gauge_integral_Fubini_curve_bounded_region_x: fixes f g :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" and g1 g2:: "'a \ 'b" and s:: "('a * 'b) set" assumes fun_lesbegue_integrable: "integrable lborel f" and x_axis_gauge_integrable: "\x. (\y. f(x, y)) integrable_on UNIV" and (*IS THIS redundant? NO IT IS NOT*) x_axis_integral_measurable: "(\x. integral UNIV (\y. f(x, y))) \ borel_measurable lborel" and f_is_g_indicator: "f = (\x. if x \ s then g x else 0)" and s_is_bounded_by_g1_and_g2: "s = {(x,y). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. (g1 x) \ i \ y \ i \ y \ i \ (g2 x) \ i)}" shows "integral s g = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)))" proof - have two_D_integral_to_one_D: "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(x,y)))" using gauge_integral_Fubini_universe_x and fun_lesbegue_integrable and x_axis_integral_measurable by auto have one_d_integral_integrable: "(\x. integral UNIV (\y. f(x,y))) integrable_on UNIV" using gauge_integral_Fubini_universe_x(2) and assms by blast have case_x_in_range: "\ x \ cbox a b. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)) = integral UNIV (\y. f(x,y))" proof fix x:: 'a assume within_range: "x \ (cbox a b)" let ?f_one_D_spec = "(\y. if y \ (cbox (g1 x) (g2 x)) then g(x,y) else 0)" have f_one_D_region: "(\y. f(x,y)) = (\y. if y \ cbox (g1 x) (g2 x) then g(x,y) else 0)" proof fix y::'b show "f (x, y) = (if y \ (cbox (g1 x) (g2 x)) then g (x, y) else 0)" apply (simp add: f_is_g_indicator s_is_bounded_by_g1_and_g2) using within_range apply (simp add: cbox_def) by smt qed have zero_out_of_bound: "\ y. y \ cbox (g1 x) (g2 x) \ f (x,y) = 0" using f_is_g_indicator and s_is_bounded_by_g1_and_g2 by (auto simp add: cbox_def) have "(\y. f(x,y)) integrable_on cbox (g1 x) (g2 x)" proof - have "?f_one_D_spec integrable_on UNIV" using f_one_D_region and x_axis_gauge_integrable by metis then have "?f_one_D_spec integrable_on cbox(g1 x) (g2 x)" using integrable_on_subcbox by blast then show ?thesis using f_one_D_region by auto qed then have f_integrale_x: "((\y. f(x,y)) has_integral (integral (cbox (g1 x) (g2 x)) (\y. f(x,y)))) (cbox (g1 x) (g2 x))" using integrable_integral and within_range and x_axis_gauge_integrable by auto have "integral (cbox (g1 x) (g2 x)) (\y. f (x, y)) = integral UNIV (\y. f (x, y))" using has_integral_on_superset[OF f_integrale_x _ Set.subset_UNIV] zero_out_of_bound by (simp add: integral_unique) then have "((\y. f(x, y)) has_integral integral UNIV (\y. f (x, y))) (cbox (g1 x) (g2 x))" using f_integrale_x by simp then have "((\y. g(x, y)) has_integral integral UNIV (\y. f (x, y))) (cbox (g1 x)(g2 x))" using Henstock_Kurzweil_Integration.has_integral_restrict [OF subset_refl ] and f_one_D_region by (smt has_integral_eq) then show "integral (cbox (g1 x) (g2 x)) (\y. g (x, y)) = integral UNIV (\y. f (x, y))" by auto qed have case_x_not_in_range: "\ x. x \ cbox a b \ integral UNIV (\y. f(x,y)) = 0" proof fix x::'a have "x \ (cbox a b) \ (\y. f(x,y) = 0)" apply (simp add: s_is_bounded_by_g1_and_g2 f_is_g_indicator cbox_def) by auto then show "x \ cbox a b \ integral UNIV (\y. f (x, y)) = 0" by (simp) qed have RHS: "integral UNIV (\x. integral UNIV (\y. f(x,y))) = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)))" proof - let ?first_integral = "(\x. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)))" let ?x_integral_cases = "(\x. if x \ cbox a b then ?first_integral x else 0)" have x_integral_cases_integral: "(\x. integral UNIV (\y. f(x,y))) = ?x_integral_cases" using case_x_in_range and case_x_not_in_range by auto have "((\x. integral UNIV (\y. f(x,y))) has_integral (integral UNIV f)) UNIV" using two_D_integral_to_one_D one_d_integral_integrable by auto then have "(?x_integral_cases has_integral (integral UNIV f)) UNIV" using x_integral_cases_integral by auto then have "(?first_integral has_integral (integral UNIV f)) (cbox a b)" using has_integral_restrict_UNIV[of "cbox a b" "?first_integral" "integral UNIV f"] by auto then show ?thesis using two_D_integral_to_one_D by (simp add: integral_unique) qed have f_integrable:"f integrable_on UNIV" using fun_lesbegue_integrable and integrable_on_lborel by auto then have LHS: "integral UNIV f = integral s g" apply (simp add: f_is_g_indicator) using integrable_restrict_UNIV integral_restrict_UNIV by auto then show ?thesis using RHS and two_D_integral_to_one_D by auto qed lemma gauge_integral_Fubini_curve_bounded_region_y: fixes f g :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" and g1 g2:: "'b \ 'a" and s:: "('a * 'b) set" assumes fun_lesbegue_integrable: "integrable lborel f" and y_axis_gauge_integrable: "\x. (\y. f(y, x)) integrable_on UNIV" and (*IS THIS redundant? NO IT IS NOT*) y_axis_integral_measurable: "(\x. integral UNIV (\y. f(y, x))) \ borel_measurable lborel" and f_is_g_indicator: "f = (\x. if x \ s then g x else 0)" and s_is_bounded_by_g1_and_g2: "s = {(y, x). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. (g1 x) \ i \ y \ i \ y \ i \ (g2 x) \ i)}" shows "integral s g = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)))" proof - have two_D_integral_to_one_D: "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using gauge_integral_Fubini_universe_y and fun_lesbegue_integrable and y_axis_integral_measurable by auto have one_d_integral_integrable: "(\x. integral UNIV (\y. f(y, x))) integrable_on UNIV" using gauge_integral_Fubini_universe_y(2) and assms by blast have case_y_in_range: "\ x \ cbox a b. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)) = integral UNIV (\y. f(y, x))" proof fix x:: 'b assume within_range: "x \ (cbox a b)" let ?f_one_D_spec = "(\y. if y \ (cbox (g1 x) (g2 x)) then g(y, x) else 0)" have f_one_D_region: "(\y. f(y, x)) = (\y. if y \ cbox (g1 x) (g2 x) then g(y, x) else 0)" proof fix y::'a show "f (y, x) = (if y \ (cbox (g1 x) (g2 x)) then g (y, x) else 0)" apply (simp add: f_is_g_indicator s_is_bounded_by_g1_and_g2) using within_range apply (simp add: cbox_def) by smt qed have zero_out_of_bound: "\ y. y \ cbox (g1 x) (g2 x) \ f (y, x) = 0" using f_is_g_indicator and s_is_bounded_by_g1_and_g2 by (auto simp add: cbox_def) have "(\y. f(y, x)) integrable_on cbox (g1 x) (g2 x)" proof - have "?f_one_D_spec integrable_on UNIV" using f_one_D_region and y_axis_gauge_integrable by metis then have "?f_one_D_spec integrable_on cbox(g1 x) (g2 x)" using integrable_on_subcbox by blast then show ?thesis using f_one_D_region by auto qed then have f_integrale_y: "((\y. f(y, x)) has_integral (integral (cbox (g1 x) (g2 x)) (\y. f(y,x)))) (cbox (g1 x) (g2 x))" using integrable_integral and within_range and y_axis_gauge_integrable by auto have "integral (cbox (g1 x) (g2 x)) (\y. f (y, x)) = integral UNIV (\y. f (y, x))" using has_integral_on_superset[OF f_integrale_y _ Set.subset_UNIV] zero_out_of_bound by (simp add: integral_unique) then have "((\y. f(y, x)) has_integral integral UNIV (\y. f (y, x))) (cbox (g1 x) (g2 x))" using f_integrale_y by simp then have "((\y. g(y, x)) has_integral integral UNIV (\y. f (y, x))) (cbox (g1 x)(g2 x))" using Henstock_Kurzweil_Integration.has_integral_restrict [OF subset_refl ] and f_one_D_region by (smt has_integral_eq) then show "integral (cbox (g1 x) (g2 x)) (\y. g (y, x)) = integral UNIV (\y. f (y, x))" by auto qed have case_y_not_in_range: "\ x. x \ cbox a b \ integral UNIV (\y. f(y, x)) = 0" proof fix x::'b have "x \ (cbox a b) \ (\y. f(y, x) = 0)" apply (simp add: s_is_bounded_by_g1_and_g2 f_is_g_indicator cbox_def) by auto then show "x \ cbox a b \ integral UNIV (\y. f (y, x)) = 0" by (simp) qed have RHS: "integral UNIV (\x. integral UNIV (\y. f(y, x))) = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)))" proof - let ?first_integral = "(\x. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)))" let ?x_integral_cases = "(\x. if x \ cbox a b then ?first_integral x else 0)" have y_integral_cases_integral: "(\x. integral UNIV (\y. f(y, x))) = ?x_integral_cases" using case_y_in_range and case_y_not_in_range by auto have "((\x. integral UNIV (\y. f(y, x))) has_integral (integral UNIV f)) UNIV" using two_D_integral_to_one_D one_d_integral_integrable by auto then have "(?x_integral_cases has_integral (integral UNIV f)) UNIV" using y_integral_cases_integral by auto then have "(?first_integral has_integral (integral UNIV f)) (cbox a b)" using has_integral_restrict_UNIV[of "cbox a b" "?first_integral" "integral UNIV f"] by auto then show ?thesis using two_D_integral_to_one_D by (simp add: integral_unique) qed have f_integrable:"f integrable_on UNIV" using fun_lesbegue_integrable and integrable_on_lborel by auto then have LHS: "integral UNIV f = integral s g" apply (simp add: f_is_g_indicator) using integrable_restrict_UNIV integral_restrict_UNIV by auto then show ?thesis using RHS and two_D_integral_to_one_D by auto qed lemma gauge_integral_by_substitution: fixes f::"(real \ real)" and g::"(real \ real)" and g'::"real \ real" and a::"real" and b::"real" assumes a_le_b: "a \ b" and ga_le_gb: "g a \ g b" and g'_derivative: "\x \ {a..b}. (g has_vector_derivative (g' x)) (at x within {a..b})" and g'_continuous: "continuous_on {a..b} g'" and f_continuous: "continuous_on (g ` {a..b}) f" shows "integral {g a..g b} (f) = integral {a..b} (\x. f(g x) * (g' x))" proof - have "\x \ {a..b}. (g has_real_derivative (g' x)) (at x within {a..b})" using has_field_derivative_iff_has_vector_derivative[of "g"] and g'_derivative by auto then have 2: "interval_lebesgue_integral lborel (ereal (a)) (ereal (b)) (\x. g' x *\<^sub>R f (g x)) = interval_lebesgue_integral lborel (ereal (g a)) (ereal (g b)) f" using interval_integral_substitution_finite[of "a" "b" "g" "g'" "f"] and g'_continuous and a_le_b and f_continuous by auto have g_continuous: "continuous_on {a .. b} g" using Derivative.differentiable_imp_continuous_on apply (simp add: differentiable_on_def differentiable_def) by (metis continuous_on_vector_derivative g'_derivative) have "set_integrable lborel {a .. b} (\x. g' x *\<^sub>R f (g x))" proof - have "continuous_on {a .. b} (\x. g' x *\<^sub>R f (g x))" proof - have "continuous_on {a .. b} (\x. f (g x))" proof - show ?thesis using Topological_Spaces.continuous_on_compose f_continuous g_continuous by auto qed then show ?thesis using Limits.continuous_on_mult g'_continuous by auto qed then show ?thesis using borel_integrable_atLeastAtMost' by blast qed then have 0: "interval_lebesgue_integral lborel (ereal (a)) (ereal (b)) (\x. g' x *\<^sub>R f (g x)) = integral {a .. b} (\x. g' x *\<^sub>R f (g x))" using a_le_b and interval_integral_eq_integral by (metis (no_types)) have "set_integrable lborel {g a .. g b} f" proof - have "continuous_on {g a .. g b} f" proof - have "{g a .. g b} \ g ` {a .. b}" using g_continuous by (metis a_le_b atLeastAtMost_iff atLeastatMost_subset_iff continuous_image_closed_interval imageI order_refl) then show "continuous_on {g a .. g b} f" using f_continuous continuous_on_subset by blast qed then show ?thesis using borel_integrable_atLeastAtMost' by blast qed then have 1: "interval_lebesgue_integral lborel (ereal (g a)) (ereal (g b)) f = integral {g a .. g b} f" using ga_le_gb and interval_integral_eq_integral by (metis (no_types)) then show ?thesis using 0 and 1 and 2 by (metis (no_types, lifting) Henstock_Kurzweil_Integration.integral_cong mult.commute real_scaleR_def) qed lemma frontier_ic: assumes "a < (b::real)" shows "frontier {a<..b} = {a,b}" apply(simp add: frontier_def) using assms by auto lemma frontier_ci: assumes "a < (b::real)" shows "frontier {a<.. closed {a<..b}" using assms frontier_subset_eq frontier_ic greaterThanAtMost_iff by blast lemma closure_ic_union_ci: assumes "a < (b::real)" "b < c" shows "closure ({a.. {b<..c}) = {a .. c}" using frontier_ic[OF assms(1)] frontier_ci[OF assms(2)] closure_Un assms apply(simp add: frontier_def) by auto lemma interior_ic_ci_union: assumes "a < (b::real)" "b < c" shows "b \ (interior ({a.. {b<..c}))" proof- have "b \ ({a.. {b<..c})" by auto then show ?thesis using interior_subset by blast qed lemma frontier_ic_union_ci: assumes "a < (b::real)" "b < c" shows "b \ frontier ({a.. {b<..c})" using closure_ic_union_ci assms interior_ic_ci_union by(simp add: frontier_def) lemma ic_union_ci_not_closed: assumes "a < (b::real)" "b < c" shows "\ closed ({a.. {b<..c})" proof- have "b \ ({a.. {b<..c})" by auto then show ?thesis using assms frontier_subset_eq frontier_ic_union_ci[OF assms] by (auto simp only: subset_iff) qed lemma integrable_continuous_: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" by (simp add: assms integrable_continuous) lemma removing_singletons_from_div: assumes "\t\S. \c d::real. c < d \ {c..d} = t" "{x} \ \S = {a..b}" "a < x" "x < b" "finite S" shows "\t\S. x \ t" proof(rule ccontr) assume "\(\t\S. x \ t)" then have "\t\S. x \ t" by auto then have "x \ \S" by auto then have i: "\S = {a..b} - {x}" using assms (2) by auto have "x \ {a..b}" using assms by auto then have "{a .. b} - {x} = {a.. {x<..b}" by auto then have 0: "\S = {a.. {x<..b}" using i by auto have 1:"closed (\S)" apply(rule closed_Union) proof- show "finite S" using assms by auto show "\T\S. closed T" using assms by auto qed show False using 0 1 ic_union_ci_not_closed assms by auto qed lemma remove_singleton_from_division_of:(*By Manuel Eberl*) assumes "A division_of {a::real..b}" "a < b" assumes "x \ {a..b}" shows "\c d. c < d \ {c..d} \ A \ x \ {c..d}" proof - from assms have "x islimpt {a..b}" by (intro connected_imp_perfect) auto also have "{a..b} = {x. {x..x} \ A} \ ({a..b} - {x. {x..x} \ A})" using assms by auto also have "x islimpt \ \ x islimpt {a..b} - {x. {x..x} \ A}" proof (intro islimpt_Un_finite) have "{x. {x..x} \ A} \ Inf ` A" proof safe fix x assume "{x..x} \ A" thus "x \ Inf ` A" by (auto intro!: bexI[of _ "{x}"] simp: image_iff) qed moreover from assms have "finite A" by (auto simp: division_of_def) hence "finite (Inf ` A)" by auto ultimately show "finite {x. {x..x} \ A}" by (rule finite_subset) qed also have "{a..b} = \A" using assms by (auto simp: division_of_def) finally have "x islimpt \(A - range (\x. {x..x}))" by (rule islimpt_subset) auto moreover have "closed (\(A - range (\x. {x..x})))" using assms by (intro closed_Union) auto ultimately have "x \ (\(A - range (\x. {x..x})))" by (auto simp: closed_limpt) then obtain X where "x \ X" "X \ A" "X \ range (\x. {x..x})" by blast moreover from division_ofD(2)[OF assms(1) this(2)] division_ofD(3)[OF assms(1) this(2)] division_ofD(4)[OF assms(1) this(2)] obtain c d where "X = cbox c d" "X \ {a..b}" "X \ {}" by blast ultimately have "c \ d" by auto have "c \ d" proof assume "c = d" with \X = cbox c d\ have "X = {c..c}" by auto hence "X \ range (\x. {x..x})" by blast with \X \ range (\x. {x..x})\ show False by contradiction qed with \c \ d\ have "c < d" by simp with \X = cbox c d\ and \x \ X\ and \X \ A\ show ?thesis by auto qed lemma remove_singleton_from_tagged_division_of: assumes "A tagged_division_of {a::real..b}" "a < b" assumes "x \ {a..b}" shows "\k c d. c < d \ (k, {c..d}) \ A \ x \ {c..d}" using remove_singleton_from_division_of[OF division_of_tagged_division[OF assms(1)] assms(2)] (*sledgehammer*) using assms(3) by fastforce lemma tagged_div_wo_singlestons: assumes "p tagged_division_of {a::real..b}" "a < b" shows "(p - {xk. \x y. xk = (x,{y})}) tagged_division_of cbox a b" using remove_singleton_from_tagged_division_of[OF assms] assms apply(auto simp add: tagged_division_of_def tagged_partial_division_of_def) apply blast apply blast apply blast by fastforce lemma tagged_div_wo_empty: assumes "p tagged_division_of {a::real..b}" "a < b" shows "(p - {xk. \x. xk = (x,{})}) tagged_division_of cbox a b" using remove_singleton_from_tagged_division_of[OF assms] assms apply(auto simp add: tagged_division_of_def tagged_partial_division_of_def) apply blast apply blast apply blast by fastforce lemma fine_diff: assumes "\ fine p" shows "\ fine (p - s)" apply (auto simp add: fine_def) using assms by auto lemma tagged_div_tage_notin_set: assumes "finite (s::real set)" "p tagged_division_of {a..b}" "\ fine p" "(\(x, K)\p. \c d::real. c < d \ K = {c..d})" "gauge \" shows "\p' \'. p' tagged_division_of {a..b} \ \' fine p' \ (\(x, K)\p'. x \ s) \ gauge \'" proof- have "(\(x::real, K)\p. \x'. x' \ s \ x'\ interior K)" proof- {fix x::real fix K assume ass: "(x::real,K) \ p" have "(\(x, K)\p. infinite (interior K))" using assms(4) infinite_Ioo interior_atLeastAtMost_real by (smt split_beta) then have i: "infinite (interior K)" using ass by auto have "\x'. x' \ s \ x'\ interior K" using infinite_imp_nonempty[OF Diff_infinite_finite[OF assms(1) i]] by auto} then show ?thesis by auto qed then obtain f where f: "(\(x::real, K)\p. (f (x,K)) \ s \ (f (x,K)) \ interior K)" using choice_iff[where ?Q = "\(x,K) x'. (x::real, K)\p \ x' \ s \ x' \ interior K"] apply (auto simp add: case_prod_beta) by metis have f': "(\(x::real, K)\p. (f (x,K)) \ s \ (f (x,K)) \ K)" using f interior_subset by (auto simp add: case_prod_beta subset_iff) let ?p' = "{m. (\xK. m = ((f xK), snd xK) \ xK \ p)}" have 0: "(\(x, K)\?p'. x \ s)" using f by (auto simp add: case_prod_beta) have i: "finite {(f (a, b), b) |a b. (a, b) \ p}" proof- have a: "{(f (a, b), b) |a b. (a, b) \ p} = (%(a,b). (f(a,b),b)) ` p" by auto have b: "finite p" using assms(2) by auto show ?thesis using a b by auto qed have 1: "?p' tagged_division_of {a..b}" using assms(2) f' apply(auto simp add: tagged_division_of_def tagged_partial_division_of_def case_prod_beta) apply(metis i) apply blast apply blast by fastforce (*f is injective becuase interiors of different K's are disjoint and f is in interior*) have f_inj: "inj_on f p" apply(simp add: inj_on_def) proof- {fix x y assume "x \ p" "y \ p" "f x = f y" then have "x = y" using f tagged_division_ofD(5)[OF assms(2)] (*sledgehammer*) by (smt case_prodE disjoint_insert(2) mk_disjoint_insert)}note * = this show "\x\p. \y\p. f x = f y \ x = y" using * by auto qed let ?\' = "\x. if (\xK \ p. f xK = x) then (\ o fst o the_inv_into p f) x else \ x" have 2: "?\' fine ?p'" using assms(3) apply(auto simp add: fine_def case_prod_beta the_inv_into_f_f[OF f_inj]) by force have 3: "gauge ?\'" using assms(5) assms(3) f' apply(auto simp add: fine_def gauge_def case_prod_beta the_inv_into_f_f[OF f_inj]) by force have "?p' tagged_division_of {a..b} \ ?\' fine ?p' \ (\(x, K)\?p'. x \ s) \ gauge ?\'" using 0 1 2 3 by auto then show ?thesis by smt qed lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed lemma has_integral_bound_: fixes f :: "real \ 'a::real_normed_vector" assumes "a < b" and "0 \ B" and f: "(f has_integral i) (cbox a b)" and "finite s" and "\x\(cbox a b)-s. norm (f x) \ B" shows "norm i \ B * content (cbox a b)" using has_integral_bound_spike_finite assms by blast corollary has_integral_bound_real': fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "finite s" and "\x\(cbox a b)-s. norm (f x) \ B" shows "norm i \ B * content {a..b}" (*sledgehammer*) by (metis assms(1) assms(3) assms(4) box_real(2) f has_integral_bound_spike_finite) lemma integral_has_vector_derivative_continuous_at': fixes f :: "real \ 'a::banach" assumes "finite s" and f: "f integrable_on {a..b}" and x: "x \ {a..b} - s" and fx: "continuous (at x within ({a..b} - s)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - s))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - s; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" if y: "y \ {a..b} - s" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" - using False x by (simp add: algebra_simps integral_combine) + using False x by (simp add: algebra_simps Henstock_Kurzweil_Integration.integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" x y] False apply (simp add: ) done show ?thesis using False apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real'[where f="(\u. f u - f x)"]) using yx False d x y \e>0\ apply (auto simp add: Idiff fux_int) proof- let ?M48= "mset_set s" show "\z. y - x < d \ (\x'. a \ x' \ x' \ b \ x' \ s \ \x' - x\ < d \ norm (f x' - f x) \ e) \ 0 < e \ z \# ?M48 \ a \ x \ x \ s \ y \ b \ y \ s \ x \ z \ z \ y \ norm (f z - f x) \ e" using assms by auto qed next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" - using True x y by (simp add: algebra_simps integral_combine) + using True x y by (simp add: algebra_simps Henstock_Kurzweil_Integration.integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" y x] True apply (simp add: ) done have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" using True apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real'[where f="(\u. f u - f x)"]) using yx True d x y \e>0\ apply (auto simp add: Idiff fux_int) proof- let ?M44= "mset_set s" show " \xa. x - y < d \ y < x \ (\x'. a \ x' \ x' \ b \ x' \ s \ \x' - x\ < d \ norm (f x' - f x) \ e) \ 0 < e \ xa \# ?M44 \ x \ b \ x \ s \ a \ y \ y \ s \ y \ xa \ xa \ x \ norm (f xa - f x) \ e" using assms by auto qed then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - s. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative': fixes f :: "real \ 'a::banach" assumes "finite s" "f integrable_on {a..b}" "x \ {a..b} - s" "continuous (at x within {a..b} - s) f" shows "((\u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b} - s)" apply (rule integral_has_vector_derivative_continuous_at') using assms apply (auto simp: continuous_on_eq_continuous_within) done lemma fundamental_theorem_of_calculus_interior_stronger: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..a < x\ \x < b\ insert.prems continuous_on_subset by force+ moreover have "(f' has_integral f b - f x) {x..b}" apply (rule insert) using \x < b\ \a < x\ insert.prems continuous_on_subset by force+ ultimately show ?thesis by (meson finite_insert fundamental_theorem_of_calculus_interior_strong insert.hyps(1) insert.prems(1) insert.prems(2) insert.prems(3)) qed qed lemma at_within_closed_interval_finite: fixes x::real assumes "a < x" "x < b" "x \ S" "finite S" shows "(at x within {a..b} - S) = at x" proof - have "interior ({a..b} - S) = {a<..finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma fundamental_theorem_of_calculus_interior_stronger': fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x within {a..b} - S)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms fundamental_theorem_of_calculus_interior_strong at_within_cbox_finite (*sledgehammer*) by (metis DiffD1 DiffD2 interior_atLeastAtMost_real interior_cbox interval_cbox) lemma has_integral_substitution_general_: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f: "f integrable_on {c..d}" "continuous_on ({c..d} - (g ` s)) f" (*and f [continuous_intros]: "continuous_on {c..d} f"*) and g : "continuous_on {a..b} g" "inj_on g ({a..b} \ s)" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g(1) subset] indefinite_integral_continuous_1 f)+ have deriv: "\x. x \ {a..b} - s \ (((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within ({a..b} - s))" apply (rule has_vector_derivative_eq_rhs) apply (rule vector_diff_chain_within) apply (subst has_field_derivative_iff_has_vector_derivative [symmetric]) proof- fix x::real assume ass: "x \ {a..b} - s" let ?f'3 = "g' x" have i:"{a..b} - s \ {a..b}" by auto have ii: " (g has_vector_derivative g' x) (at x within {a..b})" using deriv[OF ass] by (simp only: has_field_derivative_iff_has_vector_derivative) show "(g has_real_derivative ?f'3) (at x within {a..b} - s)" using has_vector_derivative_within_subset[OF ii i] by (simp only: has_field_derivative_iff_has_vector_derivative) next let ?g'3 = "f o g" show "\x. x \ {a..b} - s \ ((\x. integral {c..x} f) has_vector_derivative ?g'3 x) (at (g x) within g ` ({a..b} - s))" proof- fix x::real assume ass: "x \ {a..b} - s" have "finite (g ` s)" using s by auto then have i: "((\x. integral {c..x} f) has_vector_derivative f(g x)) (at (g x) within ({c..d} - g ` s))" apply (rule integral_has_vector_derivative') proof- show " f integrable_on {c..d}" using f by auto show "g x \ {c..d} - g ` s" using ass subset (*sledgehammer*) by (smt Diff_iff Un_upper1 Un_upper2 g(2) imageE image_subset_iff inj_onD subsetCE) show "continuous (at (g x) within {c..d} - g ` s) f" (*sledgehammer*) using \g x \ {c..d} - g ` s\ continuous_on_eq_continuous_within f(2) by blast qed have ii: "g ` ({a..b} - s) \ ({c..d} - g ` s)" using subset g(2) (*sledgehammer*) by (smt Diff_subset Un_Diff Un_commute Un_upper2 inj_on_image_set_diff subset_trans sup.order_iff) then show "((\x. integral {c..x} f) has_vector_derivative ?g'3 x) (at (g x) within g ` ({a..b} - s))" (*sledgehammer*) by (smt Diff_subset has_vector_derivative_weaken Un_upper1 Un_upper2 \finite (g ` s)\ ass comp_def continuous_on_eq_continuous_within f(1) f(2) g(2) image_diff_subset image_subset_iff inj_on_image_set_diff integral_has_vector_derivative_continuous_at' subset_trans) qed show "\x. x \ {a..b} - s \ g' x *\<^sub>R ?g'3 x = g' x *\<^sub>R f (g x)" by auto qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b} - s)" if "x \ {a<..x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using cont_int using fundamental_theorem_of_calculus_interior_stronger'[OF s le deriv] by blast also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" - using integral_combine f(1) le - by (smt atLeastatMost_subset_iff integrable_subinterval_real) + by (meson Henstock_Kurzweil_Integration.integral_combine atLeastatMost_subset_iff f(1) integrable_on_subinterval le(2) order_refl) with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" - using integral_combine f(1) le - by (smt atLeastatMost_subset_iff integrable_subinterval_real) + by (meson Henstock_Kurzweil_Integration.integral_combine atLeastatMost_subset_iff f(1) integrable_on_subinterval le(2) order_refl) with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_general__: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and s_subset: "s \ {a..b}" and subset: "g ` {a..b} \ {c..d}" and f: "f integrable_on {c..d}" "continuous_on ({c..d} - (g ` s)) f" (*and f [continuous_intros]: "continuous_on {c..d} f"*) and g : "continuous_on {a..b} g" "inj_on g {a..b}" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" using s_subset has_integral_substitution_general_[OF s le subset f g(1) _ deriv] by (simp add: g(2) sup_absorb1) lemma has_integral_substitution_general_': fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and s': "finite s'" and subset: "g ` {a..b} \ {c..d}" and f: "f integrable_on {c..d}" "continuous_on ({c..d} - s') f" and g : "continuous_on {a..b} g" "\x\s'. finite (g -` {x})" "surj_on s' g" "inj_on g ({a..b} \ ((s \ g -` s')))" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof- have a: "g -` s' = \{t. \x. t = g -` {x} \ x \ s'}" using s s' by blast have "finite (\{t. \x. t = g -` {x} \ x \ s'})" using s' by (metis (no_types, lifting) \g -` s' = \{g -` {x} |x. x \ s'}\ finite_UN_I g(2) vimage_eq_UN) then have 0: "finite (s \ (g -` s'))" using a s by simp have 1: "continuous_on ({c..d} - g ` (s \ g -` s')) f" using f(2) surj_on_image_vimage_eq by (metis Diff_mono Un_upper2 continuous_on_subset equalityE g(3) image_Un) have 2: " (\x. x \ {a..b} - (s \ g -` s') \ (g has_real_derivative g' x) (at x within {a..b}))" using deriv by auto show ?thesis using has_integral_substitution_general_[OF 0 assms(2) subset f(1) 1 g(1) g(4) 2] by auto qed end diff --git a/thys/Ordinary_Differential_Equations/Library/Interval_Integral_HK.thy b/thys/Ordinary_Differential_Equations/Library/Interval_Integral_HK.thy --- a/thys/Ordinary_Differential_Equations/Library/Interval_Integral_HK.thy +++ b/thys/Ordinary_Differential_Equations/Library/Interval_Integral_HK.thy @@ -1,305 +1,305 @@ theory Interval_Integral_HK imports Vector_Derivative_On begin subsection \interval integral\ \ \TODO: move to repo ?!\ \ \TODO: replace with Bochner Integral?! But FTC for Bochner requires continuity and euclidean space!\ definition has_ivl_integral :: "(real \ 'b::real_normed_vector) \ 'b \ real \ real \ bool"\ \TODO: generalize?\ (infixr "has'_ivl'_integral" 46) where "(f has_ivl_integral y) a b \ (if a \ b then (f has_integral y) {a .. b} else (f has_integral - y) {b .. a})" definition ivl_integral::"real \ real \ (real \ 'a) \ 'a::real_normed_vector" where "ivl_integral a b f = integral {a .. b} f - integral {b .. a} f" lemma integral_emptyI[simp]: fixes a b::real shows "a \ b \ integral {a..b} f = 0" "a > b \ integral {a..b} f = 0" by (cases "a = b") auto lemma ivl_integral_unique: "(f has_ivl_integral y) a b \ ivl_integral a b f = y" using integral_unique[of f y "{a .. b}"] integral_unique[of f "- y" "{b .. a}"] unfolding ivl_integral_def has_ivl_integral_def by (auto split: if_split_asm) lemma fundamental_theorem_of_calculus_ivl_integral: fixes f :: "real \ 'a::banach" shows "(f has_vderiv_on f') (closed_segment a b) \ (f' has_ivl_integral f b - f a) a b" by (auto simp: has_ivl_integral_def closed_segment_eq_real_ivl intro!: fundamental_theorem_of_calculus') lemma fixes f :: "real \ 'a::banach" assumes "f integrable_on (closed_segment a b)" shows indefinite_ivl_integral_continuous: "continuous_on (closed_segment a b) (\x. ivl_integral a x f)" "continuous_on (closed_segment b a) (\x. ivl_integral a x f)" using assms by (auto simp: ivl_integral_def closed_segment_eq_real_ivl split: if_split_asm intro!: indefinite_integral_continuous_1 indefinite_integral_continuous_1' continuous_intros intro: continuous_on_eq) lemma fixes f :: "real \ 'a::banach" assumes "f integrable_on (closed_segment a b)" assumes "c \ closed_segment a b" shows indefinite_ivl_integral_continuous_subset: "continuous_on (closed_segment a b) (\x. ivl_integral c x f)" proof - from assms have "f integrable_on (closed_segment c a)" "f integrable_on (closed_segment c b)" by (auto simp: closed_segment_eq_real_ivl integrable_on_subinterval integrable_on_insert_iff split: if_splits) then have "continuous_on (closed_segment a c \ closed_segment c b) (\x. ivl_integral c x f)" by (auto intro!: indefinite_ivl_integral_continuous continuous_on_closed_Un) also have "closed_segment a c \ closed_segment c b = closed_segment a b" using assms by (auto simp: closed_segment_eq_real_ivl) finally show ?thesis . qed lemma real_Icc_closed_segment: fixes a b::real shows "a \ b \ {a .. b} = closed_segment a b" by (auto simp: closed_segment_eq_real_ivl) lemma ivl_integral_zero[simp]: "ivl_integral a a f = 0" by (auto simp: ivl_integral_def) lemma ivl_integral_cong: assumes "\x. x \ closed_segment a b \ g x = f x" assumes "a = c" "b = d" shows "ivl_integral a b f = ivl_integral c d g" using assms integral_spike[of "{}" "closed_segment a b" f g] by (auto simp: ivl_integral_def closed_segment_eq_real_ivl split: if_split_asm) lemma ivl_integral_diff: "f integrable_on (closed_segment s t) \ g integrable_on (closed_segment s t) \ ivl_integral s t (\x. f x - g x) = ivl_integral s t f - ivl_integral s t g" using Henstock_Kurzweil_Integration.integral_diff[of f "closed_segment s t" g] by (auto simp: ivl_integral_def closed_segment_eq_real_ivl split: if_split_asm) lemma ivl_integral_norm_bound_ivl_integral: fixes f :: "real \ 'a::banach" assumes "f integrable_on (closed_segment a b)" and "g integrable_on (closed_segment a b)" and "\x. x \ closed_segment a b \ norm (f x) \ g x" shows "norm (ivl_integral a b f) \ abs (ivl_integral a b g)" using integral_norm_bound_integral[OF assms] by (auto simp: ivl_integral_def closed_segment_eq_real_ivl split: if_split_asm) lemma ivl_integral_norm_bound_integral: fixes f :: "real \ 'a::banach" assumes "f integrable_on (closed_segment a b)" and "g integrable_on (closed_segment a b)" and "\x. x \ closed_segment a b \ norm (f x) \ g x" shows "norm (ivl_integral a b f) \ integral (closed_segment a b) g" using integral_norm_bound_integral[OF assms] by (auto simp: ivl_integral_def closed_segment_eq_real_ivl split: if_split_asm) lemma norm_ivl_integral_le: fixes f :: "real \ real" assumes "f integrable_on (closed_segment a b)" and "g integrable_on (closed_segment a b)" and "\x. x \ closed_segment a b \ f x \ g x" and "\x. x \ closed_segment a b \ 0 \ f x" shows "abs (ivl_integral a b f) \ abs (ivl_integral a b g)" proof (cases "a = b") case True then show ?thesis by simp next case False have "0 \ integral {a..b} f" "0 \ integral {b..a} f" by (metis le_cases Henstock_Kurzweil_Integration.integral_nonneg assms(1) assms(4) closed_segment_eq_real_ivl integral_emptyI(1))+ then show ?thesis using integral_le[OF assms(1-3)] unfolding ivl_integral_def closed_segment_eq_real_ivl by (simp split: if_split_asm) qed lemma ivl_integral_const [simp]: shows "ivl_integral a b (\x. c) = (b - a) *\<^sub>R c" by (auto simp: ivl_integral_def algebra_simps) lemma ivl_integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" and "x \ closed_segment a b" shows "((\u. ivl_integral a u f) has_vector_derivative f x) (at x within closed_segment a b)" proof - have "((\x. integral {x..a} f) has_vector_derivative 0) (at x within {a..b})" if "a \ x" "x \ b" by (rule has_vector_derivative_transform) (auto simp: that) moreover have "((\x. integral {a..x} f) has_vector_derivative 0) (at x within {b..a})" if "b \ x" "x \ a" by (rule has_vector_derivative_transform) (auto simp: that) ultimately show ?thesis using assms by (auto simp: ivl_integral_def closed_segment_eq_real_ivl intro!: derivative_eq_intros integral_has_vector_derivative[of a b f] integral_has_vector_derivative[of b a "-f"] integral_has_vector_derivative'[of b a f]) qed lemma ivl_integral_has_vderiv_on: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "((\u. ivl_integral a u f) has_vderiv_on f) (closed_segment a b)" using ivl_integral_has_vector_derivative[OF assms] by (auto simp: has_vderiv_on_def) lemma ivl_integral_has_vderiv_on_subset_segment: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" and "c \ closed_segment a b" shows "((\u. ivl_integral c u f) has_vderiv_on f) (closed_segment a b)" proof - have "(closed_segment c a) \ (closed_segment a b)" "(closed_segment c b) \ (closed_segment a b)" using assms by (auto simp: closed_segment_eq_real_ivl split: if_splits) then have "((\u. ivl_integral c u f) has_vderiv_on f) ((closed_segment c a) \ (closed_segment c b))" by (auto intro!: has_vderiv_on_union_closed ivl_integral_has_vderiv_on assms intro: continuous_on_subset) also have "(closed_segment c a) \ (closed_segment c b) = (closed_segment a b)" using assms by (auto simp: closed_segment_eq_real_ivl split: if_splits) finally show ?thesis . qed lemma ivl_integral_has_vector_derivative_subset: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" and "x \ closed_segment a b" and "c \ closed_segment a b" shows "((\u. ivl_integral c u f) has_vector_derivative f x) (at x within closed_segment a b)" using ivl_integral_has_vderiv_on_subset_segment[OF assms(1)] assms(2-) by (auto simp: has_vderiv_on_def) lemma compact_interval_eq_Inf_Sup: fixes A::"real set" assumes "is_interval A" "compact A" "A \ {}" shows "A = {Inf A .. Sup A}" apply (auto simp: closed_segment_eq_real_ivl intro!: cInf_lower cSup_upper bounded_imp_bdd_below bounded_imp_bdd_above compact_imp_bounded assms) by (metis assms(1) assms(2) assms(3) cInf_eq_minimum cSup_eq_maximum compact_attains_inf compact_attains_sup mem_is_interval_1_I) lemma ivl_integral_has_vderiv_on_compact_interval: fixes f :: "real \ 'a::banach" assumes "continuous_on A f" and "c \ A" "is_interval A" "compact A" shows "((\u. ivl_integral c u f) has_vderiv_on f) A" proof - have "A = {Inf A .. Sup A}" by (rule compact_interval_eq_Inf_Sup) (use assms in auto) also have "\ = closed_segment (Inf A) (Sup A)" using assms by (auto simp add: closed_segment_eq_real_ivl intro!: cInf_le_cSup bounded_imp_bdd_below bounded_imp_bdd_above compact_imp_bounded) finally have *: "A = closed_segment (Inf A) (Sup A)" . show ?thesis apply (subst *) apply (rule ivl_integral_has_vderiv_on_subset_segment) unfolding *[symmetric] by fact+ qed lemma ivl_integral_has_vector_derivative_compact_interval: fixes f :: "real \ 'a::banach" assumes "continuous_on A f" and "is_interval A" "compact A" "x \ A" "c \ A" shows "((\u. ivl_integral c u f) has_vector_derivative f x) (at x within A)" using ivl_integral_has_vderiv_on_compact_interval[OF assms(1)] assms(2-) by (auto simp: has_vderiv_on_def) lemma ivl_integral_combine: fixes f::"real \ 'a::banach" assumes "f integrable_on (closed_segment a b)" assumes "f integrable_on (closed_segment b c)" assumes "f integrable_on (closed_segment a c)" shows "ivl_integral a b f + ivl_integral b c f = ivl_integral a c f" proof - show ?thesis using assms - integral_combine[of a b c f] - integral_combine[of a c b f] - integral_combine[of b a c f] - integral_combine[of b c a f] - integral_combine[of c a b f] - integral_combine[of c b a f] + Henstock_Kurzweil_Integration.integral_combine[of a b c f] + Henstock_Kurzweil_Integration.integral_combine[of a c b f] + Henstock_Kurzweil_Integration.integral_combine[of b a c f] + Henstock_Kurzweil_Integration.integral_combine[of b c a f] + Henstock_Kurzweil_Integration.integral_combine[of c a b f] + Henstock_Kurzweil_Integration.integral_combine[of c b a f] by (cases "a \ b"; cases "b \ c"; cases "a \ c") (auto simp: algebra_simps ivl_integral_def closed_segment_eq_real_ivl) qed lemma integral_equation_swap_initial_value: fixes x::"real\'a::banach" assumes "\t. t \ closed_segment t0 t1 \ x t = x t0 + ivl_integral t0 t (\t. f t (x t))" assumes t: "t \ closed_segment t0 t1" assumes int: "(\t. f t (x t)) integrable_on closed_segment t0 t1" shows "x t = x t1 + ivl_integral t1 t (\t. f t (x t))" proof - from t int have "(\t. f t (x t)) integrable_on closed_segment t0 t" "(\t. f t (x t)) integrable_on closed_segment t t1" by (auto intro: integrable_on_subinterval simp: closed_segment_eq_real_ivl split: if_split_asm) with assms(1)[of t] assms(2-) have "x t - x t0 = ivl_integral t0 t1 (\t. f t (x t)) + ivl_integral t1 t (\t. f t (x t))" by (subst ivl_integral_combine) (auto simp: closed_segment_commute) then have "x t + x t1 - (x t0 + ivl_integral t0 t1 (\t. f t (x t))) = x t1 + ivl_integral t1 t (\t. f t (x t))" by (simp add: algebra_simps) also have "x t0 + ivl_integral t0 t1 (\t. f t (x t)) = x t1" by (auto simp: assms(1)[symmetric]) finally show ?thesis by simp qed lemma has_integral_nonpos: fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) s" and "\x\s. f x \ 0" shows "i \ 0" by (rule has_integral_nonneg[of "-f" "-i" s, simplified]) (auto intro!: has_integral_neg simp: fun_Compl_def assms) lemma has_ivl_integral_nonneg: fixes f :: "real \ real" assumes "(f has_ivl_integral i) a b" and "\x. a \ x \ x \ b \ 0 \ f x" and "\x. b \ x \ x \ a \ f x \ 0" shows "0 \ i" using assms has_integral_nonneg[of f i "{a .. b}"] has_integral_nonpos[of f "-i" "{b .. a}"] by (auto simp: has_ivl_integral_def Ball_def not_le split: if_split_asm) lemma has_ivl_integral_ivl_integral: "f integrable_on (closed_segment a b) \ (f has_ivl_integral (ivl_integral a b f)) a b" by (auto simp: closed_segment_eq_real_ivl has_ivl_integral_def ivl_integral_def) lemma ivl_integral_nonneg: fixes f :: "real \ real" assumes "f integrable_on (closed_segment a b)" and "\x. a \ x \ x \ b \ 0 \ f x" and "\x. b \ x \ x \ a \ f x \ 0" shows "0 \ ivl_integral a b f" by (rule has_ivl_integral_nonneg[OF assms(1)[unfolded has_ivl_integral_ivl_integral] assms(2-3)]) lemma ivl_integral_bound: fixes f::"real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" assumes "\t. t \ (closed_segment a b) \ norm (f t) \ B" shows "norm (ivl_integral a b f) \ B * abs (b - a)" using integral_bound[of a b f B] integral_bound[of b a f B] assms by (auto simp: closed_segment_eq_real_ivl has_ivl_integral_def ivl_integral_def split: if_splits) lemma ivl_integral_minus_sets: fixes f::"real \ 'a::banach" shows "f integrable_on (closed_segment c a) \ f integrable_on (closed_segment c b) \ f integrable_on (closed_segment a b) \ ivl_integral c a f - ivl_integral c b f = ivl_integral b a f" using ivl_integral_combine[of f c b a] by (auto simp: algebra_simps closed_segment_commute) lemma ivl_integral_minus_sets': fixes f::"real \ 'a::banach" shows "f integrable_on (closed_segment a c) \ f integrable_on (closed_segment b c) \ f integrable_on (closed_segment a b) \ ivl_integral a c f - ivl_integral b c f = ivl_integral a b f" using ivl_integral_combine[of f a b c] by (auto simp: algebra_simps closed_segment_commute) end \ No newline at end of file diff --git a/thys/Prime_Number_Theorem/Prime_Counting_Functions.thy b/thys/Prime_Number_Theorem/Prime_Counting_Functions.thy --- a/thys/Prime_Number_Theorem/Prime_Counting_Functions.thy +++ b/thys/Prime_Number_Theorem/Prime_Counting_Functions.thy @@ -1,1530 +1,1530 @@ (* File: Prime_Counting_Functions.thy Author: Manuel Eberl (TU München) Definitions and basic properties of prime-counting functions like pi, theta, and psi *) section \Prime-Counting Functions\ theory Prime_Counting_Functions imports Prime_Number_Theorem_Library begin text \ We will now define the basic prime-counting functions \\\, \\\, and \\\. Additionally, we shall define a function M that is related to Mertens' theorems and Newman's proof of the Prime Number Theorem. Most of the results in this file are not actually required to prove the Prime Number Theorem, but are still nice to have. \ subsection \Definitions\ definition prime_sum_upto :: "(nat \ 'a) \ real \ 'a :: semiring_1" where "prime_sum_upto f x = (\p | prime p \ real p \ x. f p)" lemma prime_sum_upto_altdef1: "prime_sum_upto f x = sum_upto (\p. ind prime p * f p) x" unfolding sum_upto_def prime_sum_upto_def by (intro sum.mono_neutral_cong_left finite_subset[OF _ finite_Nats_le_real[of x]]) (auto dest: prime_gt_1_nat simp: ind_def) lemma prime_sum_upto_altdef2: "prime_sum_upto f x = (\p | prime p \ p \ nat \x\. f p)" unfolding sum_upto_altdef prime_sum_upto_altdef1 by (intro sum.mono_neutral_cong_right) (auto simp: ind_def dest: prime_gt_1_nat) lemma prime_sum_upto_altdef3: "prime_sum_upto f x = (\p\primes_upto (nat \x\). f p)" proof - have "(\p\primes_upto (nat \x\). f p) = (\p | prime p \ p \ nat \x\. f p)" by (subst sum_list_distinct_conv_sum_set) (auto simp: set_primes_upto conj_commute) thus ?thesis by (simp add: prime_sum_upto_altdef2) qed lemma prime_sum_upto_eqI: assumes "a \ b" "\k. k \ {nat \a\<..nat\b\} \ \prime k" shows "prime_sum_upto f a = prime_sum_upto f b" proof - have *: "k \ nat \a\" if "k \ nat \b\" "prime k" for k using that assms(2)[of k] by (cases "k \ nat \a\") auto from assms(1) have "nat \a\ \ nat \b\" by linarith hence "(\p | prime p \ p \ nat \a\. f p) = (\p | prime p \ p \ nat \b\. f p)" using assms by (intro sum.mono_neutral_left) (auto dest: *) thus ?thesis by (simp add: prime_sum_upto_altdef2) qed lemma prime_sum_upto_eqI': assumes "a' \ nat \a\" "a \ b" "nat \b\ \ b'" "\k. k \ {a'<..b'} \ \prime k" shows "prime_sum_upto f a = prime_sum_upto f b" by (rule prime_sum_upto_eqI) (use assms in auto) lemmas eval_prime_sum_upto = prime_sum_upto_altdef3[unfolded primes_upto_sieve] lemma of_nat_prime_sum_upto: "of_nat (prime_sum_upto f x) = prime_sum_upto (\p. of_nat (f p)) x" by (simp add: prime_sum_upto_def) lemma prime_sum_upto_mono: assumes "\n. n > 0 \ f n \ (0::real)" "x \ y" shows "prime_sum_upto f x \ prime_sum_upto f y" using assms unfolding prime_sum_upto_altdef1 sum_upto_altdef by (intro sum_mono2) (auto simp: le_nat_iff' le_floor_iff ind_def) lemma prime_sum_upto_nonneg: assumes "\n. n > 0 \ f n \ (0 :: real)" shows "prime_sum_upto f x \ 0" unfolding prime_sum_upto_altdef1 sum_upto_altdef by (intro sum_nonneg) (auto simp: ind_def assms) lemma prime_sum_upto_eq_0: assumes "x < 2" shows "prime_sum_upto f x = 0" proof - from assms have "nat \x\ = 0 \ nat \x\ = 1" by linarith thus ?thesis by (auto simp: eval_prime_sum_upto) qed lemma measurable_prime_sum_upto [measurable]: fixes f :: "'a \ nat \ real" assumes [measurable]: "\y. (\t. f t y) \ M \\<^sub>M borel" assumes [measurable]: "x \ M \\<^sub>M borel" shows "(\t. prime_sum_upto (f t) (x t)) \ M \\<^sub>M borel" unfolding prime_sum_upto_altdef1 by measurable text \ The following theorem breaks down a sum over all prime powers no greater than fixed bound into a nicer form. \ lemma sum_upto_primepows: fixes f :: "nat \ 'a :: comm_monoid_add" assumes "\n. \primepow n \ f n = 0" "\p i. prime p \ i > 0 \ f (p ^ i) = g p i" shows "sum_upto f x = (\(p, i) | prime p \ i > 0 \ real (p ^ i) \ x. g p i)" proof - let ?d = aprimedivisor have g: "g (?d n) (multiplicity (?d n) n) = f n" if "primepow n" for n using that by (subst assms(2) [symmetric]) (auto simp: primepow_decompose aprimedivisor_prime_power primepow_gt_Suc_0 intro!: aprimedivisor_nat multiplicity_aprimedivisor_gt_0_nat) have "sum_upto f x = (\n | primepow n \ real n \ x. f n)" unfolding sum_upto_def using assms by (intro sum.mono_neutral_cong_right) (auto simp: primepow_gt_0_nat) also have "\ = (\(p, i) | prime p \ i > 0 \ real (p ^ i) \ x. g p i)" (is "_ = sum _ ?S") by (rule sum.reindex_bij_witness[of _ "\(p,i). p ^ i" "\n. (?d n, multiplicity (?d n) n)"]) (auto simp: aprimedivisor_prime_power primepow_decompose primepow_gt_Suc_0 g simp del: of_nat_power intro!: aprimedivisor_nat multiplicity_aprimedivisor_gt_0_nat) finally show ?thesis . qed definition primes_pi where "primes_pi = prime_sum_upto (\p. 1 :: real)" definition primes_theta where "primes_theta = prime_sum_upto (\p. ln (real p))" definition primes_psi where "primes_psi = sum_upto (mangoldt :: nat \ real)" definition primes_M where "primes_M = prime_sum_upto (\p. ln (real p) / real p)" text \ Next, we define some nice optional notation for these functions. \ bundle prime_counting_notation begin notation primes_pi ("\") notation primes_theta ("\") notation primes_psi ("\") notation primes_M ("\") end bundle no_prime_counting_notation begin no_notation primes_pi ("\") no_notation primes_theta ("\") no_notation primes_psi ("\") no_notation primes_M ("\") end (*<*) unbundle prime_counting_notation (*>*) lemmas \_def = primes_pi_def lemmas \_def = primes_theta_def lemmas \_def = primes_psi_def lemmas eval_\ = primes_pi_def[unfolded eval_prime_sum_upto] lemmas eval_\ = primes_theta_def[unfolded eval_prime_sum_upto] lemmas eval_\ = primes_M_def[unfolded eval_prime_sum_upto] subsection \Basic properties\ text \ The proofs in this section are mostly taken from Apostol~\cite{apostol1976analytic}. \ lemma measurable_\ [measurable]: "\ \ borel \\<^sub>M borel" and measurable_\ [measurable]: "\ \ borel \\<^sub>M borel" and measurable_\ [measurable]: "\ \ borel \\<^sub>M borel" and measurable_primes_M [measurable]: "\ \ borel \\<^sub>M borel" unfolding primes_M_def \_def \_def \_def by measurable lemma \_eq_0 [simp]: "x < 2 \ \ x = 0" and \_eq_0 [simp]: "x < 2 \ \ x = 0" and primes_M_eq_0 [simp]: "x < 2 \ \ x = 0" unfolding primes_pi_def primes_theta_def primes_M_def by (rule prime_sum_upto_eq_0; simp)+ lemma \_nat_cancel [simp]: "\ (nat x) = \ x" and \_nat_cancel [simp]: "\ (nat x) = \ x" and primes_M_nat_cancel [simp]: "\ (nat x) = \ x" and \_nat_cancel [simp]: "\ (nat x) = \ x" and \_floor_cancel [simp]: "\ (of_int \y\) = \ y" and \_floor_cancel [simp]: "\ (of_int \y\) = \ y" and primes_M_floor_cancel [simp]: "\ (of_int \y\) = \ y" and \_floor_cancel [simp]: "\ (of_int \y\) = \ y" by (simp_all add: \_def \_def \_def primes_M_def prime_sum_upto_altdef2 sum_upto_altdef) lemma \_nonneg [intro]: "\ x \ 0" and \_nonneg [intro]: "\ x \ 0" and primes_M_nonneg [intro]: "\ x \ 0" unfolding primes_pi_def primes_theta_def primes_M_def by (rule prime_sum_upto_nonneg; simp)+ lemma \_mono [intro]: "x \ y \ \ x \ \ y" and \_mono [intro]: "x \ y \ \ x \ \ y" and primes_M_mono [intro]: "x \ y \ \ x \ \ y" unfolding primes_pi_def primes_theta_def primes_M_def by (rule prime_sum_upto_mono; simp)+ lemma \_pos_iff: "\ x > 0 \ x \ 2" proof assume x: "x \ 2" show "\ x > 0" by (rule less_le_trans[OF _ \_mono[OF x]]) (auto simp: eval_\) next assume "\ x > 0" hence "\(x < 2)" by auto thus "x \ 2" by simp qed lemma \_pos: "x \ 2 \ \ x > 0" by (simp add: \_pos_iff) lemma \_eq_0 [simp]: assumes "x < 2" shows "\ x = 0" proof - from assms have "nat \x\ \ 1" by linarith hence "mangoldt n = (0 :: real)" if "n \ {0<..nat \x\}" for n using that by (auto simp: mangoldt_def dest!: primepow_gt_Suc_0) thus ?thesis unfolding \_def sum_upto_altdef by (intro sum.neutral) auto qed lemma \_nonneg [intro]: "\ x \ 0" unfolding \_def sum_upto_def by (intro sum_nonneg mangoldt_nonneg) lemma \_mono: "x \ y \ \ x \ \ y" unfolding \_def sum_upto_def by (intro sum_mono2 mangoldt_nonneg) auto subsection \The $n$-th prime number\ text \ Next we define the $n$-th prime number, where counting starts from 0. In traditional mathematics, it seems that counting usually starts from 1, but it is more natural to start from 0 in HOL and the asymptotics of the function are the same. \ definition nth_prime :: "nat \ nat" where "nth_prime n = (THE p. prime p \ card {q. prime q \ q < p} = n)" lemma finite_primes_less [intro]: "finite {q::nat. prime q \ q < p}" by (rule finite_subset[of _ "{.. q < p} = n" assumes "prime p'" "card {q. prime q \ q < p'} = n" shows "p = p'" using assms proof (induction p p' rule: linorder_wlog) case (le p p') have "finite {q. prime q \ q < p'}" by (rule finite_primes_less) moreover from le have "{q. prime q \ q < p} \ {q. prime q \ q < p'}" by auto moreover from le have "card {q. prime q \ q < p} = card {q. prime q \ q < p'}" by simp ultimately have "{q. prime q \ q < p} = {q. prime q \ q < p'}" by (rule card_subset_eq) with \prime p\ have "\(p < p')" by blast with \p \ p'\ show "p = p'" by auto qed auto lemma \_smallest_prime_beyond: "\ (real (smallest_prime_beyond m)) = \ (real (m - 1)) + 1" proof (cases m) case 0 have "smallest_prime_beyond 0 = 2" by (rule smallest_prime_beyond_eq) (auto dest: prime_gt_1_nat) with 0 show ?thesis by (simp add: eval_\) next case (Suc n) define n' where "n' = smallest_prime_beyond (Suc n)" have "n < n'" using smallest_prime_beyond_le[of "Suc n"] unfolding n'_def by linarith have "prime n'" by (simp add: n'_def) have "n' \ p" if "prime p" "p > n" for p using that smallest_prime_beyond_smallest[of p "Suc n"] by (auto simp: n'_def) note n' = \n < n'\ \prime n'\ this have "\ (real n') = real (card {p. prime p \ p \ n'})" by (simp add: \_def prime_sum_upto_def) also have "Suc n \ n'" unfolding n'_def by (rule smallest_prime_beyond_le) hence "{p. prime p \ p \ n'} = {p. prime p \ p \ n} \ {p. prime p \ p \ {n<..n'}}" by auto also have "real (card \) = \ (real n) + real (card {p. prime p \ p \ {n<..n'}})" by (subst card_Un_disjoint) (auto simp: \_def prime_sum_upto_def) also have "{p. prime p \ p \ {n<..n'}} = {n'}" using n' by (auto intro: antisym) finally show ?thesis using Suc by (simp add: n'_def) qed lemma \_inverse_exists: "\n. \ (real n) = real m" proof (induction m) case 0 show ?case by (intro exI[of _ 0]) auto next case (Suc m) from Suc.IH obtain n where n: "\ (real n) = real m" by auto hence "\ (real (smallest_prime_beyond (Suc n))) = real (Suc m)" by (subst \_smallest_prime_beyond) auto thus ?case by blast qed lemma nth_prime_exists: "\p::nat. prime p \ card {q. prime q \ q < p} = n" proof - from \_inverse_exists[of n] obtain m where "\ (real m) = real n" by blast hence card: "card {q. prime q \ q \ m} = n" by (auto simp: \_def prime_sum_upto_def) define p where "p = smallest_prime_beyond (Suc m)" have "m < p" using smallest_prime_beyond_le[of "Suc m"] unfolding p_def by linarith have "prime p" by (simp add: p_def) have "p \ q" if "prime q" "q > m" for q using smallest_prime_beyond_smallest[of q "Suc m"] that by (simp add: p_def) note p = \m < p\ \prime p\ this have "{q. prime q \ q < p} = {q. prime q \ q \ m}" proof safe fix q assume "prime q" "q < p" hence "\(q > m)" using p(1,2) p(3)[of q] by auto thus "q \ m" by simp qed (insert p, auto) also have "card \ = n" by fact finally show ?thesis using \prime p\ by blast qed lemma nth_prime_exists1: "\!p::nat. prime p \ card {q. prime q \ q < p} = n" by (intro ex_ex1I nth_prime_exists) (blast intro: nth_prime_unique_aux) lemma prime_nth_prime [intro]: "prime (nth_prime n)" and card_less_nth_prime [simp]: "card {q. prime q \ q < nth_prime n} = n" using theI'[OF nth_prime_exists1[of n]] by (simp_all add: nth_prime_def) lemma card_le_nth_prime [simp]: "card {q. prime q \ q \ nth_prime n} = Suc n" proof - have "{q. prime q \ q \ nth_prime n} = insert (nth_prime n) {q. prime q \ q < nth_prime n}" by auto also have "card \ = Suc n" by simp finally show ?thesis . qed lemma \_nth_prime [simp]: "\ (real (nth_prime n)) = real n + 1" by (simp add: \_def prime_sum_upto_def) lemma nth_prime_eqI: assumes "prime p" "card {q. prime q \ q < p} = n" shows "nth_prime n = p" unfolding nth_prime_def by (rule the1_equality[OF nth_prime_exists1]) (use assms in auto) lemma nth_prime_eqI': assumes "prime p" "card {q. prime q \ q \ p} = Suc n" shows "nth_prime n = p" proof (rule nth_prime_eqI) have "{q. prime q \ q \ p} = insert p {q. prime q \ q < p}" using assms by auto also have "card \ = Suc (card {q. prime q \ q < p})" by simp finally show "card {q. prime q \ q < p} = n" using assms by simp qed (use assms in auto) lemma nth_prime_eqI'': assumes "prime p" "\ (real p) = real n + 1" shows "nth_prime n = p" proof (rule nth_prime_eqI') have "real (card {q. prime q \ q \ p}) = \ (real p)" by (simp add: \_def prime_sum_upto_def) also have "\ = real (Suc n)" by (simp add: assms) finally show "card {q. prime q \ q \ p} = Suc n" by (simp only: of_nat_eq_iff) qed fact+ lemma nth_prime_0 [simp]: "nth_prime 0 = 2" by (intro nth_prime_eqI) (auto dest: prime_gt_1_nat) lemma nth_prime_Suc: "nth_prime (Suc n) = smallest_prime_beyond (Suc (nth_prime n))" by (rule nth_prime_eqI'') (simp_all add: \_smallest_prime_beyond) lemmas nth_prime_code [code] = nth_prime_0 nth_prime_Suc lemma strict_mono_nth_prime: "strict_mono nth_prime" proof (rule strict_monoI_Suc) fix n :: nat have "Suc (nth_prime n) \ smallest_prime_beyond (Suc (nth_prime n))" by simp also have "\ = nth_prime (Suc n)" by (simp add: nth_prime_Suc) finally show "nth_prime n < nth_prime (Suc n)" by simp qed lemma nth_prime_le_iff [simp]: "nth_prime m \ nth_prime n \ m \ n" using strict_mono_less_eq[OF strict_mono_nth_prime] by blast lemma nth_prime_less_iff [simp]: "nth_prime m < nth_prime n \ m < n" using strict_mono_less[OF strict_mono_nth_prime] by blast lemma nth_prime_eq_iff [simp]: "nth_prime m = nth_prime n \ m = n" using strict_mono_eq[OF strict_mono_nth_prime] by blast lemma nth_prime_ge_2: "nth_prime n \ 2" using nth_prime_le_iff[of 0 n] by (simp del: nth_prime_le_iff) lemma nth_prime_lower_bound: "nth_prime n \ Suc (Suc n)" proof - have "n = card {q. prime q \ q < nth_prime n}" by simp also have "\ \ card {2.. = nth_prime n - 2" by simp finally show ?thesis using nth_prime_ge_2[of n] by linarith qed lemma nth_prime_at_top: "filterlim nth_prime at_top at_top" proof (rule filterlim_at_top_mono) show "filterlim (\n::nat. n + 2) at_top at_top" by real_asymp qed (auto simp: nth_prime_lower_bound) lemma \_at_top: "filterlim \ at_top at_top" unfolding filterlim_at_top proof safe fix C :: real define x0 where "x0 = real (nth_prime (nat \max 0 C\))" show "eventually (\x. \ x \ C) at_top" using eventually_ge_at_top proof eventually_elim fix x assume "x \ x0" have "C \ real (nat \max 0 C\ + 1)" by linarith also have "real (nat \max 0 C\ + 1) = \ x0" unfolding x0_def by simp also have "\ \ \ x" by (rule \_mono) fact finally show "\ x \ C" . qed qed text\ An unbounded, strictly increasing sequence $a_n$ partitions $[a_0; \infty)$ into segments of the form $[a_n; a_{n+1})$. \ lemma strict_mono_sequence_partition: assumes "strict_mono (f :: nat \ 'a :: {linorder, no_top})" assumes "x \ f 0" assumes "filterlim f at_top at_top" shows "\k. x \ {f k.. x)" { obtain n where "x \ f n" using assms by (auto simp: filterlim_at_top eventually_at_top_linorder) also have "f n < f (Suc n)" using assms by (auto simp: strict_mono_Suc_iff) finally have "\n. f (Suc n) > x" by auto } from LeastI_ex[OF this] have "x < f (Suc k)" by (simp add: k_def) moreover have "f k \ x" proof (cases k) case (Suc k') have "k \ k'" if "f (Suc k') > x" using that unfolding k_def by (rule Least_le) with Suc show "f k \ x" by (cases "f k \ x") (auto simp: not_le) qed (use assms in auto) ultimately show ?thesis by auto qed lemma nth_prime_partition: assumes "x \ 2" shows "\k. x \ {nth_prime k.. 2" shows "\k. x \ {real (nth_prime k).. nth_prime k" "n < nth_prime (Suc k)" shows "\prime n" using assms by (metis Suc_leI not_le nth_prime_Suc smallest_prime_beyond_smallest) lemma nth_prime_partition'': assumes "x \ (2 :: real)" shows "x \ {real (nth_prime (nat \\ x\ - 1))..\ x\))}" proof - obtain n where n: "x \ {nth_prime n.. (nth_prime n) = \ x" unfolding \_def using between_nth_primes_imp_nonprime n by (intro prime_sum_upto_eqI) (auto simp: le_nat_iff le_floor_iff) hence "real n = \ x - 1" by simp hence n_eq: "n = nat \\ x\ - 1" "Suc n = nat \\ x\" by linarith+ with n show ?thesis by simp qed subsection \Relations between different prime-counting functions\ text \ The \\\ function can be expressed as a sum of \\\. \ lemma \_altdef: assumes "x > 0" shows "\ x = sum_upto (\m. prime_sum_upto ln (root m x)) (log 2 x)" (is "_ = ?rhs") proof - have finite: "finite {p. prime p \ real p \ y}" for y by (rule finite_subset[of _ "{..nat \y\}"]) (auto simp: le_nat_iff' le_floor_iff) define S where "S = (SIGMA i:{i. 0 < i \ real i \ log 2 x}. {p. prime p \ real p \ root i x})" have "\ x = (\(p, i) | prime p \ 0 < i \ real (p ^ i) \ x. ln (real p))" unfolding \_def by (subst sum_upto_primepows[where g = "\p i. ln (real p)"]) (auto simp: case_prod_unfold mangoldt_non_primepow) also have "\ = (\(i, p) | prime p \ 0 < i \ real (p ^ i) \ x. ln (real p))" by (intro sum.reindex_bij_witness[of _ "\(x,y). (y,x)" "\(x,y). (y,x)"]) auto also have "{(i, p). prime p \ 0 < i \ real (p ^ i) \ x} = S" unfolding S_def proof safe fix i p :: nat assume ip: "i > 0" "real i \ log 2 x" "prime p" "real p \ root i x" hence "real (p ^ i) \ root i x ^ i" unfolding of_nat_power by (intro power_mono) auto with ip assms show "real (p ^ i) \ x" by simp next fix i p assume ip: "prime p" "i > 0" "real (p ^ i) \ x" from ip have "2 ^ i \ p ^ i" by (intro power_mono) (auto dest: prime_gt_1_nat) also have "\ \ x" using ip by simp finally show "real i \ log 2 x" using assms by (simp add: le_log_iff powr_realpow) have "root i (real p ^ i) \ root i x" using ip assms by (subst real_root_le_iff) auto also have "root i (real p ^ i) = real p" using assms ip by (subst real_root_pos2) auto finally show "real p \ root i x" . qed also have "(\(i,p)\S. ln p) = sum_upto (\m. prime_sum_upto ln (root m x)) (log 2 x)" unfolding sum_upto_def prime_sum_upto_def S_def using finite by (subst sum.Sigma) auto finally show ?thesis . qed lemma \_conv_\_sum: "x > 0 \ \ x = sum_upto (\m. \ (root m x)) (log 2 x)" by (simp add: \_altdef \_def) lemma \_minus_\: assumes x: "x \ 2" shows "\ x - \ x = (\i | 2 \ i \ real i \ log 2 x. \ (root i x))" proof - have finite: "finite {i. 2 \ i \ real i \ log 2 x}" by (rule finite_subset[of _ "{2..nat \log 2 x\}"]) (auto simp: le_nat_iff' le_floor_iff) have "\ x = (\i | 0 < i \ real i \ log 2 x. \ (root i x))" using x by (simp add: \_conv_\_sum sum_upto_def) also have "{i. 0 < i \ real i \ log 2 x} = insert 1 {i. 2 \ i \ real i \ log 2 x}" using x by (auto simp: le_log_iff) also have "(\i\\. \ (root i x)) - \ x = (\i | 2 \ i \ real i \ log 2 x. \ (root i x))" using finite by (subst sum.insert) auto finally show ?thesis . qed text \ The following theorems use summation by parts to relate different prime-counting functions to one another with an integral as a remainder term. \ lemma \_conv_\_integral: assumes "x \ 2" shows "((\t. \ t / t) has_integral (\ x * ln x - \ x)) {2..x}" proof (cases "x = 2") case False note [intro] = finite_vimage_real_of_nat_greaterThanAtMost from False and assms have x: "x > 2" by simp have "((\t. sum_upto (ind prime) t * (1 / t)) has_integral sum_upto (ind prime) x * ln x - sum_upto (ind prime) 2 * ln 2 - (\n\real -` {2<..x}. ind prime n * ln (real n))) {2..x}" using x by (intro partial_summation_strong[where X = "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp flip: has_field_derivative_iff_has_vector_derivative) hence "((\t. \ t / t) has_integral (\ x * ln x - (\ 2 * ln 2 + (\n\real -` {2<..x}. ind prime n * ln n)))) {2..x}" by (simp add: \_def prime_sum_upto_altdef1 algebra_simps) also have "\ 2 * ln 2 + (\n\real -` {2<..x}. ind prime n * ln n) = (\n\insert 2 (real -` {2<..x}). ind prime n * ln n)" by (subst sum.insert) (auto simp: eval_\) also have "\ = \ x" unfolding \_def prime_sum_upto_def using x by (intro sum.mono_neutral_cong_right) (auto simp: ind_def dest: prime_gt_1_nat) finally show ?thesis . qed (auto simp: has_integral_refl eval_\ eval_\) lemma \_conv_\_integral: assumes "x \ 2" shows "((\t. \ t / (t * ln t ^ 2)) has_integral (\ x - \ x / ln x)) {2..x}" proof (cases "x = 2") case False define b where "b = (\p. ind prime p * ln (real p))" note [intro] = finite_vimage_real_of_nat_greaterThanAtMost from False and assms have x: "x > 2" by simp have "((\t. -(sum_upto b t * (-1 / (t * (ln t)\<^sup>2)))) has_integral -(sum_upto b x * (1 / ln x) - sum_upto b 2 * (1 / ln 2) - (\n\real -` {2<..x}. b n * (1 / ln (real n))))) {2..x}" using x by (intro has_integral_neg partial_summation_strong[where X = "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp flip: has_field_derivative_iff_has_vector_derivative simp add: power2_eq_square) also have "sum_upto b = \" by (simp add: \_def b_def prime_sum_upto_altdef1 fun_eq_iff) also have "\ x * (1 / ln x) - \ 2 * (1 / ln 2) - (\n\real -` {2<..x}. b n * (1 / ln (real n))) = \ x * (1 / ln x) - (\n\insert 2 (real -` {2<..x}). b n * (1 / ln (real n)))" by (subst sum.insert) (auto simp: b_def eval_\) also have "(\n\insert 2 (real -` {2<..x}). b n * (1 / ln (real n))) = \ x" using x unfolding \_def prime_sum_upto_altdef1 sum_upto_def proof (intro sum.mono_neutral_cong_left ballI, goal_cases) case (3 p) hence "p = 1" by auto thus ?case by auto qed (auto simp: b_def) finally show ?thesis by simp qed (auto simp: has_integral_refl eval_\ eval_\) lemma integrable_weighted_\: assumes "2 \ a" "a \ x" shows "((\t. \ t / (t * ln t ^ 2)) integrable_on {a..x})" proof (cases "a < x") case True hence "((\t. \ t * (1 / (t * ln t ^ 2))) integrable_on {a..x})" using assms unfolding \_def prime_sum_upto_altdef1 by (intro partial_summation_integrable_strong[where X = "{}" and f = "\x. -1 / ln x"]) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros continuous_intros simp: power2_eq_square field_simps) thus ?thesis by simp qed (insert has_integral_refl[of _ a] assms, auto simp: has_integral_iff) lemma \_conv_\_integral: assumes "x \ 2" shows "(\ has_integral (\ x * x - \ x)) {2..x}" proof (cases "x = 2") case False with assms have x: "x > 2" by simp define b :: "nat \ real" where "b = (\p. ind prime p * ln p / p)" note [intro] = finite_vimage_real_of_nat_greaterThanAtMost have prime_le_2: "p = 2" if "p \ 2" "prime p" for p :: nat using that by (auto simp: prime_nat_iff) have "((\t. sum_upto b t * 1) has_integral sum_upto b x * x - sum_upto b 2 * 2 - (\n\real -` {2<..x}. b n * real n)) {2..x}" using x by (intro partial_summation_strong[of "{}"]) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros continuous_intros) also have "sum_upto b = \" by (simp add: fun_eq_iff primes_M_def b_def prime_sum_upto_altdef1) also have "\ x * x - \ 2 * 2 - (\n\real -` {2<..x}. b n * real n) = \ x * x - (\n\insert 2 (real -` {2<..x}). b n * real n)" by (subst sum.insert) (auto simp: eval_\ b_def) also have "(\n\insert 2 (real -` {2<..x}). b n * real n) = \ x" unfolding \_def prime_sum_upto_def using x by (intro sum.mono_neutral_cong_right) (auto simp: b_def ind_def not_less prime_le_2) finally show ?thesis by simp qed (auto simp: eval_\ eval_\) lemma \_conv_\_integral: assumes "x \ 2" shows "((\t. \ t / t\<^sup>2) has_integral (\ x - \ x / x)) {2..x}" proof (cases "x = 2") case False with assms have x: "x > 2" by simp define b :: "nat \ real" where "b = (\p. ind prime p * ln p)" note [intro] = finite_vimage_real_of_nat_greaterThanAtMost have prime_le_2: "p = 2" if "p \ 2" "prime p" for p :: nat using that by (auto simp: prime_nat_iff) have "((\t. sum_upto b t * (1 / t^2)) has_integral sum_upto b x * (-1 / x) - sum_upto b 2 * (-1 / 2) - (\n\real -` {2<..x}. b n * (-1 / real n))) {2..x}" using x by (intro partial_summation_strong[of "{}"]) (auto simp flip: has_field_derivative_iff_has_vector_derivative simp: power2_eq_square intro!: derivative_eq_intros continuous_intros) also have "sum_upto b = \" by (simp add: fun_eq_iff \_def b_def prime_sum_upto_altdef1) also have "\ x * (-1 / x) - \ 2 * (-1 / 2) - (\n\real -` {2<..x}. b n * (-1 / real n)) = -(\ x / x - (\n\insert 2 (real -` {2<..x}). b n / real n))" by (subst sum.insert) (auto simp: eval_\ b_def sum_negf) also have "(\n\insert 2 (real -` {2<..x}). b n / real n) = \ x" unfolding primes_M_def prime_sum_upto_def using x by (intro sum.mono_neutral_cong_right) (auto simp: b_def ind_def not_less prime_le_2) finally show ?thesis by simp qed (auto simp: eval_\ eval_\) lemma integrable_primes_M: "\ integrable_on {x..y}" if "2 \ x" for x y :: real proof - have "(\x. \ x * 1) integrable_on {x..y}" if "2 \ x" "x < y" for x y :: real unfolding primes_M_def prime_sum_upto_altdef1 using that by (intro partial_summation_integrable_strong[where X = "{}" and f = "\x. x"]) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros continuous_intros) thus ?thesis using that has_integral_refl(2)[of \ x] by (cases x y rule: linorder_cases) auto qed subsection \Bounds\ lemma \_upper_bound_coarse: assumes "x \ 1" shows "\ x \ x * ln x" proof - have "\ x \ sum_upto (\_. ln x) x" unfolding \_def prime_sum_upto_altdef1 sum_upto_def by (intro sum_mono) (auto simp: ind_def) also have "\ \ real_of_int \x\ * ln x" using assms by (simp add: sum_upto_altdef) also have "\ \ x * ln x" using assms by (intro mult_right_mono) auto finally show ?thesis . qed lemma \_le_\: "\ x \ \ x" proof (cases "x \ 2") case False hence "nat \x\ = 0 \ nat \x\ = 1" by linarith thus ?thesis by (auto simp: eval_\) next case True hence "\ x - \ x = (\i | 2 \ i \ real i \ log 2 x. \ (root i x))" by (rule \_minus_\) also have "\ \ 0" by (intro sum_nonneg) auto finally show ?thesis by simp qed lemma \_upper_bound_coarse: assumes "x \ 0" shows "\ x \ x / 3 + 2" proof - have "{p. prime p \ p \ nat \x\} \ {2, 3} \ {p. p \ 1 \ odd p \ \3 dvd p \ p \ nat \x\}" using primes_dvd_imp_eq[of "2 :: nat"] primes_dvd_imp_eq[of "3 :: nat"] by auto also have "\ \ {2, 3} \ ((\k. 6*k+1) ` {0<..(x+5)/6\} \ (\k. 6*k+5) ` {..(x+1)/6\})" (is "_ \ ?lhs \ _ \ ?rhs") proof (intro Un_mono subsetI) fix p :: nat assume "p \ ?lhs" hence p: "p \ 1" "odd p" "\3 dvd p" "p \ nat \x\" by auto from p (1-3) have "(\k. k > 0 \ p = 6 * k + 1 \ p = 6 * k + 5)" by presburger then obtain k where "k > 0 \ p = 6 * k + 1 \ p = 6 * k + 5" by blast hence "p = 6 * k + 1 \ k > 0 \ k < nat \(x+5)/6\ \ p = 6*k+5 \ k < nat \(x+1)/6\" unfolding add_divide_distrib using p(4) by linarith thus "p \ ?rhs" by auto qed finally have subset: "{p. prime p \ p \ nat \x\} \ \" (is "_ \ ?A") . have "\ x = real (card {p. prime p \ p \ nat \x\})" by (simp add: \_def prime_sum_upto_altdef2) also have "card {p. prime p \ p \ nat \x\} \ card ?A" by (intro card_mono subset) auto also have "\ \ 2 + (nat \(x+5)/6\ - 1 + nat \(x+1)/6\)" by (intro order.trans[OF card_Un_le] add_mono order.trans[OF card_image_le]) auto also have "\ \ x / 3 + 2" using assms unfolding add_divide_distrib by (cases "x \ 1", linarith, simp) finally show ?thesis by simp qed lemma le_numeral_iff: "m \ numeral n \ m = numeral n \ m \ pred_numeral n" using numeral_eq_Suc by presburger text \ The following nice proof for the upper bound $\theta(x) \leq \ln 4 \cdot x$ is taken from Otto Forster's lecture notes on Analytic Number Theory~\cite{forsteranalytic}. \ lemma prod_primes_upto_less: defines "F \ (\n. (\{p::nat. prime p \ p \ n}))" shows "n > 0 \ F n < 4 ^ n" proof (induction n rule: less_induct) case (less n) have "n = 0 \ n = 1 \ n = 2 \ n = 3 \ even n \ n \ 4 \ odd n \ n \ 4" by presburger then consider "n = 0" | "n = 1" | "n = 2" | "n = 3" | "even n" "n \ 4" | "odd n" "n \ 4" by metis thus ?case proof cases assume [simp]: "n = 1" have *: "{p. prime p \ p \ Suc 0} = {}" by (auto dest: prime_gt_1_nat) show ?thesis by (simp add: F_def *) next assume [simp]: "n = 2" have *: "{p. prime p \ p \ 2} = {2 :: nat}" by (auto simp: le_numeral_iff dest: prime_gt_1_nat) thus ?thesis by (simp add: F_def *) next assume [simp]: "n = 3" have *: "{p. prime p \ p \ 3} = {2, 3 :: nat}" by (auto simp: le_numeral_iff dest: prime_gt_1_nat) thus ?thesis by (simp add: F_def *) next assume n: "even n" "n \ 4" from n have "F (n - 1) < 4 ^ (n - 1)" by (intro less.IH) auto also have "prime p \ p \ n \ prime p \ p \ n - 1" for p using n prime_odd_nat[of n] by (cases "p = n") auto hence "F (n - 1) = F n" by (simp add: F_def) also have "4 ^ (n - 1) \ (4 ^ n :: nat)" by (intro power_increasing) auto finally show ?case . next assume n: "odd n" "n \ 4" then obtain k where k_eq: "n = Suc (2 * k)" by (auto elim: oddE) from n have k: "k \ 2" unfolding k_eq by presburger have prime_dvd: "p dvd (n choose k)" if p: "prime p" "p \ {k+1<..n}" for p proof - from p k n have "p dvd pochhammer (k + 2) k" unfolding pochhammer_prod by (subst prime_dvd_prod_iff) (auto intro!: bexI[of _ "p - k - 2"] simp: k_eq numeral_2_eq_2 Suc_diff_Suc) also have "pochhammer (real (k + 2)) k = real ((n choose k) * fact k)" by (simp add: binomial_gbinomial gbinomial_pochhammer' k_eq field_simps) hence "pochhammer (k + 2) k = (n choose k) * fact k" unfolding pochhammer_of_nat of_nat_eq_iff . finally show "p dvd (n choose k)" using p by (auto simp: prime_dvd_fact_iff prime_dvd_mult_nat) qed have "\{p. prime p \ p \ {k+1<..n}} dvd (n choose k)" proof (rule multiplicity_le_imp_dvd, goal_cases) case (2 p) thus ?case proof (cases "p \ {k+1<..n}") case False hence "multiplicity p (\{p. prime p \ p \ {k+1<..n}}) = 0" using 2 by (subst prime_elem_multiplicity_prod_distrib) (auto simp: prime_multiplicity_other) thus ?thesis by auto next case True hence "multiplicity p (\{p. prime p \ p \ {k+1<..n}}) = sum (multiplicity p) {p. prime p \ Suc k < p \ p \ n}" using 2 by (subst prime_elem_multiplicity_prod_distrib) auto also have "\ = sum (multiplicity p) {p}" using True 2 proof (intro sum.mono_neutral_right ballI) fix q :: nat assume "q \ {p. prime p \ Suc k < p \ p \ n} - {p}" thus "multiplicity p q = 0" using 2 by (cases "p = q") (auto simp: prime_multiplicity_other) qed auto also have "\ = 1" using 2 by simp also have "1 \ multiplicity p (n choose k)" using prime_dvd[of p] 2 True by (intro multiplicity_geI) auto finally show ?thesis . qed qed auto hence "\{p. prime p \ p \ {k+1<..n}} \ (n choose k)" by (intro dvd_imp_le) (auto simp: k_eq) also have "\ = 1 / 2 * (\i\{k, Suc k}. n choose i)" using central_binomial_odd[of n] by (simp add: k_eq) also have "(\i\{k, Suc k}. n choose i) < (\i\{0, k, Suc k}. n choose i)" using k by simp also have "\ \ (\i\n. n choose i)" by (intro sum_mono2) (auto simp: k_eq) also have "\ = (1 + 1) ^ n" using binomial[of 1 1 n] by simp also have "1 / 2 * \ = real (4 ^ k)" by (simp add: k_eq power_mult) finally have less: "(\{p. prime p \ p \ {k + 1<..n}}) < 4 ^ k" unfolding of_nat_less_iff by simp have "F n = F (Suc k) * (\{p. prime p \ p \ {k+1<..n}})" unfolding F_def by (subst prod.union_disjoint [symmetric]) (auto intro!: prod.cong simp: k_eq) also have "\ < 4 ^ Suc k * 4 ^ k" using n by (intro mult_strict_mono less less.IH) (auto simp: k_eq) also have "\ = 4 ^ (Suc k + k)" by (simp add: power_add) also have "Suc k + k = n" by (simp add: k_eq) finally show ?case . qed (insert less.prems, auto) qed lemma \_upper_bound: assumes x: "x \ 1" shows "\ x < ln 4 * x" proof - have "4 powr (\ x / ln 4) = (\p | prime p \ p \ nat \x\. 4 powr (log 4 (real p)))" by (simp add: \_def powr_sum prime_sum_upto_altdef2 sum_divide_distrib log_def) also have "\ = (\p | prime p \ p \ nat \x\. real p)" by (intro prod.cong) (auto dest: prime_gt_1_nat) also have "\ = real (\p | prime p \ p \ nat \x\. p)" by simp also have "(\p | prime p \ p \ nat \x\. p) < 4 ^ nat \x\" using x by (intro prod_primes_upto_less) auto also have "\ = 4 powr real (nat \x\)" using x by (subst powr_realpow) auto also have "\ \ 4 powr x" using x by (intro powr_mono) auto finally have "4 powr (\ x / ln 4) < 4 powr x" by simp thus "\ x < ln 4 * x" by (subst (asm) powr_less_cancel_iff) (auto simp: field_simps) qed lemma \_bigo: "\ \ O(\x. x)" by (intro le_imp_bigo_real[of "ln 4"] eventually_mono[OF eventually_ge_at_top[of 1]] less_imp_le[OF \_upper_bound]) auto lemma \_minus_\_bound: assumes x: "x \ 2" shows "\ x - \ x \ 2 * ln x * sqrt x" proof - have "\ x - \ x = (\i | 2 \ i \ real i \ log 2 x. \ (root i x))" using x by (rule \_minus_\) also have "\ \ (\i | 2 \ i \ real i \ log 2 x. ln 4 * root i x)" using x by (intro sum_mono less_imp_le[OF \_upper_bound]) auto also have "\ \ (\i | 2 \ i \ real i \ log 2 x. ln 4 * root 2 x)" using x by (intro sum_mono mult_mono) (auto simp: le_log_iff powr_realpow intro!: real_root_decreasing) also have "\ = card {i. 2 \ i \ real i \ log 2 x} * ln 4 * sqrt x" by (simp add: sqrt_def) also have "{i. 2 \ i \ real i \ log 2 x} = {2..nat \log 2 x\}" by (auto simp: le_nat_iff' le_floor_iff) also have "log 2 x \ 1" using x by (simp add: le_log_iff) hence "real (nat \log 2 x\ - 1) \ log 2 x" using x by linarith hence "card {2..nat \log 2 x\} \ log 2 x" by simp also have "ln (2 * 2 :: real) = 2 * ln 2" by (subst ln_mult) auto hence "log 2 x * ln 4 * sqrt x = 2 * ln x * sqrt x" using x by (simp add: ln_sqrt log_def power2_eq_square field_simps) finally show ?thesis using x by (simp add: mult_right_mono) qed lemma \_minus_\_bigo: "(\x. \ x - \ x) \ O(\x. ln x * sqrt x)" proof (intro bigoI[of _ "2"] eventually_mono[OF eventually_ge_at_top[of 2]]) fix x :: real assume "x \ 2" thus "norm (\ x - \ x) \ 2 * norm (ln x * sqrt x)" using \_minus_\_bound[of x] \_le_\[of x] by simp qed lemma \_bigo: "\ \ O(\x. x)" proof - have "(\x. \ x - \ x) \ O(\x. ln x * sqrt x)" by (rule \_minus_\_bigo) also have "(\x. ln x * sqrt x) \ O(\x. x)" by real_asymp finally have "(\x. \ x - \ x + \ x) \ O(\x. x)" by (rule sum_in_bigo) (fact \_bigo) thus ?thesis by simp qed text \ We shall now attempt to get some more concrete bounds on the difference between $\pi(x)$ and $\theta(x)/\ln x$ These will be essential in showing the Prime Number Theorem later. We first need some bounds on the integral \[\int\nolimits_2^x \frac{1}{\ln^2 t}\,\mathrm{d}t\] in order to bound the contribution of the remainder term. This integral actually has an antiderivative in terms of the logarithmic integral $\textrm{li}(x)$, but since we do not have a formalisation of it in Isabelle, we will instead use the following ad-hoc bound given by Apostol: \ lemma integral_one_over_log_squared_bound: assumes x: "x \ 4" shows "integral {2..x} (\t. 1 / ln t ^ 2) \ sqrt x / ln 2 ^ 2 + 4 * x / ln x ^ 2" proof - from x have "x * 1 \ x ^ 2" unfolding power2_eq_square by (intro mult_left_mono) auto with x have x': "2 \ sqrt x" "sqrt x \ x" by (auto simp: real_sqrt_le_iff' intro!: real_le_rsqrt) have "integral {2..x} (\t. 1 / ln t ^ 2) = integral {2..sqrt x} (\t. 1 / ln t ^ 2) + integral {sqrt x..x} (\t. 1 / ln t ^ 2)" (is "_ = ?I1 + ?I2") using x x' - by (intro integral_combine [symmetric] integrable_continuous_real) + by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] integrable_continuous_real) (auto intro!: continuous_intros) also have "?I1 \ integral {2..sqrt x} (\_. 1 / ln 2 ^ 2)" using x by (intro integral_le integrable_continuous_real divide_left_mono power_mono continuous_intros) auto also have "\ \ sqrt x / ln 2 ^ 2" using x' by (simp add: field_simps) also have "?I2 \ integral {sqrt x..x} (\t. 1 / ln (sqrt x) ^ 2)" using x' by (intro integral_le integrable_continuous_real divide_left_mono power_mono continuous_intros) auto also have "\ \ 4 * x / ln x ^ 2" using x' by (simp add: ln_sqrt field_simps) finally show ?thesis by simp qed lemma integral_one_over_log_squared_bigo: "(\x::real. integral {2..x} (\t. 1 / ln t ^ 2)) \ O(\x. x / ln x ^ 2)" proof - define ub where "ub = (\x::real. sqrt x / ln 2 ^ 2 + 4 * x / ln x ^ 2)" have "eventually (\x. \integral {2..x} (\t. 1 / (ln t)\<^sup>2)\ \ \ub x\) at_top" using eventually_ge_at_top[of 4] proof eventually_elim case (elim x) hence "\integral {2..x} (\t. 1 / ln t ^ 2)\ = integral {2..x} (\t. 1 / ln t ^ 2)" by (intro abs_of_nonneg integral_nonneg integrable_continuous_real continuous_intros) auto also have "\ \ \ub x\" using integral_one_over_log_squared_bound[of x] elim by (simp add: ub_def) finally show ?case . qed hence "(\x. integral {2..x} (\t. 1 / (ln t)\<^sup>2)) \ O(ub)" by (intro landau_o.bigI[of 1]) auto also have "ub \ O(\x. x / ln x ^ 2)" unfolding ub_def by real_asymp finally show ?thesis . qed lemma \_\_bound: assumes "x \ (4 :: real)" defines "ub \ 2 / ln 2 * sqrt x + 8 * ln 2 * x / ln x ^ 2" shows "\ x - \ x / ln x \ {0..ub}" proof - define r where "r = (\x. integral {2..x} (\t. \ t / (t * ln t ^ 2)))" have integrable: "(\t. c / ln t ^ 2) integrable_on {2..x}" for c by (intro integrable_continuous_real continuous_intros) auto have "r x \ integral {2..x} (\t. ln 4 / ln t ^ 2)" unfolding r_def using integrable_weighted_\[of 2 x] integrable[of "ln 4"] assms less_imp_le[OF \_upper_bound] by (intro integral_le divide_right_mono) (auto simp: field_simps) also have "\ = ln 4 * integral {2..x} (\t. 1 / ln t ^ 2)" using integrable[of 1] by (subst integral_mult) auto also have "\ \ ln 4 * (sqrt x / ln 2 ^ 2 + 4 * x / ln x ^ 2)" using assms by (intro mult_left_mono integral_one_over_log_squared_bound) auto also have "ln (4 :: real) = 2 * ln 2" using ln_realpow[of 2 2] by simp also have "\ * (sqrt x / ln 2 ^ 2 + 4 * x / ln x ^ 2) = ub" using assms by (simp add: field_simps power2_eq_square ub_def) finally have "r x \ \" . moreover have "r x \ 0" unfolding r_def using assms by (intro integral_nonneg integrable_weighted_\ divide_nonneg_pos) auto ultimately have "r x \ {0..ub}" by auto with \_conv_\_integral[of x] assms(1) show ?thesis by (simp add: r_def has_integral_iff) qed text \ The following statement already indicates that the asymptotics of \\\ and \\\ are very closely related, since through it, $\pi(x) \sim x / \ln x$ and $\theta(x) \sim x$ imply each other. \ lemma \_\_bigo: "(\x. \ x - \ x / ln x) \ O(\x. x / ln x ^ 2)" proof - define ub where "ub = (\x. 2 / ln 2 * sqrt x + 8 * ln 2 * x / ln x ^ 2)" have "(\x. \ x - \ x / ln x) \ O(ub)" proof (intro le_imp_bigo_real[of 1] eventually_mono[OF eventually_ge_at_top]) fix x :: real assume "x \ 4" from \_\_bound[OF this] show "\ x - \ x / ln x \ 0" and "\ x - \ x / ln x \ 1 * ub x" by (simp_all add: ub_def) qed auto also have "ub \ O(\x. x / ln x ^ 2)" unfolding ub_def by real_asymp finally show ?thesis . qed text \ As a foreshadowing of the Prime Number Theorem, we can already show the following upper bound on $\pi(x)$: \ lemma \_upper_bound: assumes "x \ (4 :: real)" shows "\ x < ln 4 * x / ln x + 8 * ln 2 * x / ln x ^ 2 + 2 / ln 2 * sqrt x" proof - define ub where "ub = 2 / ln 2 * sqrt x + 8 * ln 2 * x / ln x ^ 2" have "\ x \ \ x / ln x + ub" using \_\_bound[of x] assms unfolding ub_def by simp also from assms have "\ x / ln x < ln 4 * x / ln x" by (intro \_upper_bound divide_strict_right_mono) auto finally show ?thesis using assms by (simp add: algebra_simps ub_def) qed lemma \_bigo: "\ \ O(\x. x / ln x)" proof - have "(\x. \ x - \ x / ln x) \ O(\x. x / ln x ^ 2)" by (fact \_\_bigo) also have "(\x::real. x / ln x ^ 2) \ O(\x. x / ln x)" by real_asymp finally have "(\x. \ x - \ x / ln x) \ O(\x. x / ln x)" . moreover have "eventually (\x::real. ln x > 0) at_top" by real_asymp hence "eventually (\x::real. ln x \ 0) at_top" by eventually_elim auto hence "(\x. \ x / ln x) \ O(\x. x / ln x)" using \_bigo by (intro landau_o.big.divide_right) ultimately have "(\x. \ x - \ x / ln x + \ x / ln x) \ O(\x. x / ln x)" by (rule sum_in_bigo) thus ?thesis by simp qed subsection \Equivalence of various forms of the Prime Number Theorem\ text \ In this section, we show that the following forms of the Prime Number Theorem are all equivalent: \<^enum> $\pi(x) \sim x / \ln x$ \<^enum> $\pi(x) \ln \pi(x) \sim x$ \<^enum> $p_n \sim n \ln n$ \<^enum> $\vartheta(x) \sim x$ \<^enum> $\psi(x) \sim x$ We show the following implication chains: \<^item> \(1) \ (2) \ (3) \ (2) \ (1)\ \<^item> \(1) \ (4) \ (1)\ \<^item> \(4) \ (5) \ (4)\ All of these proofs are taken from Apostol's book. \ lemma PNT1_imp_PNT1': assumes "\ \[at_top] (\x. x / ln x)" shows "(\x. ln (\ x)) \[at_top] ln" proof - (* TODO: Tedious Landau sum reasoning *) from assms have "((\x. \ x / (x / ln x)) \ 1) at_top" by (rule asymp_equivD_strong[OF _ eventually_mono[OF eventually_gt_at_top[of 1]]]) auto hence "((\x. ln (\ x / (x / ln x))) \ ln 1) at_top" by (rule tendsto_ln) auto also have "?this \ ((\x. ln (\ x) - ln x + ln (ln x)) \ 0) at_top" by (intro filterlim_cong eventually_mono[OF eventually_gt_at_top[of 2]]) (auto simp: ln_div field_simps ln_mult \_pos) finally have "(\x. ln (\ x) - ln x + ln (ln x)) \ o(\_. 1)" by (intro smalloI_tendsto) auto also have "(\_::real. 1 :: real) \ o(\x. ln x)" by real_asymp finally have "(\x. ln (\ x) - ln x + ln (ln x) - ln (ln x)) \ o(\x. ln x)" by (rule sum_in_smallo) real_asymp+ thus *: "(\x. ln (\ x)) \[at_top] ln" by (simp add: asymp_equiv_altdef) qed lemma PNT1_imp_PNT2: assumes "\ \[at_top] (\x. x / ln x)" shows "(\x. \ x * ln (\ x)) \[at_top] (\x. x)" proof - have "(\x. \ x * ln (\ x)) \[at_top] (\x. x / ln x * ln x)" by (intro asymp_equiv_intros assms PNT1_imp_PNT1') also have "\ \[at_top] (\x. x)" by (intro asymp_equiv_refl_ev eventually_mono[OF eventually_gt_at_top[of 1]]) (auto simp: field_simps) finally show "(\x. \ x * ln (\ x)) \[at_top] (\x. x)" by simp qed lemma PNT2_imp_PNT3: assumes "(\x. \ x * ln (\ x)) \[at_top] (\x. x)" shows "nth_prime \[at_top] (\n. n * ln n)" proof - have "(\n. nth_prime n) \[at_top] (\n. \ (nth_prime n) * ln (\ (nth_prime n)))" using assms by (rule asymp_equiv_symI [OF asymp_equiv_compose']) (auto intro!: filterlim_compose[OF filterlim_real_sequentially nth_prime_at_top]) also have "\ = (\n. real (Suc n) * ln (real (Suc n)))" by (simp add: add_ac) also have "\ \[at_top] (\n. real n * ln (real n))" by real_asymp finally show "nth_prime \[at_top] (\n. n * ln n)" . qed lemma PNT3_imp_PNT2: assumes "nth_prime \[at_top] (\n. n * ln n)" shows "(\x. \ x * ln (\ x)) \[at_top] (\x. x)" proof (rule asymp_equiv_symI, rule asymp_equiv_sandwich_real) show "eventually (\x. x \ {real (nth_prime (nat \\ x\ - 1))..real (nth_prime (nat \\ x\))}) at_top" using eventually_ge_at_top[of 2] proof eventually_elim case (elim x) with nth_prime_partition''[of x] show ?case by auto qed next have "(\x. real (nth_prime (nat \\ x\ - 1))) \[at_top] (\x. real (nat \\ x\ - 1) * ln (real (nat \\ x\ - 1)))" by (rule asymp_equiv_compose'[OF _ \_at_top], rule asymp_equiv_compose'[OF assms]) real_asymp also have "\ \[at_top] (\x. \ x * ln (\ x))" by (rule asymp_equiv_compose'[OF _ \_at_top]) real_asymp finally show "(\x. real (nth_prime (nat \\ x\ - 1))) \[at_top] (\x. \ x * ln (\ x))" . next have "(\x. real (nth_prime (nat \\ x\))) \[at_top] (\x. real (nat \\ x\) * ln (real (nat \\ x\)))" by (rule asymp_equiv_compose'[OF _ \_at_top], rule asymp_equiv_compose'[OF assms]) real_asymp also have "\ \[at_top] (\x. \ x * ln (\ x))" by (rule asymp_equiv_compose'[OF _ \_at_top]) real_asymp finally show "(\x. real (nth_prime (nat \\ x\))) \[at_top] (\x. \ x * ln (\ x))" . qed lemma PNT2_imp_PNT1: assumes "(\x. \ x * ln (\ x)) \[at_top] (\x. x)" shows "(\x. ln (\ x)) \[at_top] (\x. ln x)" and "\ \[at_top] (\x. x / ln x)" proof - have ev: "eventually (\x. \ x > 0) at_top" "eventually (\x. ln (\ x) > 0) at_top" "eventually (\x. ln (ln (\ x)) > 0) at_top" by (rule eventually_compose_filterlim[OF _ \_at_top], real_asymp)+ let ?f = "\x. 1 + ln (ln (\ x)) / ln (\ x) - ln x / ln (\ x)" have "((\x. ln (\ x) * ?f x) \ ln 1) at_top" proof (rule Lim_transform_eventually) from assms have "((\x. \ x * ln (\ x) / x) \ 1) at_top" by (rule asymp_equivD_strong[OF _ eventually_mono[OF eventually_gt_at_top[of 1]]]) auto then show "((\x. ln (\ x * ln (\ x) / x)) \ ln 1) at_top" by (rule tendsto_ln) auto show "\\<^sub>F x in at_top. ln (\ x * ln (\ x) / x) = ln (\ x) * ?f x" using eventually_gt_at_top[of 0] ev by eventually_elim (simp add: field_simps ln_mult ln_div) qed moreover have "((\x. 1 / ln (\ x)) \ 0) at_top" by (rule filterlim_compose[OF _ \_at_top]) real_asymp ultimately have "((\x. ln (\ x) * ?f x * (1 / ln (\ x))) \ ln 1 * 0) at_top" by (rule tendsto_mult) moreover have "eventually (\x. ln (\ x) * ?f x * (1 / ln (\ x)) = ?f x) at_top" using ev by eventually_elim auto ultimately have "(?f \ ln 1 * 0) at_top" by (rule Lim_transform_eventually) hence "((\x. 1 + ln (ln (\ x)) / ln (\ x) - ?f x) \ 1 + 0 - ln 1 * 0) at_top" by (intro tendsto_intros filterlim_compose[OF _ \_at_top]) (real_asymp | simp)+ hence "((\x. ln x / ln (\ x)) \ 1) at_top" by simp thus *: "(\x. ln (\ x)) \[at_top] (\x. ln x)" by (rule asymp_equiv_symI[OF asymp_equivI']) have "eventually (\x. \ x = \ x * ln (\ x) / ln (\ x)) at_top" using ev by eventually_elim auto hence "\ \[at_top] (\x. \ x * ln (\ x) / ln (\ x))" by (rule asymp_equiv_refl_ev) also from assms and * have "(\x. \ x * ln (\ x) / ln (\ x)) \[at_top] (\x. x / ln x)" by (rule asymp_equiv_intros) finally show "\ \[at_top] (\x. x / ln x)" . qed lemma PNT4_imp_PNT5: assumes "\ \[at_top] (\x. x)" shows "\ \[at_top] (\x. x)" proof - define r where "r = (\x. \ x - \ x)" have "r \ O(\x. ln x * sqrt x)" unfolding r_def by (fact \_minus_\_bigo) also have "(\x::real. ln x * sqrt x) \ o(\x. x)" by real_asymp finally have r: "r \ o(\x. x)" . have "(\x. \ x + r x) \[at_top] (\x. x)" using assms r by (subst asymp_equiv_add_right) auto thus ?thesis by (simp add: r_def) qed lemma PNT4_imp_PNT1: assumes "\ \[at_top] (\x. x)" shows "\ \[at_top] (\x. x / ln x)" proof - have "(\x. (\ x - \ x / ln x) + ((\ x - x) / ln x)) \ o(\x. x / ln x)" proof (rule sum_in_smallo) have "(\x. \ x - \ x / ln x) \ O(\x. x / ln x ^ 2)" by (rule \_\_bigo) also have "(\x. x / ln x ^ 2) \ o(\x. x / ln x :: real)" by real_asymp finally show "(\x. \ x - \ x / ln x) \ o(\x. x / ln x)" . next have "eventually (\x::real. ln x > 0) at_top" by real_asymp hence "eventually (\x::real. ln x \ 0) at_top" by eventually_elim auto thus "(\x. (\ x - x) / ln x) \ o(\x. x / ln x)" by (intro landau_o.small.divide_right asymp_equiv_imp_diff_smallo assms) qed thus ?thesis by (simp add: diff_divide_distrib asymp_equiv_altdef) qed lemma PNT1_imp_PNT4: assumes "\ \[at_top] (\x. x / ln x)" shows "\ \[at_top] (\x. x)" proof - have "\ \[at_top] (\x. \ x * ln x)" proof (rule smallo_imp_asymp_equiv) have "(\x. \ x - \ x * ln x) \ \(\x. - ((\ x - \ x / ln x) * ln x))" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 1]]) (auto simp: field_simps) also have "(\x. - ((\ x - \ x / ln x) * ln x)) \ O(\x. x / (ln x)\<^sup>2 * ln x)" unfolding landau_o.big.uminus_in_iff by (intro landau_o.big.mult_right \_\_bigo) also have "(\x::real. x / (ln x)\<^sup>2 * ln x) \ o(\x. x / ln x * ln x)" by real_asymp also have "(\x. x / ln x * ln x) \ \(\x. \ x * ln x)" by (intro asymp_equiv_imp_bigtheta asymp_equiv_intros asymp_equiv_symI[OF assms]) finally show "(\x. \ x - \ x * ln x) \ o(\x. \ x * ln x)" . qed also have "\ \[at_top] (\x. x / ln x * ln x)" by (intro asymp_equiv_intros assms) also have "\ \[at_top] (\x. x)" by real_asymp finally show ?thesis . qed lemma PNT5_imp_PNT4: assumes "\ \[at_top] (\x. x)" shows "\ \[at_top] (\x. x)" proof - define r where "r = (\x. \ x - \ x)" have "(\x. \ x - \ x) \ O(\x. ln x * sqrt x)" by (fact \_minus_\_bigo) also have "(\x. \ x - \ x) = (\x. -r x)" by (simp add: r_def) finally have "r \ O(\x. ln x * sqrt x)" by simp also have "(\x::real. ln x * sqrt x) \ o(\x. x)" by real_asymp finally have r: "r \ o(\x. x)" . have "(\x. \ x + r x) \[at_top] (\x. x)" using assms r by (subst asymp_equiv_add_right) auto thus ?thesis by (simp add: r_def) qed subsection \The asymptotic form of Mertens' First Theorem\ text \ Mertens' first theorem states that $\mathfrak{M}(x) - \ln x$ is bounded, i.\,e.\ $\mathfrak{M}(x) = \ln x + O(1)$. With some work, one can also show some absolute bounds for $|\mathfrak{M}(x) - \ln x|$, and we will, in fact, do this later. However, this asymptotic form is somewhat easier to obtain and it is (as we shall see) enough to prove the Prime Number Theorem, so we prove the weak form here first for the sake of a smoother presentation. First of all, we need a very weak version of Stirling's formula for the logarithm of the factorial, namely: \[\ln(\lfloor x\rfloor!) = \sum\limits_{n\leq x} \ln x = x \ln x + O(x)\] We show this using summation by parts. \ lemma stirling_weak: assumes x: "x \ 1" shows "sum_upto ln x \ {x * ln x - x - ln x + 1 .. x * ln x}" proof (cases "x = 1") case True have "{0<..Suc 0} = {1}" by auto with True show ?thesis by (simp add: sum_upto_altdef) next case False with assms have x: "x > 1" by simp have "((\t. sum_upto (\_. 1) t * (1 / t)) has_integral sum_upto (\_. 1) x * ln x - sum_upto (\_. 1) 1 * ln 1 - (\n\real -` {1<..x}. 1 * ln (real n))) {1..x}" using x by (intro partial_summation_strong[of "{}"]) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros continuous_intros) hence "((\t. real (nat \t\) / t) has_integral real (nat \x\) * ln x - (\n\real -` {1<..x}. ln (real n))) {1..x}" by (simp add: sum_upto_altdef) also have "(\n\real -` {1<..x}. ln (real n)) = sum_upto ln x" unfolding sum_upto_def by (intro sum.mono_neutral_left) (auto intro!: finite_subset[OF _ finite_vimage_real_of_nat_greaterThanAtMost[of 0 x]]) finally have *: "((\t. real (nat \t\) / t) has_integral \x\ * ln x - sum_upto ln x) {1..x}" using x by simp have "0 \ real_of_int \x\ * ln x - sum_upto (\n. ln (real n)) x" using * by (rule has_integral_nonneg) auto also have "\ \ x * ln x - sum_upto ln x" using x by (intro diff_mono mult_mono) auto finally have upper: "sum_upto ln x \ x * ln x" by simp have "(x - 1) * ln x - x + 1 \ \x\ * ln x - x + 1" using x by (intro diff_mono mult_mono add_mono) auto also have "((\t. 1) has_integral (x - 1)) {1..x}" using has_integral_const_real[of "1::real" 1 x] x by simp from * and this have "\x\ * ln x - sum_upto ln x \ x - 1" by (rule has_integral_le) auto hence "\x\ * ln x - x + 1 \ sum_upto ln x" by simp finally have "sum_upto ln x \ x * ln x - x - ln x + 1" by (simp add: algebra_simps) with upper show ?thesis by simp qed lemma stirling_weak_bigo: "(\x::real. sum_upto ln x - x * ln x) \ O(\x. x)" proof - have "(\x. sum_upto ln x - x * ln x) \ O(\x. -(sum_upto ln x - x * ln x))" by (subst landau_o.big.uminus) auto also have "(\x. -(sum_upto ln x - x * ln x)) \ O(\x. x + ln x - 1)" proof (intro le_imp_bigo_real[of 2] eventually_mono[OF eventually_ge_at_top[of 1]], goal_cases) case (2 x) thus ?case using stirling_weak[of x] by (auto simp: algebra_simps) next case (3 x) thus ?case using stirling_weak[of x] by (auto simp: algebra_simps) qed auto also have "(\x. x + ln x - 1) \ O(\x::real. x)" by real_asymp finally show ?thesis . qed lemma floor_floor_div_eq: fixes x :: real and d :: nat assumes "x \ 0" shows "\nat \x\ / real d\ = \x / real d\" proof - have "\nat \x\ / real_of_int (int d)\ = \x / real_of_int (int d)\" using assms by (subst (1 2) floor_divide_real_eq_div) auto thus ?thesis by simp qed text \ The key to showing Mertens' first theorem is the function \[h(x) := \sum\limits_{n \leq x} \frac{\Lambda(d)}{d}\] where $\Lambda$ is the Mangoldt function, which is equal to $\ln p$ for any prime power $p^k$ and $0$ otherwise. As we shall see, $h(x)$ is a good approximation for $\mathfrak M(x)$, as the difference between them is bounded by a constant. \ lemma sum_upto_mangoldt_over_id_minus_phi_bounded: "(\x. sum_upto (\d. mangoldt d / real d) x - \ x) \ O(\_. 1)" proof - define f where "f = (\d. mangoldt d / real d)" define C where "C = (\p. ln (real (p + 1)) * (1 / real (p * (p - 1))))" have summable: "summable (\p::nat. ln (p + 1) * (1 / (p * (p - 1))))" proof (rule summable_comparison_test_bigo) show "summable (\p. norm (p powr (-3/2)))" by (simp add: summable_real_powr_iff) qed real_asymp have diff_bound: "sum_upto f x - \ x \ {0..C}" if x: "x \ 4" for x proof - define S where "S = {(p, i). prime p \ 0 < i \ real (p ^ i) \ x}" define S' where "S' = (SIGMA p:{2..nat \root 2 x\}. {2..nat \log 2 x\})" have "S \ {..nat \x\} \ {..nat \log 2 x\}" unfolding S_def using x primepows_le_subset[of x 1] by (auto simp: Suc_le_eq) hence "finite S" by (rule finite_subset) auto note fin = finite_subset[OF _ this, unfolded S_def] have "sum_upto f x = (\(p, i)\S. ln (real p) / real (p ^ i))" unfolding S_def by (intro sum_upto_primepows) (auto simp: f_def mangoldt_non_primepow) also have "S = {p. prime p \ p \ x} \ {1} \ {(p, i). prime p \ 1 < i \ real (p ^ i) \ x}" by (auto simp: S_def not_less le_Suc_eq not_le intro!: Suc_lessI) also have "(\(p,i)\\. ln (real p) / real (p ^ i)) = (\(p, i) \ {p. prime p \ of_nat p \ x} \ {1}. ln (real p) / real (p ^ i)) + (\(p, i) | prime p \ real (p ^ i) \ x \ i > 1. ln (real p) / real (p ^ i))" (is "_ = ?S1 + ?S2") by (subst sum.union_disjoint[OF fin fin]) (auto simp: conj_commute case_prod_unfold) also have "?S1 = \ x" by (subst sum.cartesian_product [symmetric]) (auto simp: primes_M_def prime_sum_upto_def) finally have eq: "sum_upto f x - \ x = ?S2" by simp have "?S2 \ (\(p, i)\S'. ln (real p) / real (p ^ i))" using primepows_le_subset[of x 2] x unfolding case_prod_unfold of_nat_power by (intro sum_mono2 divide_nonneg_pos zero_less_power) (auto simp: eval_nat_numeral Suc_le_eq S'_def subset_iff dest: prime_gt_1_nat)+ also have "\ = (\p=2..nat \sqrt x\. ln p * (\i\{2..nat \log 2 x\}. (1 / real p) ^ i))" by (simp add: S'_def sum.Sigma case_prod_unfold sum_distrib_left sqrt_def field_simps) also have "\ \ (\p=2..nat \sqrt x\. ln p * (1 / (p * (p - 1))))" unfolding sum_upto_def proof (intro sum_mono, goal_cases) case (1 p) from x have "nat \log 2 x\ \ 2" by (auto simp: le_nat_iff' le_log_iff) hence "(\i\{2..nat \log 2 x\}. (1 / real p) ^ i) = ((1 / p)\<^sup>2 - (1 / p) ^ nat \log 2 x\ / p) / (1 - 1 / p)" using 1 by (subst sum_gp) (auto dest!: prime_gt_1_nat simp: field_simps power2_eq_square) also have "\ \ ((1 / p) ^ 2 - 0) / (1 - 1 / p)" using 1 by (intro divide_right_mono diff_mono power_mono) (auto simp: field_simps dest: prime_gt_0_nat) also have "\ = 1 / (p * (p - 1))" by (auto simp: divide_simps power2_eq_square dest: prime_gt_0_nat) finally show ?case using 1 by (intro mult_left_mono) (auto dest: prime_gt_0_nat) qed also have "\ \ (\p=2..nat \sqrt x\. ln (p + 1) * (1 / (p * (p - 1))))" by (intro sum_mono mult_mono) auto also have "\ \ C" unfolding C_def by (intro sum_le_suminf summable) auto finally have "?S2 \ C" by simp moreover have "?S2 \ 0" by (intro sum_nonneg) (auto dest: prime_gt_0_nat) ultimately show ?thesis using eq by simp qed from diff_bound[of 4] have "C \ 0" by auto with diff_bound show "(\x. sum_upto f x - \ x) \ O(\_. 1)" by (intro le_imp_bigo_real[of C] eventually_mono[OF eventually_ge_at_top[of 4]]) auto qed text \ Next, we show that our $h(x)$ itself is close to $\ln x$, i.\,e.: \[\sum\limits_{n \leq x} \frac{\Lambda(d)}{d} = \ln x + O(1)\] \ lemma sum_upto_mangoldt_over_id_asymptotics: "(\x. sum_upto (\d. mangoldt d / real d) x - ln x) \ O(\_. 1)" proof - define r where "r = (\n::real. sum_upto (\d. mangoldt d * (n / d - real_of_int \n / d\)) n)" have r: "r \ O(\)" proof (intro landau_o.bigI[of 1] eventually_mono[OF eventually_ge_at_top[of 0]]) fix x :: real assume x: "x \ 0" have eq: "{1..nat \x\} = {0<..nat \x\}" by auto hence "r x \ 0" unfolding r_def sum_upto_def by (intro sum_nonneg mult_nonneg_nonneg mangoldt_nonneg) (auto simp: floor_le_iff) moreover have "x / real d \ 1 + real_of_int \x / real d\" for d by linarith hence "r x \ sum_upto (\d. mangoldt d * 1) x" unfolding sum_upto_altdef eq r_def using x by (intro sum_mono mult_mono mangoldt_nonneg) (auto simp: less_imp_le[OF frac_lt_1] algebra_simps) ultimately show "norm (r x) \ 1 * norm (\ x)" by (simp add: \_def) qed auto also have "\ \ O(\x. x)" by (fact \_bigo) finally have r: "r \ O(\x. x)" . define r' where "r' = (\x::real. sum_upto ln x - x * ln x)" have r'_bigo: "r' \ O(\x. x)" using stirling_weak_bigo unfolding r'_def . have ln_fact: "ln (fact n) = (\d=1..n. ln d)" for n by (induction n) (simp_all add: ln_mult) hence r': "sum_upto ln n = n * ln n + r' n" for n :: real unfolding r'_def sum_upto_altdef by (auto intro!: sum.cong) have "eventually (\n. sum_upto (\d. mangoldt d / d) n - ln n = r' n / n + r n / n) at_top" using eventually_gt_at_top proof eventually_elim fix x :: real assume x: "x > 0" have "sum_upto ln x = sum_upto (\n. mangoldt n * real (nat \x / n\)) x" unfolding sum_upto_ln_conv_sum_upto_mangoldt .. also have "\ = sum_upto (\d. mangoldt d * (x / d)) x - r x" unfolding sum_upto_def by (simp add: algebra_simps sum_subtractf r_def sum_upto_def) also have "sum_upto (\d. mangoldt d * (x / d)) x = x * sum_upto (\d. mangoldt d / d) x" unfolding sum_upto_def by (subst sum_distrib_left) (simp add: field_simps) finally have "x * sum_upto (\d. mangoldt d / real d) x = r' x + r x + x * ln x" by (simp add: r' algebra_simps) thus "sum_upto (\d. mangoldt d / d) x - ln x = r' x / x + r x / x" using x by (simp add: field_simps) qed hence "(\x. sum_upto (\d. mangoldt d / d) x - ln x) \ \(\x. r' x / x + r x / x)" by (rule bigthetaI_cong) also have "(\x. r' x / x + r x / x) \ O(\_. 1)" by (intro sum_in_bigo) (insert r r'_bigo, auto simp: landau_divide_simps) finally show ?thesis . qed text \ Combining these two gives us Mertens' first theorem. \ theorem mertens_bounded: "(\x. \ x - ln x) \ O(\_. 1)" proof - define f where "f = sum_upto (\d. mangoldt d / d)" have "(\x. (f x - ln x) - (f x - \ x)) \ O(\_. 1)" using sum_upto_mangoldt_over_id_asymptotics sum_upto_mangoldt_over_id_minus_phi_bounded unfolding f_def by (rule sum_in_bigo) thus ?thesis by simp qed lemma primes_M_bigo: "\ \ O(\x. ln x)" proof - have "(\x. \ x - ln x) \ O(\_. 1)" by (rule mertens_bounded) also have "(\_::real. 1) \ O(\x. ln x)" by real_asymp finally have "(\x. \ x - ln x + ln x) \ O(\x. ln x)" by (rule sum_in_bigo) auto thus ?thesis by simp qed (*<*) unbundle no_prime_counting_notation (*>*) end diff --git a/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy b/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy --- a/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy +++ b/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy @@ -1,2048 +1,2048 @@ section \Auxiliary material\ theory Prime_Number_Theorem_Library imports Zeta_Function.Zeta_Function "HOL-Real_Asymp.Real_Asymp" begin (* TODO: Move *) lemma asymp_equivD_strong: assumes "f \[F] g" "eventually (\x. f x \ 0 \ g x \ 0) F" shows "((\x. f x / g x) \ 1) F" proof - from assms(1) have "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" by (rule asymp_equivD) also have "?this \ ?thesis" by (intro filterlim_cong eventually_mono[OF assms(2)]) auto finally show ?thesis . qed -(* TODO: replace in library *) +(* TODO: replace in library *)thm real_root_decreasing lemma real_root_decreasing: "0 < n \ n \ N \ 1 \ x \ root N x \ root n x" by (auto simp add: order_le_less real_root_strict_decreasing) lemma real_root_increasing: "0 < n \ n \ N \ 0 \ x \ x \ 1 \ root n x \ root N x" by (auto simp add: order_le_less real_root_strict_increasing) lemma frontier_real_Ici [simp]: fixes a :: real shows "frontier {a..} = {a}" unfolding frontier_def by (auto simp: interior_real_atLeast) lemma sum_upto_ln_conv_sum_upto_mangoldt: "sum_upto (\n. ln (real n)) x = sum_upto (\n. mangoldt n * nat \x / real n\) x" proof - have "sum_upto (\n. ln (real n)) x = sum_upto (\n. \d | d dvd n. mangoldt d) x" by (intro sum_upto_cong) (simp_all add: mangoldt_sum) also have "\ = sum_upto (\k. sum_upto (\d. mangoldt k) (x / real k)) x" by (rule sum_upto_sum_divisors) also have "\ = sum_upto (\n. mangoldt n * nat \x / real n\) x" unfolding sum_upto_altdef by (simp add: mult_ac) finally show ?thesis . qed lemma ln_fact_conv_sum_upto_mangoldt: "ln (fact n) = sum_upto (\k. mangoldt k * (n div k)) n" proof - have [simp]: "{0<..Suc n} = insert (Suc n) {0<..n}" for n by auto have "ln (fact n) = sum_upto (\n. ln (real n)) n" by (induction n) (auto simp: sum_upto_altdef nat_add_distrib ln_mult) also have "\ = sum_upto (\k. mangoldt k * (n div k)) n" unfolding sum_upto_ln_conv_sum_upto_mangoldt by (intro sum_upto_cong) (auto simp: floor_divide_of_nat_eq) finally show ?thesis . qed lemma powr_sum: "x \ 0 \ finite A \ x powr sum f A = (\y\A. x powr f y)" by (simp add: powr_def exp_sum sum_distrib_right) lemma fds_abs_converges_comparison_test: fixes s :: "'a :: dirichlet_series" assumes "eventually (\n. norm (fds_nth f n) \ fds_nth g n) at_top" and "fds_converges g (s \ 1)" shows "fds_abs_converges f s" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from assms(2) show "summable (\n. fds_nth g n / n powr (s \ 1))" by (auto simp: fds_converges_def) from assms(1) eventually_gt_at_top[of 0] show "eventually (\n. norm (norm (fds_nth f n / nat_power n s)) \ fds_nth g n / real n powr (s \ 1)) at_top" by eventually_elim (auto simp: norm_divide norm_nat_power intro!: divide_right_mono) qed lemma fds_converges_scaleR [intro]: assumes "fds_converges f s" shows "fds_converges (c *\<^sub>R f) s" proof - from assms have "summable (\n. c *\<^sub>R (fds_nth f n / nat_power n s))" by (intro summable_scaleR_right) (auto simp: fds_converges_def) also have "(\n. c *\<^sub>R (fds_nth f n / nat_power n s)) = (\n. (c *\<^sub>R fds_nth f n / nat_power n s))" by (simp add: scaleR_conv_of_real) finally show ?thesis by (simp add: fds_converges_def) qed lemma fds_abs_converges_scaleR [intro]: assumes "fds_abs_converges f s" shows "fds_abs_converges (c *\<^sub>R f) s" proof - from assms have "summable (\n. abs c * norm (fds_nth f n / nat_power n s))" by (intro summable_mult) (auto simp: fds_abs_converges_def) also have "(\n. abs c * norm (fds_nth f n / nat_power n s)) = (\n. norm ((c *\<^sub>R fds_nth f n) / nat_power n s))" by (simp add: norm_divide) finally show ?thesis by (simp add: fds_abs_converges_def) qed lemma conv_abscissa_scaleR: "conv_abscissa (scaleR c f) \ conv_abscissa f" by (rule conv_abscissa_mono) auto lemma abs_conv_abscissa_scaleR: "abs_conv_abscissa (scaleR c f) \ abs_conv_abscissa f" by (rule abs_conv_abscissa_mono) auto lemma fds_converges_mult_const_left [intro]: "fds_converges f s \ fds_converges (fds_const c * f) s" by (auto simp: fds_converges_def dest: summable_mult[of _ c]) lemma fds_abs_converges_mult_const_left [intro]: "fds_abs_converges f s \ fds_abs_converges (fds_const c * f) s" by (auto simp: fds_abs_converges_def norm_mult norm_divide dest: summable_mult[of _ "norm c"]) lemma conv_abscissa_mult_const_left: "conv_abscissa (fds_const c * f) \ conv_abscissa f" by (intro conv_abscissa_mono) auto lemma abs_conv_abscissa_mult_const_left: "abs_conv_abscissa (fds_const c * f) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono) auto lemma fds_converges_mult_const_right [intro]: "fds_converges f s \ fds_converges (f * fds_const c) s" by (auto simp: fds_converges_def dest: summable_mult2[of _ c]) lemma fds_abs_converges_mult_const_right [intro]: "fds_abs_converges f s \ fds_abs_converges (f * fds_const c) s" by (auto simp: fds_abs_converges_def norm_mult norm_divide dest: summable_mult2[of _ "norm c"]) lemma conv_abscissa_mult_const_right: "conv_abscissa (f * fds_const c) \ conv_abscissa f" by (intro conv_abscissa_mono) auto lemma abs_conv_abscissa_mult_const_right: "abs_conv_abscissa (f * fds_const c) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono) auto lemma bounded_coeffs_imp_fds_abs_converges: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (fds_nth f)" "s \ 1 > 1" shows "fds_abs_converges f s" proof - from assms obtain C where C: "\n. norm (fds_nth f n) \ C" by (auto simp: Bseq_def) show ?thesis proof (rule fds_abs_converges_comparison_test) from \s \ 1 > 1\ show "fds_converges (C *\<^sub>R fds_zeta) (s \ 1)" by (intro fds_abs_converges_imp_converges) auto from C show "eventually (\n. norm (fds_nth f n) \ fds_nth (C *\<^sub>R fds_zeta) n) at_top" by (intro always_eventually) (auto simp: fds_nth_zeta) qed qed lemma bounded_coeffs_imp_fds_abs_converges': fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (\n. fds_nth f n * nat_power n s0)" "s \ 1 > 1 - s0 \ 1" shows "fds_abs_converges f s" proof - have "fds_nth (fds_shift s0 f) = (\n. fds_nth f n * nat_power n s0)" by (auto simp: fun_eq_iff) with assms have "Bseq (fds_nth (fds_shift s0 f))" by simp with assms(2) have "fds_abs_converges (fds_shift s0 f) (s + s0)" by (intro bounded_coeffs_imp_fds_abs_converges) (auto simp: algebra_simps) thus ?thesis by simp qed lemma bounded_coeffs_imp_abs_conv_abscissa_le: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" and c :: ereal assumes "Bseq (\n. fds_nth f n * nat_power n s)" "1 - s \ 1 \ c" shows "abs_conv_abscissa f \ c" proof (rule abs_conv_abscissa_leI_weak) fix x assume "c < ereal x" have "ereal (1 - s \ 1) \ c" by fact also have "\ < ereal x" by fact finally have "1 - s \ 1 < ereal x" by simp thus "fds_abs_converges f (of_real x)" by (intro bounded_coeffs_imp_fds_abs_converges'[OF assms(1)]) auto qed lemma bounded_coeffs_imp_abs_conv_abscissa_le_1: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (\n. fds_nth f n)" shows "abs_conv_abscissa f \ 1" proof - have [simp]: "fds_nth f n * nat_power n 0 = fds_nth f n" for n by (cases "n = 0") auto show ?thesis by (rule bounded_coeffs_imp_abs_conv_abscissa_le[where s = 0]) (insert assms, auto simp:) qed (* TODO: replace library version *) (* EXAMPLE: This might make a good example to illustrate real_asymp *) lemma fixes a b c :: real assumes ab: "a + b > 0" and c: "c < -1" shows set_integrable_powr_at_top: "(\x. (b + x) powr c) absolutely_integrable_on {a<..}" and set_lebesgue_integral_powr_at_top: "(\x\{a<..}. ((b + x) powr c) \lborel) = -((b + a) powr (c + 1) / (c + 1))" and powr_has_integral_at_top: "((\x. (b + x) powr c) has_integral -((b + a) powr (c + 1) / (c + 1))) {a<..}" proof - let ?f = "\x. (b + x) powr c" and ?F = "\x. (b + x) powr (c + 1) / (c + 1)" have limits: "((?F \ real_of_ereal) \ ?F a) (at_right (ereal a))" "((?F \ real_of_ereal) \ 0) (at_left \)" using c ab unfolding ereal_tendsto_simps1 by (real_asymp simp: field_simps)+ have 1: "set_integrable lborel (einterval a \) ?f" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: derivative_eq_intros) thus 2: "?f absolutely_integrable_on {a<..}" by (auto simp: set_integrable_def integrable_completion) have "LBINT x=ereal a..\. (b + x) powr c = 0 - ?F a" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: derivative_eq_intros) thus 3: "(\x\{a<..}. ((b + x) powr c) \lborel) = -((b + a) powr (c + 1) / (c + 1))" by (simp add: interval_integral_to_infinity_eq) show "(?f has_integral -((b + a) powr (c + 1) / (c + 1))) {a<..}" using set_borel_integral_eq_integral[OF 1] 3 by (simp add: has_integral_iff) qed lemma fds_converges_altdef2: "fds_converges f s \ convergent (\N. eval_fds (fds_truncate N f) s)" unfolding fds_converges_def summable_iff_convergent' eval_fds_truncate by (auto simp: not_le intro!: convergent_cong always_eventually sum.mono_neutral_right) lemma tendsto_eval_fds_truncate: assumes "fds_converges f s" shows "(\N. eval_fds (fds_truncate N f) s) \ eval_fds f s" proof - have "(\N. eval_fds (fds_truncate N f) s) \ eval_fds f s \ (\N. \i\N. fds_nth f i / nat_power i s) \ eval_fds f s" unfolding eval_fds_truncate by (intro filterlim_cong always_eventually allI sum.mono_neutral_left) (auto simp: not_le) also have \ using assms by (simp add: fds_converges_iff sums_def' atLeast0AtMost) finally show ?thesis . qed lemma linepath_translate_left: "linepath (c + a) (c + a) = (\x. c + a) \ linepath a b" by (auto simp: fun_eq_iff linepath_def algebra_simps) lemma linepath_translate_right: "linepath (a + c) (b + c) = (\x. x + c) \ linepath a b" by (auto simp: fun_eq_iff linepath_def algebra_simps) lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: field_simps vector_add_divide_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma has_contour_integral_linepath_same_Im_iff: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x + Im a * \)) has_integral I) {Re a..Re b}" proof - have deriv: "vector_derivative ((\x. x - Im a * \) \ linepath a b) (at y) = b - a" for y using linepath_translate_right[of a "-Im a * \" b, symmetric] by simp have "(f has_contour_integral I) (linepath a b) \ ((\x. f (x + Im a * \)) has_contour_integral I) (linepath (a - Im a * \) (b - Im a * \))" using linepath_translate_right[of a "-Im a * \" b] deriv by (simp add: has_contour_integral) also have "\ \ ((\x. f (x + Im a * \)) has_integral I) {Re a..Re b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (auto simp: complex_is_Real_iff) finally show ?thesis . qed lemma contour_integrable_linepath_same_Im_iff: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x + Im a * \)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_same_Im_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_same_Im: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (x + Im a * \))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_same_Im_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_same_Im_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemmas [simp del] = div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1 lemma continuous_on_compact_bound: assumes "compact A" "continuous_on A f" obtains B where "B \ 0" "\x. x \ A \ norm (f x) \ B" proof - from assms(2,1) have "compact (f ` A)" by (rule compact_continuous_image) then obtain B where "\x\A. norm (f x) \ B" by (auto dest!: compact_imp_bounded simp: bounded_iff) hence "max B 0 \ 0" and "\x\A. norm (f x) \ max B 0" by auto thus ?thesis using that by blast qed interpretation cis: periodic_fun_simple cis "2 * pi" by standard (simp_all add: complex_eq_iff) lemma open_contains_cbox: fixes x :: "'a :: euclidean_space" assumes "open A" "x \ A" obtains a b where "cbox a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" proof - from assms obtain R where R: "R > 0" "ball x R \ A" by (auto simp: open_contains_ball) define r :: real where "r = R / (2 * sqrt DIM('a))" from \R > 0\ have [simp]: "r > 0" by (auto simp: r_def) define d :: 'a where "d = r *\<^sub>R Topology_Euclidean_Space.One" have "cbox (x - d) (x + d) \ A" proof safe fix y assume y: "y \ cbox (x - d) (x + d)" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" by (subst euclidean_dist_l2) (auto simp: L2_set_def) also from y have "sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2) \ sqrt (\i\(Basis::'a set). r\<^sup>2)" by (intro real_sqrt_le_mono sum_mono power_mono) (auto simp: dist_norm d_def cbox_def algebra_simps) also have "\ = sqrt (DIM('a) * r\<^sup>2)" by simp also have "DIM('a) * r\<^sup>2 = (R / 2) ^ 2" by (simp add: r_def power_divide) also have "sqrt \ = R / 2" using \R > 0\ by simp also from \R > 0\ have "\ < R" by simp finally have "y \ ball x R" by simp with R show "y \ A" by blast qed thus ?thesis using that[of "x - d" "x + d"] by (auto simp: algebra_simps d_def box_def) qed lemma open_contains_box: fixes x :: "'a :: euclidean_space" assumes "open A" "x \ A" obtains a b where "box a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" proof - from open_contains_cbox[OF assms] guess a b . with that[of a b] box_subset_cbox[of a b] show ?thesis by auto qed lemma analytic_onE_box: assumes "f analytic_on A" "s \ A" obtains a b where "Re a < Re b" "Im a < Im b" "s \ box a b" "f analytic_on box a b" proof - from assms obtain r where r: "r > 0" "f holomorphic_on ball s r" by (auto simp: analytic_on_def) with open_contains_box[of "ball s r" s] obtain a b where "box a b \ ball s r" "s \ box a b" "\i\Basis. a \ i < b \ i" by auto moreover from r have "f analytic_on ball s r" by (simp add: analytic_on_open) ultimately show ?thesis using that[of a b] analytic_on_subset[of _ "ball s r" "box a b"] by (auto simp: Basis_complex_def) qed lemma inner_image_box: assumes "(i :: 'a :: euclidean_space) \ Basis" assumes "\i\Basis. a \ i < b \ i" shows "(\x. x \ i) ` box a b = {a \ i<.. i}" proof safe fix x assume x: "x \ {a \ i<.. i}" let ?y = "(\j\Basis. (if i = j then x else (a + b) \ j / 2) *\<^sub>R j)" from x assms have "?y \ i \ (\x. x \ i) ` box a b" by (intro imageI) (auto simp: box_def algebra_simps) also have "?y \ i = (\j\Basis. (if i = j then x else (a + b) \ j / 2) * (j \ i))" by (simp add: inner_sum_left) also have "\ = (\j\Basis. if i = j then x else 0)" by (intro sum.cong) (auto simp: inner_not_same_Basis assms) also have "\ = x" using assms by simp finally show "x \ (\x. x \ i) ` box a b" . qed (insert assms, auto simp: box_def) lemma Re_image_box: assumes "Re a < Re b" "Im a < Im b" shows "Re ` box a b = {Re a<..::complex" a b] assms by (auto simp: Basis_complex_def) lemma inner_image_cbox: assumes "(i :: 'a :: euclidean_space) \ Basis" assumes "\i\Basis. a \ i \ b \ i" shows "(\x. x \ i) ` cbox a b = {a \ i..b \ i}" proof safe fix x assume x: "x \ {a \ i..b \ i}" let ?y = "(\j\Basis. (if i = j then x else a \ j) *\<^sub>R j)" from x assms have "?y \ i \ (\x. x \ i) ` cbox a b" by (intro imageI) (auto simp: cbox_def) also have "?y \ i = (\j\Basis. (if i = j then x else a \ j) * (j \ i))" by (simp add: inner_sum_left) also have "\ = (\j\Basis. if i = j then x else 0)" by (intro sum.cong) (auto simp: inner_not_same_Basis assms) also have "\ = x" using assms by simp finally show "x \ (\x. x \ i) ` cbox a b" . qed (insert assms, auto simp: cbox_def) lemma Re_image_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "Re ` cbox a b = {Re a..Re b}" using inner_image_cbox[of "1::complex" a b] assms by (auto simp: Basis_complex_def) lemma Im_image_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "Im ` cbox a b = {Im a..Im b}" using inner_image_cbox[of "\::complex" a b] assms by (auto simp: Basis_complex_def) lemma analytic_onE_cball: assumes "f analytic_on A" "s \ A" "ub > (0::real)" obtains R where "R > 0" "R < ub" "f analytic_on cball s R" proof - from assms obtain r where "r > 0" "f holomorphic_on ball s r" by (auto simp: analytic_on_def) hence "f analytic_on ball s r" by (simp add: analytic_on_open) hence "f analytic_on cball s (min (ub / 2) (r / 2))" by (rule analytic_on_subset, subst cball_subset_ball_iff) (use \r > 0\ in auto) moreover have "min (ub / 2) (r / 2) > 0" and "min (ub / 2) (r / 2) < ub" using \r > 0\ and \ub > 0\ by (auto simp: min_def) ultimately show ?thesis using that[of "min (ub / 2) (r / 2)"] by blast qed corollary analytic_pre_zeta' [analytic_intros]: assumes "f analytic_on A" "a > 0" shows "(\x. pre_zeta a (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_pre_zeta[of a UNIV]] assms(2) by (auto simp: o_def) corollary analytic_hurwitz_zeta' [analytic_intros]: assumes "f analytic_on A" "(\x. x \ A \ f x \ 1)" "a > 0" shows "(\x. hurwitz_zeta a (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_hurwitz_zeta[of a "-{1}"]] assms(2,3) by (auto simp: o_def) corollary analytic_zeta' [analytic_intros]: assumes "f analytic_on A" "(\x. x \ A \ f x \ 1)" shows "(\x. zeta (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_zeta[of "-{1}"]] assms(2) by (auto simp: o_def) lemma logderiv_zeta_analytic: "(\s. deriv zeta s / zeta s) analytic_on {s. Re s \ 1} - {1}" using zeta_Re_ge_1_nonzero by (auto intro!: analytic_intros) lemma cis_pi_half [simp]: "cis (pi / 2) = \" by (simp add: complex_eq_iff) lemma mult_real_sqrt: "x \ 0 \ x * sqrt y = sqrt (x ^ 2 * y)" by (simp add: real_sqrt_mult) lemma arcsin_pos: "x \ {0<..1} \ arcsin x > 0" using arcsin_less_arcsin[of 0 x] by simp lemmas analytic_imp_holomorphic' = holomorphic_on_subset[OF analytic_imp_holomorphic] lemma residue_simple': assumes "open s" "0 \ s" "f holomorphic_on s" shows "residue (\w. f w / w) 0 = f 0" using residue_simple[of s 0 f] assms by simp lemma fds_converges_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" "s = s'" shows "fds_converges f s \ fds_converges g s'" unfolding fds_converges_def by (intro summable_cong eventually_mono[OF assms(1)]) (simp_all add: assms) lemma fds_abs_converges_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" "s = s'" shows "fds_abs_converges f s \ fds_abs_converges g s'" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF assms(1)]) (simp_all add: assms) lemma conv_abscissa_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" shows "conv_abscissa f = conv_abscissa g" proof - have "fds_converges f = fds_converges g" by (intro ext fds_converges_cong assms refl) thus ?thesis by (simp add: conv_abscissa_def) qed lemma abs_conv_abscissa_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" shows "abs_conv_abscissa f = abs_conv_abscissa g" proof - have "fds_abs_converges f = fds_abs_converges g" by (intro ext fds_abs_converges_cong assms refl) thus ?thesis by (simp add: abs_conv_abscissa_def) qed definition fds_remainder where "fds_remainder m = fds_subseries (\n. n > m)" lemma fds_nth_remainder: "fds_nth (fds_remainder m f) = (\n. if n > m then fds_nth f n else 0)" by (simp add: fds_remainder_def fds_subseries_def fds_nth_fds') lemma fds_converges_remainder_iff [simp]: "fds_converges (fds_remainder m f) s \ fds_converges f s" by (intro fds_converges_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma fds_abs_converges_remainder_iff [simp]: "fds_abs_converges (fds_remainder m f) s \ fds_abs_converges f s" by (intro fds_abs_converges_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma fds_converges_remainder [intro]: "fds_converges f s \ fds_converges (fds_remainder m f) s" and fds_abs_converges_remainder [intro]: "fds_abs_converges f s \ fds_abs_converges (fds_remainder m f) s" by simp_all lemma conv_abscissa_remainder [simp]: "conv_abscissa (fds_remainder m f) = conv_abscissa f" by (intro conv_abscissa_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma abs_conv_abscissa_remainder [simp]: "abs_conv_abscissa (fds_remainder m f) = abs_conv_abscissa f" by (intro abs_conv_abscissa_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma eval_fds_remainder: "eval_fds (fds_remainder m f) s = (\n. fds_nth f (n + Suc m) / nat_power (n + Suc m) s)" (is "_ = suminf (\n. ?f (n + Suc m))") proof (cases "fds_converges f s") case False hence "\fds_converges (fds_remainder m f) s" by simp hence "(\x. (\n. fds_nth (fds_remainder m f) n / nat_power n s) sums x) = (\_. False)" by (auto simp: fds_converges_def summable_def) hence "eval_fds (fds_remainder m f) s = (THE _. False)" by (simp add: eval_fds_def suminf_def) moreover from False have "\summable (\n. ?f (n + Suc m))" unfolding fds_converges_def by (subst summable_iff_shift) auto hence "(\x. (\n. ?f (n + Suc m)) sums x) = (\_. False)" by (auto simp: summable_def) hence "suminf (\n. ?f (n + Suc m)) = (THE _. False)" by (simp add: suminf_def) ultimately show ?thesis by simp next case True hence *: "fds_converges (fds_remainder m f) s" by simp have "eval_fds (fds_remainder m f) s = (\n. fds_nth (fds_remainder m f) n / nat_power n s)" unfolding eval_fds_def .. also have "\ = (\n. fds_nth (fds_remainder m f) (n + Suc m) / nat_power (n + Suc m) s)" using * unfolding fds_converges_def by (subst suminf_minus_initial_segment) (auto simp: fds_nth_remainder) also have "(\n. fds_nth (fds_remainder m f) (n + Suc m)) = (\n. fds_nth f (n + Suc m))" by (intro ext) (auto simp: fds_nth_remainder) finally show ?thesis . qed lemma fds_truncate_plus_remainder: "fds_truncate m f + fds_remainder m f = f" by (intro fds_eqI) (auto simp: fds_truncate_def fds_remainder_def fds_subseries_def) lemma holomorphic_fds_eval' [holomorphic_intros]: assumes "g holomorphic_on A" "\x. x \ A \ Re (g x) > conv_abscissa f" shows "(\x. eval_fds f (g x)) holomorphic_on A" using holomorphic_on_compose_gen[OF assms(1) holomorphic_fds_eval[OF order.refl, of f]] assms(2) by (auto simp: o_def) lemma analytic_fds_eval' [analytic_intros]: assumes "g analytic_on A" "\x. x \ A \ Re (g x) > conv_abscissa f" shows "(\x. eval_fds f (g x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_fds_eval[OF order.refl, of f]] assms(2) by (auto simp: o_def) lemma homotopic_loopsI: fixes h :: "real \ real \ _" assumes "continuous_on ({0..1} \ {0..1}) h" "h ` ({0..1} \ {0..1}) \ s" "\x. x \ {0..1} \ h (0, x) = p x" "\x. x \ {0..1} \ h (1, x) = q x" "\x. x \ {0..1} \ pathfinish (h \ Pair x) = pathstart (h \ Pair x)" shows "homotopic_loops s p q" using assms unfolding homotopic_loops by (intro exI[of _ h]) auto lemma continuous_on_linepath [continuous_intros]: assumes "continuous_on A a" "continuous_on A b" "continuous_on A f" shows "continuous_on A (\x. linepath (a x) (b x) (f x))" using assms by (auto simp: linepath_def intro!: continuous_intros assms) lemma continuous_on_part_circlepath [continuous_intros]: assumes "continuous_on A c" "continuous_on A r" "continuous_on A a" "continuous_on A b" "continuous_on A f" shows "continuous_on A (\x. part_circlepath (c x) (r x) (a x) (b x) (f x))" using assms by (auto simp: part_circlepath_def intro!: continuous_intros assms) lemma homotopic_loops_part_circlepath: assumes "sphere c r \ A" and "r \ 0" and "b1 = a1 + 2 * of_int k * pi" and "b2 = a2 + 2 * of_int k * pi" shows "homotopic_loops A (part_circlepath c r a1 b1) (part_circlepath c r a2 b2)" proof - define h where "h = (\(x,y). part_circlepath c r (linepath a1 a2 x) (linepath b1 b2 x) y)" show ?thesis proof (rule homotopic_loopsI) show "continuous_on ({0..1} \ {0..1}) h" by (auto simp: h_def case_prod_unfold intro!: continuous_intros) next from assms have "h ` ({0..1} \ {0..1}) \ sphere c r" by (auto simp: h_def part_circlepath_def dist_norm norm_mult) also have "\ \ A" by fact finally show "h ` ({0..1} \ {0..1}) \ A" . next fix x :: real assume x: "x \ {0..1}" show "h (0, x) = part_circlepath c r a1 b1 x" and "h (1, x) = part_circlepath c r a2 b2 x" by (simp_all add: h_def linepath_def) have "cis (pi * (real_of_int k * 2)) = 1" using cis.plus_of_int[of 0 k] by (simp add: algebra_simps) thus "pathfinish (h \ Pair x) = pathstart (h \ Pair x)" by (simp add: h_def o_def exp_eq_polar linepath_def algebra_simps cis_mult [symmetric] cis_divide [symmetric] assms) qed qed lemma homotopic_pathsI: fixes h :: "real \ real \ _" assumes "continuous_on ({0..1} \ {0..1}) h" assumes "h ` ({0..1} \ {0..1}) \ s" assumes "\x. x \ {0..1} \ h (0, x) = p x" assumes "\x. x \ {0..1} \ h (1, x) = q x" assumes "\x. x \ {0..1} \ pathstart (h \ Pair x) = pathstart p" assumes "\x. x \ {0..1} \ pathfinish (h \ Pair x) = pathfinish p" shows "homotopic_paths s p q" using assms unfolding homotopic_paths by (intro exI[of _ h]) auto lemma part_circlepath_conv_subpath: "part_circlepath c r a b = subpath (a / (2*pi)) (b / (2*pi)) (circlepath c r)" by (simp add: part_circlepath_def circlepath_def subpath_def linepath_def algebra_simps exp_eq_polar) lemma homotopic_paths_part_circlepath: assumes "a \ b" "b \ c" assumes "path_image (part_circlepath C r a c) \ A" "r \ 0" shows "homotopic_paths A (part_circlepath C r a c) (part_circlepath C r a b +++ part_circlepath C r b c)" (is "homotopic_paths _ ?g (?h1 +++ ?h2)") proof (cases "a = c") case False with assms have "a < c" by simp define slope where "slope = (b - a) / (c - a)" from assms and \a < c\ have slope: "slope \ {0..1}" by (auto simp: field_simps slope_def) define f :: "real \ real" where "f = linepath 0 slope +++ linepath slope 1" show ?thesis proof (rule homotopic_paths_reparametrize) fix t :: real assume t: "t \ {0..1}" show "(?h1 +++ ?h2) t = ?g (f t)" proof (cases "t \ 1 / 2") case True hence "?g (f t) = C + r * cis ((1 - f t) * a + f t * c)" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) also from True \a < c\ have "(1 - f t) * a + f t * c = (1 - 2 * t) * a + 2 * t * b" unfolding f_def slope_def linepath_def joinpaths_def by (simp add: divide_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) (simp add: algebra_simps)? also from True have "C + r * cis \ = (?h1 +++ ?h2) t" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) finally show ?thesis .. next case False hence "?g (f t) = C + r * cis ((1 - f t) * a + f t * c)" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) also from False \a < c\ have "(1 - f t) * a + f t * c = (2 - 2 * t) * b + (2 * t - 1) * c" unfolding f_def slope_def linepath_def joinpaths_def by (simp add: divide_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) (simp add: algebra_simps)? also from False have "C + r * cis \ = (?h1 +++ ?h2) t" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) finally show ?thesis .. qed next from slope have "path_image f \ {0..1}" by (auto simp: f_def path_image_join closed_segment_eq_real_ivl) thus "f ` {0..1} \ {0..1}" by (simp add: path_image_def) next have "path f" unfolding f_def by auto thus "continuous_on {0..1} f" by (simp add: path_def) qed (insert assms, auto simp: f_def joinpaths_def linepath_def) next case [simp]: True with assms have [simp]: "b = c" by auto have "part_circlepath C r c c +++ part_circlepath C r c c = part_circlepath C r c c" by (simp add: fun_eq_iff joinpaths_def part_circlepath_def) thus ?thesis using assms by simp qed lemma has_contour_integral_mirror_iff: assumes "valid_path g" shows "(f has_contour_integral I) (-g) \ ((\x. -f (- x)) has_contour_integral I) g" proof - from assms have "g piecewise_differentiable_on {0..1}" by (auto simp: valid_path_def piecewise_C1_imp_differentiable) then obtain S where S: "finite S" "\x. x \ {0..1} - S \ g differentiable at x within {0..1}" unfolding piecewise_differentiable_on_def by blast have S': "g differentiable at x" if "x \ {0..1} - ({0, 1} \ S)" for x proof - from that have "x \ interior {0..1}" by auto with S(2)[of x] that show ?thesis by (auto simp: at_within_interior[of _ "{0..1}"]) qed have "(f has_contour_integral I) (-g) \ ((\x. f (- g x) * vector_derivative (-g) (at x)) has_integral I) {0..1}" by (simp add: has_contour_integral) also have "\ \ ((\x. -f (- g x) * vector_derivative g (at x)) has_integral I) {0..1}" by (intro has_integral_spike_finite_eq[of "S \ {0, 1}"]) (insert \finite S\ S', auto simp: o_def fun_Compl_def) also have "\ \ ((\x. -f (-x)) has_contour_integral I) g" by (simp add: has_contour_integral) finally show ?thesis . qed lemma contour_integral_on_mirror_iff: assumes "valid_path g" shows "f contour_integrable_on (-g) \ (\x. -f (- x)) contour_integrable_on g" by (auto simp: contour_integrable_on_def has_contour_integral_mirror_iff assms) lemma contour_integral_mirror: assumes "valid_path g" shows "contour_integral (-g) f = contour_integral g (\x. -f (- x))" proof (cases "f contour_integrable_on (-g)") case True then obtain I where I: "(f has_contour_integral I) (-g)" by (auto simp: contour_integrable_on_def) also note has_contour_integral_mirror_iff[OF assms] finally have "((\x. - f (- x)) has_contour_integral I) g" . with I show ?thesis using contour_integral_unique by blast next case False hence "\(\x. -f (-x)) contour_integrable_on g" by (auto simp: contour_integral_on_mirror_iff assms) from False and this show ?thesis by (simp add: not_integrable_contour_integral) qed lemma contour_integrable_neg_iff: "(\x. -f x) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_neg[of f g] contour_integrable_neg[of "\x. -f x" g] by auto lemma contour_integral_neg: shows "contour_integral g (\x. -f x) = -contour_integral g f" proof (cases "f contour_integrable_on g") case True thus ?thesis by (simp add: contour_integral_neg) next case False hence "\(\x. -f x) contour_integrable_on g" by (simp add: contour_integrable_neg_iff) with False show ?thesis by (simp add: not_integrable_contour_integral) qed lemma minus_cis: "-cis x = cis (x + pi)" by (simp add: complex_eq_iff) lemma path_image_part_circlepath_subset: assumes "a \ a'" "a' \ b'" "b' \ b" shows "path_image (part_circlepath c r a' b') \ path_image (part_circlepath c r a b)" using assms by (subst (1 2) path_image_part_circlepath) auto lemma part_circlepath_mirror: assumes "a' = a + pi + 2 * pi * of_int k" "b' = b + pi + 2 * pi * of_int k" "c' = -c" shows "-part_circlepath c r a b = part_circlepath c' r a' b'" proof fix x :: real have "part_circlepath c' r a' b' x = c' + r * cis (linepath a b x + pi + k * (2 * pi))" by (simp add: part_circlepath_def exp_eq_polar assms linepath_translate_right mult_ac) also have "cis (linepath a b x + pi + k * (2 * pi)) = cis (linepath a b x + pi)" by (rule cis.plus_of_int) also have "\ = -cis (linepath a b x)" by (simp add: minus_cis) also have "c' + r * \ = -part_circlepath c r a b x" by (simp add: part_circlepath_def assms exp_eq_polar) finally show "(- part_circlepath c r a b) x = part_circlepath c' r a' b' x" by simp qed lemma path_mirror [intro]: "path (g :: _ \ 'b::topological_group_add) \ path (-g)" by (auto simp: path_def intro!: continuous_intros) lemma path_mirror_iff [simp]: "path (-g :: _ \ 'b::topological_group_add) \ path g" using path_mirror[of g] path_mirror[of "-g"] by (auto simp: fun_Compl_def) lemma valid_path_mirror [intro]: "valid_path g \ valid_path (-g)" by (auto simp: valid_path_def fun_Compl_def piecewise_C1_differentiable_neg) lemma valid_path_mirror_iff [simp]: "valid_path (-g) \ valid_path g" using valid_path_mirror[of g] valid_path_mirror[of "-g"] by (auto simp: fun_Compl_def) lemma pathstart_mirror [simp]: "pathstart (-g) = -pathstart g" and pathfinish_mirror [simp]: "pathfinish (-g) = -pathfinish g" by (simp_all add: pathstart_def pathfinish_def) lemma path_image_mirror: "path_image (-g) = uminus ` path_image g" by (auto simp: path_image_def) lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma contour_integral_spike_finite_simple_path: assumes "finite A" "simple_path g" "g = g'" "\x. x \ path_image g - A \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral proof (rule integral_spike) have "finite (g -` A \ {0<..<1})" using \simple_path g\ \finite A\ by (intro finite_vimage_IntI simple_path_inj_on) auto hence "finite ({0, 1} \ g -` A \ {0<..<1})" by auto thus "negligible ({0, 1} \ g -` A \ {0<..<1})" by (rule negligible_finite) next fix x assume "x \ {0..1} - ({0, 1} \ g -` A \ {0<..<1})" hence "g x \ path_image g - A" by (auto simp: path_image_def) from assms(4)[OF this] and assms(3) show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" by simp qed proposition contour_integral_bound_part_circlepath_strong: assumes fi: "f contour_integrable_on part_circlepath z r s t" and "finite k" and le: "0 \ B" "0 < r" "s \ t" and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod (contour_integral (part_circlepath z r s t) f) \ B * r * (t - s)" proof - from fi have "(f has_contour_integral contour_integral (part_circlepath z r s t) f) (part_circlepath z r s t)" by (rule has_contour_integral_integral) from has_contour_integral_bound_part_circlepath_strong[OF this assms(2-)] show ?thesis by auto qed lemma cos_le_zero: assumes "x \ {pi/2..3*pi/2}" shows "cos x \ 0" proof - have "cos x = -cos (x - pi)" by (simp add: cos_diff) moreover from assms have "cos (x - pi) \ 0" by (intro cos_ge_zero) auto ultimately show ?thesis by simp qed lemma cos_le_zero': "x \ {-3*pi/2..-pi/2} \ cos x \ 0" using cos_le_zero[of "-x"] by simp lemma cis_minus_pi_half [simp]: "cis (- (pi / 2)) = -\" by (simp add: complex_eq_iff) lemma winding_number_join_pos_combined': "\valid_path \1 \ z \ path_image \1 \ 0 < Re (winding_number \1 z); valid_path \2 \ z \ path_image \2 \ 0 < Re (winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) lemma Union_atLeastAtMost_real_of_nat: assumes "a < b" shows "(\n\{a.. {real a..real b}" thus "x \ (\n\{a.. real a" "x < real b" by simp_all hence "x \ real (nat \x\)" "x \ real (Suc (nat \x\))" by linarith+ moreover from x have "nat \x\ \ a" "nat \x\ < b" by linarith+ ultimately have "\n\{a.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto lemma nat_sum_has_integral_floor: fixes f :: "nat \ 'a :: banach" assumes mn: "m < n" shows "((\x. f (nat \x\)) has_integral sum f {m..i. {real i..real (Suc i)}) ` {m..x. f (nat \x\)) has_integral (\X\D. f (nat \Inf X\))) {real m..real n}" proof (rule has_integral_combine_division) fix X assume X: "X \ D" have "nat \x\ = nat \Inf X\" if "x \ X - {Sup X}" for x using that X by (auto simp: D_def nat_eq_iff floor_eq_iff) hence "((\x. f (nat \x\)) has_integral f (nat \Inf X\)) X \ ((\x. f (nat \Inf X\)) has_integral f (nat \Inf X\)) X" using X by (intro has_integral_spike_eq[of "{Sup X}"]) auto also from X have "\" using has_integral_const_real[of "f (nat \Inf X\)" "Inf X" "Sup X"] by (auto simp: D_def) finally show "((\x. f (nat \x\)) has_integral f (nat \Inf X\)) X" . qed fact+ also have "(\X\D. f (nat \Inf X\)) = (\k\{m.. 'a :: banach" assumes mn: "m < n" shows "((\x. f (nat \x\)) has_integral sum f {m<..n}) {real m..real n}" proof - define D where "D = (\i. {real i..real (Suc i)}) ` {m..x. f (nat \x\)) has_integral (\X\D. f (nat \Sup X\))) {real m..real n}" proof (rule has_integral_combine_division) fix X assume X: "X \ D" have "nat \x\ = nat \Sup X\" if "x \ X - {Inf X}" for x using that X by (auto simp: D_def nat_eq_iff ceiling_eq_iff) hence "((\x. f (nat \x\)) has_integral f (nat \Sup X\)) X \ ((\x. f (nat \Sup X\)) has_integral f (nat \Sup X\)) X" using X by (intro has_integral_spike_eq[of "{Inf X}"]) auto also from X have "\" using has_integral_const_real[of "f (nat \Sup X\)" "Inf X" "Sup X"] by (auto simp: D_def) finally show "((\x. f (nat \x\)) has_integral f (nat \Sup X\)) X" . qed fact+ also have "(\X\D. f (nat \Sup X\)) = (\k\{m.. = (\k\{m<..n}. f k)" by (intro sum.reindex_bij_witness[of _ "\x. x - 1" Suc]) auto finally show ?thesis . qed lemma zeta_partial_sum_le: fixes x :: real and m :: nat assumes x: "x \ {0<..1}" shows "(\k=1..m. real k powr (x - 1)) \ real m powr x / x" proof - consider "m = 0" | "m = 1" | "m > 1" by force thus ?thesis proof cases assume m: "m > 1" hence "{1..m} = insert 1 {1<..m}" by auto also have "(\k\\. real k powr (x - 1)) = 1 + (\k\{1<..m}. real k powr (x - 1))" by simp also have "(\k\{1<..m}. real k powr (x - 1)) \ real m powr x / x - 1 / x" proof (rule has_integral_le) show "((\t. (nat \t\) powr (x - 1)) has_integral (\n\{1<..m}. n powr (x - 1))) {real 1..m}" using m by (intro nat_sum_has_integral_ceiling) auto next have "((\t. t powr (x - 1)) has_integral (real m powr x / x - real 1 powr x / x)) {real 1..real m}" by (intro fundamental_theorem_of_calculus) (insert x m, auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros) thus "((\t. t powr (x - 1)) has_integral (real m powr x / x - 1 / x)) {real 1..real m}" by simp qed (insert x, auto intro!: powr_mono2') also have "1 + (real m powr x / x - 1 / x) \ real m powr x / x" using x by (simp add: field_simps) finally show ?thesis by simp qed (use assms in auto) qed lemma zeta_partial_sum_le': fixes x :: real and m :: nat assumes x: "x > 0" and m: "m > 0" shows "(\n=1..m. real n powr (x - 1)) \ m powr x * (1 / x + 1 / m)" proof (cases "x > 1") case False with assms have "(\n=1..m. real n powr (x - 1)) \ m powr x / x" by (intro zeta_partial_sum_le) auto also have "\ \ m powr x * (1 / x + 1 / m)" using assms by (simp add: field_simps) finally show ?thesis . next case True have "(\n\{1..m}. n powr (x - 1)) = (\n\insert m {0.. = m powr (x - 1) + (\n\{0..n\{0.. real m powr x / x" proof (rule has_integral_le) show "((\t. (nat \t\) powr (x - 1)) has_integral (\n\{0..t. t powr (x - 1)) has_integral (real m powr x / x)) {real 0..real m}" using has_integral_powr_from_0[of "x - 1"] x by auto next fix t assume "t \ {real 0..real m}" with \x > 1\ show "real (nat \t\) powr (x - 1) \ t powr (x - 1)" by (cases "t = 0") (auto intro: powr_mono2) qed also have "m powr (x - 1) + m powr x / x = m powr x * (1 / x + 1 / m)" using m x by (simp add: powr_diff field_simps) finally show ?thesis by simp qed lemma natfun_bigo_1E: assumes "(f :: nat \ _) \ O(\_. 1)" obtains C where "C \ lb" "\n. norm (f n) \ C" proof - from assms obtain C N where "\n\N. norm (f n) \ C" by (auto elim!: landau_o.bigE simp: eventually_at_top_linorder) hence *: "norm (f n) \ Max ({C, lb} \ (norm ` f ` {.. N") (subst Max_ge_iff; force simp: image_iff)+ moreover have "Max ({C, lb} \ (norm ` f ` {.. lb" by (intro Max.coboundedI) auto ultimately show ?thesis using that by blast qed lemma natfun_bigo_iff_Bseq: "f \ O(\_. 1) \ Bseq f" proof assume "Bseq f" then obtain C where "C > 0" "\n. norm (f n) \ C" by (auto simp: Bseq_def) thus "f \ O(\_. 1)" by (intro bigoI[of _ C]) auto next assume "f \ O(\_. 1)" from natfun_bigo_1E[OF this, where lb = 1] obtain C where "C \ 1" "\n. norm (f n) \ C" by auto thus "Bseq f" by (auto simp: Bseq_def intro!: exI[of _ C]) qed lemma enn_decreasing_sum_le_set_nn_integral: fixes f :: "real \ ennreal" assumes decreasing: "\x y. 0 \ x \ x \ y \ f y \ f x" shows "(\n. f (real (Suc n))) \ set_nn_integral lborel {0..} f" proof - have "(\n. (f (Suc n))) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (Suc n)) \lborel)" by (subst nn_integral_cmult_indicator) auto also have "nat \x\ = Suc n" if "x \ {real n<..real (Suc n)}" for x n using that by (auto simp: nat_eq_iff ceiling_eq_iff) hence "(\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (Suc n)) \lborel) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (real (nat \x\))) \lborel)" by (intro suminf_cong nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x\(\i. {real i<..real (Suc i)}). (f (nat \x::real\)) \lborel)" by (subst nn_integral_disjoint_family) (auto simp: disjoint_family_on_def) also have "\ \ (\\<^sup>+x\{0..}. (f x) \lborel)" by (intro nn_integral_mono) (auto simp: indicator_def intro!: decreasing) finally show ?thesis . qed (* TODO replace version in library *) lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto also have "(\x. abs (indicator \ x * f x)) = (\x. indicator \ x * f x)" using nonneg by (auto simp: indicator_def fun_eq_iff) finally show ?thesis . qed lemma decreasing_sum_le_integral: fixes f :: "real \ real" assumes nonneg: "\x. x \ 0 \ f x \ 0" assumes decreasing: "\x y. 0 \ x \ x \ y \ f y \ f x" assumes integral: "(f has_integral I) {0..}" shows "summable (\i. f (real (Suc i)))" and "suminf (\i. f (real (Suc i))) \ I" proof - have [simp]: "I \ 0" by (intro has_integral_nonneg[OF integral] nonneg) auto have "(\n. ennreal (f (Suc n))) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (Suc n)) \lborel)" by (subst nn_integral_cmult_indicator) auto also have "nat \x\ = Suc n" if "x \ {real n<..real (Suc n)}" for x n using that by (auto simp: nat_eq_iff ceiling_eq_iff) hence "(\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (Suc n)) \lborel) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (real (nat \x\))) \lborel)" by (intro suminf_cong nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x\(\i. {real i<..real (Suc i)}). ennreal (f (nat \x::real\)) \lborel)" by (subst nn_integral_disjoint_family) (auto simp: disjoint_family_on_def intro!: measurable_completion) also have "\ \ (\\<^sup>+x\{0..}. ennreal (f x) \lborel)" by (intro nn_integral_mono) (auto simp: indicator_def nonneg intro!: decreasing) also have "\ = (\\<^sup>+ x. ennreal (indicat_real {0..} x * f x) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal I" using nn_integral_has_integral_lebesgue[OF nonneg integral] by (auto simp: nonneg) finally have *: "(\n. ennreal (f (Suc n))) \ ennreal I" . from * show summable: "summable (\i. f (real (Suc i)))" by (intro summable_suminf_not_top) (auto simp: top_unique intro: nonneg) note * also from summable have "(\n. ennreal (f (Suc n))) = ennreal (\n. f (Suc n))" by (subst suminf_ennreal2) (auto simp: o_def nonneg) finally show "(\n. f (real (Suc n))) \ I" by (subst (asm) ennreal_le_iff) auto qed lemma decreasing_sum_le_integral': fixes f :: "real \ real" assumes "\x. x \ 0 \ f x \ 0" assumes "\x y. 0 \ x \ x \ y \ f y \ f x" assumes "(f has_integral I) {0..}" shows "summable (\i. f (real i))" and "suminf (\i. f (real i)) \ f 0 + I" proof - have "summable ((\i. f (real (Suc i))))" using decreasing_sum_le_integral[OF assms] by (simp add: o_def) thus *: "summable (\i. f (real i))" by (subst (asm) summable_Suc_iff) have "(\n. f (real (Suc n))) \ I" by (intro decreasing_sum_le_integral assms) thus "suminf (\i. f (real i)) \ f 0 + I" using * by (subst (asm) suminf_split_head) auto qed lemma norm_suminf_le: assumes "\n. norm (f n :: 'a :: banach) \ g n" "summable g" shows "norm (suminf f) \ suminf g" proof - have *: "summable (\n. norm (f n))" using assms by (intro summable_norm summable_comparison_test[OF _ assms(2)] exI[of _ 0]) auto hence "norm (suminf f) \ (\n. norm (f n))" by (intro summable_norm) auto also have "\ \ suminf g" by (intro suminf_le * assms allI) finally show ?thesis . qed lemma of_nat_powr_neq_1_complex [simp]: assumes "n > 1" "Re s \ 0" shows "of_nat n powr s \ (1::complex)" proof - have "norm (of_nat n powr s) = real n powr Re s" by (simp add: norm_powr_real_powr) also have "\ \ 1" using assms by (auto simp: powr_def) finally show ?thesis by auto qed lemma abs_summable_on_uminus_iff: "(\x. -f x) abs_summable_on A \ f abs_summable_on A" using abs_summable_on_uminus[of f A] abs_summable_on_uminus[of "\x. -f x" A] by auto lemma abs_summable_on_cmult_right_iff: fixes f :: "'a \ 'b :: {banach, real_normed_field, second_countable_topology}" assumes "c \ 0" shows "(\x. c * f x) abs_summable_on A \ f abs_summable_on A" using assms abs_summable_on_cmult_right[of c f A] abs_summable_on_cmult_right[of "inverse c" "\x. c * f x" A] by (auto simp: field_simps) lemma abs_summable_on_cmult_left_iff: fixes f :: "'a \ 'b :: {banach, real_normed_field, second_countable_topology}" assumes "c \ 0" shows "(\x. f x * c) abs_summable_on A \ f abs_summable_on A" using assms abs_summable_on_cmult_left[of c f A] abs_summable_on_cmult_left[of "inverse c" "\x. f x * c" A] by (auto simp: field_simps) lemma fds_logderiv_completely_multiplicative: fixes f :: "'a :: {real_normed_field} fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n)" proof - have "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n) * f / f" using completely_multiplicative_fds_deriv[of "fds_nth f"] assms by simp also have "\ = - fds (\n. fds_nth f n * mangoldt n)" using assms by (simp add: divide_fds_def fds_right_inverse) finally show ?thesis . qed lemma fds_nth_logderiv_completely_multiplicative: fixes f :: "'a :: {real_normed_field} fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "fds_nth (fds_deriv f / f) n = -fds_nth f n * mangoldt n" using assms by (subst fds_logderiv_completely_multiplicative) (simp_all add: fds_nth_fds') lemma eval_fds_logderiv_completely_multiplicative: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" defines "h \ fds_deriv f / f" assumes "completely_multiplicative_function (fds_nth f)" and [simp]: "fds_nth f 1 \ 0" assumes "s \ 1 > abs_conv_abscissa f" shows "(\p. of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1)) abs_summable_on {p. prime p}" (is ?th1) and "eval_fds h s = -(\\<^sub>ap | prime p. of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))" (is ?th2) proof - let ?P = "{p::nat. prime p}" interpret f: completely_multiplicative_function "fds_nth f" by fact have "fds_abs_converges h s" using abs_conv_abscissa_completely_multiplicative_log_deriv[OF assms(2)] assms by (intro fds_abs_converges) auto hence *: "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV" by (auto simp: h_def fds_abs_converges_altdef') note * also have "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV \ (\x. -fds_nth f x * mangoldt x / nat_power x s) abs_summable_on Collect primepow" unfolding h_def using fds_nth_logderiv_completely_multiplicative[OF assms(2)] by (intro abs_summable_on_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) finally have sum1: "(\x. -fds_nth f x * mangoldt x / nat_power x s) abs_summable_on Collect primepow" by (rule abs_summable_on_subset) auto also have "?this \ (\(p,k). -fds_nth f (p ^ Suc k) * mangoldt (p ^ Suc k) / nat_power (p ^ Suc k) s) abs_summable_on (?P \ UNIV)" using bij_betw_primepows unfolding case_prod_unfold by (intro abs_summable_on_reindex_bij_betw [symmetric]) also have "\ \ (\(p,k). -((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) abs_summable_on (?P \ UNIV)" unfolding case_prod_unfold by (intro abs_summable_on_cong, subst mangoldt_primepow) (auto simp: f.mult f.power nat_power_mult_distrib nat_power_power_left power_divide dest: prime_gt_1_nat) finally have sum2: \ . have sum4: "summable (\n. (norm (fds_nth f p / nat_power p s)) ^ Suc n)" if p: "prime p" for p proof - have "summable (\n. \ln (real p)\ * (norm (fds_nth f p / nat_power p s)) ^ Suc n)" using p abs_summable_on_Sigma_project2[OF sum2, of p] unfolding abs_summable_on_nat_iff' by (simp add: norm_power norm_mult norm_divide mult_ac del: power_Suc) thus ?thesis by (rule summable_mult_D) (insert p, auto dest: prime_gt_1_nat) qed have sums: "(\n. (fds_nth f p / nat_power p s) ^ Suc n) sums (1 / (1 - fds_nth f p / nat_power p s) - 1)" if p: "prime p" for p :: nat proof - from sum4[OF p] have "norm (fds_nth f p / nat_power p s) < 1" unfolding summable_Suc_iff by (simp add: summable_geometric_iff) from geometric_sums[OF this] show ?thesis by (subst sums_Suc_iff) auto qed have eq: "(\\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) = -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))" if p: "prime p" for p proof - have "(\\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) = (\\<^sub>ak. (fds_nth f p / nat_power p s) ^ Suc k) * of_real (-ln (real p))" using sum4[of p] p by (subst infsetsum_cmult_left [symmetric]) (auto simp: abs_summable_on_nat_iff' norm_power simp del: power_Suc) also have "(\\<^sub>ak. (fds_nth f p / nat_power p s) ^ Suc k) = (1 / (1 - fds_nth f p / nat_power p s) - 1)" using sum4[OF p] sums[OF p] by (subst infsetsum_nat') (auto simp: sums_iff abs_summable_on_nat_iff' norm_power simp del: power_Suc) finally show ?thesis by (simp add: mult_ac) qed have sum3: "(\x. \\<^sub>ay. - ((fds_nth f x / nat_power x s) ^ Suc y * of_real (ln (real x)))) abs_summable_on {p. prime p}" using sum2 by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\p. -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))) abs_summable_on {p. prime p}" by (intro abs_summable_on_cong eq) auto also have "\ \ ?th1" by (subst abs_summable_on_uminus_iff) auto finally show ?th1 . have "eval_fds h s = (\\<^sub>an. fds_nth h n / nat_power n s)" using * unfolding eval_fds_def by (subst infsetsum_nat') auto also have "\ = (\\<^sub>an \ {n. primepow n}. -fds_nth f n * mangoldt n / nat_power n s)" unfolding h_def using fds_nth_logderiv_completely_multiplicative[OF assms(2)] by (intro infsetsum_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). -fds_nth f (p ^ Suc k) * mangoldt (p ^ Suc k) / nat_power (p ^ Suc k) s)" using bij_betw_primepows unfolding case_prod_unfold by (intro infsetsum_reindex_bij_betw [symmetric]) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). -((fds_nth f p / nat_power p s) ^ Suc k) * of_real (ln (real p)))" by (intro infsetsum_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps simp del: power_Suc) also have "\ = (\\<^sub>ap | prime p. \\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k) * of_real (ln (real p)))" using sum2 by (subst infsetsum_Times) (auto simp: case_prod_unfold) also have "\ = (\\<^sub>ap | prime p. -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1)))" using eq by (intro infsetsum_cong) auto finally show ?th2 by (subst (asm) infsetsum_uminus) qed lemma eval_fds_logderiv_zeta: assumes "Re s > 1" shows "(\p. of_real (ln (real p)) / (p powr s - 1)) abs_summable_on {p. prime p}" (is ?th1) and "deriv zeta s / zeta s = -(\\<^sub>ap | prime p. of_real (ln (real p)) / (p powr s - 1))" (is ?th2) proof - have *: "completely_multiplicative_function (fds_nth fds_zeta :: _ \ complex)" by standard auto note abscissa = le_less_trans[OF abs_conv_abscissa_completely_multiplicative_log_deriv[OF *]] have "(\p. ln (real p) * (1 / (1 - fds_nth fds_zeta p / p powr s) - 1)) abs_summable_on {p. prime p}" using eval_fds_logderiv_completely_multiplicative[OF *, of s] assms by auto also have "?this \ (\p. ln (real p) / (p powr s - 1)) abs_summable_on {p. prime p}" using assms by (intro abs_summable_on_cong) (auto simp: fds_nth_zeta divide_simps dest: prime_gt_1_nat) finally show ?th1 . from assms have ev: "eventually (\z. z \ {z. Re z > 1}) (nhds s)" by (intro eventually_nhds_in_open open_halfspace_Re_gt) auto have "deriv zeta s = deriv (eval_fds fds_zeta) s" by (intro deriv_cong_ev[OF eventually_mono[OF ev]]) (auto simp: eval_fds_zeta) also have "deriv (eval_fds fds_zeta) s / zeta s = eval_fds (fds_deriv fds_zeta / fds_zeta) s" using assms zeta_Re_gt_1_nonzero[of s] by (subst eval_fds_log_deriv) (auto simp: eval_fds_zeta eval_fds_deriv intro!: abscissa) also have "eval_fds (fds_deriv fds_zeta / fds_zeta) s = -(\\<^sub>ap | prime p. ln (real p) * (1 / (1 - fds_nth fds_zeta p / p powr s) - 1))" (is "_ = -?S") using eval_fds_logderiv_completely_multiplicative[OF *, of s] assms by auto also have "?S = (\\<^sub>ap | prime p. ln (real p) / (p powr s - 1))" using assms by (intro infsetsum_cong) (auto simp: fds_nth_zeta divide_simps dest: prime_gt_1_nat) finally show ?th2 . qed lemma sums_logderiv_zeta: assumes "Re s > 1" shows "(\p. if prime p then of_real (ln (real p)) / (of_nat p powr s - 1) else 0) sums -(deriv zeta s / zeta s)" (is "?f sums _") proof - note * = eval_fds_logderiv_zeta[OF assms] from sums_infsetsum_nat[OF *(1)] and *(2) show ?thesis by simp qed lemma abs_conv_abscissa_diff_le: "abs_conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" using abs_conv_abscissa_add_le[of f "-g"] by auto lemma abs_conv_abscissa_diff_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f - g) \ d" using abs_conv_abscissa_diff_le[of f g] by (auto simp: le_max_iff_disj) lemma range_add_nat: "range (\n. n + c) = {(c::nat)..}" proof safe fix x assume "x \ c" hence "x = x - c + c" by simp thus "x \ range (\n. n + c)" by blast qed auto lemma abs_summable_hurwitz_zeta: assumes "Re s > 1" "a + real b > 0" shows "(\n. 1 / (of_nat n + a) powr s) abs_summable_on {b..}" proof - from assms have "summable (\n. cmod (1 / (of_nat (n + b) + a) powr s))" using summable_hurwitz_zeta_real[of "Re s" "a + b"] by (auto simp: norm_divide powr_minus field_simps norm_powr_real_powr) hence "(\n. 1 / (of_nat (n + b) + a) powr s) abs_summable_on UNIV" by (auto simp: abs_summable_on_nat_iff' add_ac) also have "?this \ (\n. 1 / (of_nat n + a) powr s) abs_summable_on range (\n. n + b)" by (rule abs_summable_on_reindex_iff) auto also have "range (\n. n + b) = {b..}" by (rule range_add_nat) finally show ?thesis . qed lemma hurwitz_zeta_nat_conv_infsetsum: assumes "a > 0" and "Re s > 1" shows "hurwitz_zeta (real a) s = (\\<^sub>an. of_nat (n + a) powr -s)" "hurwitz_zeta (real a) s = (\\<^sub>an\{a..}. of_nat n powr -s)" proof - have "hurwitz_zeta (real a) s = (\n. of_nat (n + a) powr -s)" using assms by (subst hurwitz_zeta_conv_suminf) auto also have "\ = (\\<^sub>an. of_nat (n + a) powr -s)" using abs_summable_hurwitz_zeta[of s a 0] assms by (intro infsetsum_nat' [symmetric]) (auto simp: powr_minus field_simps) finally show "hurwitz_zeta (real a) s = (\\<^sub>an. of_nat (n + a) powr -s)" . also have "\ = (\\<^sub>an\range (\n. n + a). of_nat n powr -s)" by (rule infsetsum_reindex [symmetric]) auto also have "range (\n. n + a) = {a..}" by (rule range_add_nat) finally show "hurwitz_zeta (real a) s = (\\<^sub>an\{a..}. of_nat n powr -s)" . qed lemma continuous_on_pre_zeta [continuous_intros]: assumes "continuous_on A f" "a > 0" shows "continuous_on A (\x. pre_zeta a (f x))" proof - from assms have "continuous_on UNIV (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on[OF holomorphic_pre_zeta]) auto from continuous_on_compose2[OF this assms(1)] show ?thesis by simp qed lemma continuous_pre_zeta [continuous_intros]: assumes "continuous (at x within A) f" "a > 0" shows "continuous (at x within A) (\x. pre_zeta a (f x))" proof - have "continuous (at z) (pre_zeta a)" for z by (rule continuous_on_interior[of UNIV]) (insert assms, auto intro!: continuous_intros) from continuous_within_compose3[OF this assms(1)] show ?thesis . qed lemma pre_zeta_bound: assumes "0 < Re s" and a: "a > 0" shows "norm (pre_zeta a s) \ (1 + norm s / Re s) / 2 * a powr -Re s" proof - let ?f = "\x. - (s * (x + a) powr (-1-s))" let ?g' = "\x. norm s * (x + a) powr (-1-Re s)" let ?g = "\x. -norm s / Re s * (x + a) powr (-Re s)" define R where "R = EM_remainder 1 ?f 0" have [simp]: "-Re s - 1 = -1 - Re s" by (simp add: algebra_simps) have "\frac x - 1 / 2\ \ 1 / 2" for x :: real unfolding frac_def by linarith hence "\pbernpoly (Suc 0) x\ \ 1 / 2" for x by (simp add: pbernpoly_def bernpoly_def) moreover have "((\b. cmod s * (b + a) powr - Re s / Re s) \ 0) at_top" using \Re s > 0\ \a > 0\ by real_asymp ultimately have *: "\x. x \ real 0 \ norm (EM_remainder 1 ?f (int x)) \ (1 / 2) / fact 1 * (-?g (real x))" using \a > 0\ \Re s > 0\ by (intro norm_EM_remainder_le_strong_nat'[where g' = ?g' and Y = "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp: field_simps norm_mult norm_powr_real_powr add_eq_0_iff) have R: "norm R \ norm s / (2 * Re s) * a powr -Re s" unfolding R_def using spec[OF *, of 0] by simp from assms have "pre_zeta a s = a powr -s / 2 + R" by (simp add: pre_zeta_def pre_zeta_aux_def R_def) also have "norm \ \ a powr -Re s / 2 + norm s / (2 * Re s) * a powr -Re s" using a by (intro order.trans[OF norm_triangle_ineq] add_mono R) (auto simp: norm_powr_real_powr) also have "\ = (1 + norm s / Re s) / 2 * a powr -Re s" by (simp add: field_simps) finally show ?thesis . qed lemma pre_zeta_bound': assumes "0 < Re s" and a: "a > 0" shows "norm (pre_zeta a s) \ norm s / (Re s * a powr Re s)" proof - from assms have "norm (pre_zeta a s) \ (1 + norm s / Re s) / 2 * a powr -Re s" by (intro pre_zeta_bound) auto also have "\ = (Re s + norm s) / 2 / (Re s * a powr Re s)" using assms by (auto simp: field_simps powr_minus) also have "Re s + norm s \ norm s + norm s" by (intro add_right_mono complex_Re_le_cmod) also have "(norm s + norm s) / 2 = norm s" by simp finally show "norm (pre_zeta a s) \ norm s / (Re s * a powr Re s)" using assms by (simp add: divide_right_mono) qed lemma summable_comparison_test_bigo: fixes f :: "nat \ real" assumes "summable (\n. norm (g n))" "f \ O(g)" shows "summable f" proof - from \f \ O(g)\ obtain C where C: "eventually (\x. norm (f x) \ C * norm (g x)) at_top" by (auto elim: landau_o.bigE) thus ?thesis by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult) qed lemma deriv_zeta_eq: assumes s: "s \ 1" shows "deriv zeta s = deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2" proof - from s have ev: "eventually (\z. z \ 1) (nhds s)" by (intro t1_space_nhds) have [derivative_intros]: "(pre_zeta 1 has_field_derivative deriv (pre_zeta 1) s) (at s)" by (intro holomorphic_derivI[of _ UNIV] holomorphic_intros) auto have "((\s. pre_zeta 1 s + 1 / (s - 1)) has_field_derivative (deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2)) (at s)" using s by (auto intro!: derivative_eq_intros simp: power2_eq_square) also have "?this \ (zeta has_field_derivative (deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2)) (at s)" by (intro has_field_derivative_cong_ev eventually_mono[OF ev]) (auto simp: zeta_def hurwitz_zeta_def) finally show ?thesis by (rule DERIV_imp_deriv) qed lemma zeta_remove_zero: assumes "Re s \ 1" shows "(s - 1) * pre_zeta 1 s + 1 \ 0" proof (cases "s = 1") case False hence "(s - 1) * pre_zeta 1 s + 1 = (s - 1) * zeta s" by (simp add: zeta_def hurwitz_zeta_def divide_simps) also from False assms have "\ \ 0" using zeta_Re_ge_1_nonzero[of s] by auto finally show ?thesis . qed auto lemma eval_fds_deriv_zeta: assumes "Re s > 1" shows "eval_fds (fds_deriv fds_zeta) s = deriv zeta s" proof - have ev: "eventually (\z. z \ {z. Re z > 1}) (nhds s)" using assms by (intro eventually_nhds_in_open open_halfspace_Re_gt) auto from assms have "eval_fds (fds_deriv fds_zeta) s = deriv (eval_fds fds_zeta) s" by (subst eval_fds_deriv) auto also have "\ = deriv zeta s" by (intro deriv_cong_ev eventually_mono[OF ev]) (auto simp: eval_fds_zeta) finally show ?thesis . qed lemma length_sorted_list_of_set [simp]: "finite A \ length (sorted_list_of_set A) = card A" by (metis length_remdups_card_conv length_sort set_sorted_list_of_set sorted_list_of_set_sort_remdups) lemma le_nat_iff': "x \ nat y \ x = 0 \ y \ 0 \ int x \ y" by auto lemma sum_upto_plus1: assumes "x \ 0" shows "sum_upto f (x + 1) = sum_upto f x + f (Suc (nat \x\))" proof - have "sum_upto f (x + 1) = sum f {0<..Suc (nat \x\)}" using assms by (simp add: sum_upto_altdef nat_add_distrib) also have "{0<..Suc (nat \x\)} = insert (Suc (nat \x\)) {0<..nat \x\}" by auto also have "sum f \ = sum_upto f x + f (Suc (nat \x\))" by (subst sum.insert) (auto simp: sum_upto_altdef add_ac) finally show ?thesis . qed lemma sum_upto_minus1: assumes "x \ 1" shows "sum_upto f (x - 1) = (sum_upto f x - f (nat \x\) :: 'a :: ab_group_add)" using sum_upto_plus1[of "x - 1" f] assms by (simp add: algebra_simps nat_diff_distrib) lemma integral_smallo: fixes f g g' :: "real \ real" assumes "f \ o(g')" and "filterlim g at_top at_top" assumes "\a' x. a \ a' \ a' \ x \ f integrable_on {a'..x}" assumes deriv: "\x. x \ a \ (g has_field_derivative g' x) (at x)" assumes cont: "continuous_on {a..} g'" assumes nonneg: "\x. x \ a \ g' x \ 0" shows "(\x. integral {a..x} f) \ o(g)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" note [continuous_intros] = continuous_on_subset[OF cont] define c' where "c' = c / 2" from c have c': "c' > 0" by (simp add: c'_def) from landau_o.smallD[OF assms(1) this] obtain b where b: "\x. x \ b \ norm (f x) \ c' * norm (g' x)" unfolding eventually_at_top_linorder by blast define b' where "b' = max a b" define D where "D = norm (integral {a..b'} f)" have "filterlim (\x. c' * g x) at_top at_top" using c' by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] assms) hence "eventually (\x. c' * g x \ D - c' * g b') at_top" by (auto simp: filterlim_at_top) thus "eventually (\x. norm (integral {a..x} f) \ c * norm (g x)) at_top" using eventually_ge_at_top[of b'] proof eventually_elim case (elim x) have b': "a \ b'" "b \ b'" by (auto simp: b'_def) from elim b' have integrable: "(\x. \g' x\) integrable_on {b'..x}" by (intro integrable_continuous_real continuous_intros) auto have "integral {a..x} f = integral {a..b'} f + integral {b'..x} f" - using elim b' by (intro integral_combine [symmetric] assms) auto + using elim b' by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] assms) auto also have "norm \ \ D + norm (integral {b'..x} f)" unfolding D_def by (rule norm_triangle_ineq) also have "norm (integral {b'..x} f) \ integral {b'..x} (\x. c' * norm (g' x))" using b' elim assms c' integrable by (intro integral_norm_bound_integral b assms) auto also have "\ = c' * integral {b'..x} (\x. \g' x\)" by simp also have "integral {b'..x} (\x. \g' x\) = integral {b'..x} g'" using assms b' by (intro integral_cong) auto also have "(g' has_integral (g x - g b')) {b'..x}" using b' elim by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: has_field_derivative_at_within[OF deriv]) hence "integral {b'..x} g' = g x - g b'" by (simp add: has_integral_iff) also have "D + c' * (g x - g b') \ c * g x" using elim by (simp add: field_simps c'_def) also have "\ \ c * norm (g x)" using c by (intro mult_left_mono) auto finally show ?case by simp qed qed lemma integral_bigo: fixes f g g' :: "real \ real" assumes "f \ O(g')" and "filterlim g at_top at_top" assumes "\a' x. a \ a' \ a' \ x \ f integrable_on {a'..x}" assumes deriv: "\x. x \ a \ (g has_field_derivative g' x) (at x within {a..})" assumes cont: "continuous_on {a..} g'" assumes nonneg: "\x. x \ a \ g' x \ 0" shows "(\x. integral {a..x} f) \ O(g)" proof - note [continuous_intros] = continuous_on_subset[OF cont] from landau_o.bigE[OF assms(1)] obtain c b where c: "c > 0" and b: "\x. x \ b \ norm (f x) \ c * norm (g' x)" unfolding eventually_at_top_linorder by metis define c' where "c' = c / 2" define b' where "b' = max a b" define D where "D = norm (integral {a..b'} f)" have "filterlim (\x. c * g x) at_top at_top" using c by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] assms) hence "eventually (\x. c * g x \ D - c * g b') at_top" by (auto simp: filterlim_at_top) hence "eventually (\x. norm (integral {a..x} f) \ 2 * c * norm (g x)) at_top" using eventually_ge_at_top[of b'] proof eventually_elim case (elim x) have b': "a \ b'" "b \ b'" by (auto simp: b'_def) from elim b' have integrable: "(\x. \g' x\) integrable_on {b'..x}" by (intro integrable_continuous_real continuous_intros) auto have "integral {a..x} f = integral {a..b'} f + integral {b'..x} f" - using elim b' by (intro integral_combine [symmetric] assms) auto + using elim b' by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] assms) auto also have "norm \ \ D + norm (integral {b'..x} f)" unfolding D_def by (rule norm_triangle_ineq) also have "norm (integral {b'..x} f) \ integral {b'..x} (\x. c * norm (g' x))" using b' elim assms c integrable by (intro integral_norm_bound_integral b assms) auto also have "\ = c * integral {b'..x} (\x. \g' x\)" by simp also have "integral {b'..x} (\x. \g' x\) = integral {b'..x} g'" using assms b' by (intro integral_cong) auto also have "(g' has_integral (g x - g b')) {b'..x}" using b' elim by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: DERIV_subset[OF deriv]) hence "integral {b'..x} g' = g x - g b'" by (simp add: has_integral_iff) also have "D + c * (g x - g b') \ 2 * c * g x" using elim by (simp add: field_simps c'_def) also have "\ \ 2 * c * norm (g x)" using c by (intro mult_left_mono) auto finally show ?case by simp qed thus ?thesis by (rule bigoI) qed lemma primepows_le_subset: assumes x: "x > 0" and l: "l > 0" shows "{(p, i). prime p \ l \ i \ real (p ^ i) \ x} \ {..nat \root l x\} \ {..nat \log 2 x\}" proof safe fix p i :: nat assume pi: "prime p" "i \ l" "real (p ^ i) \ x" have "real p ^ l \ real p ^ i" using pi x l by (intro power_increasing) (auto dest: prime_gt_0_nat) also have "\ \ x" using pi by simp finally have "root l (real p ^ l) \ root l x" using x pi l by (subst real_root_le_iff) auto also have "root l (real p ^ l) = real p" using pi l by (subst real_root_pos2) auto finally show "p \ nat \root l x\" using pi l x by (simp add: le_nat_iff' le_floor_iff) from pi have "2 ^ i \ real p ^ i" using l by (intro power_mono) (auto dest: prime_gt_1_nat) also have "\ \ x" using pi by simp finally show "i \ nat \log 2 x\" using pi x by (auto simp: le_nat_iff' le_floor_iff le_log_iff powr_realpow) qed lemma mangoldt_non_primepow: "\primepow n \ mangoldt n = 0" by (auto simp: mangoldt_def) lemma le_imp_bigo_real: assumes "c \ 0" "eventually (\x. f x \ c * (g x :: real)) F" "eventually (\x. 0 \ f x) F" shows "f \ O[F](g)" proof - have "eventually (\x. norm (f x) \ c * norm (g x)) F" using assms(2,3) proof eventually_elim case (elim x) have "norm (f x) \ c * g x" using elim by simp also have "\ \ c * norm (g x)" by (intro mult_left_mono assms) auto finally show ?case . qed thus ?thesis by (intro bigoI[of _ c]) auto qed (* TODO: unneeded. But why does real_asymp not work? *) lemma ln_minus_ln_floor_bigo: "(\x. ln x - ln (real (nat \x\))) \ O(\_. 1)" proof (intro le_imp_bigo_real[of 1] eventually_mono[OF eventually_ge_at_top[of 1]]) fix x :: real assume x: "x \ 1" from x have *: "x - real (nat \x\) \ 1" by linarith from x have "ln x - ln (real (nat \x\)) \ (x - real (nat \x\)) / real (nat \x\)" by (intro ln_diff_le) auto also have "\ \ 1 / 1" using x * by (intro frac_le) auto finally show "ln x - ln (real (nat \x\)) \ 1 * 1" by simp qed auto lemma cos_geD: assumes "cos x \ cos a" "0 \ a" "a \ pi" "-pi \ x" "x \ pi" shows "x \ {-a..a}" proof (cases "x \ 0") case True with assms show ?thesis by (subst (asm) cos_mono_le_eq) auto next case False with assms show ?thesis using cos_mono_le_eq[of a "-x"] by auto qed (* TODO: Could be generalised *) lemma path_image_part_circlepath_same_Re: assumes "0 \ b" "b \ pi" "a = -b" "r \ 0" shows "path_image (part_circlepath c r a b) = sphere c r \ {s. Re s \ Re c + r * cos a}" proof safe fix z assume "z \ path_image (part_circlepath c r a b)" with assms obtain t where t: "t \ {a..b}" "z = c + of_real r * cis t" by (auto simp: path_image_part_circlepath exp_eq_polar) from t and assms show "z \ sphere c r" by (auto simp: dist_norm norm_mult) from t and assms show "Re z \ Re c + r * cos a" using cos_monotone_0_pi_le[of t b] cos_monotone_minus_pi_0'[of a t] by (cases "t \ 0") (auto intro!: mult_left_mono) next fix z assume z: "z \ sphere c r" "Re z \ Re c + r * cos a" show "z \ path_image (part_circlepath c r a b)" proof (cases "r = 0") case False with assms have r: "r > 0" by simp with z have z_eq: "z = c + r * cis (Arg (z - c))" using Arg_eq[of "z - c"] by (auto simp: dist_norm exp_eq_polar norm_minus_commute) moreover from z(2) r assms have "cos b \ cos (Arg (z - c))" by (subst (asm) z_eq) auto with assms have "Arg (z - c) \ {-b..b}" using Arg_le_pi[of "z - c"] mpi_less_Arg[of "z - c"] by (intro cos_geD) auto ultimately show "z \ path_image (part_circlepath c r a b)" using assms by (subst path_image_part_circlepath) (auto simp: exp_eq_polar) qed (insert assms z, auto simp: path_image_part_circlepath) qed lemma part_circlepath_rotate_left: "part_circlepath c r (x + a) (x + b) = (\z. c + cis x * (z - c)) \ part_circlepath c r a b" by (simp add: part_circlepath_def exp_eq_polar fun_eq_iff linepath_translate_left linepath_translate_right cis_mult add_ac) lemma part_circlepath_rotate_right: "part_circlepath c r (a + x) (b + x) = (\z. c + cis x * (z - c)) \ part_circlepath c r a b" by (simp add: part_circlepath_def exp_eq_polar fun_eq_iff linepath_translate_left linepath_translate_right cis_mult add_ac) lemma path_image_semicircle_Re_ge: assumes "r \ 0" shows "path_image (part_circlepath c r (-pi/2) (pi/2)) = sphere c r \ {s. Re s \ Re c}" by (subst path_image_part_circlepath_same_Re) (simp_all add: assms) lemma sphere_rotate: "(\z. c + cis x * (z - c)) ` sphere c r = sphere c r" proof safe fix z assume z: "z \ sphere c r" hence "z = c + cis x * (c + cis (-x) * (z - c) - c)" "c + cis (-x) * (z - c) \ sphere c r" by (auto simp: dist_norm norm_mult norm_minus_commute cis_conv_exp exp_minus field_simps norm_divide) with z show "z \ (\z. c + cis x * (z - c)) ` sphere c r" by blast qed (auto simp: dist_norm norm_minus_commute norm_mult) lemma path_image_semicircle_Re_le: assumes "r \ 0" shows "path_image (part_circlepath c r (pi/2) (3/2*pi)) = sphere c r \ {s. Re s \ Re c}" proof - let ?f = "(\z. c + cis pi * (z - c))" have *: "part_circlepath c r (pi/2) (3/2*pi) = part_circlepath c r (pi + (-pi/2)) (pi + pi/2)" by simp have "path_image (part_circlepath c r (pi/2) (3/2*pi)) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "?f ` {s. Re c \ Re s} = {s. Re c \ Re s}" by (auto simp: image_iff intro!: exI[of _ "2 * c - x" for x]) finally show ?thesis . qed lemma path_image_semicircle_Im_ge: assumes "r \ 0" shows "path_image (part_circlepath c r 0 pi) = sphere c r \ {s. Im s \ Im c}" proof - let ?f = "(\z. c + cis (pi/2) * (z - c))" have *: "part_circlepath c r 0 pi = part_circlepath c r (pi / 2 + (-pi/2)) (pi / 2 + pi/2)" by simp have "path_image (part_circlepath c r 0 pi) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "?f ` {s. Re c \ Re s} = {s. Im c \ Im s}" by (auto simp: image_iff intro!: exI[of _ "c - \ * (x - c)" for x]) finally show ?thesis . qed lemma path_image_semicircle_Im_le: assumes "r \ 0" shows "path_image (part_circlepath c r pi (2 * pi)) = sphere c r \ {s. Im s \ Im c}" proof - let ?f = "(\z. c + cis (3*pi/2) * (z - c))" have *: "part_circlepath c r pi (2*pi) = part_circlepath c r (3*pi/2 + (-pi/2)) (3*pi/2 + pi/2)" by simp have "path_image (part_circlepath c r pi (2 * pi)) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "cis (3 * pi / 2) = -\" using cis_mult[of pi "pi / 2"] by simp hence "?f ` {s. Re c \ Re s} = {s. Im c \ Im s}" by (auto simp: image_iff intro!: exI[of _ "c + \ * (x - c)" for x]) finally show ?thesis . qed lemma powr_numeral [simp]: "x \ 0 \ (x::real) powr numeral y = x ^ numeral y" using powr_numeral[of x y] by (cases "x = 0") auto lemma eval_fds_logderiv_zeta_real: assumes "x > (1 :: real)" shows "(\p. ln (real p) / (p powr x - 1)) abs_summable_on {p. prime p}" (is ?th1) and "deriv zeta (of_real x) / zeta (of_real x) = -of_real (\\<^sub>ap | prime p. ln (real p) / (p powr x - 1))" (is ?th2) proof - have "(\p. Re (of_real (ln (real p)) / (of_nat p powr of_real x - 1))) abs_summable_on {p. prime p}" using assms by (intro abs_summable_Re eval_fds_logderiv_zeta) auto also have "?this \ ?th1" by (intro abs_summable_on_cong) (auto simp: powr_Reals_eq) finally show ?th1 . show ?th2 using assms by (subst eval_fds_logderiv_zeta) (auto simp: infsetsum_of_real [symmetric] powr_Reals_eq) qed lemma fixes a b c d :: real assumes ab: "d * a + b \ 1" and c: "c < -1" and d: "d > 0" defines "C \ - ((ln (d * a + b) - 1 / (c + 1)) * (d * a + b) powr (c + 1) / (d * (c + 1)))" shows set_integrable_ln_powr_at_top: "(\x. (ln (d * x + b) * ((d * x + b) powr c))) absolutely_integrable_on {a<..}" (is ?th1) and set_lebesgue_integral_ln_powr_at_top: "(\x\{a<..}. (ln (d * x + b) * ((d * x + b) powr c)) \lborel) = C" (is ?th2) and ln_powr_has_integral_at_top: "((\x. ln (d * x + b) * (d * x + b) powr c) has_integral C) {a<..}" (is ?th3) proof - define f where "f = (\x. ln (d * x + b) * (d * x + b) powr c)" define F where "F = (\x. (ln (d * x + b) - 1 / (c + 1)) * (d * x + b) powr (c + 1) / (d * (c + 1)))" have *: "(F has_field_derivative f x) (at x)" "isCont f x" "f x \ 0" if "x > a" for x proof - have "1 \ d * a + b" by fact also have "\ < d * x + b" using that assms by (intro add_strict_right_mono mult_strict_left_mono) finally have gt_1: "d * x + b > 1" . show "(F has_field_derivative f x) (at x)" "isCont f x" using ab c d gt_1 by (auto simp: F_def f_def divide_simps intro!: derivative_eq_intros continuous_intros) (auto simp: algebra_simps powr_add)? show "f x \ 0" using gt_1 by (auto simp: f_def) qed have limits: "((F \ real_of_ereal) \ F a) (at_right (ereal a))" "((F \ real_of_ereal) \ 0) (at_left \)" using c ab d unfolding ereal_tendsto_simps1 F_def by (real_asymp; simp add: field_simps)+ have 1: "set_integrable lborel (einterval a \) f" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: * AE_I2) thus 2: "f absolutely_integrable_on {a<..}" by (auto simp: set_integrable_def integrable_completion) have "(LBINT x=ereal a..\. f x) = 0 - F a" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: *) thus 3: ?th2 by (simp add: interval_integral_to_infinity_eq F_def f_def C_def) show ?th3 using set_borel_integral_eq_integral[OF 1] 3 by (simp add: has_integral_iff f_def C_def) qed lemma ln_fact_conv_sum_upto: "ln (fact n) = sum_upto ln n" by (induction n) (auto simp: sum_upto_plus1 add.commute[of 1] ln_mult) lemma sum_upto_ln_conv_ln_fact: "sum_upto ln x = ln (fact (nat \x\))" by (simp add: ln_fact_conv_sum_upto sum_upto_altdef) lemma real_of_nat_div: "real (a div b) = real_of_int \real a / real b\" by (subst floor_divide_of_nat_eq) auto lemma integral_subset_negligible: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" "negligible (T - S)" shows "integral S f = integral T f" proof - have "integral T f = integral T (\x. if x \ S then f x else 0)" by (rule integral_spike[of "T - S"]) (use assms in auto) also have "\ = integral (S \ T) f" by (subst integral_restrict_Int) auto also have "S \ T = S" using assms by auto finally show ?thesis .. qed lemma integrable_on_cong [cong]: assumes "\x. x \ A \ f x = g x" "A = B" shows "f integrable_on A \ g integrable_on B" using has_integral_cong[of A f g, OF assms(1)] assms(2) by (auto simp: integrable_on_def) lemma measurable_sum_upto [measurable]: fixes f :: "'a \ nat \ real" assumes [measurable]: "\y. (\t. f t y) \ M \\<^sub>M borel" assumes [measurable]: "x \ M \\<^sub>M borel" shows "(\t. sum_upto (f t) (x t)) \ M \\<^sub>M borel" proof - have meas: "(\t. set_lebesgue_integral lborel {y. y \ 0 \ y - real (nat \x t\) \ 0} (\y. f t (nat \y\))) \ M \\<^sub>M borel" (is "?f \ _") unfolding set_lebesgue_integral_def by measurable also have "?f = (\t. sum_upto (f t) (x t))" proof fix t :: 'a show "?f t = sum_upto (f t) (x t)" proof (cases "x t < 1") case True hence "{y. y \ 0 \ y - real (nat \x t\) \ 0} = {0}" by auto thus ?thesis using True by (simp add: set_integral_at_point sum_upto_altdef) next case False define n where "n = nat \x t\" from False have "n > 0" by (auto simp: n_def) have *: "((\x. f t (nat \x\)) has_integral sum (f t) {0<..n}) {real 0..real n}" using \n > 0\ by (intro nat_sum_has_integral_ceiling) auto have **: "(\x. f t (nat \x\)) absolutely_integrable_on {real 0..real n}" proof (rule absolutely_integrable_absolutely_integrable_ubound) show "(\_. MAX n\{0..n}. \f t n\) absolutely_integrable_on {real 0..real n}" using \n > 0\ by (subst absolutely_integrable_on_iff_nonneg) (auto simp: Max_ge_iff intro!: exI[of _ "f t 0"]) show "(\x. f t (nat \x\)) integrable_on {real 0..real n}" using * by (simp add: has_integral_iff) next fix y :: real assume y: "y \ {real 0..real n}" have "f t (nat \y\) \ \f t (nat \y\)\" by simp also have "\ \ (MAX n\{0..n}. \f t n\)" using y by (intro Max.coboundedI) auto finally show "f t (nat \y\) \ (MAX n\{0..n}. \f t n\)" . qed have "sum (f t) {0<..n} = (\x\{real 0..real n}. f t (nat \x\) \lebesgue)" using has_integral_set_lebesgue[OF **] * by (simp add: has_integral_iff) also have "\ = (\x\{real 0..real n}. f t (nat \x\) \lborel)" unfolding set_lebesgue_integral_def by (subst integral_completion) auto also have "{real 0..real n} = {y. 0 \ y \ y - real (nat \x t\) \ 0}" by (auto simp: n_def) also have "sum (f t) {0<..n} = sum_upto (f t) (x t)" by (simp add: sum_upto_altdef n_def) finally show ?thesis .. qed qed finally show ?thesis . qed end \ No newline at end of file diff --git a/thys/Zeta_Function/Zeta_Function.thy b/thys/Zeta_Function/Zeta_Function.thy --- a/thys/Zeta_Function/Zeta_Function.thy +++ b/thys/Zeta_Function/Zeta_Function.thy @@ -1,3632 +1,3632 @@ (* File: Zeta_Function.thy Author: Manuel Eberl, TU München *) section \The Hurwitz and Riemann $\zeta$ functions\ theory Zeta_Function imports "Euler_MacLaurin.Euler_MacLaurin" "Bernoulli.Bernoulli_Zeta" "Dirichlet_Series.Dirichlet_Series_Analysis" "Winding_Number_Eval.Winding_Number_Eval" "HOL-Real_Asymp.Real_Asymp" begin subsection \Preliminary facts\ (* TODO Move once Landau symbols are in the distribution *) lemma holomorphic_on_extend: assumes "f holomorphic_on S - {\}" "\ \ interior S" "f \ O[at \](\_. 1)" shows "(\g. g holomorphic_on S \ (\z\S - {\}. g z = f z))" by (subst holomorphic_on_extend_bounded) (insert assms, auto elim!: landau_o.bigE) lemma removable_singularities: assumes "finite X" "X \ interior S" "f holomorphic_on (S - X)" assumes "\z. z \ X \ f \ O[at z](\_. 1)" shows "\g. g holomorphic_on S \ (\z\S-X. g z = f z)" using assms proof (induction arbitrary: f rule: finite_induct) case empty thus ?case by auto next case (insert z0 X f) from insert.prems and insert.hyps have z0: "z0 \ interior (S - X)" by (auto simp: interior_diff finite_imp_closed) hence "\g. g holomorphic_on (S - X) \ (\z\S - X - {z0}. g z = f z)" using insert.prems insert.hyps by (intro holomorphic_on_extend) auto then obtain g where g: "g holomorphic_on (S - X)" "\z\S - X - {z0}. g z = f z" by blast have "\h. h holomorphic_on S \ (\z\S - X. h z = g z)" proof (rule insert.IH) fix z0' assume z0': "z0' \ X" hence "eventually (\z. z \ interior S - (X - {z0'}) - {z0}) (nhds z0')" using insert.prems insert.hyps by (intro eventually_nhds_in_open open_Diff finite_imp_closed) auto hence ev: "eventually (\z. z \ S - X - {z0}) (at z0')" unfolding eventually_at_filter by eventually_elim (insert z0' insert.hyps interior_subset[of S], auto) have "g \ \[at z0'](f)" by (intro bigthetaI_cong eventually_mono[OF ev]) (insert g, auto) also have "f \ O[at z0'](\_. 1)" using z0' by (intro insert.prems) auto finally show "g \ \" . qed (insert insert.prems g, auto) then obtain h where "h holomorphic_on S" "\z\S - X. h z = g z" by blast with g have "h holomorphic_on S" "\z\S - insert z0 X. h z = f z" by auto thus ?case by blast qed lemma continuous_imp_bigo_1: assumes "continuous (at x within A) f" shows "f \ O[at x within A](\_. 1)" proof (rule bigoI_tendsto) from assms show "((\x. f x / 1) \ f x) (at x within A)" by (auto simp: continuous_within) qed auto lemma taylor_bigo_linear: assumes "f field_differentiable at x0 within A" shows "(\x. f x - f x0) \ O[at x0 within A](\x. x - x0)" proof - from assms obtain f' where "(f has_field_derivative f') (at x0 within A)" by (auto simp: field_differentiable_def) hence "((\x. (f x - f x0) / (x - x0)) \ f') (at x0 within A)" by (auto simp: has_field_derivative_iff) thus ?thesis by (intro bigoI_tendsto[where c = f']) (auto simp: eventually_at_filter) qed (* END TODO *) (* TODO Move? *) lemma powr_add_minus_powr_asymptotics: fixes a z :: complex shows "((\z. ((1 + z) powr a - 1) / z) \ a) (at 0)" proof (rule Lim_transform_eventually) have "eventually (\z::complex. z \ ball 0 1 - {0}) (at 0)" using eventually_at_ball'[of 1 "0::complex" UNIV] by (simp add: dist_norm) thus "eventually (\z. (\n. (a gchoose (Suc n)) * z ^ n) = ((1 + z) powr a - 1) / z) (at 0)" proof eventually_elim case (elim z) hence "(\n. (a gchoose n) * z ^ n) sums (1 + z) powr a" by (intro gen_binomial_complex) auto hence "(\n. (a gchoose (Suc n)) * z ^ (Suc n)) sums ((1 + z) powr a - 1)" by (subst sums_Suc_iff) simp_all also have "(\n. (a gchoose (Suc n)) * z ^ (Suc n)) = (\n. z * ((a gchoose (Suc n)) * z ^ n))" by (simp add: algebra_simps) finally have "(\n. (a gchoose (Suc n)) * z ^ n) sums (((1 + z) powr a - 1) / z)" by (rule sums_mult_D) (use elim in auto) thus ?case by (simp add: sums_iff) qed next have "conv_radius (\n. a gchoose (n + 1)) = conv_radius (\n. a gchoose n)" using conv_radius_shift[of "\n. a gchoose n" 1] by simp hence "continuous_on (cball 0 (1/2)) (\z. \n. (a gchoose (Suc n)) * (z - 0) ^ n)" using conv_radius_gchoose[of a] by (intro powser_continuous_suminf) (simp_all) hence "isCont (\z. \n. (a gchoose (Suc n)) * z ^ n) 0" by (auto intro: continuous_on_interior) thus "(\z. \n. (a gchoose Suc n) * z ^ n) \0\ a" by (auto simp: isCont_def) qed lemma complex_powr_add_minus_powr_asymptotics: fixes s :: complex assumes a: "a > 0" and s: "Re s < 1" shows "filterlim (\x. of_real (x + a) powr s - of_real x powr s) (nhds 0) at_top" proof (rule Lim_transform_eventually) show "eventually (\x. ((1 + of_real (a / x)) powr s - 1) / of_real (a / x) * of_real x powr (s - 1) * a = of_real (x + a) powr s - of_real x powr s) at_top" (is "eventually (\x. ?f x / ?g x * ?h x * _ = _) _") using eventually_gt_at_top[of a] proof eventually_elim case (elim x) have "?f x / ?g x * ?h x * a = ?f x * (a * ?h x / ?g x)" by simp also have "a * ?h x / ?g x = of_real x powr s" using elim a by (simp add: powr_diff) also have "?f x * \ = of_real (x + a) powr s - of_real x powr s" using a elim by (simp add: algebra_simps powr_times_real [symmetric]) finally show ?case . qed have "filterlim (\x. complex_of_real (a / x)) (nhds (complex_of_real 0)) at_top" by (intro tendsto_of_real real_tendsto_divide_at_top[OF tendsto_const] filterlim_ident) hence "filterlim (\x. complex_of_real (a / x)) (at 0) at_top" using a by (intro filterlim_atI) auto hence "((\x. ?f x / ?g x * ?h x * a) \ s * 0 * a) at_top" using s by (intro tendsto_mult filterlim_compose[OF powr_add_minus_powr_asymptotics] tendsto_const tendsto_neg_powr_complex_of_real filterlim_ident) auto thus "((\x. ?f x / ?g x * ?h x * a) \ 0) at_top" by simp qed (* END TODO *) lemma summable_zeta: assumes "Re s > 1" shows "summable (\n. of_nat (Suc n) powr -s)" proof - have "summable (\n. exp (complex_of_real (ln (real (Suc n))) * - s))" (is "summable ?f") by (subst summable_Suc_iff, rule summable_complex_powr_iff) (use assms in auto) also have "?f = (\n. of_nat (Suc n) powr -s)" by (simp add: powr_def algebra_simps del: of_nat_Suc) finally show ?thesis . qed lemma summable_zeta_real: assumes "x > 1" shows "summable (\n. real (Suc n) powr -x)" proof - have "summable (\n. of_nat (Suc n) powr -complex_of_real x)" using assms by (intro summable_zeta) simp_all also have "(\n. of_nat (Suc n) powr -complex_of_real x) = (\n. of_real (real (Suc n) powr -x))" by (subst powr_Reals_eq) simp_all finally show ?thesis by (subst (asm) summable_complex_of_real) qed lemma summable_hurwitz_zeta: assumes "Re s > 1" "a > 0" shows "summable (\n. (of_nat n + of_real a) powr -s)" proof - have "summable (\n. (of_nat (Suc n) + of_real a) powr -s)" proof (rule summable_comparison_test' [OF summable_zeta_real [OF assms(1)]] ) fix n :: nat have "norm ((of_nat (Suc n) + of_real a) powr -s) = (real (Suc n) + a) powr - Re s" (is "?N = _") using assms by (simp add: norm_powr_real_powr) also have "\ \ real (Suc n) powr -Re s" using assms by (intro powr_mono2') auto finally show "?N \ \" . qed thus ?thesis by (subst (asm) summable_Suc_iff) qed lemma summable_hurwitz_zeta_real: assumes "x > 1" "a > 0" shows "summable (\n. (real n + a) powr -x)" proof - have "summable (\n. (of_nat n + of_real a) powr -complex_of_real x)" using assms by (intro summable_hurwitz_zeta) simp_all also have "(\n. (of_nat n + of_real a) powr -complex_of_real x) = (\n. of_real ((real n + a) powr -x))" using assms by (subst powr_Reals_eq) simp_all finally show ?thesis by (subst (asm) summable_complex_of_real) qed subsection \Definitions\ text \ We use the Euler--MacLaurin summation formula to express $\zeta(s,a) - \frac{a^{1-s}}{s-1}$ as a polynomial plus some remainder term, which is an integral over a function of order $O(-1-2n-\mathfrak{R}(s))$. It is then clear that this integral converges uniformly to an analytic function in $s$ for all $s$ with $\mathfrak{R}(s) > -2n$. \ definition pre_zeta_aux :: "nat \ real \ complex \ complex" where "pre_zeta_aux N a s = a powr - s / 2 + (\i=1..N. (bernoulli (2 * i) / fact (2 * i)) *\<^sub>R (pochhammer s (2*i - 1) * of_real a powr (- s - of_nat (2*i - 1)))) + EM_remainder (Suc (2*N)) (\x. -(pochhammer s (Suc (2*N)) * of_real (x + a) powr (- 1 - 2*N - s))) 0" text \ By iterating the above construction long enough, we can extend this to the entire complex plane. \ definition pre_zeta :: "real \ complex \ complex" where "pre_zeta a s = pre_zeta_aux (nat (1 - \Re s / 2\)) a s" text \ We can then obtain the Hurwitz $\zeta$ function by adding back the pole at 1. Note that it is not necessary to trust that this somewhat complicated definition is, in fact, the correct one, since we will later show that this Hurwitz zeta function fulfils \[\zeta(s,a) = \sum_{n=0}^\infty \frac{1}{(n + a)^s}\] and is analytic on $\mathbb{C}\setminus \{1\}$, which uniquely defines the function due to analytic continuation. It is therefore obvious that any alternative definition that is analytic on $\mathbb{C}\setminus \{1\}$ and satisfies the above equation must be equal to our Hurwitz $\zeta$ function. \ definition hurwitz_zeta :: "real \ complex \ complex" where "hurwitz_zeta a s = (if s = 1 then 0 else pre_zeta a s + of_real a powr (1 - s) / (s - 1))" text \ The Riemann $\zeta$ function is simply the Hurwitz $\zeta$ function with $a = 1$. \ definition zeta :: "complex \ complex" where "zeta = hurwitz_zeta 1" text \ We define the $\zeta$ functions as 0 at their poles. To avoid confusion, these facts are not added as simplification rules by default. \ lemma hurwitz_zeta_1: "hurwitz_zeta c 1 = 0" by (simp add: hurwitz_zeta_def) lemma zeta_1: "zeta 1 = 0" by (simp add: zeta_def hurwitz_zeta_1) context begin private lemma holomorphic_pre_zeta_aux': assumes "a > 0" "bounded U" "open U" "U \ {s. Re s > \}" and \: "\ > - 2 * real n" shows "pre_zeta_aux n a holomorphic_on U" unfolding pre_zeta_aux_def proof (intro holomorphic_intros) define C :: real where "C = max 0 (Sup ((\s. norm (pochhammer s (Suc (2 * n)))) ` closure U))" have "compact (closure U)" using assms by (auto simp: compact_eq_bounded_closed) hence "compact ((\s. norm (pochhammer s (Suc (2 * n)))) ` closure U)" by (rule compact_continuous_image [rotated]) (auto intro!: continuous_intros) hence "bounded ((\s. norm (pochhammer s (Suc (2 * n)))) ` closure U)" by (simp add: compact_eq_bounded_closed) hence C: "cmod (pochhammer s (Suc (2 * n))) \ C" if "s \ U" for s using that closure_subset[of U] unfolding C_def by (intro max.coboundedI2 cSup_upper bounded_imp_bdd_above) (auto simp: image_iff) have C' [simp]: "C \ 0" by (simp add: C_def) let ?g = "\(x::real). C * (x + a) powr (- 1 - 2 * of_nat n - \)" let ?G = "\(x::real). C / (- 2 * of_nat n - \) * (x + a) powr (- 2 * of_nat n - \)" define poch' where "poch' = deriv (\z::complex. pochhammer z (Suc (2 * n)))" have [derivative_intros]: "((\z. pochhammer z (Suc (2 * n))) has_field_derivative poch' z) (at z within A)" for z :: complex and A unfolding poch'_def by (rule holomorphic_derivI [OF holomorphic_pochhammer [of _ UNIV]]) auto have A: "continuous_on A poch'" for A unfolding poch'_def by (rule continuous_on_subset[OF _ subset_UNIV], intro holomorphic_on_imp_continuous_on holomorphic_deriv) (auto intro: holomorphic_pochhammer) note [continuous_intros] = continuous_on_compose2[OF this _ subset_UNIV] define f' where "f' = (\z t. - (poch' z * complex_of_real (t + a) powr (- 1 - 2 * of_nat n - z) - Ln (complex_of_real (t + a)) * complex_of_real (t + a) powr (- 1 - 2 * of_nat n - z) * pochhammer z (Suc (2 * n))))" show "(\z. EM_remainder (Suc (2 * n)) (\x. - (pochhammer z (Suc (2 * n)) * complex_of_real (x + a) powr (- 1 - 2 * of_nat n - z))) 0) holomorphic_on U" unfolding pre_zeta_aux_def proof (rule holomorphic_EM_remainder[of _ ?G ?g _ _ f'], goal_cases) case (1 x) show ?case by (insert 1 \ \a > 0\, rule derivative_eq_intros refl | simp)+ (auto simp: field_simps powr_diff powr_add powr_minus) next case (2 z t x) note [derivative_intros] = has_field_derivative_powr_right [THEN DERIV_chain2] show ?case by (insert 2 \ \a > 0\, (rule derivative_eq_intros refl | (simp add: add_eq_0_iff; fail))+) (simp add: f'_def) next case 3 hence *: "complex_of_real x + complex_of_real a \ \\<^sub>\\<^sub>0" if "x \ 0" for x using nonpos_Reals_of_real_iff[of "x+a", unfolded of_real_add] that \a > 0\ by auto show ?case using \a > 0\ and * unfolding f'_def by (auto simp: case_prod_unfold add_eq_0_iff intro!: continuous_intros) next case (4 b c z e) have "- 2 * real n < \" by (fact \) also from 4 assms have "\ < Re z" by auto finally show ?case using assms 4 by (intro integrable_continuous_real continuous_intros) (auto simp: add_eq_0_iff) next case (5 t x s) thus ?case using \a > 0\ by (intro integrable_EM_remainder') (auto intro!: continuous_intros simp: add_eq_0_iff) next case 6 from \ have "(\y. C / (-2 * real n - \) * (a + y) powr (-2 * real n - \)) \ 0" by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially filterlim_tendsto_add_at_top [OF tendsto_const]) auto thus ?case unfolding convergent_def by (auto simp: add_ac) next case 7 show ?case proof (intro eventually_mono [OF eventually_ge_at_top[of 1]] ballI) fix x :: real and s :: complex assume x: "x \ 1" and s: "s \ U" have "norm (- (pochhammer s (Suc (2 * n)) * of_real (x + a) powr (- 1 - 2 * of_nat n - s))) = norm (pochhammer s (Suc (2 * n))) * (x + a) powr (-1 - 2 * of_nat n - Re s)" (is "?N = _") using 7 \a > 0\ x by (simp add: norm_mult norm_powr_real_powr) also have "\ \ ?g x" using 7 assms x s \a > 0\ by (intro mult_mono C powr_mono) auto finally show "?N \ ?g x" . qed qed (insert assms, auto) qed (insert assms, auto) lemma analytic_pre_zeta_aux: assumes "a > 0" shows "pre_zeta_aux n a analytic_on {s. Re s > - 2 * real n}" unfolding analytic_on_def proof fix s assume s: "s \ {s. Re s > - 2 * real n}" define \ where "\ = (Re s - 2 * real n) / 2" with s have \: "\ > - 2 * real n" by (simp add: \_def field_simps) from s have s': "s \ {s. Re s > \}" by (auto simp: \_def field_simps) have "open {s. Re s > \}" by (rule open_halfspace_Re_gt) with s' obtain \ where "\ > 0" "ball s \ \ {s. Re s > \}" unfolding open_contains_ball by blast with \ have "pre_zeta_aux n a holomorphic_on ball s \" by (intro holomorphic_pre_zeta_aux' [OF assms, of _ \]) auto with \\ > 0\ show "\e>0. pre_zeta_aux n a holomorphic_on ball s e" by blast qed end context fixes s :: complex and N :: nat and \ :: "complex \ complex" and a :: real assumes s: "Re s > 1" and a: "a > 0" defines "\ \ (\s. \n. (of_nat n + of_real a) powr -s)" begin interpretation \: euler_maclaurin_nat' "\x. of_real (x + a) powr (1 - s) / (1 - s)" "\x. of_real (x + a) powr -s" "\n x. (-1) ^ n * pochhammer s n * of_real (x + a) powr -(s + n)" 0 N "\ s" "{}" proof (standard, goal_cases) case 2 show ?case by (simp add: powr_minus field_simps) next case (3 k) have "complex_of_real x + complex_of_real a = 0 \ x = -a" for x by (simp only: of_real_add [symmetric] of_real_eq_0_iff add_eq_0_iff2) with a s show ?case by (intro continuous_intros) (auto simp: add_nonneg_nonneg) next case (4 k x) with a have "0 < x + a" by simp hence *: "complex_of_real x + complex_of_real a \ \\<^sub>\\<^sub>0" using nonpos_Reals_of_real_iff[of "x+a", unfolded of_real_add] by auto have **: "pochhammer z (Suc n) = - pochhammer z n * (-z - of_nat n :: complex)" for z n by (simp add: pochhammer_rec' field_simps) show "((\x. (- 1) ^ k * pochhammer s k * of_real (x + a) powr - (s + of_nat k)) has_vector_derivative (- 1) ^ Suc k * pochhammer s (Suc k) * of_real (x + a) powr - (s + of_nat (Suc k))) (at x)" by (insert 4 *, (rule has_vector_derivative_real_field derivative_eq_intros refl | simp)+) (auto simp: divide_simps powr_add powr_diff powr_minus **) next case 5 with s a show ?case by (auto intro!: continuous_intros simp: minus_equation_iff add_eq_0_iff) next case (6 x) with a have "0 < x + a" by simp hence *: "complex_of_real x + complex_of_real a \ \\<^sub>\\<^sub>0" using nonpos_Reals_of_real_iff[of "x+a", unfolded of_real_add] by auto show ?case unfolding of_real_add by (insert 6 s *, (rule has_vector_derivative_real_field derivative_eq_intros refl | force simp add: minus_equation_iff)+) next case 7 from s a have "(\k. (of_nat k + of_real a) powr -s) sums \ s" unfolding \_def by (intro summable_sums summable_hurwitz_zeta) auto hence 1: "(\b. (\k=0..b. (of_nat k + of_real a) powr -s)) \ \ s" by (simp add: sums_def') { fix z assume "Re z < 0" hence "((\b. (a + real b) powr Re z) \ 0) at_top" by (intro tendsto_neg_powr filterlim_tendsto_add_at_top filterlim_real_sequentially) auto also have "(\b. (a + real b) powr Re z) = (\b. norm ((of_nat b + a) powr z))" using a by (subst norm_powr_real_powr) (auto simp: add_ac) finally have "((\b. (of_nat b + a) powr z) \ 0) at_top" by (subst (asm) tendsto_norm_zero_iff) simp } note * = this have "(\b. (of_nat b + a) powr (1 - s) / (1 - s)) \ 0 / (1 - s)" using s by (intro tendsto_divide tendsto_const *) auto hence 2: "(\b. (of_nat b + a) powr (1 - s) / (1 - s)) \ 0" by simp have "(\b. (\i<2 * N + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ((- 1) ^ i * pochhammer s i * (of_nat b + a) powr -(s + of_nat i)))) \ (\i<2 * N + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ((- 1) ^ i * pochhammer s i * 0))" using s by (intro tendsto_intros *) auto hence 3: "(\b. (\i<2 * N + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ((- 1) ^ i * pochhammer s i * (of_nat b + a) powr -(s + of_nat i)))) \ 0" by simp from tendsto_diff[OF tendsto_diff[OF 1 2] 3] show ?case by simp qed simp_all text \ The pre-$\zeta$ functions agree with the infinite sum that is used to define the $\zeta$ function for $\mathfrak{R}(s)>1$. \ lemma pre_zeta_aux_conv_zeta: "pre_zeta_aux N a s = \ s + a powr (1 - s) / (1 - s)" proof - let ?R = "(\i=1..N. ((bernoulli (2*i) / fact (2*i)) *\<^sub>R pochhammer s (2*i - 1) * of_real a powr (-s - (2*i-1))))" let ?S = "EM_remainder (Suc (2 * N)) (\x. - (pochhammer s (Suc (2*N)) * of_real (x + a) powr (- 1 - 2 * of_nat N - s))) 0" from \.euler_maclaurin_strong_nat'[OF le_refl, simplified] have "of_real a powr -s = a powr (1 - s) / (1 - s) + \ s + a powr -s / 2 + (-?R) - ?S" unfolding sum_negf [symmetric] by (simp add: scaleR_conv_of_real pre_zeta_aux_def mult_ac) thus ?thesis unfolding pre_zeta_aux_def (* TODO: Field_as_Ring causes some problems with field_simps vs. div_mult_self *) by (simp add: field_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) qed end text \ Since all of the partial pre-$\zeta$ functions are analytic and agree in the halfspace with $\mathfrak R(s)>0$, they must agree in their entire domain. \ lemma pre_zeta_aux_eq: assumes "m \ n" "a > 0" "Re s > -2 * real m" shows "pre_zeta_aux m a s = pre_zeta_aux n a s" proof - have "pre_zeta_aux n a s - pre_zeta_aux m a s = 0" proof (rule analytic_continuation[of "\s. pre_zeta_aux n a s - pre_zeta_aux m a s"]) show "(\s. pre_zeta_aux n a s - pre_zeta_aux m a s) holomorphic_on {s. Re s > -2 * real m}" using assms by (intro holomorphic_intros analytic_imp_holomorphic analytic_on_subset[OF analytic_pre_zeta_aux]) auto next fix s assume "s \ {s. Re s > 1}" with \a > 0\ show "pre_zeta_aux n a s - pre_zeta_aux m a s = 0" by (simp add: pre_zeta_aux_conv_zeta) next have "2 \ {s. Re s > 1}" by simp also have "\ = interior \" by (intro interior_open [symmetric] open_halfspace_Re_gt) finally show "2 islimpt {s. Re s > 1}" by (rule interior_limit_point) next show "connected {s. Re s > -2 * real m}" using convex_halfspace_gt[of "-2 * real m" "1::complex"] by (intro convex_connected) auto qed (insert assms, auto simp: open_halfspace_Re_gt) thus ?thesis by simp qed lemma pre_zeta_aux_eq': assumes "a > 0" "Re s > -2 * real m" "Re s > -2 * real n" shows "pre_zeta_aux m a s = pre_zeta_aux n a s" proof (cases m n rule: linorder_cases) case less with assms show ?thesis by (intro pre_zeta_aux_eq) auto next case greater with assms show ?thesis by (subst eq_commute, intro pre_zeta_aux_eq) auto qed auto lemma pre_zeta_aux_eq_pre_zeta: assumes "Re s > -2 * real n" and "a > 0" shows "pre_zeta_aux n a s = pre_zeta a s" unfolding pre_zeta_def proof (intro pre_zeta_aux_eq') from assms show "- 2 * real (nat (1 - \Re s / 2\)) < Re s" by linarith qed (insert assms, simp_all) text \ This means that the idea of iterating that construction infinitely does yield a well-defined entire function. \ lemma analytic_pre_zeta: assumes "a > 0" shows "pre_zeta a analytic_on A" unfolding analytic_on_def proof fix s assume "s \ A" let ?B = "{s'. Re s' > of_int \Re s\ - 1}" have s: "s \ ?B" by simp linarith? moreover have "open ?B" by (rule open_halfspace_Re_gt) ultimately obtain \ where \: "\ > 0" "ball s \ \ ?B" unfolding open_contains_ball by blast define C where "C = ball s \" note analytic = analytic_on_subset[OF analytic_pre_zeta_aux] have "pre_zeta_aux (nat \- Re s\ + 2) a holomorphic_on C" proof (intro analytic_imp_holomorphic analytic subsetI assms, goal_cases) case (1 w) with \ have "w \ ?B" by (auto simp: C_def) thus ?case by (auto simp: ceiling_minus) qed also have "?this \ pre_zeta a holomorphic_on C" proof (intro holomorphic_cong refl pre_zeta_aux_eq_pre_zeta assms) fix w assume "w \ C" with \ have w: "w \ ?B" by (auto simp: C_def) thus " - 2 * real (nat \- Re s\ + 2) < Re w" by (simp add: ceiling_minus) qed finally show "\e>0. pre_zeta a holomorphic_on ball s e" using \\ > 0\ unfolding C_def by blast qed lemma holomorphic_pre_zeta [holomorphic_intros]: "f holomorphic_on A \ a > 0 \ (\z. pre_zeta a (f z)) holomorphic_on A" using holomorphic_on_compose [OF _ analytic_imp_holomorphic [OF analytic_pre_zeta], of f] by (simp add: o_def) corollary continuous_on_pre_zeta: "a > 0 \ continuous_on A (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto corollary continuous_on_pre_zeta' [continuous_intros]: "continuous_on A f \ a > 0 \ continuous_on A (\x. pre_zeta a (f x))" using continuous_on_compose2 [OF continuous_on_pre_zeta, of a A f "f ` A"] by (auto simp: image_iff) corollary continuous_pre_zeta [continuous_intros]: "a > 0 \ continuous (at s within A) (pre_zeta a)" by (rule continuous_within_subset[of _ UNIV]) (insert continuous_on_pre_zeta[of a UNIV], auto simp: continuous_on_eq_continuous_at open_Compl) corollary continuous_pre_zeta' [continuous_intros]: "a > 0 \ continuous (at s within A) f \ continuous (at s within A) (\s. pre_zeta a (f s))" using continuous_within_compose3[OF continuous_pre_zeta, of a s A f] by auto text \ It is now obvious that $\zeta$ is holomorphic everywhere except 1, where it has a simple pole with residue 1, which we can simply read off. \ theorem holomorphic_hurwitz_zeta: assumes "a > 0" "1 \ A" shows "hurwitz_zeta a holomorphic_on A" proof - have "(\s. pre_zeta a s + complex_of_real a powr (1 - s) / (s - 1)) holomorphic_on A" using assms by (auto intro!: holomorphic_intros) also from assms have "?this \ ?thesis" by (intro holomorphic_cong) (auto simp: hurwitz_zeta_def) finally show ?thesis . qed corollary holomorphic_hurwitz_zeta' [holomorphic_intros]: assumes "f holomorphic_on A" and "a > 0" and "\z. z \ A \ f z \ 1" shows "(\x. hurwitz_zeta a (f x)) holomorphic_on A" proof - have "hurwitz_zeta a \ f holomorphic_on A" using assms by (intro holomorphic_on_compose_gen[of _ _ _ "f ` A"] holomorphic_hurwitz_zeta assms) auto thus ?thesis by (simp add: o_def) qed theorem holomorphic_zeta: "1 \ A\ zeta holomorphic_on A" unfolding zeta_def by (auto intro: holomorphic_intros) corollary holomorphic_zeta' [holomorphic_intros]: assumes "f holomorphic_on A" and "\z. z \ A \ f z \ 1" shows "(\x. zeta (f x)) holomorphic_on A" using assms unfolding zeta_def by (auto intro: holomorphic_intros) corollary analytic_hurwitz_zeta: assumes "a > 0" "1 \ A" shows "hurwitz_zeta a analytic_on A" proof - from assms(1) have "hurwitz_zeta a holomorphic_on -{1}" by (rule holomorphic_hurwitz_zeta) auto also have "?this \ hurwitz_zeta a analytic_on -{1}" by (intro analytic_on_open [symmetric]) auto finally show ?thesis by (rule analytic_on_subset) (insert assms, auto) qed corollary analytic_zeta: "1 \ A \ zeta analytic_on A" unfolding zeta_def by (rule analytic_hurwitz_zeta) auto corollary continuous_on_hurwitz_zeta: "a > 0 \ 1 \ A \ continuous_on A (hurwitz_zeta a)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto corollary continuous_on_hurwitz_zeta' [continuous_intros]: "continuous_on A f \ a > 0 \ (\x. x \ A \ f x \ 1) \ continuous_on A (\x. hurwitz_zeta a (f x))" using continuous_on_compose2 [OF continuous_on_hurwitz_zeta, of a "f ` A" A f] by (auto simp: image_iff) corollary continuous_on_zeta: "1 \ A \ continuous_on A zeta" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto corollary continuous_on_zeta' [continuous_intros]: "continuous_on A f \ (\x. x \ A \ f x \ 1) \ continuous_on A (\x. zeta (f x))" using continuous_on_compose2 [OF continuous_on_zeta, of "f ` A" A f] by (auto simp: image_iff) corollary continuous_hurwitz_zeta [continuous_intros]: "a > 0 \ s \ 1 \ continuous (at s within A) (hurwitz_zeta a)" by (rule continuous_within_subset[of _ UNIV]) (insert continuous_on_hurwitz_zeta[of a "-{1}"], auto simp: continuous_on_eq_continuous_at open_Compl) corollary continuous_hurwitz_zeta' [continuous_intros]: "a > 0 \ f s \ 1 \ continuous (at s within A) f \ continuous (at s within A) (\s. hurwitz_zeta a (f s))" using continuous_within_compose3[OF continuous_hurwitz_zeta, of a f s A] by auto corollary continuous_zeta [continuous_intros]: "s \ 1 \ continuous (at s within A) zeta" unfolding zeta_def by (intro continuous_intros) auto corollary continuous_zeta' [continuous_intros]: "f s \ 1 \ continuous (at s within A) f \ continuous (at s within A) (\s. zeta (f s))" unfolding zeta_def by (intro continuous_intros) auto corollary field_differentiable_at_zeta: assumes "s \ 1" shows "zeta field_differentiable at s" proof - have "zeta holomorphic_on (- {1})" using holomorphic_zeta by force moreover have "open (-{1} :: complex set)" by (intro open_Compl) auto ultimately show ?thesis using assms by (auto simp add: holomorphic_on_open open_halfspace_Re_gt open_Diff field_differentiable_def) qed theorem is_pole_hurwitz_zeta: assumes "a > 0" shows "is_pole (hurwitz_zeta a) 1" proof - from assms have "continuous_on UNIV (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on analytic_imp_holomorphic analytic_pre_zeta) hence "isCont (pre_zeta a) 1" by (auto simp: continuous_on_eq_continuous_at) hence *: "pre_zeta a \1\ pre_zeta a 1" by (simp add: isCont_def) from assms have "isCont (\s. complex_of_real a powr (1 - s)) 1" by (intro isCont_powr_complex) auto with assms have **: "(\s. complex_of_real a powr (1 - s)) \1\ 1" by (simp add: isCont_def) have "(\s::complex. s - 1) \1\ 1 - 1" by (intro tendsto_intros) hence "filterlim (\s::complex. s - 1) (at 0) (at 1)" by (auto simp: filterlim_at eventually_at_filter) hence ***: "filterlim (\s :: complex. a powr (1 - s) / (s - 1)) at_infinity (at 1)" by (intro filterlim_divide_at_infinity [OF **]) auto have "is_pole (\s. pre_zeta a s + complex_of_real a powr (1 - s) / (s - 1)) 1" unfolding is_pole_def hurwitz_zeta_def by (rule tendsto_add_filterlim_at_infinity * ***)+ also have "?this \ ?thesis" unfolding is_pole_def by (intro filterlim_cong refl) (auto simp: eventually_at_filter hurwitz_zeta_def) finally show ?thesis . qed corollary is_pole_zeta: "is_pole zeta 1" unfolding zeta_def by (rule is_pole_hurwitz_zeta) auto theorem zorder_hurwitz_zeta: assumes "a > 0" shows "zorder (hurwitz_zeta a) 1 = -1" proof (rule zorder_eqI[of UNIV]) fix w :: complex assume "w \ UNIV" "w \ 1" thus "hurwitz_zeta a w = (pre_zeta a w * (w - 1) + a powr (1 - w)) * (w - 1) powr (of_int (-1))" apply (subst (1) powr_of_int) by (auto simp add: hurwitz_zeta_def field_simps) qed (insert assms, auto intro!: holomorphic_intros) corollary zorder_zeta: "zorder zeta 1 = - 1" unfolding zeta_def by (rule zorder_hurwitz_zeta) auto theorem residue_hurwitz_zeta: assumes "a > 0" shows "residue (hurwitz_zeta a) 1 = 1" proof - note holo = analytic_imp_holomorphic[OF analytic_pre_zeta] have "residue (hurwitz_zeta a) 1 = residue (\z. pre_zeta a z + a powr (1 - z) / (z - 1)) 1" by (intro residue_cong) (auto simp: eventually_at_filter hurwitz_zeta_def) also have "\ = residue (\z. a powr (1 - z) / (z - 1)) 1" using assms by (subst residue_add [of UNIV]) (auto intro!: holomorphic_intros holo intro: residue_holo[of UNIV, OF _ _ holo]) also have "\ = complex_of_real a powr (1 - 1)" using assms by (intro residue_simple [of UNIV]) (auto intro!: holomorphic_intros) also from assms have "\ = 1" by simp finally show ?thesis . qed corollary residue_zeta: "residue zeta 1 = 1" unfolding zeta_def by (rule residue_hurwitz_zeta) auto lemma zeta_bigo_at_1: "zeta \ O[at 1 within A](\x. 1 / (x - 1))" proof - have "zeta \ \[at 1 within A](\s. pre_zeta 1 s + 1 / (s - 1))" by (intro bigthetaI_cong) (auto simp: eventually_at_filter zeta_def hurwitz_zeta_def) also have "(\s. pre_zeta 1 s + 1 / (s - 1)) \ O[at 1 within A](\s. 1 / (s - 1))" proof (rule sum_in_bigo) have "continuous_on UNIV (pre_zeta 1)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto hence "isCont (pre_zeta 1) 1" by (auto simp: continuous_on_eq_continuous_at) hence "continuous (at 1 within A) (pre_zeta 1)" by (rule continuous_within_subset) auto hence "pre_zeta 1 \ O[at 1 within A](\_. 1)" by (intro continuous_imp_bigo_1) auto also have ev: "eventually (\s. s \ ball 1 1 \ s \ 1 \ s \ A) (at 1 within A)" by (intro eventually_at_ball') auto have "(\_. 1) \ O[at 1 within A](\s. 1 / (s - 1))" by (intro landau_o.bigI[of 1] eventually_mono[OF ev]) (auto simp: eventually_at_filter norm_divide dist_norm norm_minus_commute field_simps) finally show "pre_zeta 1 \ O[at 1 within A](\s. 1 / (s - 1))" . qed simp_all finally show ?thesis . qed theorem assumes "a > 0" "Re s > 1" shows hurwitz_zeta_conv_suminf: "hurwitz_zeta a s = (\n. (of_nat n + of_real a) powr -s)" and sums_hurwitz_zeta: "(\n. (of_nat n + of_real a) powr -s) sums hurwitz_zeta a s" proof - from assms have [simp]: "s \ 1" by auto from assms have "hurwitz_zeta a s = pre_zeta_aux 0 a s + of_real a powr (1 - s) / (s - 1)" by (simp add: hurwitz_zeta_def pre_zeta_def) also from assms have "pre_zeta_aux 0 a s = (\n. (of_nat n + of_real a) powr -s) + of_real a powr (1 - s) / (1 - s)" by (intro pre_zeta_aux_conv_zeta) also have "\ + a powr (1 - s) / (s - 1) = (\n. (of_nat n + of_real a) powr -s) + a powr (1 - s) * (1 / (1 - s) + 1 / (s - 1))" by (simp add: algebra_simps) also have "1 / (1 - s) + 1 / (s - 1) = 0" by (simp add: divide_simps) finally show "hurwitz_zeta a s = (\n. (of_nat n + of_real a) powr -s)" by simp moreover have "(\n. (of_nat n + of_real a) powr -s) sums (\n. (of_nat n + of_real a) powr -s)" by (intro summable_sums summable_hurwitz_zeta assms) ultimately show "(\n. (of_nat n + of_real a) powr -s) sums hurwitz_zeta a s" by simp qed corollary assumes "Re s > 1" shows zeta_conv_suminf: "zeta s = (\n. of_nat (Suc n) powr -s)" and sums_zeta: "(\n. of_nat (Suc n) powr -s) sums zeta s" using hurwitz_zeta_conv_suminf[of 1 s] sums_hurwitz_zeta[of 1 s] assms by (simp_all add: zeta_def add_ac) corollary assumes "n > 1" shows zeta_nat_conv_suminf: "zeta (of_nat n) = (\k. 1 / of_nat (Suc k) ^ n)" and sums_zeta_nat: "(\k. 1 / of_nat (Suc k) ^ n) sums zeta (of_nat n)" proof - have "(\k. of_nat (Suc k) powr -of_nat n) sums zeta (of_nat n)" using assms by (intro sums_zeta) auto also have "(\k. of_nat (Suc k) powr -of_nat n) = (\k. 1 / of_nat (Suc k) ^ n :: complex)" by (simp add: powr_minus divide_simps del: of_nat_Suc) finally show "(\k. 1 / of_nat (Suc k) ^ n) sums zeta (of_nat n)" . thus "zeta (of_nat n) = (\k. 1 / of_nat (Suc k) ^ n)" by (simp add: sums_iff) qed lemma pre_zeta_aux_cnj [simp]: assumes "a > 0" shows "pre_zeta_aux n a (cnj z) = cnj (pre_zeta_aux n a z)" proof - have "cnj (pre_zeta_aux n a z) = of_real a powr -cnj z / 2 + (\x=1..n. (bernoulli (2 * x) / fact (2 * x)) *\<^sub>R a powr (-cnj z - (2*x-1)) * pochhammer (cnj z) (2*x-1)) + EM_remainder (2*n+1) (\x. -(pochhammer (cnj z) (Suc (2 * n)) * cnj (of_real (x + a) powr (-1 - 2 * of_nat n - z)))) 0" (is "_ = _ + ?A + ?B") unfolding pre_zeta_aux_def complex_cnj_add using assms by (subst EM_remainder_cnj [symmetric]) (auto intro!: continuous_intros simp: cnj_powr add_eq_0_iff mult_ac) also have "?B = EM_remainder (2*n+1) (\x. -(pochhammer (cnj z) (Suc (2 * n)) * of_real (x + a) powr (-1 - 2 * of_nat n - cnj z))) 0" using assms by (intro EM_remainder_cong) (auto simp: cnj_powr) also have "of_real a powr -cnj z / 2 + ?A + \ = pre_zeta_aux n a (cnj z)" by (simp add: pre_zeta_aux_def mult_ac) finally show ?thesis .. qed lemma pre_zeta_cnj [simp]: "a > 0 \ pre_zeta a (cnj z) = cnj (pre_zeta a z)" by (simp add: pre_zeta_def) theorem hurwitz_zeta_cnj [simp]: "a > 0 \ hurwitz_zeta a (cnj z) = cnj (hurwitz_zeta a z)" proof - assume "a > 0" moreover have "cnj z = 1 \ z = 1" by (simp add: complex_eq_iff) ultimately show ?thesis by (auto simp: hurwitz_zeta_def cnj_powr) qed theorem zeta_cnj [simp]: "zeta (cnj z) = cnj (zeta z)" by (simp add: zeta_def) corollary hurwitz_zeta_real: "a > 0 \ hurwitz_zeta a (of_real x) \ \" using hurwitz_zeta_cnj [of a "of_real x"] by (simp add: Reals_cnj_iff del: zeta_cnj) corollary zeta_real: "zeta (of_real x) \ \" unfolding zeta_def by (rule hurwitz_zeta_real) auto subsection \Connection to Dirichlet series\ lemma eval_fds_zeta: "Re s > 1 \ eval_fds fds_zeta s = zeta s" using sums_zeta [of s] by (intro eval_fds_eqI) (auto simp: powr_minus divide_simps) theorem euler_product_zeta: assumes "Re s > 1" shows "(\n. \p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1) \ zeta s" using euler_product_fds_zeta[of s] assms unfolding nat_power_complex_def by (simp add: eval_fds_zeta) corollary euler_product_zeta': assumes "Re s > 1" shows "(\n. \p | prime p \ p \ n. inverse (1 - 1 / of_nat p powr s)) \ zeta s" proof - note euler_product_zeta [OF assms] also have "(\n. \p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1) = (\n. \p | prime p \ p \ n. inverse (1 - 1 / of_nat p powr s))" by (intro ext prod.mono_neutral_cong_right refl) auto finally show ?thesis . qed theorem zeta_Re_gt_1_nonzero: "Re s > 1 \ zeta s \ 0" using eval_fds_zeta_nonzero[of s] by (simp add: eval_fds_zeta) theorem tendsto_zeta_Re_going_to_at_top: "(zeta \ 1) (Re going_to at_top)" proof (rule Lim_transform_eventually) have "eventually (\x::real. x > 1) at_top" by (rule eventually_gt_at_top) hence "eventually (\s. Re s > 1) (Re going_to at_top)" by blast thus "eventually (\z. eval_fds fds_zeta z = zeta z) (Re going_to at_top)" by eventually_elim (simp add: eval_fds_zeta) next have "conv_abscissa (fds_zeta :: complex fds) \ 1" proof (rule conv_abscissa_leI) fix c' assume "ereal c' > 1" thus "\s. s \ 1 = c' \ fds_converges fds_zeta (s::complex)" by (auto intro!: exI[of _ "of_real c'"]) qed hence "(eval_fds fds_zeta \ fds_nth fds_zeta 1) (Re going_to at_top)" by (intro tendsto_eval_fds_Re_going_to_at_top') auto thus "(eval_fds fds_zeta \ 1) (Re going_to at_top)" by simp qed lemma conv_abscissa_zeta [simp]: "conv_abscissa (fds_zeta :: complex fds) = 1" and abs_conv_abscissa_zeta [simp]: "abs_conv_abscissa (fds_zeta :: complex fds) = 1" proof - let ?z = "fds_zeta :: complex fds" have A: "conv_abscissa ?z \ 1" proof (intro conv_abscissa_geI) fix c' assume "ereal c' < 1" hence "\summable (\n. real n powr -c')" by (subst summable_real_powr_iff) auto hence "\summable (\n. of_real (real n powr -c') :: complex)" by (subst summable_of_real_iff) also have "summable (\n. of_real (real n powr -c') :: complex) \ fds_converges fds_zeta (of_real c' :: complex)" unfolding fds_converges_def by (intro summable_cong eventually_mono [OF eventually_gt_at_top[of 0]]) (simp add: fds_nth_zeta powr_Reals_eq powr_minus divide_simps) finally show "\s::complex. s \ 1 = c' \ \fds_converges fds_zeta s" by (intro exI[of _ "of_real c'"]) auto qed have B: "abs_conv_abscissa ?z \ 1" proof (intro abs_conv_abscissa_leI) fix c' assume "1 < ereal c'" thus "\s::complex. s \ 1 = c' \ fds_abs_converges fds_zeta s" by (intro exI[of _ "of_real c'"]) auto qed have "conv_abscissa ?z \ abs_conv_abscissa ?z" by (rule conv_le_abs_conv_abscissa) also note B finally show "conv_abscissa ?z = 1" using A by (intro antisym) note A also have "conv_abscissa ?z \ abs_conv_abscissa ?z" by (rule conv_le_abs_conv_abscissa) finally show "abs_conv_abscissa ?z = 1" using B by (intro antisym) qed theorem deriv_zeta_sums: assumes s: "Re s > 1" shows "(\n. -of_real (ln (real (Suc n))) / of_nat (Suc n) powr s) sums deriv zeta s" proof - from s have "fds_converges (fds_deriv fds_zeta) s" by (intro fds_converges_deriv) simp_all with s have "(\n. -of_real (ln (real (Suc n))) / of_nat (Suc n) powr s) sums deriv (eval_fds fds_zeta) s" unfolding fds_converges_altdef by (simp add: fds_nth_deriv scaleR_conv_of_real eval_fds_deriv eval_fds_zeta) also from s have "eventually (\s. s \ {s. Re s > 1}) (nhds s)" by (intro eventually_nhds_in_open) (auto simp: open_halfspace_Re_gt) hence "eventually (\s. eval_fds fds_zeta s = zeta s) (nhds s)" by eventually_elim (auto simp: eval_fds_zeta) hence "deriv (eval_fds fds_zeta) s = deriv zeta s" by (intro deriv_cong_ev refl) finally show ?thesis . qed theorem inverse_zeta_sums: assumes s: "Re s > 1" shows "(\n. moebius_mu (Suc n) / of_nat (Suc n) powr s) sums inverse (zeta s)" proof - have "fds_converges (fds moebius_mu) s" using assms by (auto intro!: fds_abs_converges_moebius_mu) hence "(\n. moebius_mu (Suc n) / of_nat (Suc n) powr s) sums eval_fds (fds moebius_mu) s" by (simp add: fds_converges_altdef) also have "fds moebius_mu = inverse (fds_zeta :: complex fds)" by (rule fds_moebius_inverse_zeta) also from s have "eval_fds \ s = inverse (zeta s)" by (subst eval_fds_inverse) (auto simp: fds_moebius_inverse_zeta [symmetric] eval_fds_zeta intro!: fds_abs_converges_moebius_mu) finally show ?thesis . qed text \ The following gives an extension of the $\zeta$ functions to the critical strip. \ lemma hurwitz_zeta_critical_strip: fixes s :: complex and a :: real defines "S \ (\n. \i (\n. of_nat n powr (1 - s) / (1 - s))" assumes "Re s > 0" "s \ 1" and "a > 0" shows "(\n. S n - I' n) \ hurwitz_zeta a s" proof - from assms have [simp]: "s \ 1" by auto let ?f = "\x. of_real (x + a) powr -s" let ?fs = "\n x. (-1) ^ n * pochhammer s n * of_real (x + a) powr (-s - of_nat n)" have minus_commute: "-a - b = -b - a" for a b :: complex by (simp add: algebra_simps) define I where "I = (\n. (of_nat n + a) powr (1 - s) / (1 - s))" define R where "R = (\n. EM_remainder' 1 (?fs 1) (real 0) (real n))" define R_lim where "R_lim = EM_remainder 1 (?fs 1) 0" define C where "C = - (a powr -s / 2)" define D where "D = (\n. (1/2) * (of_real (a + real n) powr - s))" define D' where "D' = (\n. of_real (a + real n) powr - s)" define C' where "C' = a powr (1 - s) / (1 - s)" define C'' where "C'' = of_real a powr - s" { fix n :: nat assume n: "n > 0" have "((\x. of_real (x + a) powr -s) has_integral (of_real (real n + a) powr (1-s) / (1 - s) - of_real (0 + a) powr (1 - s) / (1 - s))) {0..real n}" using n assms by (intro fundamental_theorem_of_calculus) (auto intro!: continuous_intros has_vector_derivative_real_field derivative_eq_intros simp: complex_nonpos_Reals_iff) hence I: "((\x. of_real (x + a) powr -s) has_integral (I n - C')) {0..n}" by (auto simp: divide_simps C'_def I_def) have "(\i\{0<..n}. ?f (real i)) - integral {real 0..real n} ?f = (\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (?fs k (real n) - ?fs k (real 0))) + R n" using n assms unfolding R_def by (intro euler_maclaurin_strong_raw_nat[where Y = "{0}"]) (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_field simp: pochhammer_rec' algebra_simps complex_nonpos_Reals_iff add_eq_0_iff) also have "(\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (?fs k (real n) - ?fs k (real 0))) = ((n + a) powr - s - a powr - s) / 2" by (simp add: lessThan_nat_numeral scaleR_conv_of_real numeral_2_eq_2 [symmetric]) also have "\ = C + D n" by (simp add: C_def D_def field_simps) also have "integral {real 0..real n} (\x. complex_of_real (x + a) powr - s) = I n - C'" using I by (simp add: has_integral_iff) also have "(\i\{0<..n}. of_real (real i + a) powr - s) = (\i=0..n. of_real (real i + a) powr - s) - of_real a powr -s" using assms by (subst sum.head) auto also have "(\i=0..n. of_real (real i + a) powr - s) = S n + of_real (real n + a) powr -s" unfolding S_def by (subst sum.last_plus) (auto simp: atLeast0LessThan) finally have "C - C' + C'' - D' n + D n + R n + (I n - I' n) = S n - I' n" by (simp add: algebra_simps S_def D'_def C''_def) } hence ev: "eventually (\n. C - C' + C'' - D' n + D n + R n + (I n - I' n) = S n - I' n) at_top" by (intro eventually_mono[OF eventually_gt_at_top[of 0]]) auto have [simp]: "-1 - s = -s - 1" by simp { let ?C = "norm (pochhammer s 1)" have "R \ R_lim" unfolding R_def R_lim_def of_nat_0 proof (subst of_int_0 [symmetric], rule tendsto_EM_remainder) show "eventually (\x. norm (?fs 1 x) \ ?C * (x + a) powr (-Re s - 1)) at_top" using eventually_ge_at_top[of 0] by eventually_elim (insert assms, auto simp: norm_mult norm_powr_real_powr) next fix x assume x: "x \ real_of_int 0" have [simp]: "-numeral n - (x :: real) = -x - numeral n" for x n by (simp add: algebra_simps) show "((\x. ?C / (-Re s) * (x + a) powr (-Re s)) has_real_derivative ?C * (x + a) powr (- Re s - 1)) (at x within {real_of_int 0..})" using assms x by (auto intro!: derivative_eq_intros) next have "(\y. ?C / (- Re s) * (a + real y) powr (- Re s)) \ 0" by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially filterlim_tendsto_add_at_top[OF tendsto_const]) (use assms in auto) thus "convergent (\y. ?C / (- Re s) * (real y + a) powr (- Re s))" by (auto simp: add_ac convergent_def) qed (intro integrable_EM_remainder' continuous_intros, insert assms, auto simp: add_eq_0_iff) } moreover have "(\n. I n - I' n) \ 0" proof - have "(\n. (complex_of_real (real n + a) powr (1 - s) - of_real (real n) powr (1 - s)) / (1 - s)) \ 0 / (1 - s)" using assms(3-5) by (intro filterlim_compose[OF _ filterlim_real_sequentially] tendsto_divide complex_powr_add_minus_powr_asymptotics) auto thus "(\n. I n - I' n) \ 0" by (simp add: I_def I'_def divide_simps) qed ultimately have "(\n. C - C' + C'' - D' n + D n + R n + (I n - I' n)) \ C - C' + C'' - 0 + 0 + R_lim + 0" unfolding D_def D'_def using assms by (intro tendsto_add tendsto_diff tendsto_const tendsto_mult_right_zero tendsto_neg_powr_complex_of_real filterlim_tendsto_add_at_top filterlim_real_sequentially) auto also have "C - C' + C'' - 0 + 0 + R_lim + 0 = (a powr - s / 2) + a powr (1 - s) / (s - 1) + R_lim" by (simp add: C_def C'_def C''_def field_simps) also have "\ = hurwitz_zeta a s" using assms by (simp add: hurwitz_zeta_def pre_zeta_def pre_zeta_aux_def R_lim_def scaleR_conv_of_real) finally have "(\n. C - C' + C'' - D' n + D n + R n + (I n - I' n)) \ hurwitz_zeta a s" . with ev show ?thesis by (blast intro: Lim_transform_eventually) qed lemma zeta_critical_strip: fixes s :: complex and a :: real defines "S \ (\n. \i=1..n. (of_nat i) powr - s)" defines "I \ (\n. of_nat n powr (1 - s) / (1 - s))" assumes s: "Re s > 0" "s \ 1" shows "(\n. S n - I n) \ zeta s" proof - from hurwitz_zeta_critical_strip[OF s zero_less_one] have "(\n. (\i hurwitz_zeta 1 s" by (simp add: add_ac) also have "(\n. (\in. (\i=1..n. of_nat i powr -s))" by (intro ext sum.reindex_bij_witness[of _ "\x. x - 1" Suc]) auto finally show ?thesis by (simp add: zeta_def S_def I_def) qed subsection \The non-vanishing of $\zeta$ for $\mathfrak{R}(s) \geq 1$\ text \ This proof is based on a sketch by Newman~\cite{newman1998analytic}, which was previously formalised in HOL Light by Harrison~\cite{harrison2009pnt}, albeit in a much more concrete and low-level style. Our aim here is to reproduce Newman's proof idea cleanly and on the same high level of abstraction. \ theorem zeta_Re_ge_1_nonzero: fixes s assumes "Re s \ 1" "s \ 1" shows "zeta s \ 0" proof (cases "Re s > 1") case False define a where "a = -Im s" from False assms have s [simp]: "s = 1 - \ * a" and a: "a \ 0" by (auto simp: complex_eq_iff a_def) show ?thesis proof assume "zeta s = 0" hence zero: "zeta (1 - \ * a) = 0" by simp with zeta_cnj[of "1 - \ * a"] have zero': "zeta (1 + \ * a) = 0" by simp \ \We define the function $Q(s) = \zeta(s)^2\zeta(s+ia)\zeta(s-ia)$ and its Dirichlet series. The objective will be to show that this function is entire and its Dirichlet series converges everywhere. Of course, $Q(s)$ has singularities at $1$ and $1\pm ia$, so we need to show they can be removed.\ define Q Q_fds where "Q = (\s. zeta s ^ 2 * zeta (s + \ * a) * zeta (s - \ * a))" and "Q_fds = fds_zeta ^ 2 * fds_shift (\ * a) fds_zeta * fds_shift (-\ * a) fds_zeta" let ?sings = "{1, 1 + \ * a, 1 - \ * a}" \ \We show that @{term Q} is locally bounded everywhere. This is the case because the poles of $\zeta(s)$ cancel with the zeros of $\zeta(s\pm ia)$ and vice versa. This boundedness is then enough to show that @{term Q} has only removable singularities.\ have Q_bigo_1: "Q \ O[at s](\_. 1)" for s proof - have Q_eq: "Q = (\s. (zeta s * zeta (s + \ * a)) * (zeta s * zeta (s - \ * a)))" by (simp add: Q_def power2_eq_square mult_ac) \ \The singularity of $\zeta(s)$ at 1 gets cancelled by the zero of $\zeta(s-ia)$:\ have bigo1: "(\s. zeta s * zeta (s - \ * a)) \ O[at 1](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real proof - have "(\s. zeta (s - \ * a) - zeta (1 - \ * a)) \ O[at 1](\s. s - 1)" using that by (intro taylor_bigo_linear holomorphic_on_imp_differentiable_at[of _ "-{1 + \ * a}"] holomorphic_intros) (auto simp: complex_eq_iff) hence "(\s. zeta s * zeta (s - \ * a)) \ O[at 1](\s. 1 / (s - 1) * (s - 1))" using that by (intro landau_o.big.mult zeta_bigo_at_1) simp_all also have "(\s. 1 / (s - 1) * (s - 1)) \ \[at 1](\_. 1)" by (intro bigthetaI_cong) (auto simp: eventually_at_filter) finally show ?thesis . qed \ \The analogous result for $\zeta(s) \zeta(s+ia)$:\ have bigo1': "(\s. zeta s * zeta (s + \ * a)) \ O[at 1](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real using bigo1[of "-a"] that zeta_cnj[of "1 - \ * a"] by simp \ \The singularity of $\zeta(s-ia)$ gets cancelled by the zero of $\zeta(s)$:\ have bigo2: "(\s. zeta s * zeta (s - \ * a)) \ O[at (1 + \ * a)](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real proof - have "(\s. zeta s * zeta (s - \ * a)) \ O[filtermap (\s. s + \ * a) (at 1)](\_. 1)" using bigo1'[of a] that by (simp add: mult.commute landau_o.big.in_filtermap_iff) also have "filtermap (\s. s + \ * a) (at 1) = at (1 + \ * a)" using filtermap_at_shift[of "-\ * a" 1] by simp finally show ?thesis . qed \ \Again, the analogous result for $\zeta(s) \zeta(s+ia)$:\ have bigo2': "(\s. zeta s * zeta (s + \ * a)) \ O[at (1 - \ * a)](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real using bigo2[of "-a"] that zeta_cnj[of "1 - \ * a"] by simp \ \Now the final case distinction to show $Q(s)\in O(1)$ for all $s\in\mathbb{C}$:\ consider "s = 1" | "s = 1 + \ * a" | "s = 1 - \ * a" | "s \ ?sings" by blast thus ?thesis proof cases case 1 thus ?thesis unfolding Q_eq using zero zero' a by (auto intro: bigo1 bigo1' landau_o.big.mult_in_1) next case 2 from a have "isCont (\s. zeta s * zeta (s + \ * a)) (1 + \ * a)" by (auto intro!: continuous_intros) with 2 show ?thesis unfolding Q_eq using zero zero' a by (auto intro: bigo2 landau_o.big.mult_in_1 continuous_imp_bigo_1) next case 3 from a have "isCont (\s. zeta s * zeta (s - \ * a)) (1 - \ * a)" by (auto intro!: continuous_intros) with 3 show ?thesis unfolding Q_eq using zero zero' a by (auto intro: bigo2' landau_o.big.mult_in_1 continuous_imp_bigo_1) qed (auto intro!: continuous_imp_bigo_1 continuous_intros simp: Q_def complex_eq_iff) qed \ \Thus, we can remove the singularities from @{term Q} and extend it to an entire function.\ have "\Q'. Q' holomorphic_on UNIV \ (\z\UNIV - ?sings. Q' z = Q z)" by (intro removable_singularities Q_bigo_1) (auto simp: Q_def complex_eq_iff intro!: holomorphic_intros) then obtain Q' where Q': "Q' holomorphic_on UNIV" "\z. z \ ?sings \ Q' z = Q z" by blast \ \@{term Q'} constitutes an analytic continuation of the Dirichlet series of @{term Q}.\ have eval_Q_fds: "eval_fds Q_fds s = Q' s" if "Re s > 1" for s proof - have "eval_fds Q_fds s = Q s" using that by (simp add: Q_fds_def Q_def eval_fds_mult eval_fds_power fds_abs_converges_mult fds_abs_converges_power eval_fds_zeta) also from that have "\ = Q' s" by (subst Q') auto finally show ?thesis . qed \ \Since $\zeta(s)$ and $\zeta(s\pm ia)$ are completely multiplicative Dirichlet series, the logarithm of their product can be rewritten into the following nice form:\ have ln_Q_fds_eq: "fds_ln 0 Q_fds = fds (\k. of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k))))" proof - note simps = fds_ln_mult[where l' = 0 and l'' = 0] fds_ln_power[where l' = 0] fds_ln_prod[where l' = "\_. 0"] have "fds_ln 0 Q_fds = 2 * fds_ln 0 fds_zeta + fds_shift (\ * a) (fds_ln 0 fds_zeta) + fds_shift (-\ * a) (fds_ln 0 fds_zeta)" by (auto simp: Q_fds_def simps) also have "completely_multiplicative_function (fds_nth (fds_zeta :: complex fds))" by standard auto hence "fds_ln (0 :: complex) fds_zeta = fds (\n. mangoldt n /\<^sub>R ln (real n))" by (subst fds_ln_completely_multiplicative) (auto simp: fds_eq_iff) also have "2 * \ + fds_shift (\ * a) \ + fds_shift (-\ * a) \ = fds (\k. of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k))))" (is "?a = ?b") proof (intro fds_eqI, goal_cases) case (1 n) then consider "n = 1" | "n > 1" by force hence "fds_nth ?a n = mangoldt n / ln (real n) * (2 + (n powr (\ * a) + n powr (-\ * a)))" by cases (auto simp: field_simps scaleR_conv_of_real numeral_fds) also have "n powr (\ * a) + n powr (-\ * a) = 2 * cos (of_real (a * ln n))" using 1 by (subst cos_exp_eq) (simp_all add: powr_def algebra_simps) also have "mangoldt n / ln (real n) * (2 + \) = of_real (2 * mangoldt n / ln n * (1 + cos (a * ln n)))" by (subst cos_of_real) simp_all finally show ?case by (simp add: fds_nth_fds') qed finally show ?thesis . qed \ \It is then obvious that this logarithm series has non-negative real coefficients.\ also have "nonneg_dirichlet_series \" proof (standard, goal_cases) case (1 n) from cos_ge_minus_one[of "a * ln n"] have "1 + cos (a * ln (real n)) \ 0" by linarith thus ?case using 1 by (cases "n = 0") (auto simp: complex_nonneg_Reals_iff fds_nth_fds' mangoldt_nonneg intro!: divide_nonneg_nonneg mult_nonneg_nonneg) qed \ \Therefore, the original series also has non-negative real coefficients.\ finally have nonneg: "nonneg_dirichlet_series Q_fds" by (rule nonneg_dirichlet_series_lnD) (auto simp: Q_fds_def) \ \By the Pringsheim--Landau theorem, a Dirichlet series with non-negative coefficnets that can be analytically continued to the entire complex plane must converge everywhere, i.\,e.\ its abscissa of (absolute) convergence is $-\infty$:\ have abscissa_Q_fds: "abs_conv_abscissa Q_fds \ 1" unfolding Q_fds_def by (auto intro!: abs_conv_abscissa_mult_leI abs_conv_abscissa_power_leI) with nonneg and eval_Q_fds and \Q' holomorphic_on UNIV\ have abscissa: "abs_conv_abscissa Q_fds = -\" by (intro entire_continuation_imp_abs_conv_abscissa_MInfty[where c = 1 and g = Q']) (auto simp: one_ereal_def) \ \This now leads to a contradiction in a very obvious way. If @{term Q_fds} is absolutely convergent, then the subseries corresponding to powers of 2 (\i.e. we delete all summands $a_n / n ^ s$ where $n$ is not a power of 2 from the sum) is also absolutely convergent. We denote this series with $R$.\ define R_fds where "R_fds = fds_primepow_subseries 2 Q_fds" have "conv_abscissa R_fds \ abs_conv_abscissa R_fds" by (rule conv_le_abs_conv_abscissa) also have "abs_conv_abscissa R_fds \ abs_conv_abscissa Q_fds" unfolding R_fds_def by (rule abs_conv_abscissa_restrict) also have "\ = -\" by (simp add: abscissa) finally have abscissa': "conv_abscissa R_fds = -\" by simp \ \Since $\zeta(s)$ and $\zeta(s \pm ia)$ have an Euler product expansion for $\mathfrak{R}(s) > 1$, we have \[R(s) = (1 - 2^{-s})^-2 (1 - 2^{-s+ia})^{-1} (1 - 2^{-s-ia})^{-1}\] there, and since $R$ converges everywhere and the right-hand side is holomorphic for $\mathfrak{R}(s) > 0$, the equation is also valid for all $s$ with $\mathfrak{R}(s) > 0$ by analytic continuation.\ have eval_R: "eval_fds R_fds s = 1 / ((1 - 2 powr -s) ^ 2 * (1 - 2 powr (-s + \ * a)) * (1 - 2 powr (-s - \ * a)))" (is "_ = ?f s") if "Re s > 0" for s proof - show ?thesis proof (rule analytic_continuation_open[where f = "eval_fds R_fds"]) show "?f holomorphic_on {s. Re s > 0}" by (intro holomorphic_intros) (auto simp: powr_def exp_eq_1 Ln_Reals_eq) next fix z assume z: "z \ {s. Re s > 1}" have [simp]: "completely_multiplicative_function (fds_nth fds_zeta)" by standard auto thus "eval_fds R_fds z = ?f z" using z by (simp add: R_fds_def Q_fds_def eval_fds_mult eval_fds_power fds_abs_converges_mult fds_abs_converges_power fds_primepow_subseries_euler_product_cm divide_simps powr_minus powr_diff powr_add fds_abs_summable_zeta) qed (insert that abscissa', auto intro!: exI[of _ 2] convex_connected open_halfspace_Re_gt convex_halfspace_Re_gt holomorphic_intros) qed \ \We now clearly have a contradiction: $R(s)$, being entire, is continuous everywhere, while the function on the right-hand side clearly has a pole at $0$.\ show False proof (rule not_tendsto_and_filterlim_at_infinity) have "((\b. (1-2 powr - b)\<^sup>2 * (1 - 2 powr (-b+\*a)) * (1 - 2 powr (-b-\*a))) \ 0) (at 0 within {s. Re s > 0})" (is "filterlim ?f' _ _") by (intro tendsto_eq_intros) (auto) moreover have "eventually (\s. s \ {s. Re s > 0}) (at 0 within {s. Re s > 0})" by (auto simp: eventually_at_filter) hence "eventually (\s. ?f' s \ 0) (at 0 within {s. Re s > 0})" by eventually_elim (auto simp: powr_def exp_eq_1 Ln_Reals_eq) ultimately have "filterlim ?f' (at 0) (at 0 within {s. Re s > 0})" by (simp add: filterlim_at) hence "filterlim ?f at_infinity (at 0 within {s. Re s > 0})" (is ?lim) by (intro filterlim_divide_at_infinity[OF tendsto_const] tendsto_mult_filterlim_at_infinity) auto also have ev: "eventually (\s. Re s > 0) (at 0 within {s. Re s > 0})" by (auto simp: eventually_at intro!: exI[of _ 1]) have "?lim \ filterlim (eval_fds R_fds) at_infinity (at 0 within {s. Re s > 0})" by (intro filterlim_cong refl eventually_mono[OF ev]) (auto simp: eval_R) finally show \ . next have "continuous (at 0 within {s. Re s > 0}) (eval_fds R_fds)" by (intro continuous_intros) (auto simp: abscissa') thus "((eval_fds R_fds \ eval_fds R_fds 0)) (at 0 within {s. Re s > 0})" by (auto simp: continuous_within) next have "0 \ {s. Re s \ 0}" by simp also have "{s. Re s \ 0} = closure {s. Re s > 0}" using closure_halfspace_gt[of "1::complex" 0] by (simp add: inner_commute) finally have "0 \ \" . thus "at 0 within {s. Re s > 0} \ bot" by (subst at_within_eq_bot_iff) auto qed qed qed (fact zeta_Re_gt_1_nonzero) subsection \Special values of the $\zeta$ functions\ theorem hurwitz_zeta_neg_of_nat: assumes "a > 0" shows "hurwitz_zeta a (-of_nat n) = -bernpoly (Suc n) a / of_nat (Suc n)" proof - have "-of_nat n \ (1::complex)" by (simp add: complex_eq_iff) hence "hurwitz_zeta a (-of_nat n) = pre_zeta a (-of_nat n) + a powr real (Suc n) / (-of_nat (Suc n))" unfolding zeta_def hurwitz_zeta_def using assms by (simp add: powr_of_real [symmetric]) also have "a powr real (Suc n) / (-of_nat (Suc n)) = - (a powr real (Suc n) / of_nat (Suc n))" by (simp add: divide_simps del: of_nat_Suc) also have "a powr real (Suc n) = a ^ Suc n" using assms by (intro powr_realpow) also have "pre_zeta a (-of_nat n) = pre_zeta_aux (Suc n) a (- of_nat n)" using assms by (intro pre_zeta_aux_eq_pre_zeta [symmetric]) auto also have "\ = of_real a powr of_nat n / 2 + (\i = 1..Suc n. (bernoulli (2 * i) / fact (2 * i)) *\<^sub>R (pochhammer (- of_nat n) (2 * i - 1) * of_real a powr (of_nat n - of_nat (2 * i - 1)))) + EM_remainder (Suc (2 * Suc n)) (\x. - (pochhammer (- of_nat n) (2*n+3) * of_real (x + a) powr (- of_nat n - 3))) 0" (is "_ = ?B + sum (\n. ?f (2 * n)) _ + _") unfolding pre_zeta_aux_def by (simp add: add_ac eval_nat_numeral) also have "?B = of_real (a ^ n) / 2" using assms by (subst powr_Reals_eq) (auto simp: powr_realpow) also have "pochhammer (-of_nat n :: complex) (2*n+3) = 0" by (subst pochhammer_eq_0_iff) auto finally have "hurwitz_zeta a (-of_nat n) = - (a ^ Suc n / of_nat (Suc n)) + (a ^ n / 2 + sum (\n. ?f (2 * n)) {1..Suc n})" by simp also have "sum (\n. ?f (2 * n)) {1..Suc n} = sum ?f ((*) 2 ` {1..Suc n})" by (intro sum.reindex_bij_witness[of _ "\i. i div 2" "\i. 2*i"]) auto also have "\ = (\i=2..2*n+2. ?f i)" proof (intro sum.mono_neutral_left ballI, goal_cases) case (3 i) hence "odd i" "i \ 1" by (auto elim!: evenE) thus ?case by (simp add: bernoulli_odd_eq_0) qed auto also have "\ = (\i=2..Suc n. ?f i)" proof (intro sum.mono_neutral_right ballI, goal_cases) case (3 i) hence "pochhammer (-of_nat n :: complex) (i - 1) = 0" by (subst pochhammer_eq_0_iff) auto thus ?case by simp qed auto also have "\ = (\i=Suc 1..Suc n. -of_real (real (Suc n choose i) * bernoulli i * a ^ (Suc n - i)) / of_nat (Suc n))" (is "sum ?lhs _ = sum ?f _") proof (intro sum.cong, goal_cases) case (2 i) hence "of_nat n - of_nat (i - 1) = (of_nat (Suc n - i) :: complex)" by (auto simp: of_nat_diff) also have "of_real a powr \ = of_real (a ^ (Suc n - i))" using 2 assms by (subst powr_nat) auto finally have A: "of_real a powr (of_nat n - of_nat (i - 1)) = \" . have "pochhammer (-of_nat n) (i - 1) = complex_of_real (pochhammer (-real n) (i - 1))" by (simp add: pochhammer_of_real [symmetric]) also have "pochhammer (-real n) (i - 1) = pochhammer (-of_nat (Suc n)) i / (-1 - real n)" using 2 by (subst (2) pochhammer_rec_if) auto also have "-1 - real n = -real (Suc n)" by simp finally have B: "pochhammer (-of_nat n) (i - 1) = -complex_of_real (pochhammer (-real (Suc n)) i / real (Suc n))" by (simp del: of_nat_Suc) have "?lhs i = -complex_of_real (bernoulli i * pochhammer (-real (Suc n)) i / fact i * a ^ (Suc n - i)) / of_nat (Suc n)" by (simp only: A B) (simp add: scaleR_conv_of_real) also have "bernoulli i * pochhammer (-real (Suc n)) i / fact i = (real (Suc n) gchoose i) * bernoulli i" using 2 by (subst gbinomial_pochhammer) (auto simp: minus_one_power_iff bernoulli_odd_eq_0) also have "real (Suc n) gchoose i = Suc n choose i" by (subst binomial_gbinomial) auto finally show ?case by simp qed auto also have "a ^ n / 2 + sum ?f {Suc 1..Suc n} = sum ?f {1..Suc n}" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: scaleR_conv_of_real del: of_nat_Suc) also have "-(a ^ Suc n / of_nat (Suc n)) + sum ?f {1..Suc n} = sum ?f {0..Suc n}" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: scaleR_conv_of_real) also have "\ = - bernpoly (Suc n) a / of_nat (Suc n)" unfolding sum_negf sum_divide_distrib [symmetric] by (simp add: bernpoly_def atLeast0AtMost) finally show ?thesis . qed lemma hurwitz_zeta_0 [simp]: "a > 0 \ hurwitz_zeta a 0 = 1 / 2 - a" using hurwitz_zeta_neg_of_nat[of a 0] by (simp add: bernpoly_def) lemma zeta_0 [simp]: "zeta 0 = -1 / 2" by (simp add: zeta_def) theorem zeta_neg_of_nat: "zeta (-of_nat n) = -of_real (bernoulli' (Suc n)) / of_nat (Suc n)" unfolding zeta_def by (simp add: hurwitz_zeta_neg_of_nat bernpoly_1') corollary zeta_trivial_zero: assumes "even n" "n \ 0" shows "zeta (-of_nat n) = 0" using zeta_neg_of_nat[of n] assms by (simp add: bernoulli'_odd_eq_0) theorem zeta_even_nat: "zeta (2 * of_nat n) = of_real ((-1) ^ Suc n * bernoulli (2 * n) * (2 * pi) ^ (2 * n) / (2 * fact (2 * n)))" proof (cases "n = 0") case False hence "(\k. 1 / of_nat (Suc k) ^ (2 * n)) sums zeta (of_nat (2 * n))" by (intro sums_zeta_nat) auto from sums_unique2 [OF this nat_even_power_sums_complex] False show ?thesis by simp qed (insert zeta_neg_of_nat[of 0], simp_all) corollary zeta_even_numeral: "zeta (numeral (Num.Bit0 n)) = of_real ((- 1) ^ Suc (numeral n) * bernoulli (numeral (num.Bit0 n)) * (2 * pi) ^ numeral (num.Bit0 n) / (2 * fact (numeral (num.Bit0 n))))" (is "_ = ?rhs") proof - have *: "(2 * numeral n :: nat) = numeral (Num.Bit0 n)" by (subst numeral.numeral_Bit0, subst mult_2, rule refl) have "numeral (Num.Bit0 n) = (2 * numeral n :: complex)" by (subst numeral.numeral_Bit0, subst mult_2, rule refl) also have "\ = 2 * of_nat (numeral n)" by (simp only: of_nat_numeral of_nat_mult) also have "zeta \ = ?rhs" by (subst zeta_even_nat) (simp only: *) finally show ?thesis . qed corollary zeta_neg_even_numeral [simp]: "zeta (-numeral (Num.Bit0 n)) = 0" proof - have "-numeral (Num.Bit0 n) = (-of_nat (numeral (Num.Bit0 n)) :: complex)" by simp also have "zeta \ = 0" proof (rule zeta_trivial_zero) have "numeral (Num.Bit0 n) = (2 * numeral n :: nat)" by (subst numeral.numeral_Bit0, subst mult_2, rule refl) also have "even \" by (rule dvd_triv_left) finally show "even (numeral (Num.Bit0 n) :: nat)" . qed auto finally show ?thesis . qed corollary zeta_neg_numeral: "zeta (-numeral n) = -of_real (bernoulli' (numeral (Num.inc n)) / numeral (Num.inc n))" proof - have "-numeral n = (- of_nat (numeral n) :: complex)" by simp also have "zeta \ = -of_real (bernoulli' (numeral (Num.inc n)) / numeral (Num.inc n))" by (subst zeta_neg_of_nat) (simp add: numeral_inc) finally show ?thesis . qed corollary zeta_neg1: "zeta (-1) = - 1 / 12" using zeta_neg_of_nat[of 1] by (simp add: eval_bernoulli) corollary zeta_neg3: "zeta (-3) = 1 / 120" by (simp add: zeta_neg_numeral) corollary zeta_neg5: "zeta (-5) = - 1 / 252" by (simp add: zeta_neg_numeral) corollary zeta_2: "zeta 2 = pi ^ 2 / 6" by (simp add: zeta_even_numeral) corollary zeta_4: "zeta 4 = pi ^ 4 / 90" by (simp add: zeta_even_numeral fact_num_eq_if) corollary zeta_6: "zeta 6 = pi ^ 6 / 945" by (simp add: zeta_even_numeral fact_num_eq_if) corollary zeta_8: "zeta 8 = pi ^ 8 / 9450" by (simp add: zeta_even_numeral fact_num_eq_if) subsection \Integral relation between $\Gamma$ and $\zeta$ function\ lemma assumes z: "Re z > 0" and a: "a > 0" shows Gamma_hurwitz_zeta_aux_integral: "Gamma z / (of_nat n + of_real a) powr z = (\s\{0<..}. (s powr (z - 1) / exp ((n+a) * s)) \lebesgue)" and Gamma_hurwitz_zeta_aux_integrable: "set_integrable lebesgue {0<..} (\s. s powr (z - 1) / exp ((n+a) * s))" proof - note integrable = absolutely_integrable_Gamma_integral' [OF z] let ?INT = "set_lebesgue_integral lebesgue {0<..} :: (real \ complex) \ complex" let ?INT' = "set_lebesgue_integral lebesgue {0<..} :: (real \ real) \ real" have meas1: "set_borel_measurable lebesgue {0<..} (\x. of_real x powr (z - 1) / of_real (exp ((n+a) * x)))" unfolding set_borel_measurable_def by (intro measurable_completion, subst measurable_lborel2, intro borel_measurable_continuous_on_indicator) (auto intro!: continuous_intros) show integrable1: "set_integrable lebesgue {0<..} (\s. s powr (z - 1) / exp ((n+a) * s))" using assms by (intro absolutely_integrable_Gamma_integral) auto from assms have pos: "0 < real n + a" by (simp add: add_nonneg_pos) hence "complex_of_real 0 \ of_real (real n + a)" by (simp only: of_real_eq_iff) also have "complex_of_real (real n + a) = of_nat n + of_real a" by simp finally have nz: "\ \ 0" by auto have "((\t. complex_of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}" by (rule Gamma_integral_complex') fact+ hence "Gamma z = ?INT (\t. t powr (z - 1) / of_real (exp t))" using set_lebesgue_integral_eq_integral(2) [OF integrable] by (simp add: has_integral_iff exp_of_real) also have "lebesgue = density (distr lebesgue lebesgue (\t. (real n+a) * t)) (\x. ennreal (real n+a))" using lebesgue_real_scale[of "real n + a"] pos by auto also have "set_lebesgue_integral \ {0<..} (\t. of_real t powr (z - 1) / of_real (exp t)) = set_lebesgue_integral (distr lebesgue lebesgue (\t. (real n + a) * t)) {0<..} (\t. (real n + a) * t powr (z - 1) / exp t)" using integrable pos unfolding set_lebesgue_integral_def by (subst integral_density) (simp_all add: exp_of_real algebra_simps scaleR_conv_of_real set_integrable_def) also have "\ = ?INT (\s. (n + a) * (of_real (n+a) * of_real s) powr (z - 1) / of_real (exp ((n+a) * s)))" unfolding set_lebesgue_integral_def proof (subst integral_distr) show "(*) (real n + a) \ lebesgue \\<^sub>M lebesgue" using lebesgue_measurable_scaling[of "real n + a", where ?'a = real] unfolding real_scaleR_def . next have "(\x. (n+a) * (indicator {0<..} x *\<^sub>R (of_real x powr (z - 1) / of_real (exp x)))) \ lebesgue \\<^sub>M borel" using integrable unfolding set_integrable_def by (intro borel_measurable_times) simp_all thus "(\x. indicator {0<..} x *\<^sub>R (complex_of_real (real n + a) * complex_of_real x powr (z - 1) / exp x)) \ borel_measurable lebesgue" by simp qed (intro Bochner_Integration.integral_cong refl, insert pos, auto simp: indicator_def zero_less_mult_iff) also have "\ = ?INT (\s. ((n+a) powr z) * (s powr (z - 1) / exp ((n+a) * s)))" using pos by (intro set_lebesgue_integral_cong refl allI impI, simp, subst powr_times_real) (auto simp: powr_diff) also have "\ = (n + a) powr z * ?INT (\s. s powr (z - 1) / exp ((n+a) * s))" unfolding set_lebesgue_integral_def by (subst integral_mult_right_zero [symmetric]) simp_all finally show "Gamma z / (of_nat n + of_real a) powr z = ?INT (\s. s powr (z - 1) / exp ((n+a) * s))" using nz by (auto simp add: field_simps) qed lemma assumes x: "x > 0" and "a > 0" shows Gamma_hurwitz_zeta_aux_integral_real: "Gamma x / (real n + a) powr x = set_lebesgue_integral lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" and Gamma_hurwitz_zeta_aux_integrable_real: "set_integrable lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" proof - show "set_integrable lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" using absolutely_integrable_Gamma_integral[of "of_real x" "real n + a"] unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound, goal_cases) case 3 have "set_integrable lebesgue {0<..} (\xa. complex_of_real xa powr (of_real x - 1) / of_real (exp ((n + a) * xa)))" using assms by (intro Gamma_hurwitz_zeta_aux_integrable) auto also have "?this \ integrable lebesgue (\s. complex_of_real (indicator {0<..} s *\<^sub>R (s powr (x - 1) / (exp ((n+a) * s)))))" unfolding set_integrable_def by (intro Bochner_Integration.integrable_cong refl) (auto simp: powr_Reals_eq indicator_def) finally have "set_integrable lebesgue {0<..} (\s. s powr (x - 1) / exp ((n+a) * s))" unfolding set_integrable_def complex_of_real_integrable_eq . thus ?case by (simp add: set_integrable_def) qed (insert assms, auto intro!: AE_I2 simp: indicator_def norm_divide norm_powr_real_powr) from Gamma_hurwitz_zeta_aux_integral[of "of_real x" a n] and assms have "of_real (Gamma x / (real n + a) powr x) = set_lebesgue_integral lebesgue {0<..} (\s. complex_of_real s powr (of_real x - 1) / of_real (exp ((n+a) * s)))" (is "_ = ?I") by (auto simp: Gamma_complex_of_real powr_Reals_eq) also have "?I = lebesgue_integral lebesgue (\s. of_real (indicator {0<..} s *\<^sub>R (s powr (x - 1) / exp ((n+a) * s))))" unfolding set_lebesgue_integral_def using assms by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def powr_Reals_eq) also have "\ = of_real (set_lebesgue_integral lebesgue {0<..} (\s. s powr (x - 1) / exp ((n+a) * s)))" unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_complex_of_real) finally show "Gamma x / (real n + a) powr x = set_lebesgue_integral lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" by (subst (asm) of_real_eq_iff) qed theorem assumes "Re z > 1" and "a > (0::real)" shows Gamma_times_hurwitz_zeta_integral: "Gamma z * hurwitz_zeta a z = (\x\{0<..}. (of_real x powr (z - 1) * of_real (exp (-a*x) / (1 - exp (-x)))) \lebesgue)" and Gamma_times_hurwitz_zeta_integrable: "set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) * of_real (exp (-a*x) / (1 - exp (-x))))" proof - from assms have z: "Re z > 0" by simp let ?INT = "set_lebesgue_integral lebesgue {0<..} :: (real \ complex) \ complex" let ?INT' = "set_lebesgue_integral lebesgue {0<..} :: (real \ real) \ real" have 1: "complex_set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) / of_real (exp ((real n + a) * x)))" for n by (rule Gamma_hurwitz_zeta_aux_integrable) (use assms in simp_all) have 2: "summable (\n. norm (indicator {0<..} s *\<^sub>R (of_real s powr (z - 1) / of_real (exp ((n + a) * s)))))" for s proof (cases "s > 0") case True hence "summable (\n. norm (of_real s powr (z - 1)) * exp (-a * s) * exp (-s) ^ n)" using assms by (intro summable_mult summable_geometric) simp_all with True show ?thesis by (simp add: norm_mult norm_divide exp_add exp_diff exp_minus field_simps exp_of_nat_mult [symmetric]) qed simp_all have 3: "summable (\n. \x. norm (indicator {0<..} x *\<^sub>R (complex_of_real x powr (z - 1) / complex_of_real (exp ((n + a) * x)))) \lebesgue)" proof - have "summable (\n. Gamma (Re z) * (real n + a) powr -Re z)" using assms by (intro summable_mult summable_hurwitz_zeta_real) simp_all also have "?this \ summable (\n. ?INT' (\s. norm (of_real s powr (z - 1) / of_real (exp ((n+a) * s)))))" proof (intro summable_cong always_eventually allI, goal_cases) case (1 n) have "Gamma (Re z) * (real n + a) powr -Re z = Gamma (Re z) / (real n + a) powr Re z" by (subst powr_minus) (simp_all add: field_simps) also from assms have "\ = (\x\{0<..}. (x powr (Re z-1) / exp ((n+a) * x)) \lebesgue)" by (subst Gamma_hurwitz_zeta_aux_integral_real) simp_all also have "\ = (\xa\{0<..}. norm (of_real xa powr (z-1) / of_real (exp ((n+a) * xa))) \lebesgue)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def norm_divide norm_powr_real_powr) finally show ?case . qed finally show ?thesis by (simp add: set_lebesgue_integral_def) qed have sum_eq: "(\n. indicator {0<..} s *\<^sub>R (of_real s powr (z - 1) / of_real (exp ((n+a) * s)))) = indicator {0<..} s *\<^sub>R (of_real s powr (z - 1) * of_real (exp (-a * s) / (1 - exp (-s))))" for s proof (cases "s > 0") case True hence "(\n. indicator {0..} s *\<^sub>R (of_real s powr (z - 1) / of_real (exp ((n+a) * s)))) = (\n. of_real s powr (z - 1) * of_real (exp (-a * s)) * of_real (exp (-s)) ^ n)" by (intro suminf_cong) (auto simp: exp_add exp_minus exp_of_nat_mult [symmetric] field_simps of_real_exp) also have "(\n. of_real s powr (z - 1) * of_real (exp (-a * s)) * of_real (exp (-s)) ^ n) = of_real s powr (z - 1) * of_real (exp (-a * s)) * (\n. of_real (exp (-s)) ^ n)" using True by (intro suminf_mult summable_geometric) simp_all also have "(\n. complex_of_real (exp (-s)) ^ n) = 1 / (1 - of_real (exp (-s)))" using True by (intro suminf_geometric) auto also have "of_real s powr (z - 1) * of_real (exp (-a * s)) * \ = of_real s powr (z - 1) * of_real (exp (-a * s) / (1 - exp (-s)))" using \a > 0\ by (auto simp add: divide_simps exp_minus) finally show ?thesis using True by simp qed simp_all show "set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) * of_real (exp (-a*x) / (1 - exp (-x))))" using 1 unfolding sum_eq [symmetric] set_integrable_def by (intro integrable_suminf[OF _ AE_I2] 2 3) have "(\n. ?INT (\s. s powr (z - 1) / exp ((n+a) * s))) sums lebesgue_integral lebesgue (\s. \n. indicator {0<..} s *\<^sub>R (s powr (z - 1) / exp ((n+a) * s)))" (is "?A sums ?B") using 1 unfolding set_lebesgue_integral_def set_integrable_def by (rule sums_integral[OF _ AE_I2[OF 2] 3]) also have "?A = (\n. Gamma z * (n + a) powr -z)" using assms by (subst Gamma_hurwitz_zeta_aux_integral [symmetric]) (simp_all add: powr_minus divide_simps) also have "?B = ?INT (\s. of_real s powr (z - 1) * of_real (exp (-a * s) / (1 - exp (-s))))" unfolding sum_eq set_lebesgue_integral_def .. finally have "(\n. Gamma z * (of_nat n + of_real a) powr -z) sums ?INT (\x. of_real x powr (z - 1) * of_real (exp (-a * x) / (1 - exp (-x))))" by simp moreover have "(\n. Gamma z * (of_nat n + of_real a) powr -z) sums (Gamma z * hurwitz_zeta a z)" using assms by (intro sums_mult sums_hurwitz_zeta) simp_all ultimately show "Gamma z * hurwitz_zeta a z = (\x\{0<..}. (of_real x powr (z - 1) * of_real (exp (-a * x) / (1 - exp (-x)))) \lebesgue)" by (rule sums_unique2 [symmetric]) qed corollary assumes "Re z > 1" shows Gamma_times_zeta_integral: "Gamma z * zeta z = (\x\{0<..}. (of_real x powr (z - 1) / of_real (exp x - 1)) \lebesgue)" (is ?th1) and Gamma_times_zeta_integrable: "set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) / of_real (exp x - 1))" (is ?th2) proof - have *: "(\x. indicator {0<..} x *\<^sub>R (of_real x powr (z - 1) * of_real (exp (-x) / (1 - exp (-x))))) = (\x. indicator {0<..} x *\<^sub>R (of_real x powr (z - 1) / of_real (exp x - 1)))" by (intro ext) (simp add: field_simps exp_minus indicator_def) from Gamma_times_hurwitz_zeta_integral [OF assms zero_less_one] and * show ?th1 by (simp add: zeta_def set_lebesgue_integral_def) from Gamma_times_hurwitz_zeta_integrable [OF assms zero_less_one] and * show ?th2 by (simp add: zeta_def set_integrable_def) qed corollary hurwitz_zeta_integral_Gamma_def: assumes "Re z > 1" "a > 0" shows "hurwitz_zeta a z = rGamma z * (\x\{0<..}. (of_real x powr (z - 1) * of_real (exp (-a * x) / (1 - exp (-x)))) \lebesgue)" proof - from assms have "Gamma z \ 0" by (subst Gamma_eq_zero_iff) (auto elim!: nonpos_Ints_cases) with Gamma_times_hurwitz_zeta_integral [OF assms] show ?thesis by (simp add: rGamma_inverse_Gamma field_simps) qed corollary zeta_integral_Gamma_def: assumes "Re z > 1" shows "zeta z = rGamma z * (\x\{0<..}. (of_real x powr (z - 1) / of_real (exp x - 1)) \lebesgue)" proof - from assms have "Gamma z \ 0" by (subst Gamma_eq_zero_iff) (auto elim!: nonpos_Ints_cases) with Gamma_times_zeta_integral [OF assms] show ?thesis by (simp add: rGamma_inverse_Gamma field_simps) qed subsection \An analytic proof of the infinitude of primes\ text \ We can now also do an analytic proof of the infinitude of primes. \ lemma primes_infinite_analytic: "infinite {p :: nat. prime p}" proof \ \Suppose the set of primes were finite.\ define P :: "nat set" where "P = {p. prime p}" assume fin: "finite P" \ \Then the Euler product form of the $\zeta$ function ranges over a finite set, and since each factor is holomorphic in the positive real half-space, the product is, too.\ define zeta' :: "complex \ complex" where "zeta' = (\s. (\p\P. inverse (1 - 1 / of_nat p powr s)))" have holo: "zeta' holomorphic_on A" if "A \ {s. Re s > 0}" for A proof - { fix p :: nat and s :: complex assume p: "p \ P" and s: "s \ A" from p have p': "real p > 1" by (subst of_nat_1 [symmetric], subst of_nat_less_iff) (simp add: prime_gt_Suc_0_nat P_def) have "norm (of_nat p powr s) = real p powr Re s" by (simp add: norm_powr_real_powr) also have "\ > real p powr 0" using p p' s that by (subst powr_less_cancel_iff) (auto simp: prime_gt_1_nat) finally have "of_nat p powr s \ 1" using p by (auto simp: P_def) } thus ?thesis by (auto simp: zeta'_def P_def intro!: holomorphic_intros) qed \ \Since the Euler product expansion of $\zeta(s)$ is valid for all $s$ with real value at least 1, and both $\zeta(s)$ and the Euler product must be equal in the positive real half-space punctured at 1 by analytic continuation.\ have eq: "zeta s = zeta' s" if "Re s > 0" "s \ 1" for s proof (rule analytic_continuation_open[of "{s. Re s > 1}" "{s. Re s > 0} - {1}" zeta zeta']) fix s assume s: "s \ {s. Re s > 1}" let ?f = "(\n. \p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1)" have "eventually (\n. ?f n = zeta' s) sequentially" using eventually_ge_at_top[of "Max P"] proof eventually_elim case (elim n) have "P \ {}" by (auto simp: P_def intro!: exI[of _ 2]) with elim have "P \ {..n}" using fin by auto thus ?case unfolding zeta'_def by (intro prod.mono_neutral_cong_right) (auto simp: P_def) qed moreover from s have "?f \ zeta s" by (intro euler_product_zeta) auto ultimately have "(\_. zeta' s) \ zeta s" by (blast intro: Lim_transform_eventually) thus "zeta s = zeta' s" by (simp add: LIMSEQ_const_iff) qed (auto intro!: exI[of _ 2] open_halfspace_Re_gt connected_open_delete convex_halfspace_Re_gt holomorphic_intros holo that intro: convex_connected) \ \However, since the Euler product is holomorphic on the entire positive real half-space, it cannot have a pole at 1, while $\zeta(s)$ does have a pole at 1. Since they are equal in the punctured neighbourhood of 1, this is a contradiction.\ have ev: "eventually (\s. s \ {s. Re s > 0} - {1}) (at 1)" by (auto simp: eventually_at_filter intro!: open_halfspace_Re_gt eventually_mono[OF eventually_nhds_in_open[of "{s. Re s > 0}"]]) have "\is_pole zeta' 1" by (rule not_is_pole_holomorphic [of "{s. Re s > 0}"]) (auto intro!: holo open_halfspace_Re_gt) also have "is_pole zeta' 1 \ is_pole zeta 1" unfolding is_pole_def by (intro filterlim_cong refl eventually_mono [OF ev] eq [symmetric]) auto finally show False using is_pole_zeta by contradiction qed subsection \The periodic zeta function\ text \ The periodic zeta function $F(s, q)$ (as described e.\,g.\ by Apostol~\cite{apostol1976analytic} is related to the Hurwitz zeta function. It is periodic in \q\ with period 1 and it can be represented by a Dirichlet series that is absolutely convergent for $\mathfrak{R}(s) > 1$. If \q \ \\, it furthermore convergent for $\mathfrak{R}(s) > 0$. It is clear that for integer \q\, we have $F(s, q) = F(s, 0) = \zeta(s)$. Moreover, for non-integer \q\, $F(s, q)$ can be analytically continued to an entire function. \ definition fds_perzeta :: "real \ complex fds" where "fds_perzeta q = fds (\m. exp (2 * pi * \ * m * q))" text \ The definition of the periodic zeta function on the full domain is a bit unwieldy. The precise reasoning for this definition will be given later, and, in any case, it is probably more instructive to look at the derived ``alternative'' definitions later. \ definition perzeta :: "real \ complex \ complex" where "perzeta q' s = (if q' \ \ then zeta s else let q = frac q' in if s = 0 then \ / (2 * pi) * (pre_zeta q 1 - pre_zeta (1 - q) 1 + ln (1 - q) - ln q + pi * \) else if s \ \ then eval_fds (fds_perzeta q) s else complex_of_real (2 * pi) powr (s - 1) * \ * Gamma (1 - s) * (\ powr (-s) * hurwitz_zeta q (1 - s) - \ powr s * hurwitz_zeta (1 - q) (1 - s)))" interpretation fds_perzeta: periodic_fun_simple' fds_perzeta by standard (simp_all add: fds_perzeta_def exp_add ring_distribs exp_eq_polar cis_mult [symmetric] cis_multiple_2pi) interpretation perzeta: periodic_fun_simple' perzeta proof - have [simp]: "n + 1 \ \ \ n \ \" for n :: real by (simp flip: frac_eq_0_iff add: frac.plus_1) show "periodic_fun_simple' perzeta" by standard (auto simp: fun_eq_iff perzeta_def Let_def frac.plus_1) qed lemma perzeta_frac [simp]: "perzeta (frac q) = perzeta q" by (auto simp: perzeta_def fun_eq_iff Let_def) lemma fds_perzeta_frac [simp]: "fds_perzeta (frac q) = fds_perzeta q" using fds_perzeta.plus_of_int[of "frac q" "\q\"] by (simp add: frac_def) lemma abs_conv_abscissa_perzeta: "abs_conv_abscissa (fds_perzeta q) \ 1" proof (rule abs_conv_abscissa_leI) fix c assume c: "ereal c > 1" hence "summable (\n. n powr -c)" by (simp add: summable_real_powr_iff) also have "?this \ fds_abs_converges (fds_perzeta q) (of_real c)" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) (auto simp: norm_divide norm_powr_real_powr fds_perzeta_def powr_minus field_simps) finally show "\s. s \ 1 = c \ fds_abs_converges (fds_perzeta q) s" by (intro exI[of _ "of_real c"]) auto qed lemma conv_abscissa_perzeta: "conv_abscissa (fds_perzeta q) \ 1" by (rule order.trans[OF conv_le_abs_conv_abscissa abs_conv_abscissa_perzeta]) lemma fds_perzeta_0 [simp]: "fds_perzeta 0 = fds_zeta" by (simp add: fds_perzeta_def fds_zeta_def) lemma perzeta_0 [simp]: "Re s > 1 \ perzeta 0 s = zeta s" by (simp add: perzeta_def eval_fds_zeta) lemma perzeta_int: "q \ \ \ perzeta q = zeta" by (simp add: perzeta_def fun_eq_iff) lemma fds_perzeta_int: "q \ \ \ fds_perzeta q = fds_zeta" by (auto simp: fds_perzeta.of_int elim!: Ints_cases) lemma sums_fds_perzeta: assumes "Re s > 1" shows "(\m. exp (2 * pi * \ * Suc m * q) / of_nat (Suc m) powr s) sums eval_fds (fds_perzeta q) s" proof - have "conv_abscissa (fds_perzeta q) \ 1" by (rule conv_abscissa_perzeta) also have "\ < ereal (Re s)" using assms by (simp add: one_ereal_def) finally have "fds_converges (fds_perzeta q) s" by (intro fds_converges) auto hence "(\n. fds_nth (fds_perzeta q) (Suc n) / nat_power (Suc n) s) sums eval_fds (fds_perzeta q) s" by (subst sums_Suc_iff) (auto simp: fds_converges_iff) thus ?thesis by (simp add: fds_perzeta_def) qed lemma sum_tendsto_fds_perzeta: assumes "Re s > 1" shows "(\n. \k\{0<..n}. exp (2 * real k * pi * q * \) * of_nat k powr - s) \ eval_fds (fds_perzeta q) s" proof - have "(\m. exp (2 * pi * \ * Suc m * q) / of_nat (Suc m) powr s) sums eval_fds (fds_perzeta q) s" by (intro sums_fds_perzeta assms) hence "(\n. \k) * of_nat (Suc k) powr -s) \ eval_fds (fds_perzeta q) s" (is "filterlim ?f _ _") by (simp add: sums_def powr_minus field_simps) also have "?f = (\n. \k\{0<..n}. exp (2 * real k * pi * q * \) * of_nat k powr -s)" by (intro ext sum.reindex_bij_betw sum.reindex_bij_witness[of _ "\n. n - 1" Suc]) auto finally show ?thesis by simp qed text \ Using the geometric series, it is easy to see that the Dirichlet series for $F(s, q)$ has bounded partial sums for non-integer \q\, so it must converge for any \s\ with $\mathfrak{R}(s) > 0$. \ lemma conv_abscissa_perzeta': assumes "q \ \" shows "conv_abscissa (fds_perzeta q) \ 0" proof (rule conv_abscissa_leI) fix c :: real assume c: "ereal c > 0" have "fds_converges (fds_perzeta q) (of_real c)" proof (rule bounded_partial_sums_imp_fps_converges) define \ where "\ = exp (2 * pi * \ * q)" have [simp]: "norm \ = 1" by (simp add: \_def) define M where "M = 2 / norm (1 - \)" from \q \ \\ have "\ \ 1" by (auto simp: \_def exp_eq_1) hence "M > 0" by (simp add: M_def) show "Bseq (\n. \k\n. fds_nth (fds_perzeta q) k / nat_power k 0)" unfolding Bseq_def proof (rule exI, safe) fix n :: nat show "norm (\k\n. fds_nth (fds_perzeta q) k / nat_power k 0) \ M" proof (cases "n = 0") case False have "(\k\n. fds_nth (fds_perzeta q) k / nat_power k 0) = (\k\{1..1 + (n - 1)}. \ ^ k)" using False by (intro sum.mono_neutral_cong_right) (auto simp: fds_perzeta_def fds_nth_fds exp_of_nat_mult [symmetric] mult_ac \_def) also have "\ = \ * (1 - \ ^ n) / (1 - \)" using \\ \ 1\ and False by (subst sum_gp_offset) simp also have "norm \ \ 1 * (norm (1::complex) + norm (\ ^ n)) / norm (1 - \)" unfolding norm_mult norm_divide by (intro mult_mono divide_right_mono norm_triangle_ineq4) auto also have "\ = M" by (simp add: M_def norm_power) finally show ?thesis . qed (use \M > 0\ in simp_all) qed fact+ qed (insert c, auto) thus "\s. s \ 1 = c \ fds_converges (fds_perzeta q) s" by (intro exI[of _ "of_real c"]) auto qed subsection \Hurwitz's formula\ text \ We now move on to prove Hurwitz's formula relating the Hurwitz zeta function and the periodic zeta function. We mostly follow Apostol's proof, although we do make some small changes in order to make the proof more amenable to Isabelle's complex analysis library. The big difference is that Apostol integrates along a circle with a slit, where the two sides of the slit lie on different branches of the integrand. This makes sense when looking as the integrand as a Riemann surface, but we do not have a notion of Riemann surfaces in Isabelle. It is therefore much easier to simply cut the circle into an upper and a lower half. In fact, the integral on the lower half can be reduced to the one on the upper half easily by symmetry, so we really only need to handle the integral on the upper half. The integration contour that we will use is therefore a semi-annulus in the upper half of the complex plane, centred around the origin. Now, first of all, we prove the existence of an important improper integral that we will need later. \ (* TODO: Move *) lemma set_integrable_bigo: fixes f g :: "real \ 'a :: {banach, real_normed_field, second_countable_topology}" assumes "f \ O(\x. g x)" and "set_integrable lborel {a..} g" assumes "\b. b \ a \ set_integrable lborel {a.. 0" "\x. x \ x0 \ norm (f x) \ C * norm (g x)" by (fastforce elim!: landau_o.bigE simp: eventually_at_top_linorder) define x0' where "x0' = max a x0" have "set_integrable lborel {a..x. C *\<^sub>R (indicator {x0'..} x *\<^sub>R g x))" unfolding set_integrable_def by (intro integrable_scaleR_right) (simp add: abs_mult norm_mult) next from assms(4) have "set_borel_measurable borel {x0'..} f" by (rule set_borel_measurable_subset) (auto simp: x0'_def) thus "(\x. indicator {x0'..} x *\<^sub>R f x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next show "AE x in lborel. norm (indicator {x0'..} x *\<^sub>R f x) \ norm (C *\<^sub>R (indicator {x0'..} x *\<^sub>R g x))" using C by (intro AE_I2) (auto simp: abs_mult indicator_def x0'_def) qed ultimately have "set_integrable lborel ({a.. {x0'..}) f" by (rule set_integrable_Un) auto also have "{a.. {x0'..} = {a..}" by (auto simp: x0'_def) finally show ?thesis . qed lemma set_integrable_Gamma_hurwitz_aux2_real: fixes s a :: real assumes "r > 0" and "a > 0" shows "set_integrable lborel {r..} (\x. x powr s * (exp (-a * x)) / (1 - exp (-x)))" (is "set_integrable _ _ ?g") proof (rule set_integrable_bigo) have "(\x. exp (-(a/2) * x)) integrable_on {r..}" using assms by (intro integrable_on_exp_minus_to_infinity) auto hence "set_integrable lebesgue {r..} (\x. exp (-(a/2) * x))" by (intro nonnegative_absolutely_integrable) simp_all thus "set_integrable lborel {r..} (\x. exp (-(a/2) * x))" by (simp add: set_integrable_def integrable_completion) next fix y :: real have "set_integrable lborel {r..y} ?g" using assms by (intro borel_integrable_atLeastAtMost') (auto intro!: continuous_intros) thus "set_integrable lborel {r.. O(\x. exp (-(a/2) * x))" by real_asymp qed (auto simp: set_borel_measurable_def) lemma set_integrable_Gamma_hurwitz_aux2: fixes s :: complex and a :: real assumes "r > 0" and "a > 0" shows "set_integrable lborel {r..} (\x. x powr -s * (exp (- a * x)) / (1 - exp (- x)))" (is "set_integrable _ _ ?g") proof - have "set_integrable lborel {r..} (\x. x powr -Re s * (exp (-a * x)) / (1 - exp (-x)))" (is "set_integrable _ _ ?g'") by (rule set_integrable_Gamma_hurwitz_aux2_real) (use assms in auto) also have "?this \ integrable lborel (\x. indicator {r..} x *\<^sub>R ?g' x)" by (simp add: set_integrable_def) also have "(\x. indicator {r..} x *\<^sub>R ?g' x) = (\x. norm (indicator {r..} x *\<^sub>R ?g x))" using assms by (auto simp: indicator_def norm_mult norm_divide norm_powr_real_powr fun_eq_iff exp_of_real exp_minus norm_inverse in_Reals_norm power2_eq_square divide_simps) finally show ?thesis unfolding set_integrable_def by (subst (asm) integrable_norm_iff) auto qed lemma closed_neg_Im_slit: "closed {z. Re z = 0 \ Im z \ 0}" proof - have "closed ({z. Re z = 0} \ {z. Im z \ 0})" by (intro closed_Int closed_halfspace_Re_eq closed_halfspace_Im_le) also have "{z. Re z = 0} \ {z. Im z \ 0} = {z. Re z = 0 \ Im z \ 0}" by blast finally show ?thesis . qed text \ We define tour semi-annulus path. When this path is mirrored into the lower half of the complex plane and subtracted from the original path and the outer radius tends to \\\, this becomes a Hankel contour extending to \-\\. \ definition hankel_semiannulus :: "real \ nat \ real \ complex" where "hankel_semiannulus r N = (let R = (2 * N + 1) * pi in part_circlepath 0 R 0 pi +++ \\Big half circle\ linepath (of_real (-R)) (of_real (-r)) +++ \\Line on the negative real axis\ part_circlepath 0 r pi 0 +++ \\Small half circle\ linepath (of_real r) (of_real R)) \\Line on the positive real axis\" lemma path_hankel_semiannulus [simp, intro]: "path (hankel_semiannulus r R)" and valid_path_hankel_semiannulus [simp, intro]: "valid_path (hankel_semiannulus r R)" and pathfinish_hankel_semiannulus [simp, intro]: "pathfinish (hankel_semiannulus r R) = pathstart (hankel_semiannulus r R)" by (simp_all add: hankel_semiannulus_def Let_def) text \ We set the stage for an application of the Residue Theorem. We define a function \[f(s, z) = z^{-s} \frac{\exp(az)}{1-\exp(-z)}\ ,\] which will be the integrand. However, the principal branch of $z^{-s}$ has a branch cut along the non-positive real axis, which is bad because a part of our integration path also lies on the non-positive real axis. We therefore choose a slightly different branch of $z^{-s}$ by moving the logarithm branch along by $90^{\circ}$ so that the branch cut lies on the non-positive imaginary axis instead. \ context fixes a :: real fixes f :: "complex \ complex \ complex" and g :: "complex \ real \ complex" and h :: "real \ complex \ real \ complex" and Res :: "complex \ nat \ complex" and Ln' :: "complex \ complex" and F :: "real \ complex \ complex" assumes a: "a \ {0<..1}" \\Our custom branch of the logarithm\ defines "Ln' \ (\z. ln (-\ * z) + \ * pi / 2)" \\The integrand\ defines "f \ (\s z. exp (Ln' z * (-s) + of_real a * z) / (1 - exp z))" \\The integrand on the negative real axis\ defines "g \ (\s x. complex_of_real x powr -s * of_real (exp (-a*x)) / of_real (1 - exp (-x)))" \\The integrand on the circular arcs\ defines "h \ (\r s t. r * \ * cis t * exp (a * (r * cis t) - (ln r + \ * t) * s) / (1 - exp (r * cis t)))" \\The interesting part of the residues\ defines "Res \ (\s k. exp (of_real (2 * real k * pi * a) * \) * of_real (2 * real k * pi) powr (-s))" \\The periodic zeta function (at least on $\mathfrak{R}(s) > 1$ half-plane)\ defines "F \ (\q. eval_fds (fds_perzeta q))" begin text \ First, some basic properties of our custom branch of the logarithm: \ lemma Ln'_i: "Ln' \ = \ * pi / 2" by (simp add: Ln'_def) lemma Ln'_of_real_pos: assumes "x > 0" shows "Ln' (of_real x) = of_real (ln x)" proof - have "Ln' (of_real x) = Ln (of_real x * (-\)) + \ * pi / 2" by (simp add: Ln'_def mult_ac) also have "\ = of_real (ln x)" using assms by (subst Ln_times_of_real) (auto simp: Ln_of_real) finally show ?thesis . qed lemma Ln'_of_real_neg: assumes "x < 0" shows "Ln' (of_real x) = of_real (ln (-x)) + \ * pi" proof - have "Ln' (of_real x) = Ln (of_real (-x) * \) + \ * pi / 2" by (simp add: Ln'_def mult_ac) also have "\ = of_real (ln (-x)) + \ * pi" using assms by (subst Ln_times_of_real) (auto simp: Ln_Reals_eq) finally show ?thesis . qed lemma Ln'_times_of_real: "Ln' (of_real x * z) = of_real (ln x) + Ln' z" if "x > 0" "z \ 0" for z x proof - have "Ln' (of_real x * z) = Ln (of_real x * (- \ * z)) + \ * pi / 2" by (simp add: Ln'_def mult_ac) also have "\ = of_real (ln x) + Ln' z" using that by (subst Ln_times_of_real) (auto simp: Ln'_def Ln_of_real) finally show ?thesis . qed lemma Ln'_cis: assumes "t \ {-pi / 2<..3 / 2 * pi}" shows "Ln' (cis t) = \ * t" proof - have "exp (\ * pi / 2) = \" by (simp add: exp_eq_polar) hence "Ln (- (\ * cis t)) = \ * (t - pi / 2)" using assms by (intro Ln_unique) (auto simp: algebra_simps exp_diff cis_conv_exp) thus ?thesis by (simp add: Ln'_def algebra_simps) qed text \ Next, we show that the line and circle integrals are holomorphic using Leibniz's rule: \ lemma contour_integral_part_circlepath_h: assumes r: "r > 0" shows "contour_integral (part_circlepath 0 r 0 pi) (f s) = integral {0..pi} (h r s)" proof - have "contour_integral (part_circlepath 0 r 0 pi) (f s) = integral {0..pi} (\t. f s (r * cis t) * r * \ * cis t)" by (simp add: contour_integral_part_circlepath_eq) also have "integral {0..pi} (\t. f s (r * cis t) * r * \ * cis t) = integral {0..pi} (h r s)" proof (intro integral_cong) fix t assume t: "t \ {0..pi}" have "-(pi / 2) < 0" by simp also have "0 \ t" using t by simp finally have "t \ {-(pi/2)<..3/2*pi}" using t by auto thus "f s (r * cis t) * r * \ * cis t = h r s t" using r by (simp add: f_def Ln'_times_of_real Ln'_cis h_def Ln_Reals_eq) qed finally show ?thesis . qed lemma integral_g_holomorphic: assumes "b > 0" shows "(\s. integral {b..c} (g s)) holomorphic_on A" proof - define g' where "g' = (\s t. - (of_real t powr (-s)) * complex_of_real (ln t) * of_real (exp (- (a * t))) / of_real (1 - exp (- t)))" have "(\s. integral (cbox b c) (g s)) holomorphic_on UNIV" proof (rule leibniz_rule_holomorphic) fix s :: complex and t :: real assume "t \ cbox b c" thus "((\s. g s t) has_field_derivative g' s t) (at s)" using assms by (auto simp: g_def g'_def powr_def Ln_Reals_eq intro!: derivative_eq_intros) next fix s show "g s integrable_on cbox b c" for s unfolding g_def using assms by (intro integrable_continuous continuous_intros) auto next show "continuous_on (UNIV \ cbox b c) (\(s, t). g' s t)" using assms by (auto simp: g'_def case_prod_unfold intro!: continuous_intros) qed auto thus ?thesis by auto qed lemma integral_h_holomorphic: assumes r: "r \ {0<..<2}" shows "(\s. integral {b..c} (h r s)) holomorphic_on A" proof - have no_sing: "exp (r * cis t) \ 1" for t proof define z where "z = r * cis t" assume "exp z = 1" then obtain n where "norm z = 2 * pi * of_int \n\" by (auto simp: exp_eq_1 cmod_def abs_mult) moreover have "norm z = r" using r by (simp add: z_def norm_mult) ultimately have r_eq: "r = 2 * pi * of_int \n\" by simp with r have "n \ 0" by auto moreover from r have "r < 2 * pi" using pi_gt3 by simp with r_eq have "\n\ < 1" by simp ultimately show False by simp qed define h' where "h' = (\s t. exp (a * r * cis t - (ln r + \ * t) * s) * (-ln r - \ * t) * (r * \ * cis t) / (1 - exp (r * cis t)))" have "(\s. integral (cbox b c) (h r s)) holomorphic_on UNIV" proof (rule leibniz_rule_holomorphic) fix s t assume "t \ cbox b c" thus "((\s. h r s t) has_field_derivative h' s t) (at s)" using no_sing r by (auto intro!: derivative_eq_intros simp: h_def h'_def mult_ac Ln_Reals_eq) next fix s show "h r s integrable_on cbox b c" using no_sing unfolding h_def by (auto intro!: integrable_continuous_real continuous_intros) next show "continuous_on (UNIV \ cbox b c) (\(s, t). h' s t)" using no_sing by (auto simp: h'_def case_prod_unfold intro!: continuous_intros) qed auto thus ?thesis by auto qed text \ We now move on to the core result, which uses the Residue Theorem to relate a contour integral along a semi-annulus to a partial sum of the periodic zeta function. \ lemma hurwitz_formula_integral_semiannulus: fixes N :: nat and r :: real and s :: complex defines "R \ real (2 * N + 1) * pi" assumes "r > 0" and "r < 2" shows "exp (-\ * pi * s) * integral {r..R} (\x. x powr (-s) * exp (-a * x) / (1 - exp (-x))) + integral {r..R} (\x. x powr (-s) * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 R 0 pi) (f s) + contour_integral (part_circlepath 0 r pi 0) (f s) = -2 * pi * \ * exp (- s * of_real pi * \ / 2) * (\k\{0<..N}. Res s k)" (is ?thesis1) and "f s contour_integrable_on hankel_semiannulus r N" proof - have "2 < 1 * pi" using pi_gt3 by simp also have "\ \ R" unfolding R_def by (intro mult_right_mono) auto finally have "R > 2" by (auto simp: R_def) note r_R = \r > 0\ \r < 2\ this \\We integrate along the edge of a semi-annulus in the upper half of the complex plane. It consists of a big semicircle, a small semicircle, and two lines connecting the two circles, one on the positive real axis and one on the negative real axis. The integral along the big circle will vanish as the radius of the circle tends to \\\, whereas the remaining path becomes a Hankel contour, and the integral along that Hankel contour is what we are interested in, since it is connected to the Hurwitz zeta function.\ define big_circle where "big_circle = part_circlepath 0 R 0 pi" define small_circle where "small_circle = part_circlepath 0 r pi 0" define neg_line where "neg_line = linepath (complex_of_real (-R)) (complex_of_real (-r))" define pos_line where "pos_line = linepath (complex_of_real r) (complex_of_real R)" define \ where "\ = hankel_semiannulus r N" have \_altdef: "\ = big_circle +++ neg_line +++ small_circle +++ pos_line" by (simp add: \_def R_def add_ac hankel_semiannulus_def big_circle_def neg_line_def small_circle_def pos_line_def) have [simp]: "path \" "valid_path \" "pathfinish \ = pathstart \" by (simp_all add: \_def) \\The integrand has a branch cut along the non-positive imaginary axis and additional simple poles at $2n\pi i$ for any \n \ \\<^sub>>\<^sub>0\. The radius of the smaller circle will always be less than $2\pi$ and the radius of the bigger circle of the form $(2N+1)\pi$, so we always have precisely the first $N$ poles inside our path.\ define sngs where "sngs = (\n. of_real (2 * pi * real n) * \) ` {0<..}" define sngs' where "sngs' = (\n. of_real (2 * pi * real n) * \) ` {0<..N}" have sngs_subset: "sngs' \ sngs" unfolding sngs_def sngs'_def by (intro image_mono) auto have closed_sngs [intro]: "closed (sngs - sngs')" unfolding sngs_def proof (rule discrete_imp_closed[of "2 * pi"]; safe?) fix m n :: nat assume "dist (of_real (2 * pi * real m) * \) (of_real (2 * pi * real n) * \) < 2 * pi" also have "dist (of_real (2 * pi * real m) * \) (of_real (2 * pi * real n) * \) = norm (of_real (2 * pi * real m) * \ - of_real (2 * pi * real n) * \)" by (simp add: dist_norm) also have "of_real (2 * pi * real m) * \ - of_real (2 * pi * real n) * \ = of_real (2 * pi) * \ * of_int (int m - int n)" by (simp add: algebra_simps) also have "norm \ = 2 * pi * of_int \int m - int n\" unfolding norm_mult norm_of_int by (simp add: norm_mult) finally have "\real m - real n\ < 1" by simp hence "m = n" by linarith thus "of_real (2 * pi * real m) * \ = of_real (2 * pi * real n) * \" by simp qed auto \\We define an area within which the integrand is holomorphic. Choosing this area as big as possible makes things easier later on, so we only remove the branch cut and the poles.\ define S where "S = - {z. Re z = 0 \ Im z \ 0} - (sngs - sngs')" define S' where "S' = - {z. Re z = 0 \ Im z \ 0}" have sngs: "exp z = 1 \ z \ sngs" if "Re z \ 0 \ Im z > 0" for z proof assume "exp z = 1" then obtain n where n: "z = 2 * pi * of_int n * \" unfolding exp_eq_1 by (auto simp: complex_eq_iff mult_ac) moreover from n and pi_gt_zero and that have "n > 0" by (auto simp: zero_less_mult_iff) ultimately have "z = of_real (2 * pi * nat n) * \" and "nat n \ {0<..}" by auto thus "z \ sngs" unfolding sngs_def by blast qed (insert that, auto simp: sngs_def exp_eq_polar) \\We show that the path stays within the well-behaved area.\ have "path_image neg_line = of_real ` {-R..-r}" using r_R by (auto simp: neg_line_def closed_segment_Reals closed_segment_eq_real_ivl) hence "path_image neg_line \ S - sngs'" using r_R sngs_subset by (auto simp: S_def sngs_def complex_eq_iff) have "path_image pos_line = of_real ` {r..R}" using r_R by (auto simp: pos_line_def closed_segment_Reals closed_segment_eq_real_ivl) hence "path_image pos_line \ S - sngs'" using r_R sngs_subset by (auto simp: S_def sngs_def complex_eq_iff) have part_circlepath_in_S: "z \ S - sngs'" if "z \ path_image (part_circlepath 0 r' 0 pi) \ z \ path_image (part_circlepath 0 r' pi 0)" and "r' > 0" "r' \ (\n. 2 * pi * real n) ` {0<..}" for z r' proof - have z: "norm z = r' \ Im z \ 0" using that by (auto simp: path_image_def part_circlepath_def norm_mult Im_exp linepath_def intro!: mult_nonneg_nonneg sin_ge_zero) hence "Re z \ 0 \ Im z > 0" using that by (auto simp: cmod_def) moreover from z and that have "z \ sngs" by (auto simp: sngs_def norm_mult image_iff) ultimately show "z \ S - sngs'" using sngs_subset by (auto simp: S_def) qed { fix n :: nat assume n: "n > 0" have "r < 2 * pi * 1" using pi_gt3 r_R by simp also have "\ \ 2 * pi * real n" using n by (intro mult_left_mono) auto finally have "r < \" . } hence "r \ (\n. 2 * pi * real n) ` {0<..}" using r_R by auto from part_circlepath_in_S[OF _ _ this] r_R have "path_image small_circle \ S - sngs'" by (auto simp: small_circle_def) { fix n :: nat assume n: "n > 0" "2 * pi * real n = real (2 * N + 1) * pi" hence "real (2 * n) = real (2 * N + 1)" unfolding of_nat_mult by simp hence False unfolding of_nat_eq_iff by presburger } hence "R \ (\n. 2 * pi * real n) ` {0<..}" unfolding R_def by force from part_circlepath_in_S[OF _ _ this] r_R have "path_image big_circle \ S - sngs'" by (auto simp: big_circle_def) note path_images = \path_image small_circle \ S - sngs'\ \path_image big_circle \ S - sngs'\ \path_image neg_line \ S - sngs'\ \path_image pos_line \ S - sngs'\ have "path_image \ \ S - sngs'" using path_images by (auto simp: \_altdef path_image_join big_circle_def neg_line_def small_circle_def pos_line_def) \\We need to show that the complex plane is still connected after we removed the branch cut and the poles. We do this by showing that the complex plane with the branch cut removed is starlike and therefore connected. Then we remove the (countably many) poles, which does not break connectedness either.\ have "open S" using closed_neg_Im_slit by (auto simp: S_def) have "starlike (UNIV - {z. Re z = 0 \ Im z \ 0})" (is "starlike ?S'") unfolding starlike_def proof (rule bexI ballI)+ have "1 \ 1 * pi" using pi_gt3 by simp also have "\ < (2 + 2 * real N) * pi" by (intro mult_strict_right_mono) auto finally show *: "\ \ ?S'" by (auto simp: S_def) fix z assume z: "z \ ?S'" have "closed_segment \ z \ {z. Re z = 0 \ Im z \ 0} = {}" proof safe fix s assume s: "s \ closed_segment \ z" "Re s = 0" "Im s \ 0" then obtain t where t: "t \ {0..1}" "s = linepath \ z t" using linepath_image_01 by blast with z s t have z': "Re z = 0" "Im z > 0" by (auto simp: Re_linepath' S_def linepath_0') with s have "Im s \ closed_segment 1 (Im z) \ Im s \ 0" by (subst (asm) closed_segment_same_Re) auto with z' show "s \ {}" by (auto simp: closed_segment_eq_real_ivl split: if_splits) qed thus "closed_segment \ z \ ?S'" by (auto simp: S_def) qed hence "connected ?S'" by (rule starlike_imp_connected) hence "connected S'" by (simp add: Compl_eq_Diff_UNIV S'_def) have "connected S" unfolding S_def by (rule connected_open_diff_countable) (insert \connected S'\, auto simp: sngs_def closed_neg_Im_slit S'_def) \\The integrand is now clearly holomorphic on @{term "S - sngs'"} and we can apply the Residue Theorem.\ have holo: "f s holomorphic_on (S - sngs')" unfolding f_def Ln'_def S_def using sngs by (auto intro!: holomorphic_intros simp: complex_nonpos_Reals_iff) have "contour_integral \ (f s) = of_real (2 * pi) * \ * (\z\sngs'. winding_number \ z * residue (f s) z)" proof (rule Residue_theorem) show "\z. z \ S \ winding_number \ z = 0" proof safe fix z assume "z \ S" hence "Re z = 0 \ Im z \ 0 \ z \ sngs - sngs'" by (auto simp: S_def) thus "winding_number \ z = 0" proof define x where "x = -Im z" assume "Re z = 0 \ Im z \ 0" hence x: "z = -of_real x * \" "x \ 0" unfolding complex_eq_iff by (simp_all add: x_def) obtain B where "\z. norm z \ B \ winding_number \ z = 0" using winding_number_zero_at_infinity[of \] by auto hence "winding_number \ (-of_real (max B 0) * \) = 0" by (auto simp: norm_mult) also have "winding_number \ (-of_real (max B 0) * \) = winding_number \ z" proof (rule winding_number_eq) from x have "closed_segment (-of_real (max B 0) * \) z \ {z. Re z = 0 \ Im z \ 0}" by (auto simp: closed_segment_same_Re closed_segment_eq_real_ivl) with \path_image \ \ S - sngs'\ show "closed_segment (-of_real (max B 0) * \) z \ path_image \ = {}" by (auto simp: S_def) qed auto finally show "winding_number \ z = 0" . next assume z: "z \ sngs - sngs'" show "winding_number \ z = 0" proof (rule winding_number_zero_outside) have "path_image \ = path_image big_circle \ path_image neg_line \ path_image small_circle \ path_image pos_line" unfolding \_altdef small_circle_def big_circle_def pos_line_def neg_line_def by (simp add: path_image_join Un_assoc) also have "\ \ cball 0 ((2 * N + 1) * pi)" using r_R by (auto simp: small_circle_def big_circle_def pos_line_def neg_line_def path_image_join norm_mult R_def path_image_part_circlepath' in_Reals_norm closed_segment_Reals closed_segment_eq_real_ivl) finally show "path_image \ \ \" . qed (insert z, auto simp: sngs_def sngs'_def norm_mult) qed qed qed (insert \path_image \ \ S - sngs'\ \connected S\ \open S\ holo, auto simp: sngs'_def) \\We can use Wenda Li's framework to compute the winding numbers at the poles and show that they are all 1.\ also have "winding_number \ z = 1" if "z \ sngs'" for z proof - have "r < 2 * pi * 1" using pi_gt3 r_R by simp also have "\ \ 2 * pi * real n" if "n > 0" for n using that by (intro mult_left_mono) auto finally have norm_z: "norm z > r" "norm z < R" using that r_R by (auto simp: sngs'_def norm_mult R_def) have "cindex_pathE big_circle z = -1" using r_R that unfolding big_circle_def by (subst cindex_pathE_circlepath_upper(1)) (auto simp: sngs'_def norm_mult R_def) have "cindex_pathE small_circle z = -1" using r_R that norm_z unfolding small_circle_def by (subst cindex_pathE_reversepath', subst reversepath_part_circlepath, subst cindex_pathE_circlepath_upper(2)) (auto simp: sngs'_def norm_mult) have "cindex_pathE neg_line z = 0" "cindex_pathE pos_line z = 0" unfolding neg_line_def pos_line_def using r_R that by (subst cindex_pathE_linepath; force simp: neg_line_def cindex_pathE_linepath closed_segment_Reals closed_segment_eq_real_ivl sngs'_def complex_eq_iff)+ note indices = \cindex_pathE big_circle z = -1\ \cindex_pathE small_circle z = -1\ \cindex_pathE neg_line z = 0\ \cindex_pathE pos_line z = 0\ show ?thesis unfolding \_altdef big_circle_def small_circle_def pos_line_def neg_line_def by eval_winding (insert indices path_images that, auto simp: big_circle_def small_circle_def pos_line_def neg_line_def) qed hence "(\z\sngs'. winding_number \ z * residue (f s) z) = (\z\sngs'. residue (f s) z)" by simp also have "\ = (\k\{0<..N}. residue (f s) (2 * pi * of_nat k * \))" unfolding sngs'_def by (subst sum.reindex) (auto intro!: inj_onI simp: o_def) \\Next, we compute the residues at each pole.\ also have "residue (f s) (2 * pi * of_nat k * \) = -exp (- s * of_real pi * \ / 2) * Res s k" if "k \ {0<..N}" for k unfolding f_def proof (subst residue_simple_pole_deriv) show "open S'" using closed_neg_Im_slit by (auto simp: S'_def) show "connected S'" by fact show "(\z. exp (Ln' z * (-s) + of_real a * z)) holomorphic_on S'" "(\z. 1 - exp z) holomorphic_on S'" by (auto simp: S'_def Ln'_def complex_nonpos_Reals_iff intro!: holomorphic_intros) have "((\z. 1 - exp z) has_field_derivative -exp (2 * pi * k * \)) (at (of_real (2 * pi * real k) * \))" by (auto intro!: derivative_eq_intros) also have "-exp (2 * pi * k * \) = -1" by (simp add: exp_eq_polar) finally show "((\z. 1 - exp z) has_field_derivative -1) (at (of_real (2 * pi * real k) * \))" . have "Im (of_real (2 * pi * real k) * \) > 0" using pi_gt_zero that by auto thus "of_real (2 * pi * real k) * \ \ S'" by (simp add: S'_def) have "exp (\ * pi / 2) = \" by (simp add: exp_eq_polar) hence "exp (Ln' (complex_of_real (2 * pi * real k) * \) * -s + of_real a * (of_real (2 * pi * real k) * \)) / -1 = - exp (2 * k * a * pi * \ - s * pi * \ / 2 - s * ln (2 * k * pi))" (is "?R = _") using that by (subst Ln'_times_of_real) (simp_all add: Ln'_i algebra_simps exp_diff) also have "\ = -exp (- s * of_real pi * \ / 2) * Res s k" using that by (simp add: Res_def exp_diff powr_def exp_minus inverse_eq_divide Ln_Reals_eq mult_ac) finally show "?R = -exp (- s * of_real pi * \ / 2) * Res s k" . qed (insert that, auto simp: S'_def exp_eq_polar) hence "(\k\{0<..N}. residue (f s) (2 * pi * of_nat k * \)) = -exp (- s * of_real pi * \ / 2) * (\k\{0<..N}. Res s k)" by (simp add: sum_distrib_left) \\This gives us the final result:\ finally have "contour_integral \ (f s) = -2 * pi * \ * exp (- s * of_real pi * \ / 2) * (\k\{0<..N}. Res s k)" by simp \\Lastly, we decompose the contour integral into its four constituent integrals because this makes them somewhat nicer to work with later on.\ also show "f s contour_integrable_on \" proof (rule contour_integrable_holomorphic_simple) show "path_image \ \ S - sngs'" by fact have "closed sngs'" by (intro finite_imp_closed) (auto simp: sngs'_def) with \open S\ show "open (S - sngs')" by auto qed (insert holo, auto) hence eq: "contour_integral \ (f s) = contour_integral big_circle (f s) + contour_integral neg_line (f s) + contour_integral small_circle (f s) + contour_integral pos_line (f s)" unfolding \_altdef big_circle_def neg_line_def small_circle_def pos_line_def by simp also have "contour_integral neg_line (f s) = integral {-R..-r} (\x. f s (complex_of_real x))" unfolding neg_line_def using r_R by (subst contour_integral_linepath_Reals_eq) auto also have "\ = exp (- \ * pi * s) * integral {r..R} (\x. exp (-ln x * s) * exp (-a * x) / (1 - exp (-x)))" (is "_ = _ * ?I") unfolding integral_mult_right [symmetric] using r_R - by (subst integral_reflect_real [symmetric], intro integral_cong) + by (subst Henstock_Kurzweil_Integration.integral_reflect_real [symmetric], intro integral_cong) (auto simp: f_def exp_of_real Ln'_of_real_neg exp_minus exp_Reals_eq exp_diff exp_add field_simps) also have "?I = integral {r..R} (\x. x powr (-s) * exp (-a * x) / (1 - exp (-x)))" using r_R by (intro integral_cong) (auto simp: powr_def Ln_Reals_eq exp_minus exp_diff field_simps) also have "contour_integral pos_line (f s) = integral {r..R} (\x. f s (complex_of_real x))" unfolding pos_line_def using r_R by (subst contour_integral_linepath_Reals_eq) auto also have "\ = integral {r..R} (\x. x powr (-s) * exp (a * x) / (1 - exp x))" using r_R by (intro integral_cong) (simp add: f_def Ln'_of_real_pos exp_diff exp_minus exp_Reals_eq field_simps powr_def Ln_Reals_eq) finally show ?thesis1 by (simp only: add_ac big_circle_def small_circle_def) qed text \ Next, we need bounds on the integrands of the two semicircles. \ lemma hurwitz_formula_bound1: defines "H \ \z. exp (complex_of_real a * z) / (1 - exp z)" assumes "r > 0" obtains C where "C \ 0" and "\z. z \ (\n::int. ball (2 * n * pi * \) r) \ norm (H z) \ C" proof - define A where "A = cbox (-1 - pi * \) (1 + pi * \) - ball 0 r" { fix z assume "z \ A" have "exp z \ 1" proof assume "exp z = 1" then obtain n :: int where [simp]: "z = 2 * n * pi * \" by (subst (asm) exp_eq_1) (auto simp: complex_eq_iff) from \z \ A\ have "(2 * n) * pi \ (-1) * pi" and "(2 * n) * pi \ 1 * pi" by (auto simp: A_def in_cbox_complex_iff) hence "n = 0" by (subst (asm) (1 2) mult_le_cancel_right) auto with \z \ A\ and \r > 0\ show False by (simp add: A_def) qed } hence "continuous_on A H" by (auto simp: A_def H_def intro!: continuous_intros) moreover have "compact A" by (auto simp: A_def compact_eq_bounded_closed) ultimately have "compact (H ` A)" by (rule compact_continuous_image) hence "bounded (H ` A)" by (rule compact_imp_bounded) then obtain C where bound_inside: "\z. z \ A \ norm (H z) \ C" by (auto simp: bounded_iff) have bound_outside: "norm (H z) \ exp 1 / (exp 1 - 1)" if "\Re z\ > 1" for z proof - have "norm (H z) = exp (a * Re z) / norm (1 - exp z)" by (simp add: H_def norm_divide) also have "\1 - exp (Re z)\ \ norm (1 - exp z)" by (rule order.trans[OF _ norm_triangle_ineq3]) simp hence "exp (a * Re z) / norm (1 - exp z) \ exp (a * Re z) / \1 - exp (Re z)\" using that by (intro divide_left_mono mult_pos_pos) auto also have "\ \ exp 1 / (exp 1 - 1)" proof (cases "Re z > 1") case True hence "exp (a * Re z) / \1 - exp (Re z)\ = exp (a * Re z) / (exp (Re z) - 1)" by simp also have "\ \ exp (Re z) / (exp (Re z) - 1)" using a True by (intro divide_right_mono) auto also have "\ = 1 / (1 - exp (-Re z))" by (simp add: exp_minus field_simps) also have "\ \ 1 / (1 - exp (-1))" using True by (intro divide_left_mono diff_mono) auto also have "\ = exp 1 / (exp 1 - 1)" by (simp add: exp_minus field_simps) finally show ?thesis . next case False with that have "Re z < -1" by simp hence "exp (a * Re z) / \1 - exp (Re z)\ = exp (a * Re z) / (1 - exp (Re z))" by simp also have "\ \ 1 / (1 - exp (Re z))" using a and \Re z < -1\ by (intro divide_right_mono) (auto intro: mult_nonneg_nonpos) also have "\ \ 1 / (1 - exp (-1))" using \Re z < -1\ by (intro divide_left_mono) auto also have "\ = exp 1 / (exp 1 - 1)" by (simp add: exp_minus field_simps) finally show ?thesis . qed finally show ?thesis . qed define D where "D = max C (exp 1 / (exp 1 - 1))" have "D \ 0" by (simp add: D_def max.coboundedI2) have "norm (H z) \ D" if "z \ (\n::int. ball (2 * n * pi * \) r)" for z proof (cases "\Re z\ \ 1") case False with bound_outside[of z] show ?thesis by (simp add: D_def) next case True define n where "n = \Im z / (2 * pi) + 1 / 2\" have "Im (z - 2 * n * pi * \) = frac (Im z / (2 * pi) + 1 / 2) * (2 * pi) - pi" by (simp add: n_def frac_def algebra_simps) also have "\ \ {-pi..)) \ C" using True that by (intro bound_inside) (auto simp: A_def in_cbox_complex_iff dist_norm n_def) also have "exp (2 * pi * n * \) = 1" by (simp add: exp_eq_polar) hence "norm (H (z - 2 * n * pi * \)) = norm (H z)" by (simp add: H_def norm_divide exp_diff mult_ac) also have "C \ D" by (simp add: D_def) finally show ?thesis . qed from \D \ 0\ and this show ?thesis by (rule that) qed lemma hurwitz_formula_bound2: obtains C where "C \ 0" and "\r z. r > 0 \ r < pi \ z \ sphere 0 r \ norm (f s z) \ C * r powr (-Re s - 1)" proof - have "2 * pi > 0" by auto have nz: "1 - exp z \ 0" if "z \ ball 0 (2 * pi) - {0}" for z :: complex proof assume "1 - exp z = 0" then obtain n where "z = 2 * pi * of_int n * \" by (auto simp: exp_eq_1 complex_eq_iff[of z]) moreover have "\real_of_int n\ < 1 \ n = 0" by linarith ultimately show False using that by (auto simp: norm_mult) qed have ev: "eventually (\z::complex. 1 - exp z \ 0) (at 0)" using eventually_at_ball'[OF \2 * pi > 0\] by eventually_elim (use nz in auto) have [simp]: "subdegree (1 - fps_exp (1 :: complex)) = 1" by (intro subdegreeI) auto hence "(\z. exp (a * z) * (if z = 0 then -1 else z / (1 - exp z :: complex))) has_fps_expansion fps_exp a * (fps_X / (fps_const 1 - fps_exp 1))" by (auto intro!: fps_expansion_intros) hence "(\z::complex. exp (a * z) * (if z = 0 then -1 else z / (1 - exp z))) \ O[at 0](\z. 1)" using continuous_imp_bigo_1 has_fps_expansion_imp_continuous by blast also have "?this \ (\z::complex. exp (a * z) * (z / (1 - exp z))) \ O[at 0](\z. 1)" by (intro landau_o.big.in_cong eventually_mono[OF ev]) auto finally have "\g. g holomorphic_on ball 0 (2 * pi) \ (\z\ball 0 (2 * pi) - {0}. g z = exp (of_real a * z) * (z / (1 - exp z)))" using nz by (intro holomorphic_on_extend holomorphic_intros) auto then guess g by (elim exE conjE) note g = this hence "continuous_on (ball 0 (2 * pi)) g" by (auto dest: holomorphic_on_imp_continuous_on) hence "continuous_on (cball 0 pi) g" by (rule continuous_on_subset) (subst cball_subset_ball_iff, use pi_gt_zero in auto) hence "compact (g ` cball 0 pi)" by (intro compact_continuous_image) auto hence "bounded (g ` cball 0 pi)" by (auto simp: compact_imp_bounded) then obtain C where C: "\x\cball 0 pi. norm (g x) \ C" by (auto simp: bounded_iff) { fix r :: real assume r: "r > 0" "r < pi" fix z :: complex assume z: "z \ sphere 0 r" define x where "x = (if Arg z \ -pi / 2 then Arg z + 2 * pi else Arg z)" have "exp (\ * (2 * pi)) = 1" by (simp add: exp_eq_polar) with z have "z = r * exp (\ * x)" using r pi_gt_zero Arg_eq[of z] by (auto simp: x_def exp_add distrib_left) have "x > - pi / 2" "x \ 3 / 2 * pi" using Arg_le_pi[of z] mpi_less_Arg[of z] by (auto simp: x_def) note x = \z = r * exp (\ * x)\ this from x r have z': "z \ cball 0 pi - {0}" using pi_gt3 by (auto simp: norm_mult) also have "cball 0 pi \ ball (0::complex) (2 * pi)" by (subst cball_subset_ball_iff) (use pi_gt_zero in auto) hence "cball 0 pi - {0} \ ball 0 (2 * pi) - {0::complex}" by blast finally have z'': "z \ ball 0 (2 * pi) - {0}" . hence bound: "norm (exp (a * z) * (z / (1 - exp z))) \ C" using C and g and z' by force have "exp z \ 1" using nz z'' by auto with bound z'' have bound': "norm (exp (a * z) / (1 - exp z)) \ C / norm z" by (simp add: norm_divide field_simps norm_mult) have "Ln' z = of_real (ln r) + Ln' (exp (\ * of_real x))" using x r by (simp add: Ln'_times_of_real) also have "exp (\ * pi / 2) = \" by (simp add: exp_eq_polar) hence "Ln' (exp (\ * of_real x)) = Ln (exp (\ * of_real (x - pi / 2))) + \ * pi / 2" by (simp add: algebra_simps Ln'_def exp_diff) also have "\ = \ * x" using x pi_gt3 by (subst Ln_exp) (auto simp: algebra_simps) finally have "norm (exp (-Ln' z * s)) = exp (x * Im s - ln r * Re s)" by simp also { have "x * Im s \ \x * Im s\" by (rule abs_ge_self) also have "\ \ (3/2 * pi) * \Im s\" unfolding abs_mult using x by (intro mult_right_mono) auto finally have "exp (x * Im s - ln r * Re s) \ exp (3 / 2 * pi * \Im s\ - ln r * Re s)" by simp } finally have "norm (exp (-Ln' z * s) * (exp (a * z) / (1 - exp z))) \ exp (3 / 2 * pi * \Im s\ - ln r * Re s) * (C / norm z)" unfolding norm_mult[of "exp t" for t] by (intro mult_mono bound') simp_all also have "norm z = r" using \r > 0\ by (simp add: x norm_mult) also have "exp (3 / 2 * pi * \Im s\ - ln r * Re s) = exp (3 / 2 * pi * \Im s\) * r powr (-Re s)" using r by (simp add: exp_diff powr_def exp_minus inverse_eq_divide) finally have "norm (f s z) \ C * exp (3 / 2 * pi * \Im s\) * r powr (-Re s - 1)" using r by (simp add: f_def exp_diff exp_minus field_simps powr_diff) also have "\ \ max 0 (C * exp (3 / 2 * pi * \Im s\)) * r powr (-Re s - 1)" by (intro mult_right_mono max.coboundedI2) auto finally have "norm (f s z) \ \" . } with that[of "max 0 (C * exp (3 / 2 * pi * \Im s\))"] show ?thesis by auto qed text \ We can now relate the integral along a partial Hankel contour that is cut off at $-\pi$ to $\zeta(1 - s, a) / \Gamma(s)$. \ lemma rGamma_hurwitz_zeta_eq_contour_integral: fixes s :: complex and r :: real assumes "s \ 0" and r: "r \ {1..<2}" and a: "a > 0" defines "err1 \ (\s r. contour_integral (part_circlepath 0 r pi 0) (f s))" defines "err2 \ (\s r. cnj (contour_integral (part_circlepath 0 r pi 0) (f (cnj s))))" shows "2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s) = err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)" (is "?f s = ?g s") proof (rule analytic_continuation_open[where f = ?f]) fix s :: complex assume s: "s \ {s. Re s < 0}" \\We first show that the integrals along the Hankel contour cut off at $-\pi$ all have the same value, no matter what the radius of the circle is (as long as it is small enough). We call this value \C\. This argument could be done by a homotopy argument, but it is easier to simply re-use the above result about the contour integral along the annulus where we fix the radius of the outer circle to $\pi$.\ define C where "C = -contour_integral (part_circlepath 0 pi 0 pi) (f s) + cnj (contour_integral (part_circlepath 0 pi 0 pi) (f (cnj s)))" have integrable: "set_integrable lborel A (g s)" if "A \ sets lborel" "A \ {0<..}" for A proof (rule set_integrable_subset) show "set_integrable lborel {0<..} (g s)" using Gamma_times_hurwitz_zeta_integrable[of "1 - s" a] s a by (simp add: g_def exp_of_real exp_minus integrable_completion set_integrable_def) qed (insert that, auto) { fix r' :: real assume r': "r' \ {0<..<2}" from hurwitz_formula_integral_semiannulus(2)[of r' s 0] and r' have "f s contour_integrable_on part_circlepath 0 r' pi 0" by (auto simp: hankel_semiannulus_def add_ac) } note integrable_circle = this { fix r' :: real assume r': "r' \ {0<..<2}" from hurwitz_formula_integral_semiannulus(2)[of r' "cnj s" 0] and r' have "f (cnj s) contour_integrable_on part_circlepath 0 r' pi 0" by (auto simp: hankel_semiannulus_def add_ac) } note integrable_circle' = this have eq: "-2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r) = C" if r: "r \ {0<..<2}" for r :: real proof - have eq1: "integral {r..pi} (\x. cnj (x powr - cnj s) * (exp (- (a * x))) / (1 - (exp (- x)))) = integral {r..pi} (g s)" using r by (intro integral_cong) (auto simp: cnj_powr g_def exp_of_real exp_minus) have eq2: "integral {r..pi} (\x. cnj (x powr - cnj s) * (exp (a * x)) / (1 - (exp x))) = integral {r..pi} (\x. x powr - s * (exp (a * x)) / (1 - (exp x)))" using r by (intro integral_cong) (auto simp: cnj_powr) from hurwitz_formula_integral_semiannulus(1)[of r s 0] hurwitz_formula_integral_semiannulus(1)[of r "cnj s" 0] have "exp (-\*pi * s) * integral {r..real (2*0+1) * pi} (g s) + integral {r..real (2*0+1) * pi} (\x. x powr -s * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 (real (2 * 0 + 1) * pi) 0 pi) (f s) + contour_integral (part_circlepath 0 r pi 0) (f s) - cnj ( exp (-\*pi * cnj s) * integral {r..real (2*0+1) * pi} (\x. x powr - cnj s * exp (-a*x) / (1 - exp (-x))) + integral {r..real (2*0+1) * pi} (\x. x powr - cnj s * exp (a*x) / (1 - exp x)) + contour_integral (part_circlepath 0 (real (2 * 0 + 1) * pi) 0 pi) (f (cnj s)) + contour_integral (part_circlepath 0 r pi 0) (f (cnj s))) = 0" (is "?lhs = _") unfolding g_def using r by (subst (1 2) hurwitz_formula_integral_semiannulus) auto also have "?lhs = -2 * \ * sin (pi * s) * integral {r..pi} (g s) + err1 s r - err2 s r - C" using eq1 eq2 by (auto simp: integral_cnj exp_cnj err1_def err2_def sin_exp_eq algebra_simps C_def) also have "integral {r..pi} (g s) = (CLBINT x:{r..pi}. g s x)" using r by (intro set_borel_integral_eq_integral(2) [symmetric] integrable) auto finally show "-2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r) = C" by (simp add: algebra_simps) qed \\Next, compute the value of @{term C} by letting the radius tend to 0 so that the contribution of the circle vanishes.\ have "((\r. -2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r)) \ -2 * \ * sin (pi * s) * (CLBINT x:{0<..pi}. g s x) + 0) (at_right 0)" proof (intro tendsto_intros tendsto_set_lebesgue_integral_at_right integrable) from hurwitz_formula_bound2[of s] guess C1 . note C1 = this from hurwitz_formula_bound2[of "cnj s"] guess C2 . note C2 = this have ev: "eventually (\r::real. r \ {0<..<2}) (at_right 0)" by (intro eventually_at_right_real) auto show "((\r. err1 s r - err2 s r) \ 0) (at_right 0)" proof (rule Lim_null_comparison[OF eventually_mono[OF ev]]) fix r :: real assume r: "r \ {0<..<2}" have "norm (err1 s r - err2 s r) \ norm (err1 s r) + norm (err2 s r)" by (rule norm_triangle_ineq4) also have "norm (err1 s r) \ C1 * r powr (- Re s - 1) * r * \0 - pi\" unfolding err1_def using C1(1) C1(2)[of r] pi_gt3 integrable_circle[of r] path_image_part_circlepath_subset'[of r 0 pi 0] r by (intro contour_integral_bound_part_circlepath) auto also have "\ = C1 * r powr (-Re s) * pi" using r by (simp add: powr_diff field_simps) also have "norm (err2 s r) \ C2 * r powr (- Re s - 1) * r * \0 - pi\" unfolding err2_def complex_mod_cnj using C2(1) C2(2)[of r] r pi_gt3 integrable_circle'[of r] path_image_part_circlepath_subset'[of r 0 pi 0] by (intro contour_integral_bound_part_circlepath) auto also have "\ = C2 * r powr (-Re s) * pi" using r by (simp add: powr_diff field_simps) also have "C1 * r powr (-Re s) * pi + C2 * r powr (-Re s) * pi = (C1 + C2) * pi * r powr (-Re s)" by (simp add: algebra_simps) finally show "norm (err1 s r - err2 s r) \ (C1 + C2) * pi * r powr - Re s" by simp next show "((\x. (C1 + C2) * pi * x powr - Re s) \ 0) (at_right 0)" using s by (auto intro!: tendsto_eq_intros simp: eventually_at exI[of _ 1]) qed qed auto moreover have "eventually (\r::real. r \ {0<..<2}) (at_right 0)" by (intro eventually_at_right_real) auto hence "eventually (\r. -2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r) = C) (at_right 0)" by eventually_elim (use eq in auto) hence "((\r. -2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r)) \ C) (at_right 0)" by (rule tendsto_eventually) ultimately have [simp]: "C = -2 * \ * sin (pi * s) * (CLBINT x:{0<..pi}. g s x)" using tendsto_unique by force \\We now rearrange everything and obtain the result.\ have "2 * \ * sin (pi * s) * ((CLBINT x:{0<..pi}. g s x) - (CLBINT x:{r..pi}. g s x)) = err2 s r - err1 s r" using eq[of r] r by (simp add: algebra_simps) also have "{0<..pi} = {0<.. {r..pi}" using r pi_gt3 by auto also have "(CLBINT x:\. g s x) - (CLBINT x:{r..pi}. g s x) = (CLBINT x:{0<.. {r..}. g s x) - (CLBINT x:{r..}. g s x)" using r pi_gt3 by (subst set_integral_Un[OF _ integrable integrable]) auto also have "{0<.. {r..} = {0<..}" using r by auto also have "(CLBINT x:{0<..}. g s x) = Gamma (1 - s) * hurwitz_zeta a (1 - s)" using Gamma_times_hurwitz_zeta_integral[of "1 - s" a] s a by (simp add: g_def exp_of_real exp_minus integral_completion set_lebesgue_integral_def) finally have "2 * \ * (sin (pi * s) * Gamma (1 - s)) * hurwitz_zeta a (1 - s) = err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)" by (simp add: algebra_simps) also have "sin (pi * s) * Gamma (1 - s) = pi * rGamma s" proof (cases "s \ \") case False with Gamma_reflection_complex[of s] show ?thesis by (auto simp: divide_simps sin_eq_0 Ints_def rGamma_inverse_Gamma mult_ac split: if_splits) next case True with s have "rGamma s = 0" by (auto simp: rGamma_eq_zero_iff nonpos_Ints_def Ints_def) moreover from True have "sin (pi * s) = 0" by (subst sin_eq_0) (auto elim!: Ints_cases) ultimately show ?thesis by simp qed finally show "2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s) = err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)" by (simp add: mult_ac) next \\By analytic continuation, we lift the result to the case of any non-zero @{term s}.\ show "(\s. 2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s)) holomorphic_on - {0}" using a by (auto intro!: holomorphic_intros) show "(\s. err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)) holomorphic_on -{0}" proof (intro holomorphic_intros) have "(\s. err2 s r) = (\s. - cnj (integral {0..pi} (h r (cnj s))))" using r by (simp add: err2_def contour_integral_part_circlepath_reverse' contour_integral_part_circlepath_h) also have "(\s. - cnj (integral {0..pi} (h r (cnj s)))) = (\s. (integral {0..pi} (\x. h r s (-x))))" using r by (simp add: integral_cnj h_def exp_cnj cis_cnj Ln_Reals_eq) also have "\ = (\s. integral {-pi..0} (h r s))" - by (subst integral_reflect_real [symmetric]) simp + by (subst Henstock_Kurzweil_Integration.integral_reflect_real [symmetric]) simp finally have "(\s. err2 s r) = \" . moreover have "(\s. integral {-pi..0} (h r s)) holomorphic_on -{0}" using r by (intro integral_h_holomorphic) auto ultimately show "(\s. err2 s r) holomorphic_on -{0}" by simp next have "(\s. - integral {0..pi} (h r s)) holomorphic_on -{0}" using r by (intro holomorphic_intros integral_h_holomorphic) auto also have "(\s. - integral {0..pi} (h r s)) = (\s. err1 s r)" unfolding err1_def using r by (simp add: contour_integral_part_circlepath_reverse' contour_integral_part_circlepath_h) finally show "(\s. err1 s r) holomorphic_on -{0}" . next show "(\s. CLBINT x:{r..}. g s x) holomorphic_on -{0}" proof (rule holomorphic_on_balls_imp_entire') fix R :: real have "eventually (\b. b > r) at_top" by (rule eventually_gt_at_top) hence 1: "eventually (\b. continuous_on (cball 0 R) (\s. CLBINT x:{r..b}. g s x) \ (\s. CLBINT x:{r..b}. g s x) holomorphic_on ball 0 R) at_top" proof eventually_elim case (elim b) have integrable: "set_integrable lborel {r..b} (g s)" for s unfolding g_def using r by (intro borel_integrable_atLeastAtMost' continuous_intros) auto have "(\s. integral {r..b} (g s)) holomorphic_on UNIV" using r by (intro integral_g_holomorphic) auto also have "(\s. integral {r..b} (g s)) = (\s. CLBINT x:{r..b}. g s x)" by (intro ext set_borel_integral_eq_integral(2)[symmetric] integrable) finally have "\ holomorphic_on UNIV" . thus ?case by (auto intro!: holomorphic_on_imp_continuous_on) qed have 2: "uniform_limit (cball 0 R) (\b s. CLBINT x:{r..b}. g s x) (\s. CLBINT x:{r..}. g s x) at_top" proof (rule uniform_limit_set_lebesgue_integral_at_top) fix s :: complex and x :: real assume s: "s \ cball 0 R" and x: "x \ r" have "norm (g s x) = x powr -Re s * exp (-a * x) / (1 - exp (-x))" using x r by (simp add: g_def norm_mult norm_divide in_Reals_norm norm_powr_real_powr) also have "\ \ x powr R * exp (-a * x) / (1 - exp (-x))" using r s x abs_Re_le_cmod[of s] by (intro mult_right_mono divide_right_mono powr_mono) auto finally show "norm (g s x) \ x powr R * exp (- a * x) / (1 - exp (- x))" . next show "set_integrable lborel {r..} (\x. x powr R * exp (-a * x) / (1 - exp (-x)))" using r a by (intro set_integrable_Gamma_hurwitz_aux2_real) auto qed (simp_all add: set_borel_measurable_def g_def) show "(\s. CLBINT x:{r..}. g s x) holomorphic_on ball 0 R" using holomorphic_uniform_limit[OF 1 2] by auto qed qed qed (insert \s \ 0\, auto simp: connected_punctured_universe open_halfspace_Re_lt intro: exI[of _ "-1"]) text \ Finally, we obtain Hurwitz's formula by letting the radius of the outer circle tend to \\\. \ lemma hurwitz_zeta_formula_aux: fixes s :: complex assumes s: "Re s > 1" shows "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr -s * (\ powr (-s) * F a s + \ powr s * F (-a) s)" proof - from s have [simp]: "s \ 0" by auto define r where "r = (1 :: real)" have r: "r \ {0<..<2}" by (simp add: r_def) define R where "R = (\n. real (2 * n + 1) * pi)" define bigc where "bigc = (\n. contour_integral (part_circlepath 0 (R n) 0 pi) (f s) - cnj (contour_integral (part_circlepath 0 (R n) 0 pi) (f (cnj s))))" define smallc where "smallc = contour_integral (part_circlepath 0 r pi 0) (f s) - cnj (contour_integral (part_circlepath 0 r pi 0) (f (cnj s)))" define I where "I = (\n. CLBINT x:{r..R n}. g s x)" define F1 and F2 where "F1 = (\n. exp (-s * pi * \ / 2) * (\k\{0<..n}. exp (2 * real k * pi * a * \) * k powr (-s)))" "F2 = (\n. exp (s * pi * \ / 2) * (\k\{0<..n}. exp (2 * real k * pi * (-a) * \) * k powr (-s)))" have R: "R n \ pi" for n using r by (auto simp: R_def field_simps) have [simp]: "\(pi \ 0)" using pi_gt_zero by linarith have integrable: "set_integrable lborel A (g s)" if "A \ sets lborel" "A \ {r..}" for A proof - have "set_integrable lborel {r..} (g s)" using set_integrable_Gamma_hurwitz_aux2[of r a s] a r by (simp add: g_def exp_of_real exp_minus) thus ?thesis by (rule set_integrable_subset) (use that in auto) qed { fix n :: nat from hurwitz_formula_integral_semiannulus(2)[of "r" s n] and r R[of n] have "f s contour_integrable_on part_circlepath 0 (R n) 0 pi" by (auto simp: hankel_semiannulus_def R_def add_ac) } note integrable_circle = this { fix n :: nat from hurwitz_formula_integral_semiannulus(2)[of "r" "cnj s" n] and r R[of n] have "f (cnj s) contour_integrable_on part_circlepath 0 (R n) 0 pi" by (auto simp: hankel_semiannulus_def R_def add_ac) } note integrable_circle' = this { fix n :: nat have "(exp (-\ * pi * s) * integral {r..R n} (g s) + integral {r..R n} (\x. x powr (-s) * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 (R n) 0 pi) (f s) + contour_integral (part_circlepath 0 r pi 0) (f s)) - cnj ( exp (-\ * pi * cnj s) * integral {r..R n} (g (cnj s)) + integral {r..R n} (\x. x powr (-cnj s) * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 (R n) 0 pi) (f (cnj s)) + contour_integral (part_circlepath 0 r pi 0) (f (cnj s))) = -2 * pi * \ * exp (- s * of_real pi * \ / 2) * (\k\{0<..n}. Res s k) - cnj (-2 * pi * \ * exp (- cnj s * of_real pi * \ / 2) * (\k\{0<..n}. Res (cnj s) k))" (is "?lhs = ?rhs") unfolding R_def g_def using r by (subst (1 2) hurwitz_formula_integral_semiannulus) auto also have "?rhs = -2 * pi * \ * (exp (- s * pi * \ / 2) * (\k\{0<..n}. Res s k) + exp (s * pi * \ / 2) * (\k\{0<..n}. cnj (Res (cnj s) k)))" by (simp add: exp_cnj sum.distrib algebra_simps sum_distrib_left sum_distrib_right sum_negf) also have "(\k\{0<..n}. Res s k) = (2 * pi) powr (-s) * (\k\{0<..n}. exp (2 * k * pi * a * \) * k powr (-s))" (is "_ = ?S1") by (simp add: Res_def powr_times_real algebra_simps sum_distrib_left) also have "(\k\{0<..n}. cnj (Res (cnj s) k)) = (2 * pi) powr (-s) * (\k\{0<..n}. exp (-2 * k * pi * a * \) * k powr (-s))" by (simp add: Res_def cnj_powr powr_times_real algebra_simps exp_cnj sum_distrib_left) also have "exp (-s * pi * \ / 2) * ?S1 + exp (s * pi * \ / 2) * \ = (2 * pi) powr (-s) * (exp (-s * pi * \ / 2) * (\k\{0<..n}. exp (2 * k * pi * a * \) * k powr (-s)) + exp (s * pi * \ / 2) * (\k\{0<..n}. exp (-2 * k * pi * a * \) * k powr (-s)))" by (simp add: algebra_simps) also have 1: "integral {r..R n} (g s) = I n" unfolding I_def by (intro set_borel_integral_eq_integral(2) [symmetric] integrable) auto have 2: "cnj (integral {r..R n} (g (cnj s))) = integral {r..R n} (g s)" using r unfolding integral_cnj by (intro integral_cong) (auto simp: g_def cnj_powr) have 3: "integral {r..R n} (\x. exp (x * a) * cnj (x powr - cnj s) / (1 - exp x)) = integral {r..R n} (\x. exp (x * a) * of_real x powr - s / (1 - exp x))" unfolding I_def g_def using r R[of n] by (intro integral_cong; force simp: cnj_powr)+ from 1 2 3 have "?lhs = (exp (-\ * s * pi) - exp (\ * s * pi)) * I n + bigc n + smallc" by (simp add: integral_cnj cnj_powr algebra_simps exp_cnj bigc_def smallc_def g_def) also have "exp (-\ * s * pi) - exp (\ * s * pi) = -2 * \ * sin (s * pi)" by (simp add: sin_exp_eq' algebra_simps) finally have "(- 2 * \ * sin (s * pi) * I n + smallc) + bigc n = -2 * \ * pi * (2 * pi) powr (-s) * (F1 n + F2 n)" by (simp add: F1_F2_def algebra_simps) } note eq = this have "(\n. - 2 * \ * sin (s * pi) * I n + smallc + bigc n) \ (-2 * \ * sin (s * pi)) * (CLBINT x:{r..}. g s x) + smallc + 0" unfolding I_def proof (intro tendsto_intros filterlim_compose[OF tendsto_set_lebesgue_integral_at_top] integrable) show "filterlim R at_top sequentially" unfolding R_def by (intro filterlim_at_top_mult_tendsto_pos[OF tendsto_const] pi_gt_zero filterlim_compose[OF filterlim_real_sequentially] filterlim_subseq) (auto simp: strict_mono_Suc_iff) from hurwitz_formula_bound1[OF pi_gt_zero] guess C . note C = this define D where "D = C * exp (3 / 2 * pi * \Im s\)" from \C \ 0\ have "D \ 0" by (simp add: D_def) show "bigc \ 0" proof (rule Lim_null_comparison[OF always_eventually[OF allI]]) fix n :: nat have bound: "norm (f s' z) \ D * R n powr (-Re s')" if z: "z \ sphere 0 (R n)" "Re s' = Re s" "\Im s'\ = \Im s\" for z s' proof - from z and r R[of n] have [simp]: "z \ 0" by auto have not_in_ball: "z \ ball (2 * m * pi * \) pi" for m :: int proof - have "dist z (2 * m * pi * \) \ \dist z 0 - dist 0 (2 * m * pi * \)\" by (rule abs_dist_diff_le) also have "dist 0 (2 * m * pi * \) = 2 * \m\ * pi" by (simp add: norm_mult) also from z have "dist z 0 = R n" by simp also have "R n - 2 * \m\ * pi = (int (2 * n + 1) - 2 * \m\) * pi" by (simp add: R_def algebra_simps) also have "\\\ = \int (2 * n + 1) - 2 * \m\\ * pi" by (subst abs_mult) simp_all also have "\int (2 * n + 1) - 2 * \m\\ \ 1" by presburger hence "\ * pi \ 1 * pi" by (intro mult_right_mono) auto finally show ?thesis by (simp add: dist_commute) qed have "norm (f s' z) = norm (exp (-Ln' z * s')) * norm (exp (a * z) / (1 - exp z))" by (simp add: f_def exp_diff norm_mult norm_divide mult_ac exp_minus norm_inverse divide_simps del: norm_exp_eq_Re) also have "\ \ norm (exp (-Ln' z * s')) * C" using not_in_ball by (intro mult_left_mono C) auto also have "norm (exp (-Ln' z * s')) = exp (Im s' * (Im (Ln (- (\ * z))) + pi / 2)) / exp (Re s' * ln (R n))" using z r R[of n] pi_gt_zero by (simp add: Ln'_def norm_mult norm_divide exp_add exp_diff exp_minus norm_inverse algebra_simps inverse_eq_divide) also have "\ \ exp (3/2 * pi * \Im s'\) / exp (Re s' * ln (R n))" proof (intro divide_right_mono, subst exp_le_cancel_iff) have "Im s' * (Im (Ln (- (\ * z))) + pi / 2) \ \Im s' * (Im (Ln (- (\ * z))) + pi / 2)\" by (rule abs_ge_self) also have "\ \ \Im s'\ * (pi + pi / 2)" unfolding abs_mult using mpi_less_Im_Ln[of "- (\ * z)"] Im_Ln_le_pi[of "- (\ * z)"] by (intro mult_left_mono order.trans[OF abs_triangle_ineq] add_mono) auto finally show "Im s' * (Im (Ln (- (\ * z))) + pi / 2) \ 3/2 * pi * \Im s'\" by (simp add: algebra_simps) qed auto also have "exp (Re s' * ln (R n)) = R n powr Re s'" using r R[of n] by (auto simp: powr_def) finally show "norm (f s' z) \ D * R n powr (-Re s')" using \C \ 0\ by (simp add: that D_def powr_minus mult_right_mono mult_left_mono field_simps) qed have "norm (bigc n) \ norm (contour_integral (part_circlepath 0 (R n) 0 pi) (f s)) + norm (cnj (contour_integral (part_circlepath 0 (R n) 0 pi) (f (cnj s))))" (is "_ \ norm ?err1 + norm ?err2") unfolding bigc_def by (rule norm_triangle_ineq4) also have "norm ?err1 \ D * R n powr (-Re s) * R n * \pi - 0\" using \D \ 0\ and r R[of n] and pi_gt3 and integrable_circle and path_image_part_circlepath_subset[of 0 pi "R n" 0] and bound[of _ s] by (intro contour_integral_bound_part_circlepath) auto also have "\ = D * pi * R n powr (1 - Re s)" using r R[of n] pi_gt3 by (simp add: powr_diff field_simps powr_minus) also have "norm ?err2 \ D * R n powr (-Re s) * R n * \pi - 0\" unfolding complex_mod_cnj using \D \ 0\ and r R[of n] and pi_gt3 and integrable_circle'[of n] and path_image_part_circlepath_subset[of 0 pi "R n" 0] and bound[of _ "cnj s"] by (intro contour_integral_bound_part_circlepath) auto also have "\ = D * pi * R n powr (1 - Re s)" using r R[of n] pi_gt3 by (simp add: powr_diff field_simps powr_minus) finally show "norm (bigc n) \ 2 * D * pi * R n powr (1 - Re s)" by simp next have "filterlim R at_top at_top" by fact hence "(\x. 2 * D * pi * R x powr (1 - Re s)) \ 2 * D * pi * 0" using s unfolding R_def by (intro tendsto_intros tendsto_neg_powr) auto thus "(\x. 2 * D * pi * R x powr (1 - Re s)) \ 0" by simp qed qed auto also have "(\n. - 2 * \ * sin (s * pi) * I n + smallc + bigc n) = (\n. -2 * \ * pi * (2 * pi) powr -s * (F1 n + F2 n))" by (subst eq) auto finally have "\ \ (-2 * \ * sin (s * pi)) * (CLBINT x:{r..}. g s x) + smallc" by simp moreover have "(\n. -2 * \ * pi * (2 * pi) powr -s * (F1 n + F2 n)) \ -2 * \ * pi * (2 * pi) powr -s * (exp (-s * pi * \ / 2) * F a s + exp (s * pi * \ / 2) * F (-a) s)" unfolding F1_F2_def F_def using s by (intro tendsto_intros sum_tendsto_fds_perzeta) ultimately have "-2 * \ * pi * (2 * pi) powr -s * (exp (-s * pi * \ / 2) * F a s + exp (s * pi * \ / 2) * F (-a) s) = (-2 * \ * sin (s * pi)) * (CLBINT x:{r..}. g s x) + smallc" by (force intro: tendsto_unique) also have "\ = -2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s)" using s r a using rGamma_hurwitz_zeta_eq_contour_integral[of s r] by (simp add: r_def smallc_def algebra_simps) also have "exp (- s * complex_of_real pi * \ / 2) = \ powr (-s)" by (simp add: powr_def field_simps) also have "exp (s * complex_of_real pi * \ / 2) = \ powr s" by (simp add: powr_def field_simps) finally show "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr -s * (\ powr (-s) * F a s + \ powr s * F (-a) s)" by simp qed end text \ We can now use Hurwitz's formula to prove the following nice formula that expresses the periodic zeta function in terms of the Hurwitz zeta function: \[F(s, a) = (2\pi)^{s-1} i \Gamma(1 - s) \left(i^{-s} \zeta(1 - s, a) - i^{s} \zeta(1 - s, 1 - a)\right)\] This holds for all \s\ with \\mathfrak{R}(s) > 0\ as long as \a \ \\. For convenience, we move the \\\ function to the left-hand side in order to avoid having to account for its poles. \ lemma perzeta_conv_hurwitz_zeta_aux: fixes a :: real and s :: complex assumes a: "a \ {0<..<1}" and s: "Re s > 0" shows "rGamma (1 - s) * eval_fds (fds_perzeta a) s = (2 * pi) powr (s - 1) * \ * (\ powr -s * hurwitz_zeta a (1 - s) - \ powr s * hurwitz_zeta (1 - a) (1 - s))" (is "?lhs s = ?rhs s") proof (rule analytic_continuation_open[where f = ?lhs]) show "connected {s. Re s > 0}" by (intro convex_connected convex_halfspace_Re_gt) show "{s. Re s > 1} \ {}" by (auto intro: exI[of _ 2]) show "(\s. rGamma (1 - s) * eval_fds (fds_perzeta a) s) holomorphic_on {s. 0 < Re s}" unfolding perzeta_def using a by (auto intro!: holomorphic_intros le_less_trans[OF conv_abscissa_perzeta'] elim!: Ints_cases) show "?rhs holomorphic_on {s. 0 < Re s}" using assms by (auto intro!: holomorphic_intros) next fix s assume s: "s \ {s. Re s > 1}" have [simp]: "fds_perzeta (1 - a) = fds_perzeta (-a)" using fds_perzeta.plus_of_nat[of "-a" 1] by simp have [simp]: "fds_perzeta (a - 1) = fds_perzeta a" using fds_perzeta.minus_of_nat[of a 1] by simp from s have [simp]: "Gamma s \ 0" by (auto simp: Gamma_eq_zero_iff elim!: nonpos_Ints_cases) have "(2 * pi) powr (-s) * (\ * (\ powr (-s) * (rGamma s * hurwitz_zeta a (1 - s)) - \ powr s * (rGamma s * hurwitz_zeta (1 - a) (1 - s)))) = (2 * pi) powr (-s) * ((\ powr (1 - s) * (rGamma s * hurwitz_zeta a (1 - s)) + \ powr (s - 1) * (rGamma s * hurwitz_zeta (1 - a) (1 - s))))" by (simp add: powr_diff field_simps powr_minus) also have "\ = ((2 * pi) powr (-s)) ^ 2 * ( eval_fds (fds_perzeta a) s * (\ powr s * \ powr (s - 1) + \ powr (-s) * \ powr (1 - s)) + eval_fds (fds_perzeta (-a)) s * (\ powr s * \ powr (1 - s) + \ powr (-s) * \ powr (s - 1)))" using s a by (subst (1 2) hurwitz_zeta_formula_aux) (auto simp: algebra_simps power2_eq_square) also have "(\ powr s * \ powr (1 - s) + \ powr (-s) * \ powr (s - 1)) = exp (\ * complex_of_real pi / 2) + exp (- (\ * complex_of_real pi / 2))" by (simp add: powr_def exp_add [symmetric] field_simps) also have "\ = 0" by (simp add: exp_eq_polar) also have "\ powr s * \ powr (s - 1) = \ powr (2 * s - 1)" by (simp add: powr_def exp_add [symmetric] field_simps) also have "\ powr (-s) * \ powr (1 - s) = \ powr (1 - 2 * s)" by (simp add: powr_def exp_add [symmetric] field_simps) also have "\ powr (2 * s - 1) + \ powr (1 - 2 * s) = 2 * cos ((2 * s - 1) * pi / 2)" by (simp add: powr_def cos_exp_eq algebra_simps minus_divide_left cos_sin_eq) also have "\ = 2 * sin (pi - s * pi)" by (simp add: cos_sin_eq field_simps) also have "\ = 2 * sin (s * pi)" by (simp add: sin_diff) finally have "\ * (rGamma s * \ powr (-s) * hurwitz_zeta a (1 - s) - rGamma s * \ powr s * hurwitz_zeta (1 - a) (1 - s)) = 2 * (2 * pi) powr -s * sin (s * pi) * eval_fds (fds_perzeta a) s" by (simp add: power2_eq_square mult_ac) hence "(2 * pi) powr s / 2 * \ * (\ powr (-s) * hurwitz_zeta a (1 - s) - \ powr s * hurwitz_zeta (1 - a) (1 - s)) = Gamma s * sin (s * pi) * eval_fds (fds_perzeta a) s" by (subst (asm) (2) powr_minus) (simp add: field_simps rGamma_inverse_Gamma) also have "Gamma s * sin (s * pi) = pi * rGamma (1 - s)" using Gamma_reflection_complex[of s] by (auto simp: divide_simps rGamma_inverse_Gamma mult_ac split: if_splits) finally show "?lhs s = ?rhs s" by (simp add: powr_diff) qed (insert s, auto simp: open_halfspace_Re_gt) text \ We can now use the above equation as a defining equation to continue the periodic zeta function $F$ to the entire complex plane except at non-negative integer values for \s\. However, the positive integers are already covered by the original Dirichlet series definition of $F$, so we only need to take care of \s = 0\. We do this by cancelling the pole of \\\ at \0\ with the zero of $i^{-s} \zeta(1 - s, a) - i^s \zeta(1 - s, 1 - a)$. \ lemma assumes "q' \ \" shows holomorphic_perzeta': "perzeta q' holomorphic_on A" and perzeta_altdef2: "Re s > 0 \ perzeta q' s = eval_fds (fds_perzeta q') s" proof - define q where "q = frac q'" from assms have q: "q \ {0<..<1}" by (auto simp: q_def frac_lt_1) hence [simp]: "q \ \" by (auto elim!: Ints_cases) have [simp]: "frac q = q" by (simp add: q_def frac_def) define f where "f = (\s. complex_of_real (2 * pi) powr (s - 1) * \ * Gamma (1 - s) * (\ powr (-s) * hurwitz_zeta q (1 - s) - \ powr s * hurwitz_zeta (1 - q) (1 - s)))" { fix s :: complex assume "1 - s \ \\<^sub>\\<^sub>0" then obtain n where "1 - s = of_int n" "n \ 0" by (auto elim!: nonpos_Ints_cases) hence "s = 1 - of_int n" by (simp add: algebra_simps) also have "\ \ \" using \n \ 0\ by (auto simp: Nats_altdef1 intro: exI[of _ "1 - n"]) finally have "s \ \" . } note * = this hence "f holomorphic_on -\" using q by (auto simp: f_def Nats_altdef2 nonpos_Ints_altdef not_le intro!: holomorphic_intros) also have "?this \ perzeta q holomorphic_on -\" using assms by (intro holomorphic_cong refl) (auto simp: perzeta_def Let_def f_def) finally have holo: "perzeta q holomorphic_on -\" . have f_altdef: "f s = eval_fds (fds_perzeta q) s" if "Re s > 0" and "s \ \" for s using perzeta_conv_hurwitz_zeta_aux[OF q, of s] that * by (auto simp: rGamma_inverse_Gamma Gamma_eq_zero_iff divide_simps f_def perzeta_def split: if_splits) show "perzeta q' s = eval_fds (fds_perzeta q') s" if "Re s > 0" for s using f_altdef[of s] that assms by (auto simp: f_def perzeta_def Let_def q_def) have cont: "isCont (perzeta q) s" if "s \ \" for s proof (cases "s = 0") case False with that obtain n where [simp]: "s = of_nat n" and n: "n > 0" by (auto elim!: Nats_cases) have *: "open ({s. Re s > 0} - (\ - {of_nat n}))" using Nats_subset_Ints by (intro open_Diff closed_subset_Ints open_halfspace_Re_gt) auto have "eventually (\s. s \ {s. Re s > 0} - (\ - {of_nat n})) (nhds (of_nat n))" using \n > 0\ by (intro eventually_nhds_in_open *) auto hence ev: "eventually (\s. eval_fds (fds_perzeta q) s = perzeta q s) (nhds (of_nat n))" proof eventually_elim case (elim s) thus ?case using q f_altdef[of s] by (auto simp: perzeta_def dist_of_nat f_def elim!: Nats_cases Ints_cases) qed have "isCont (eval_fds (fds_perzeta q)) (of_nat n)" using q and \n > 0\ by (intro continuous_eval_fds le_less_trans[OF conv_abscissa_perzeta']) (auto elim!: Ints_cases) also have "?this \ isCont (perzeta q) (of_nat n)" using ev by (intro isCont_cong ev) finally show ?thesis by simp next assume [simp]: "s = 0" define a where "a = Complex (ln q) (-pi / 2)" define b where "b = Complex (ln (1 - q)) (pi / 2)" have "eventually (\s::complex. s \ \) (at 0)" unfolding eventually_at_topological using Nats_subset_Ints by (intro exI[of _ "-(\-{0})"] conjI open_Compl closed_subset_Ints) auto hence ev: "eventually (\s. perzeta q s = (2 * pi) powr (s - 1) * Gamma (1 - s) * \ * (\ powr - s * pre_zeta q (1 - s) -\ powr s * pre_zeta (1 - q) (1 - s) + (exp (b * s) - exp (a * s)) / s)) (at (0::complex))" (is "eventually (\s. _ = ?f s) _") proof eventually_elim case (elim s) have "perzeta q s = (2 * pi) powr (s - 1) * Gamma (1 - s) * \ * (\ powr (-s) * hurwitz_zeta q (1 - s) - \ powr s * hurwitz_zeta (1 - q) (1 - s))" (is "_ = _ * ?T") using elim by (auto simp: perzeta_def powr_diff powr_minus field_simps) also have "?T = \ powr (-s) * pre_zeta q (1 - s) - \ powr s * pre_zeta (1 - q) (1 - s) + (\ powr s * (1 - q) powr s - \ powr (-s) * q powr s) / s" using elim by (auto simp: hurwitz_zeta_def field_simps) also have "\ powr s * (1 - q) powr s = exp (b * s)" using q by (simp add: powr_def exp_add algebra_simps Ln_Reals_eq Complex_eq b_def) also have "\ powr (-s) * q powr s = exp (a * s)" using q by (simp add: powr_def exp_add Ln_Reals_eq exp_diff exp_minus diff_divide_distrib ring_distribs inverse_eq_divide mult_ac Complex_eq a_def) finally show ?case . qed have [simp]: "\(pi \ 0)" using pi_gt_zero by (simp add: not_le) have "(\s::complex. if s = 0 then b - a else (exp (b * s) - exp (a * s)) / s) has_fps_expansion (fps_exp b - fps_exp a) / fps_X" (is "?f' has_fps_expansion _") by (rule fps_expansion_intros)+ (auto intro!: subdegree_geI simp: Ln_Reals_eq a_def b_def) hence "isCont ?f' 0" by (rule has_fps_expansion_imp_continuous) hence "?f' \0\ b - a" by (simp add: isCont_def) also have "?this \ (\s. (exp (b * s) - exp (a * s)) / s) \0\ b - a" by (intro filterlim_cong refl) (auto simp: eventually_at intro: exI[of _ 1]) finally have "?f \0\ of_real (2 * pi) powr (0 - 1) * Gamma (1 - 0) * \ * (\ powr -0 * pre_zeta q (1 - 0) -\ powr 0 * pre_zeta (1 - q) (1 - 0) + (b - a))" (is "filterlim _ (nhds ?c) _") using q by (intro tendsto_intros isContD) (auto simp: complex_nonpos_Reals_iff intro!: continuous_intros) also have "?c = perzeta q 0" using q by (simp add: powr_minus perzeta_def Ln_Reals_eq a_def b_def Complex_eq mult_ac inverse_eq_divide) also have "?f \0\ \ \ perzeta q \0\ \" by (rule sym, intro filterlim_cong refl ev) finally show "isCont (perzeta q) s" by (simp add: isCont_def) qed have "perzeta q field_differentiable at s" for s proof (cases "s \ \") case False with holo have "perzeta q field_differentiable at s within -\" unfolding holomorphic_on_def by blast also have "at s within -\ = at s" using False by (intro at_within_open) auto finally show ?thesis . next case True hence *: "perzeta q holomorphic_on (ball s 1 - {s})" by (intro holomorphic_on_subset[OF holo]) (auto elim!: Nats_cases simp: dist_of_nat) have "perzeta q holomorphic_on ball s 1" using cont True by (intro no_isolated_singularity'[OF _ *]) (auto simp: at_within_open[of _ "ball s 1"] isCont_def) hence "perzeta q field_differentiable at s within ball s 1" unfolding holomorphic_on_def by auto thus ?thesis by (simp add: at_within_open[of _ "ball s 1"]) qed hence "perzeta q holomorphic_on UNIV" by (auto simp: holomorphic_on_def) also have "perzeta q = perzeta q'" by (simp add: q_def) finally show "perzeta q' holomorphic_on A" by auto qed lemma perzeta_altdef1: "Re s > 1 \ perzeta q' s = eval_fds (fds_perzeta q') s" by (cases "q' \ \") (auto simp: perzeta_int eval_fds_zeta fds_perzeta_int perzeta_altdef2) lemma holomorphic_perzeta: "q \ \ \ 1 \ A \ perzeta q holomorphic_on A" by (cases "q \ \") (auto simp: perzeta_int intro: holomorphic_perzeta' holomorphic_zeta) lemma holomorphic_perzeta'' [holomorphic_intros]: assumes "f holomorphic_on A" and "q \ \ \ (\x\A. f x \ 1)" shows "(\x. perzeta q (f x)) holomorphic_on A" proof - have "perzeta q \ f holomorphic_on A" using assms by (intro holomorphic_on_compose holomorphic_perzeta) auto thus ?thesis by (simp add: o_def) qed text \ Using this analytic continuation of the periodic zeta function, Hurwitz's formula now holds (almost) on the entire complex plane. \ theorem hurwitz_zeta_formula: fixes a :: real and s :: complex assumes "a \ {0<..1}" and "s \ 0" and "a \ 1 \ s \ 1" shows "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr - s * (\ powr - s * perzeta a s + \ powr s * perzeta (-a) s)" (is "?f s = ?g s") proof - define A where "A = UNIV - (if a \ \ then {0, 1} else {0 :: complex})" show ?thesis proof (rule analytic_continuation_open[where f = ?f]) show "?f holomorphic_on A" using assms by (auto intro!: holomorphic_intros simp: A_def) show "?g holomorphic_on A" using assms by (auto intro!: holomorphic_intros simp: A_def minus_in_Ints_iff) next fix s assume "s \ {s. Re s > 1}" thus "?f s = ?g s" using hurwitz_zeta_formula_aux[of a s] assms by (simp add: perzeta_altdef1) qed (insert assms, auto simp: open_halfspace_Re_gt A_def elim!: Ints_cases intro: connected_open_delete_finite exI[of _ 2]) qed text \ The equation expressing the periodic zeta function in terms of the Hurwitz zeta function can be extened similarly. \ theorem perzeta_conv_hurwitz_zeta: fixes a :: real and s :: complex assumes "a \ {0<..<1}" and "s \ 0" shows "rGamma (1 - s) * perzeta a s = (2 * pi) powr (s - 1) * \ * (\ powr (-s) * hurwitz_zeta a (1 - s) - \ powr s * hurwitz_zeta (1 - a) (1 - s))" (is "?f s = ?g s") proof (rule analytic_continuation_open[where f = ?f]) show "?f holomorphic_on -{0}" using assms by (auto intro!: holomorphic_intros elim: Ints_cases) show "?g holomorphic_on -{0}" using assms by (auto intro!: holomorphic_intros) next fix s assume "s \ {s. Re s > 1}" thus "?f s = ?g s" using perzeta_conv_hurwitz_zeta_aux[of a s] assms by (simp add: perzeta_altdef1) qed (insert assms, auto simp: open_halfspace_Re_gt connected_punctured_universe intro: exI[of _ 2]) text \ As a simple corollary, we derive the reflection formula for the Riemann zeta function: \ corollary zeta_reflect: fixes s :: complex assumes "s \ 0" "s \ 1" shows "rGamma s * zeta (1 - s) = 2 * (2 * pi) powr -s * cos (s * pi / 2) * zeta s" using hurwitz_zeta_formula[of 1 s] assms by (simp add: zeta_def cos_exp_eq powr_def perzeta_int algebra_simps) corollary zeta_reflect': fixes s :: complex assumes "s \ 0" "s \ 1" shows "rGamma (1 - s) * zeta s = 2 * (2 * pi) powr (s - 1) * sin (s * pi / 2) * zeta (1 - s)" using zeta_reflect[of "1 - s"] assms by (simp add: cos_sin_eq field_simps) text \ It is now easy to see that all the non-trivial zeroes of the Riemann zeta function must lie the critical strip $(0;1)$, and they must be symmetric around the $\mathfrak{R}(z) = \frac{1}{2}$ line. \ corollary zeta_zeroD: assumes "zeta s = 0" "s \ 1" shows "Re s \ {0<..<1} \ (\n::nat. n > 0 \ even n \ s = -real n)" proof (cases "Re s \ 0") case False with zeta_Re_ge_1_nonzero[of s] assms have "Re s < 1" by (cases "Re s < 1") auto with False show ?thesis by simp next case True { assume *: "\n. n > 0 \ even n \ s \ -real n" have "s \ of_int n" for n :: int proof assume [simp]: "s = of_int n" show False proof (cases n "0::int" rule: linorder_cases) assume "n < 0" show False proof (cases "even n") case True hence "nat (-n) > 0" "even (nat (-n))" using \n < 0\ by (auto simp: even_nat_iff) with * have "s \ -real (nat (-n))" . with \n < 0\ and True show False by auto next case False with \n < 0\ have "of_int n = (-of_nat (nat (-n)) :: complex)" by simp also have "zeta \ = -(bernoulli' (Suc (nat (-n)))) / of_nat (Suc (nat (-n)))" using \n < 0\ by (subst zeta_neg_of_nat) (auto) finally have "bernoulli' (Suc (nat (-n))) = 0" using assms by (auto simp del: of_nat_Suc) with False and \n < 0\ show False by (auto simp: bernoulli'_zero_iff even_nat_iff) qed qed (insert assms True, auto) qed hence "rGamma s \ 0" by (auto simp: rGamma_eq_zero_iff nonpos_Ints_def) moreover from assms have [simp]: "s \ 0" by auto ultimately have "zeta (1 - s) = 0" using zeta_reflect[of s] and assms by auto with True zeta_Re_ge_1_nonzero[of "1 - s"] have "Re s > 0" by auto } with True show ?thesis by auto qed lemma zeta_zero_reflect: assumes "Re s \ {0<..<1}" and "zeta s = 0" shows "zeta (1 - s) = 0" proof - from assms have "rGamma s \ 0" by (auto simp: rGamma_eq_zero_iff elim!: nonpos_Ints_cases) moreover from assms have "s \ 0" and "s \ 1" by auto ultimately show ?thesis using zeta_reflect[of s] and assms by auto qed corollary zeta_zero_reflect_iff: assumes "Re s \ {0<..<1}" shows "zeta (1 - s) = 0 \ zeta s = 0" using zeta_zero_reflect[of s] zeta_zero_reflect[of "1 - s"] assms by auto subsection \More functional equations\ lemma perzeta_conv_hurwitz_zeta_multiplication: fixes k :: nat and a :: int and s :: complex assumes "k > 0" "s \ 1" shows "k powr s * perzeta (a / k) s = (\n=1..k. exp (2 * pi * n * a / k * \) * hurwitz_zeta (n / k) s)" (is "?lhs s = ?rhs s") proof (rule analytic_continuation_open[where ?f = ?lhs and ?g = ?rhs]) show "connected (-{1::complex})" by (rule connected_punctured_universe) auto show "{s. Re s > 1} \ {}" by (auto intro!: exI[of _ 2]) next fix s assume s: "s \ {s. Re s > 1}" let ?f = "\n. exp (2 * pi * n * a / k * \)" show "?lhs s = ?rhs s" proof (rule sums_unique2) have "(\m. \n=1..k. ?f n * (of_nat m + of_real (real n / real k)) powr -s) sums (\n=1..k. ?f n * hurwitz_zeta (real n / real k) s)" using assms s by (intro sums_sum sums_mult sums_hurwitz_zeta) auto also have "(\m. \n=1..k. ?f n * (of_nat m + of_real (real n / real k)) powr -s) = (\m. of_nat k powr s * (\n=1..k. ?f n * of_nat (m * k + n) powr -s))" unfolding sum_distrib_left proof (intro ext sum.cong, goal_cases) case (2 m n) hence "m * k + n > 0" by (intro add_nonneg_pos) auto hence "of_nat 0 \ (of_nat (m * k + n) :: complex)" by (simp only: of_nat_eq_iff) also have "of_nat (m * k + n) = of_nat m * of_nat k + (of_nat n :: complex)" by simp finally have nz: "\ \ 0" by auto have "of_nat m + of_real (real n / real k) = (inverse (of_nat k) * of_nat (m * k + n) :: complex)" using assms (* TODO: Field_as_Ring messing up things again *) by (simp add: field_simps del: div_mult_self1 div_mult_self2 div_mult_self3 div_mult_self4) also from nz have "\ powr -s = of_nat k powr s * of_nat (m * k + n) powr -s" by (subst powr_times_real) (auto simp: add_eq_0_iff powr_def exp_minus Ln_inverse) finally show ?case by simp qed auto finally show "\ sums (\n=1..k. ?f n * hurwitz_zeta (real n / real k) s)" . next define g where "g = (\m. exp (2 * pi * \ * m * (real_of_int a / real k)))" have "(\m. g (Suc m) / (Suc m) powr s) sums eval_fds (fds_perzeta (a / k)) s" unfolding g_def using s by (intro sums_fds_perzeta) auto also have "(\m. g (Suc m) / (Suc m) powr s) = (\m. ?f (Suc m) * (Suc m) powr -s)" by (simp add: powr_minus field_simps g_def) also have "eval_fds (fds_perzeta (a / k)) s = perzeta (a / k) s" using s by (simp add: perzeta_altdef1) finally have "(\m. \n=m*k..k > 0\ by (rule sums_group) also have "(\m. \n=m*k..m. \n=1..k. ?f (m * k + n) * of_nat (m * k + n) powr -s)" proof (rule ext, goal_cases) case (1 m) show ?case using assms by (intro ext sum.reindex_bij_witness[of _ "\n. m * k + n - 1" "\n. Suc n - m * k"]) auto qed also have "(\m n. ?f (m * k + n)) = (\m n. ?f n)" proof (intro ext) fix m n :: nat have "?f (m * k + n) / ?f n = exp (2 * pi * m * a * \)" using \k > 0\ by (auto simp: ring_distribs add_divide_distrib exp_add mult_ac) also have "\ = cis (2 * pi * (m * a))" by (simp add: exp_eq_polar mult_ac) also have "\ = 1" by (rule cis_multiple_2pi) auto finally show "?f (m * k + n) = ?f n" by simp qed finally show "(\m. of_nat k powr s * (\n=1..k. ?f n * of_nat (m * k + n) powr -s)) sums (of_nat k powr s * perzeta (a / k) s)" by (rule sums_mult) qed qed (use assms in \auto intro!: holomorphic_intros simp: finite_imp_closed open_halfspace_Re_gt\) lemma perzeta_conv_hurwitz_zeta_multiplication': fixes k :: nat and a :: int and s :: complex assumes "k > 0" "s \ 1" shows "perzeta (a / k) s = k powr -s * (\n=1..k. exp (2 * pi * n * a / k * \) * hurwitz_zeta (n / k) s)" using perzeta_conv_hurwitz_zeta_multiplication[of k s a] assms by (simp add: powr_minus field_simps) lemma zeta_conv_hurwitz_zeta_multiplication: fixes k a :: nat and s :: complex assumes "k > 0" "s \ 1" shows "k powr s * zeta s = (\n=1..k. hurwitz_zeta (n / k) s)" using perzeta_conv_hurwitz_zeta_multiplication[of k s 0] using assms by (simp add: perzeta_int) lemma hurwitz_zeta_one_half_left: assumes "s \ 1" shows "hurwitz_zeta (1 / 2) s = (2 powr s - 1) * zeta s" using zeta_conv_hurwitz_zeta_multiplication[of 2 s] assms by (simp add: eval_nat_numeral zeta_def field_simps) theorem hurwitz_zeta_functional_equation: fixes h k :: nat and s :: complex assumes hk: "k > 0" "h \ {0<..k}" and s: "s \ {0, 1}" defines "a \ real h / real k" shows "rGamma s * hurwitz_zeta a (1 - s) = 2 * (2 * pi * k) powr -s * (\n=1..k. cos (s*pi/2 - 2*pi*n*h/k) * hurwitz_zeta (n / k) s)" proof - from hk have a: "a \ {0<..1}" by (auto simp: a_def) have "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr - s * (\ powr - s * perzeta a s + \ powr s * perzeta (- a) s)" using s a by (intro hurwitz_zeta_formula) auto also have "\ = (2 * pi) powr - s * (\ powr - s * perzeta (of_int (int h) / k) s + \ powr s * perzeta (of_int (-int h) / k) s)" by (simp add: a_def) also have "\ = (2 * pi) powr -s * k powr -s * ((\n=1..k. \ powr -s * cis (2 * pi * n * h / k) * hurwitz_zeta (n / k) s) + (\n=1..k. \ powr s * cis (-2 * pi * n * h / k) * hurwitz_zeta (n / k) s))" (is "_ = _ * (?S1 + ?S2)") using hk a s by (subst (1 2) perzeta_conv_hurwitz_zeta_multiplication') (auto simp: field_simps sum_distrib_left sum_distrib_right exp_eq_polar) also have "(2 * pi) powr -s * k powr -s = (2 * k * pi) powr -s" using hk pi_gt_zero by (simp add: powr_def Ln_times_Reals field_simps exp_add exp_diff exp_minus) also have "?S1 + ?S2 = (\n=1..k. (\ powr -s * cis (2*pi*n*h/k) + \ powr s * cis (-2*pi*n*h/k)) * hurwitz_zeta (n / k) s)" (is "_ = (\n\_. ?c n * _)") by (simp add: algebra_simps sum.distrib) also have "?c = (\n. 2 * cos (s*pi/2 - 2*pi*n*h/k))" proof fix n :: nat have "\ powr -s * cis (2*pi*n*h/k) = exp (-s*pi/2*\ + 2*pi*n*h/k*\)" unfolding exp_add by (simp add: powr_def cis_conv_exp mult_ac) moreover have "\ powr s * cis (-2*pi*n*h/k) = exp (s*pi/2*\ + -2*pi*n*h/k*\)" unfolding exp_add by (simp add: powr_def cis_conv_exp mult_ac) ultimately have "?c n = exp (\ * (s*pi/2 - 2*pi*n*h/k)) + exp (-(\ * (s*pi/2 - 2*pi*n*h/k)))" by (simp add: mult_ac ring_distribs) also have "\ / 2 = cos (s*pi/2 - 2*pi*n*h/k)" by (rule cos_exp_eq [symmetric]) finally show "?c n = 2 * cos (s*pi/2 - 2*pi*n*h/k)" by simp qed also have "(2 * k * pi) powr -s * (\n=1..k. \ n * hurwitz_zeta (n / k) s) = 2 * (2 * pi * k) powr -s * (\n=1..k. cos (s*pi/2 - 2*pi*n*h/k) * hurwitz_zeta (n / k) s)" by (simp add: sum_distrib_left sum_distrib_right mult_ac) finally show ?thesis . qed end