diff --git a/thys/Actuarial_Mathematics/Preliminaries.thy b/thys/Actuarial_Mathematics/Preliminaries.thy --- a/thys/Actuarial_Mathematics/Preliminaries.thy +++ b/thys/Actuarial_Mathematics/Preliminaries.thy @@ -1,183 +1,183 @@ theory Preliminaries imports "HOL-Analysis.Analysis" begin notation powr (infixr ".^" 80) section \Preliminary Definitions and Lemmas\ lemma seq_part_multiple: fixes m n :: nat assumes "m \ 0" defines "A \ \i::nat. {i*m ..< (i+1)*m}" shows "\i j. i \ j \ A i \ A j = {}" and "(\i j \ A i \ A j = {}" proof (erule contrapos_np) assume "A i \ A j \ {}" then obtain k where "k \ A i \ A j" by blast hence "i*m < (j+1)*m \ j*m < (i+1)*m" unfolding A_def by force hence "i < j+1 \ j < i+1" using mult_less_cancel2 by blast thus "i = j" by force qed } thus "\i j. i \ j \ A i \ A j = {}" by blast next show "(\ii {..< n*m}" proof fix x::nat assume "x \ (\i n" by linarith hence "x < n*m" by (meson less_le_trans mult_le_cancel2 i_x) thus "x \ {..< n*m}" using diff_mult_distrib mult_1 i_n by auto qed next show "{..< n*m} \ (\i {..< n*m}" hence "?i < n" by (simp add: less_mult_imp_div_less) moreover have "?i*m \ x \ x < (?i+1)*m" using assms div_times_less_eq_dividend dividend_less_div_times by auto ultimately show "x \ (\i 0" shows "a / b * b = a" by (simp add: assms) lemma inverse_powr: "(1/a).^b = a.^-b" if "a > 0" for a b :: real by (smt that powr_divide powr_minus_divide powr_one_eq_one) lemma powr_eq_one_iff_gen[simp]: "a.^x = 1 \ x = 0" if "a > 0" "a \ 1" for a x :: real by (metis powr_eq_0_iff powr_inj powr_zero_eq_one that) lemma powr_less_cancel2: "0 < a \ 0 < x \ 0 < y \ x.^a < y.^a \ x < y" for a x y ::real proof - assume a_pos: "0 < a" and x_pos: "0 < x" and y_pos: "0 < y" show "x.^a < y.^a \ x < y" proof (erule contrapos_pp) assume "\ x < y" hence "x \ y" by fastforce hence "x.^a \ y.^a" proof (cases "x = y") case True thus ?thesis by simp next case False hence "x.^a > y.^a" using \x \ y\ powr_less_mono2 a_pos y_pos by auto thus ?thesis by auto qed thus "\ x.^a < y.^a" by fastforce qed qed lemma geometric_increasing_sum_aux: "(1-r)^2 * (\kk 1" for n::nat and r::real by (subst geometric_increasing_sum_aux[THEN sym], simp add: that) lemma Reals_UNIV[simp]: "\ = {x::real. True}" unfolding Reals_def by auto lemma DERIV_fun_powr2: fixes a::real assumes a_pos: "a > 0" and f: "DERIV f x :> r" shows "DERIV (\x. a.^(f x)) x :> a.^(f x) * r * ln a" proof - let ?g = "(\x. a)" have g: "DERIV ?g x :> 0" by simp have pos: "?g x > 0" by (simp add: a_pos) show ?thesis using DERIV_powr[OF g pos f] a_pos by (auto simp add: field_simps) qed lemma has_real_derivative_powr2: assumes a_pos: "a > 0" shows "((\x. a.^x) has_real_derivative a.^x * ln a) (at x)" proof - let ?f = "(\x. x::real)" have f: "DERIV ?f x :> 1" by simp thus ?thesis using DERIV_fun_powr2[OF a_pos f] by simp qed lemma has_integral_powr2_from_0: fixes a c :: real assumes a_pos: "a > 0" and a_neq_1: "a \ 1" and c_nneg: "c \ 0" shows "((\x. a.^x) has_integral ((a.^c - 1) / (ln a))) {0..c}" proof - have "((\x. a.^x) has_integral ((a.^c)/(ln a) - (a.^0)/(ln a))) {0..c}" proof (rule fundamental_theorem_of_calculus[OF c_nneg]) fix x::real assume "x \ {0..c}" show "((\y. a.^y / ln a) has_vector_derivative a.^x) (at x within {0..c})" using has_real_derivative_powr2[OF a_pos, of x] apply - apply (drule DERIV_cdivide[where c = "ln a"], simp add: assms) apply (rule has_vector_derivative_within_subset[where S=UNIV and T="{0..c}"], auto) - by (rule iffD1[OF has_field_derivative_iff_has_vector_derivative]) + by (rule iffD1[OF has_real_derivative_iff_has_vector_derivative]) qed thus ?thesis using assms powr_zero_eq_one by (simp add: field_simps) qed lemma integrable_on_powr2_from_0: fixes a c :: real assumes a_pos: "a > 0" and a_neq_1: "a \ 1" and c_nneg: "c \ 0" shows "(\x. a.^x) integrable_on {0..c}" using has_integral_powr2_from_0[OF assms] unfolding integrable_on_def by blast lemma integrable_on_powr2_from_0_general: fixes a c :: real assumes a_pos: "a > 0" and c_nneg: "c \ 0" shows "(\x. a.^x) integrable_on {0..c}" proof (cases "a = 1") case True thus ?thesis using has_integral_const_real by auto next case False thus ?thesis using has_integral_powr2_from_0 False assms by auto qed lemma has_integral_null_interval: fixes a b :: real and f::"real \ real" assumes "a \ b" shows "(f has_integral 0) {a..b}" using assms content_real_eq_0 by blast lemma has_integral_interval_reverse: fixes f :: "real \ real" and a b :: real assumes "a \ b" and "continuous_on {a..b} f" shows "((\x. f (a+b-x)) has_integral (integral {a..b} f)) {a..b}" proof - let ?g = "\x. a + b - x" let ?g' = "\x. -1" have g_C0: "continuous_on {a..b} ?g" using continuous_on_op_minus by simp have Dg_g': "\x. x\{a..b} \ (?g has_field_derivative ?g' x) (at x within {a..b})" by (auto intro!: derivative_eq_intros) show ?thesis using has_integral_substitution_general [of "{}" a b ?g a b f, simplified, OF assms g_C0 Dg_g', simplified] apply (simp add: has_integral_null_interval[OF assms(1), THEN integral_unique]) by (simp add: has_integral_neg_iff) qed end diff --git a/thys/Akra_Bazzi/Master_Theorem.thy b/thys/Akra_Bazzi/Master_Theorem.thy --- a/thys/Akra_Bazzi/Master_Theorem.thy +++ b/thys/Akra_Bazzi/Master_Theorem.thy @@ -1,465 +1,465 @@ (* File: Master_Theorem.thy Author: Manuel Eberl The Master theorem in a generalised form as derived from the Akra-Bazzi theorem. *) section \The Master theorem\ theory Master_Theorem imports "HOL-Analysis.Equivalence_Lebesgue_Henstock_Integration" Akra_Bazzi_Library Akra_Bazzi begin lemma fundamental_theorem_of_calculus_real: "a \ b \ \x\{a..b}. (f has_real_derivative f' x) (at x within {a..b}) \ (f' has_integral (f b - f a)) {a..b}" by (intro fundamental_theorem_of_calculus ballI) - (simp_all add: has_field_derivative_iff_has_vector_derivative[symmetric]) + (simp_all add: has_real_derivative_iff_has_vector_derivative[symmetric]) lemma integral_powr: "y \ -1 \ a \ b \ a > 0 \ integral {a..b} (\x. x powr y :: real) = inverse (y + 1) * (b powr (y + 1) - a powr (y + 1))" by (subst right_diff_distrib, intro integral_unique fundamental_theorem_of_calculus_real) (auto intro!: derivative_eq_intros) lemma integral_ln_powr_over_x: "y \ -1 \ a \ b \ a > 1 \ integral {a..b} (\x. ln x powr y / x :: real) = inverse (y + 1) * (ln b powr (y + 1) - ln a powr (y + 1))" by (subst right_diff_distrib, intro integral_unique fundamental_theorem_of_calculus_real) (auto intro!: derivative_eq_intros) lemma integral_one_over_x_ln_x: "a \ b \ a > 1 \ integral {a..b} (\x. inverse (x * ln x) :: real) = ln (ln b) - ln (ln a)" by (intro integral_unique fundamental_theorem_of_calculus_real) (auto intro!: derivative_eq_intros simp: field_simps) lemma akra_bazzi_integral_kurzweil_henstock: "akra_bazzi_integral (\f a b. f integrable_on {a..b}) (\f a b. integral {a..b} f)" apply unfold_locales apply (rule integrable_const_ivl) apply simp apply (erule integrable_subinterval_real, simp) apply (blast intro!: integral_le) apply (rule integral_combine, simp_all) [] done locale master_theorem_function = akra_bazzi_recursion + fixes g :: "nat \ real" assumes f_nonneg_base: "x \ x\<^sub>0 \ x < x\<^sub>1 \ f x \ 0" and f_rec: "x \ x\<^sub>1 \ f x = g x + (\i x\<^sub>1 \ g x \ 0" and ex_pos_a: "\a\set as. a > 0" begin interpretation akra_bazzi_integral "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" by (rule akra_bazzi_integral_kurzweil_henstock) sublocale akra_bazzi_function x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g using f_nonneg_base f_rec g_nonneg ex_pos_a by unfold_locales context begin private lemma g_nonneg': "eventually (\x. g x \ 0) at_top" using g_nonneg by (force simp: eventually_at_top_linorder) private lemma g_pos: assumes "g \ \(h)" assumes "eventually (\x. h x > 0) at_top" shows "eventually (\x. g x > 0) at_top" proof- from landau_omega.bigE_nonneg_real[OF assms(1) g_nonneg'] guess c . note c = this from assms(2) c(2) show ?thesis by eventually_elim (rule less_le_trans[OF mult_pos_pos[OF c(1)]], simp_all) qed private lemma f_pos: assumes "g \ \(h)" assumes "eventually (\x. h x > 0) at_top" shows "eventually (\x. f x > 0) at_top" using g_pos[OF assms(1,2)] eventually_ge_at_top[of x\<^sub>1] by (eventually_elim) (subst f_rec, insert step_ge_x0, auto intro!: add_pos_nonneg sum_nonneg mult_nonneg_nonneg[OF a_ge_0] f_nonneg) lemma bs_lower_bound: "\C>0. \b\set bs. C < b" proof (intro exI conjI ballI) from b_pos show A: "Min (set bs) / 2 > 0" by auto fix b assume b: "b \ set bs" from A have "Min (set bs) / 2 < Min (set bs)" by simp also from b have "... \ b" by simp finally show "Min (set bs) / 2 < b" . qed private lemma powr_growth2: "\C c2. 0 < c2 \ C < Min (set bs) \ eventually (\x. \u\{C * x..x}. c2 * x powr p' \ u powr p') at_top" proof (intro exI conjI allI ballI) define C where "C = Min (set bs) / 2" from b_bounds bs_nonempty have C_pos: "C > 0" unfolding C_def by auto thus "C < Min (set bs)" unfolding C_def by simp show "max (C powr p') 1 > 0" by simp show "eventually (\x. \u\{C * x..x}. max ((Min (set bs)/2) powr p') 1 * x powr p' \ u powr p') at_top" using eventually_gt_at_top[of "0::real"] apply eventually_elim proof clarify fix x u assume x: "x > 0" and "u \ {C*x..x}" hence u: "u \ C*x" "u \ x" unfolding C_def by simp_all from u have "u powr p' \ max ((C*x) powr p') (x powr p')" using C_pos x by (intro powr_upper_bound mult_pos_pos) simp_all also from u x C_pos have "max ((C*x) powr p') (x powr p') = x powr p' * max (C powr p') 1" by (subst max_mult_left) (simp_all add: powr_mult algebra_simps) finally show "u powr p' \ max ((Min (set bs)/2) powr p') 1 * x powr p'" by (simp add: C_def algebra_simps) qed qed private lemma powr_growth1: "\C c1. 0 < c1 \ C < Min (set bs) \ eventually (\x. \u\{C * x..x}. c1 * x powr p' \ u powr p') at_top" proof (intro exI conjI allI ballI) define C where "C = Min (set bs) / 2" from b_bounds bs_nonempty have C_pos: "C > 0" unfolding C_def by auto thus "C < Min (set bs)" unfolding C_def by simp from C_pos show "min (C powr p') 1 > 0" by simp show "eventually (\x. \u\{C * x..x}. min ((Min (set bs)/2) powr p') 1 * x powr p' \ u powr p') at_top" using eventually_gt_at_top[of "0::real"] apply eventually_elim proof clarify fix x u assume x: "x > 0" and "u \ {C*x..x}" hence u: "u \ C*x" "u \ x" unfolding C_def by simp_all from u x C_pos have "x powr p' * min (C powr p') 1 = min ((C*x) powr p') (x powr p')" by (subst min_mult_left) (simp_all add: powr_mult algebra_simps) also from u have "u powr p' \ min ((C*x) powr p') (x powr p')" using C_pos x by (intro powr_lower_bound mult_pos_pos) simp_all finally show "u powr p' \ min ((Min (set bs)/2) powr p') 1 * x powr p'" by (simp add: C_def algebra_simps) qed qed private lemma powr_ln_powr_lower_bound: "a > 1 \ a \ x \ x \ b \ min (a powr p) (b powr p) * min (ln a powr p') (ln b powr p') \ x powr p * ln x powr p'" by (intro mult_mono powr_lower_bound) (auto intro: min.coboundedI1) private lemma powr_ln_powr_upper_bound: "a > 1 \ a \ x \ x \ b \ max (a powr p) (b powr p) * max (ln a powr p') (ln b powr p') \ x powr p * ln x powr p'" by (intro mult_mono powr_upper_bound) (auto intro: max.coboundedI1) private lemma powr_ln_powr_upper_bound': "eventually (\a. \b>a. \c. \x\{a..b}. x powr p * ln x powr p' \ c) at_top" by (subst eventually_at_top_dense) (force intro: powr_ln_powr_upper_bound) private lemma powr_upper_bound': "eventually (\a::real. \b>a. \c. \x\{a..b}. x powr p' \ c) at_top" by (subst eventually_at_top_dense) (force intro: powr_upper_bound) lemmas bounds = powr_ln_powr_lower_bound powr_ln_powr_upper_bound powr_ln_powr_upper_bound' powr_upper_bound' private lemma eventually_ln_const: assumes "(C::real) > 0" shows "eventually (\x. ln (C*x) / ln x > 1/2) at_top" proof- from tendstoD[OF tendsto_ln_over_ln[of C 1], of "1/2"] assms have "eventually (\x. \ln (C*x) / ln x - 1\ < 1/2) at_top" by (simp add: dist_real_def) thus ?thesis by eventually_elim linarith qed private lemma powr_ln_powr_growth1: "\C c1. 0 < c1 \ C < Min (set bs) \ eventually (\x. \u\{C * x..x}. c1 * (x powr r * ln x powr r') \ u powr r * ln u powr r') at_top" proof (intro exI conjI) let ?C = "Min (set bs) / 2" and ?f = "\x. x powr r * ln x powr r'" define C where "C = ?C" from b_bounds have C_pos: "C > 0" unfolding C_def by simp let ?T = "min (C powr r) (1 powr r) * min ((1/2) powr r') (1 powr r')" from C_pos show "?T > 0" unfolding min_def by (auto split: if_split) from bs_nonempty b_bounds have C_pos: "C > 0" unfolding C_def by simp thus "C < Min (set bs)" by (simp add: C_def) show "eventually (\x. \u\{C*x..x}. ?T * ?f x \ ?f u) at_top" using eventually_gt_at_top[of "max 1 (inverse C)"] eventually_ln_const[OF C_pos] apply eventually_elim proof clarify fix x u assume x: "x > max 1 (inverse C)" and u: "u \ {C*x..x}" hence x': "x > 1" by (simp add: field_simps) with C_pos have x_pos: "x > 0" by (simp add: field_simps) from x u C_pos have u': "u > 1" by (simp add: field_simps) assume A: "ln (C*x) / ln x > 1/2" have "min (C powr r) (1 powr r) \ (u/x) powr r" using x u u' C_pos by (intro powr_lower_bound) (simp_all add: field_simps) moreover { note A also from C_pos x' u u' have "ln (C*x) \ ln u" by (subst ln_le_cancel_iff) simp_all with x' have "ln (C*x) / ln x \ ln u / ln x" by (simp add: field_simps) finally have "min ((1/2) powr r') (1 powr r') \ (ln u / ln x) powr r'" using x u u' C_pos A by (intro powr_lower_bound) simp_all } ultimately have "?T \ (u/x) powr r * (ln u / ln x) powr r'" using x_pos by (intro mult_mono) simp_all also from x u u' have "... = ?f u / ?f x" by (simp add: powr_divide) finally show "?T * ?f x \ ?f u" using x' by (simp add: field_simps) qed qed private lemma powr_ln_powr_growth2: "\C c1. 0 < c1 \ C < Min (set bs) \ eventually (\x. \u\{C * x..x}. c1 * (x powr r * ln x powr r') \ u powr r * ln u powr r') at_top" proof (intro exI conjI) let ?C = "Min (set bs) / 2" and ?f = "\x. x powr r * ln x powr r'" define C where "C = ?C" let ?T = "max (C powr r) (1 powr r) * max ((1/2) powr r') (1 powr r')" show "?T > 0" by simp from b_bounds bs_nonempty have C_pos: "C > 0" unfolding C_def by simp thus "C < Min (set bs)" by (simp add: C_def) show "eventually (\x. \u\{C*x..x}. ?T * ?f x \ ?f u) at_top" using eventually_gt_at_top[of "max 1 (inverse C)"] eventually_ln_const[OF C_pos] apply eventually_elim proof clarify fix x u assume x: "x > max 1 (inverse C)" and u: "u \ {C*x..x}" hence x': "x > 1" by (simp add: field_simps) with C_pos have x_pos: "x > 0" by (simp add: field_simps) from x u C_pos have u': "u > 1" by (simp add: field_simps) assume A: "ln (C*x) / ln x > 1/2" from x u u' have "?f u / ?f x = (u/x) powr r * (ln u/ln x) powr r'" by (simp add: powr_divide) also { have "(u/x) powr r \ max (C powr r) (1 powr r)" using x u u' C_pos by (intro powr_upper_bound) (simp_all add: field_simps) moreover { note A also from C_pos x' u u' have "ln (C*x) \ ln u" by (subst ln_le_cancel_iff) simp_all with x' have "ln (C*x) / ln x \ ln u / ln x" by (simp add: field_simps) finally have "(ln u / ln x) powr r' \ max ((1/2) powr r') (1 powr r')" using x u u' C_pos A by (intro powr_upper_bound) simp_all } ultimately have "(u/x) powr r * (ln u / ln x) powr r' \ ?T" using x_pos by (intro mult_mono) simp_all } finally show "?T * ?f x \ ?f u" using x' by (simp add: field_simps) qed qed lemmas growths = powr_growth1 powr_growth2 powr_ln_powr_growth1 powr_ln_powr_growth2 private lemma master_integrable: "\a::real. \b\a. (\u. u powr r * ln u powr s / u powr t) integrable_on {a..b}" "\a::real. \b\a. (\u. u powr r / u powr s) integrable_on {a..b}" by (rule exI[of _ 2], force intro!: integrable_continuous_real continuous_intros)+ private lemma master_integral: fixes a p p' :: real assumes p: "p \ p'" and a: "a > 0" obtains c d where "c \ 0" "p > p' \ d \ 0" "(\x::nat. x powr p * (1 + integral {a..x} (\u. u powr p' / u powr (p+1)))) \ \(\x::nat. d * x powr p + c * x powr p')" proof- define e where "e = a powr (p' - p)" from assms have e: "e \ 0" by (simp add: e_def) define c where "c = inverse (p' - p)" define d where "d = 1 - inverse (p' - p) * e" have "c \ 0" and "p > p' \ d \ 0" using e p a unfolding c_def d_def by (auto simp: field_simps) thus ?thesis apply (rule that) apply (rule bigtheta_real_nat_transfer, rule bigthetaI_cong) using eventually_ge_at_top[of a] proof eventually_elim fix x assume x: "x \ a" hence "integral {a..x} (\u. u powr p' / u powr (p+1)) = integral {a..x} (\u. u powr (p' - (p + 1)))" by (intro Henstock_Kurzweil_Integration.integral_cong) (simp_all add: powr_diff [symmetric] ) also have "... = inverse (p' - p) * (x powr (p' - p) - a powr (p' - p))" using p x0_less_x1 a x by (simp add: integral_powr) also have "x powr p * (1 + ...) = d * x powr p + c * x powr p'" using p unfolding c_def d_def by (simp add: algebra_simps powr_diff e_def) finally show "x powr p * (1 + integral {a..x} (\u. u powr p' / u powr (p+1))) = d * x powr p + c * x powr p'" . qed qed private lemma master_integral': fixes a p p' :: real assumes p': "p' \ 0" and a: "a > 1" obtains c d :: real where "p' < 0 \ c \ 0" "d \ 0" "(\x::nat. x powr p * (1 + integral {a..x} (\u. u powr p * ln u powr (p'-1) / u powr (p+1)))) \ \(\x::nat. c * x powr p + d * x powr p * ln x powr p')" proof- define e where "e = ln a powr p'" from assms have e: "e > 0" by (simp add: e_def) define c where "c = 1 - inverse p' * e" define d where "d = inverse p'" from assms e have "p' < 0 \ c \ 0" "d \ 0" unfolding c_def d_def by (auto simp: field_simps) thus ?thesis apply (rule that) apply (rule landau_real_nat_transfer, rule bigthetaI_cong) using eventually_ge_at_top[of a] proof eventually_elim fix x :: real assume x: "x \ a" have "integral {a..x} (\u. u powr p * ln u powr (p' - 1) / u powr (p + 1)) = integral {a..x} (\u. ln u powr (p' - 1) / u)" using x a x0_less_x1 by (intro Henstock_Kurzweil_Integration.integral_cong) (simp_all add: powr_add) also have "... = inverse p' * (ln x powr p' - ln a powr p')" using p' x0_less_x1 a(1) x by (simp add: integral_ln_powr_over_x) also have "x powr p * (1 + ...) = c * x powr p + d * x powr p * ln x powr p'" using p' by (simp add: algebra_simps c_def d_def e_def) finally show "x powr p * (1+integral {a..x} (\u. u powr p * ln u powr (p'-1) / u powr (p+1))) = c * x powr p + d * x powr p * ln x powr p'" . qed qed private lemma master_integral'': fixes a p p' :: real assumes a: "a > 1" shows "(\x::nat. x powr p * (1 + integral {a..x} (\u. u powr p * ln u powr - 1/u powr (p+1)))) \ \(\x::nat. x powr p * ln (ln x))" proof (rule landau_real_nat_transfer) have "(\x::real. x powr p * (1 + integral {a..x} (\u. u powr p * ln u powr - 1/u powr (p+1)))) \ \(\x::real. (1 - ln (ln a)) * x powr p + x powr p * ln (ln x))" (is "?f \ _") apply (rule bigthetaI_cong) using eventually_ge_at_top[of a] proof eventually_elim fix x assume x: "x \ a" have "integral {a..x} (\u. u powr p * ln u powr -1 / u powr (p + 1)) = integral {a..x} (\u. inverse (u * ln u))" using x a x0_less_x1 by (intro Henstock_Kurzweil_Integration.integral_cong) (simp_all add: powr_add powr_minus field_simps) also have "... = ln (ln x) - ln (ln a)" using x0_less_x1 a(1) x by (subst integral_one_over_x_ln_x) simp_all also have "x powr p * (1 + ...) = (1 - ln (ln a)) * x powr p + x powr p * ln (ln x)" by (simp add: algebra_simps) finally show "x powr p * (1 + integral {a..x} (\u. u powr p * ln u powr - 1 / u powr (p+1))) = (1 - ln (ln a)) * x powr p + x powr p * ln (ln x)" . qed also have "(\x. (1 - ln (ln a)) * x powr p + x powr p * ln (ln x)) \ \(\x. x powr p * ln (ln x))" by simp finally show "?f \ \(\a. a powr p * ln (ln a))" . qed lemma master1_bigo: assumes g_bigo: "g \ O(\x. real x powr p')" assumes less_p': "(\i 1" shows "f \ O(\x. real x powr p)" proof- interpret akra_bazzi_upper x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g "\x. x powr p'" using assms growths g_bigo master_integrable by unfold_locales (assumption | simp)+ from less_p' have less_p: "p' < p" by (rule p_greaterI) from bigo_f[of "0"] guess a . note a = this note a(2) also from a(1) less_p x0_less_x1 have "p \ p'" by simp_all from master_integral[OF this a(1)] guess c d . note cd = this note cd(3) also from cd(1,2) less_p have "(\x::nat. d * real x powr p + c * real x powr p') \ \(\x. real x powr p)" by force finally show "f \ O(\x::nat. x powr p)" . qed lemma master1: assumes g_bigo: "g \ O(\x. real x powr p')" assumes less_p': "(\i 1" assumes f_pos: "eventually (\x. f x > 0) at_top" shows "f \ \(\x. real x powr p)" proof (rule bigthetaI) interpret akra_bazzi_lower x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g "\_. 0" using assms(1,3) bs_lower_bound by unfold_locales (auto intro: always_eventually) from bigomega_f show "f \ \(\x. real x powr p)" by force qed (fact master1_bigo[OF g_bigo less_p']) lemma master2_3: assumes g_bigtheta: "g \ \(\x. real x powr p * ln (real x) powr (p' - 1))" assumes p': "p' > 0" shows "f \ \(\x. real x powr p * ln (real x) powr p')" proof- have "eventually (\x::real. x powr p * ln x powr (p' - 1) > 0) at_top" using eventually_gt_at_top[of "1::real"] by eventually_elim simp hence "eventually (\x. f x > 0) at_top" by (rule f_pos[OF bigthetaD2[OF g_bigtheta] eventually_nat_real]) then interpret akra_bazzi x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g "\x. x powr p * ln x powr (p' - 1)" using assms growths bounds master_integrable by unfold_locales (assumption | simp)+ from bigtheta_f[of "1"] guess a . note a = this note a(2) also from a(1) p' have "p' \ 0" by simp_all from master_integral'[OF this a(1), of p] guess c d . note cd = this note cd(3) also have "(\x::nat. c * real x powr p + d * real x powr p * ln (real x) powr p') \ \(\x::nat. x powr p * ln x powr p')" using cd(1,2) p' by force finally show "f \ \(\x. real x powr p * ln (real x) powr p')" . qed lemma master2_1: assumes g_bigtheta: "g \ \(\x. real x powr p * ln (real x) powr p')" assumes p': "p' < -1" shows "f \ \(\x. real x powr p)" proof- have "eventually (\x::real. x powr p * ln x powr p' > 0) at_top" using eventually_gt_at_top[of "1::real"] by eventually_elim simp hence "eventually (\x. f x > 0) at_top" by (rule f_pos[OF bigthetaD2[OF g_bigtheta] eventually_nat_real]) then interpret akra_bazzi x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g "\x. x powr p * ln x powr p'" using assms growths bounds master_integrable by unfold_locales (assumption | simp)+ from bigtheta_f[of "1"] guess a . note a = this note a(2) also from a(1) p' have A: "p' + 1 \ 0" by simp_all obtain c d :: real where cd: "c \ 0" "d \ 0" and "(\x::nat. x powr p * (1 + integral {a..x} (\u. u powr p * ln u powr p'/ u powr (p+1)))) \ \(\x::nat. c * x powr p + d * x powr p * ln x powr (p' + 1))" by (rule master_integral'[OF A a(1), of p]) (insert p', simp) note this(3) also have "(\x::nat. c * real x powr p + d * real x powr p * ln (real x) powr (p' + 1)) \ \(\x::nat. x powr p)" using cd(1,2) p' by force finally show "f \ \(\x::nat. x powr p)" . qed lemma master2_2: assumes g_bigtheta: "g \ \(\x. real x powr p / ln (real x))" shows "f \ \(\x. real x powr p * ln (ln (real x)))" proof- have "eventually (\x::real. x powr p / ln x > 0) at_top" using eventually_gt_at_top[of "1::real"] by eventually_elim simp hence "eventually (\x. f x > 0) at_top" by (rule f_pos[OF bigthetaD2[OF g_bigtheta] eventually_nat_real]) moreover from g_bigtheta have g_bigtheta': "g \ \(\x. real x powr p * ln (real x) powr -1)" by (rule landau_theta.trans, intro landau_real_nat_transfer) simp ultimately interpret akra_bazzi x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g "\x. x powr p * ln x powr -1" using assms growths bounds master_integrable by unfold_locales (assumption | simp)+ from bigtheta_f[of 1] guess a . note a = this note a(2) also note master_integral''[OF a(1)] finally show "f \ \(\x::nat. x powr p * ln (ln x))" . qed lemma master3: assumes g_bigtheta: "g \ \(\x. real x powr p')" assumes p'_greater': "(\i \(\x. real x powr p')" proof- have "eventually (\x::real. x powr p' > 0) at_top" using eventually_gt_at_top[of "1::real"] by eventually_elim simp hence "eventually (\x. f x > 0) at_top" by (rule f_pos[OF bigthetaD2[OF g_bigtheta] eventually_nat_real]) then interpret akra_bazzi x\<^sub>0 x\<^sub>1 k as bs ts f "\f a b. f integrable_on {a..b}" "\f a b. integral {a..b} f" g "\x. x powr p'" using assms growths bounds master_integrable by unfold_locales (assumption | simp)+ from p'_greater' have p'_greater: "p' > p" by (rule p_lessI) from bigtheta_f[of 0] guess a . note a = this note a(2) also from p'_greater have "p \ p'" by simp from master_integral[OF this a(1)] guess c d . note cd = this note cd(3) also have "(\x::nat. d * x powr p + c * x powr p') \ \(\x::real. x powr p')" using p'_greater cd(1,2) by force finally show "f \ \(\x. real x powr p')" . qed end end end 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,534 +1,534 @@ (* 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)) 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_raw \ \begin{figure} \begin{center} \begin{tikzpicture} \coordinate (lefttick) at (-3,0); \coordinate (righttick) at (3,0); \draw (lefttick) -- (righttick); \draw [thick] (lefttick) ++ (0,0.4) -- ++(0,3); \draw [thick] (righttick) ++ (0,0.4) -- ++(0,3); \coordinate (needle) at (1,2); \newcommand{\needleangle}{55} \newcommand{\needlelength}{{1}} \newcommand{\needlethickness}{0.6pt} \draw ($(lefttick)+(0,4pt)$) -- ($(lefttick)-(0,4pt)$); \draw ($(righttick)+(0,4pt)$) -- ($(righttick)-(0,4pt)$); \draw (0,4pt) -- (0,-4pt); \draw [densely dashed, thin] let \p1 = (needle) in (\x1, 0) -- (needle); \draw [densely dashed, thin] let \p1 = (needle) in (needle) -- (3, \y1); \draw (needle) ++ (15pt,0) arc(0:\needleangle:15pt); \path (needle) -- ++(15pt,0) node [above, midway, yshift=-1.9pt, xshift=1.8pt] {$\scriptstyle\varphi$}; \node [below, xshift=-3.5pt] at ($(lefttick)-(0,4pt)$) {$-\nicefrac{d}{2}$}; \node [below] at ($(righttick)-(0,4pt)$) {$\nicefrac{d}{2}$}; \node [below,yshift=-1pt] at (0,-4pt) {$0$}; \node [below,yshift=-2pt] at (needle |- 0,-4pt) {$x$}; \draw[<->] (needle) ++({\needleangle+90}:5pt) ++(\needleangle:{-\needlelength}) -- ++(\needleangle:2) node [midway, above, rotate=\needleangle] {$\scriptstyle l$}; \draw [line width=0.7pt,fill=white] (needle) ++({\needleangle+90}:\needlethickness) -- ++(\needleangle:\needlelength) arc({\needleangle+90}:{\needleangle-90}:\needlethickness) -- ++(\needleangle:-\needlelength) -- ++(\needleangle:-\needlelength) arc({\needleangle+270}:{\needleangle+90}:\needlethickness) -- ++(\needleangle:\needlelength); \end{tikzpicture} \end{center} \caption{A sketch of the situation in Buffon's needle experiment. There is a needle of length $l$ with its centre at a certain $x$ coordinate, angled at an angle $\varphi$ off the horizontal axis. The two vertical lines are a distance of $d$ apart, each being $\nicefrac{d}{2}$ away from the origin.} \label{fig:buffon} \end{figure} \definecolor{myred}{HTML}{cc2428} \begin{figure}[h] \begin{center} \begin{tikzpicture} \begin{axis}[ xmin=0, xmax=7, ymin=0, ymax=1, width=\textwidth, height=0.6\textwidth, xlabel={$l/d$}, ylabel={$\mathcal P$}, tick style={thin,black}, ylabel style = {rotate=270,anchor=west}, ] \addplot [color=myred, line width=1pt, mark=none,domain=0:1,samples=200] ({x}, {2/pi*x}); \addplot [color=myred, line width=1pt, mark=none,domain=1:7,samples=200] ({x}, {2/pi*(x-sqrt(x*x-1)+acos(1/x)/180*pi)}); \end{axis} \end{tikzpicture} \caption{The probability $\mathcal P$ of the needle hitting one of the lines, as a function of the quotient $l/d$ (where $l$ is the length of the needle and $d$ the horizontal distance between the lines).} \label{fig:buffonplot} \end{center} \end{figure} \ 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$)? We will show that, if we let $x := \nicefrac{l}{d}$, the probability of this is \[ \mathcal P_{l,d} = \begin{cases} \nicefrac{2}{\pi} \cdot x & \text{if}\ l \leq d\\ \nicefrac{2}{\pi}\cdot(x - \sqrt{x^2 - 1} + \arccos (\nicefrac{1}{x})) & \text{if}\ l \geq d \end{cases} \] A plot of this function can be found in Figure~\ref{fig:buffonplot}. \ locale Buffon = fixes d l :: real assumes d: "d > 0" and l: "l > 0" begin definition Buffon :: "(real \ real) measure" where "Buffon = uniform_measure lborel ({-d/2..d/2} \ {-pi..pi})" lemma space_Buffon [simp]: "space Buffon = UNIV" by (simp add: Buffon_def) definition Buffon_set :: "(real \ real) set" where "Buffon_set = {(x,\) \ {-d/2..d/2} \ {-pi..pi}. needle l x \ \ {-d/2, d/2} \ {}}" subsection \Derivation of the solution\ text \ The following form is a bit easier to handle. \ lemma Buffon_set_altdef1: "Buffon_set = {(x,\) \ {-d/2..d/2} \ {-pi..pi}. 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 - 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 show ?thesis unfolding Buffon_set_def case_prod_unfold by (intro Collect_cong conj_cong refl) meson qed lemma Buffon_set_altdef2: "Buffon_set = {(x,\) \ {-d/2..d/2} \ {-pi..pi}. abs x \ d / 2 - abs (sin \) * l / 2}" unfolding Buffon_set_altdef1 proof (intro Collect_cong prod.case_cong refl conj_cong) fix x \ assume *: "(x, \) \ {-d/2..d/2} \ {-pi..pi}" 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" show "?P x \ \ (d / 2 - \sin \\ * l / 2 \ \x\)" proof (cases "\ \ 0") case True have "x - l / 2 * sin \ \ x + l / 2 * sin \" using l True * by (auto simp: sin_ge_zero) moreover from True and * have "sin \ \ 0" by (auto simp: sin_ge_zero) ultimately show ?thesis using * True by (force simp: field_simps Let_def min_def max_def case_prod_unfold abs_if) next case False with * have "x - l / 2 * sin \ \ x + l / 2 * sin \" using l by (auto simp: sin_le_zero' mult_nonneg_nonpos) moreover from False and * have "sin \ \ 0" by (auto simp: sin_le_zero') ultimately show ?thesis using * False l d by (force simp: field_simps Let_def min_def max_def case_prod_unfold abs_if) qed qed 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) set" where "Buffon_set' = {(x,\) \ {0..d/2} \ {0..pi}. x \ d / 2 - sin \ * l / 2}" lemma closed_buffon_set [simp, intro, measurable]: "closed Buffon_set" proof - have "Buffon_set = ({-d/2..d/2} \ {-pi..pi}) \ (\z. abs (fst z) + abs (sin (snd z)) * l / 2 - d / 2) -` {0..}" (is "_ = ?A") unfolding Buffon_set_altdef2 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'" proof - have "Buffon_set' = ({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 \ sets borel" by measurable lemma measurable_buffon_set' [measurable]: "Buffon_set' \ sets borel" by measurable sublocale prob_space Buffon unfolding Buffon_def proof - have "emeasure lborel ({- d / 2..d / 2} \ {- pi..pi}) = ennreal (2 * d * pi)" unfolding lborel_prod [symmetric] using d by (subst lborel.emeasure_pair_measure_Times) (auto simp: ennreal_mult mult_ac simp flip: ennreal_numeral) also have "\ \ 0 \ \ \ \" using d by auto finally show "prob_space (uniform_measure lborel ({- d / 2..d / 2} \ {- pi..pi}))" by (intro prob_space_uniform_measure) auto qed lemma buffon_prob_aux: "emeasure Buffon {(x,\). needle l x \ \ {-d/2, d/2} \ {}} = emeasure lborel Buffon_set / 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 "{(x, \). needle l x \ \ {- d / 2, d / 2} \ {}} \ sets borel" by (intro pred_Collect_borel) (simp add: borel_prod [symmetric] needle_def closed_segment_eq_real_ivl case_prod_unfold) hence "emeasure Buffon {(x,\). needle l x \ \ {-d/2, d/2} \ {}} = emeasure lborel (({-d/2..d/2} \ {- pi..pi}) \ {(x,\). needle l x \ \ {-d/2, d/2} \ {}}) / emeasure lborel ({-(d/2)..d/2} \ {-pi..pi})" unfolding Buffon_def Buffon_set_def by (subst emeasure_uniform_measure) simp_all also have "({-d/2..d/2} \ {- pi..pi}) \ {(x, \). needle l x \ \ {-d/2, d/2} \ {}} = Buffon_set" unfolding Buffon_set_def by auto also have "emeasure lborel ({-(d/2)..d/2} \ {-pi..pi}) = ennreal (2 * d * pi)" using d by (simp flip: lborel_prod ennreal_mult add: lborel.emeasure_pair_measure_Times) finally show ?thesis . qed lemma emeasure_buffon_set_conv_buffon_set': "emeasure lborel Buffon_set = 4 * emeasure lborel Buffon_set'" 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'" 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 = 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_altdef2 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 \ B" by (auto split: prod.splits simp: Buffon_set_altdef2 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_altdef2 Buffon_set'_def sin_le_zero' A_def B_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_altdef2 Buffon_set'_def sin_le_zero' A_def B_def D_def) thus ?thesis by blast qed next case (2 z) thus ?case using d l by (auto simp: Buffon_set_altdef2 Buffon_set'_def sin_ge_zero sin_le_zero' A_def B_def C_def D_def) qed 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 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 = 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' = 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' = ennreal l" proof - have "emeasure lborel Buffon_set' = 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]) + (auto intro!: derivative_eq_intros simp: has_real_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 = 4 * ennreal l" by (simp add: emeasure_buffon_set_conv_buffon_set' emeasure_buffon_set'_short l_le_d) lemma prob_short_aux: "Buffon {(x, \). needle l x \ \ {- d / 2, d / 2} \ {}} = ennreal (2 * l / (d * pi))" unfolding buffon_prob_aux emeasure_buffon_set_short using d l by (simp flip: ennreal_mult ennreal_numeral add: divide_ennreal) lemma prob_short: "\

((x,\) in Buffon. needle l x \ \ {-d/2, d/2} \ {}) = 2 * l / (d * pi)" using prob_short_aux unfolding emeasure_eq_measure using l d by (subst (asm) ennreal_inj) auto 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: shows "l * (1 - sqrt (1 - (d / l)\<^sup>2)) + arccos (d / l) * d \ 0" and "emeasure lborel Buffon_set' = 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) have "cos \' = sqrt (1 - (d / l)^2)" unfolding \'_def by (rule cos_arcsin) (insert d l l_ge_d, auto simp: field_simps) have "l * (1 - cos \') + arccos (d / l) * d \ 0" using l d l_ge_d by (intro add_nonneg_nonneg mult_nonneg_nonneg arccos_lbound) (auto simp: field_simps) thus "l * (1 - sqrt (1 - (d / l)\<^sup>2)) + arccos (d / l) * d \ 0" by (simp add: \cos \' = sqrt (1 - (d / l)^2)\) let ?f = "(\x. min (d / 2) (sin x * l / 2))" have "emeasure lborel Buffon_set' = 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 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 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 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) + (auto simp: has_real_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 note \cos \' = sqrt (1 - (d / l)^2)\ 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 "emeasure lborel Buffon_set' = ennreal (l * (1 - sqrt (1 - (d / l)\<^sup>2)) + arccos (d / l) * d)" . qed lemma emeasure_set_long: "emeasure lborel Buffon_set = 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) lemma prob_long_aux: shows "2 / pi * ((l / d) - sqrt ((l / d)\<^sup>2 - 1) + arccos (d / l)) \ 0" and "Buffon {(x, \). needle l x \ \ {- d / 2, d / 2} \ {}} = ennreal (2 / pi * ((l / d) - sqrt ((l / d)\<^sup>2 - 1) + arccos (d / l)))" using emeasure_buffon_set'_long(1) 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 "l / d \ sqrt ((l / d)\<^sup>2 - 1)" using l d l_ge_d by (intro real_le_lsqrt) (auto simp: field_simps) thus "2 / pi * ((l / d) - sqrt ((l / d)\<^sup>2 - 1) + arccos (d / l)) \ 0" using d l l_ge_d by (intro mult_nonneg_nonneg add_nonneg_nonneg arccos_lbound) (auto simp: field_simps) have "emeasure Buffon {(x,\). needle l x \ \ {-d/2, d/2} \ {}} = 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_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 "emeasure Buffon {(x,\). needle l x \ \ {-d/2, d/2} \ {}} = ennreal (2 / pi * ((l / d) - sqrt ((l / d)\<^sup>2 - 1) + arccos (d / l)))" . qed lemma prob_long: "\

((x,\) in Buffon. needle l x \ \ {-d/2, d/2} \ {}) = 2 / pi * ((l / d) - sqrt ((l / d)\<^sup>2 - 1) + arccos (d / l))" using prob_long_aux unfolding emeasure_eq_measure by (subst (asm) ennreal_inj) simp_all end theorem prob_eq: defines "x \ l / d" shows "\

((x,\) in Buffon. needle l x \ \ {-d/2, d/2} \ {}) = (if l \ d then 2 / pi * x else 2 / pi * (x - sqrt (x\<^sup>2 - 1) + arccos (1 / x)))" using prob_short prob_long unfolding x_def by auto end end \ No newline at end of file diff --git a/thys/Catalan_Numbers/Catalan_Auxiliary_Integral.thy b/thys/Catalan_Numbers/Catalan_Auxiliary_Integral.thy --- a/thys/Catalan_Numbers/Catalan_Auxiliary_Integral.thy +++ b/thys/Catalan_Numbers/Catalan_Auxiliary_Integral.thy @@ -1,159 +1,159 @@ section \Catalan numbers\ theory Catalan_Auxiliary_Integral imports "HOL-Analysis.Equivalence_Lebesgue_Henstock_Integration" begin subsection \Auxiliary integral\ text \ First, we will prove the integral $$\int\limits_0^4 \sqrt{\frac{4-x}{x}}\,\textrm{d}x = 2\pi$$ which occurs in the proof for the integral formula for the Catalan numbers. \ context begin text \ First of all, we prove two limits that we will require. \ private lemma limit1: "filterlim (\x::real. sqrt (4/x - 1) / (x - 4)) at_bot (at_left 4)" proof (rule lhopital_left) show "((\x::real. sqrt (4 / x - 1)) \ 0) (at_left 4)" by (force intro: tendsto_eq_intros) show "((\x::real. x - 4) \ 0) (at_left 4)" by (force intro: tendsto_eq_intros) show "\\<^sub>F x in at_left 4. ((\x. x - 4) has_real_derivative 1) (at x)" by (intro always_eventually) (auto intro!: derivative_eq_intros) show "\\<^sub>F x in at_left 4. ((\x. sqrt (4/x - 1)) has_real_derivative -2 / (sqrt (4/x - 1) * x^2)) (at x)" by (intro eventually_at_leftI[of 0]) (auto intro!: derivative_eq_intros simp: field_simps power2_eq_square) have *: "((\x::real. sqrt (4 / x - 1) * x\<^sup>2) \ 0) (at_left 4)" by (force intro!: tendsto_eq_intros) have "filterlim (\x. - (2 / (sqrt (4 / x - 1) * x\<^sup>2))) at_bot (at_left 4)" unfolding filterlim_uminus_at_top [symmetric] by (rule LIM_at_top_divide tendsto_const eventually_at_leftI[of 0] * | simp)+ thus "filterlim (\x::real. - 2 / (sqrt (4 / x - 1) * x\<^sup>2) / 1) at_bot (at_left 4)" by simp qed (auto intro: eventually_at_leftI[of 0]) private lemma limit2: "((\x. sqrt (4/x - 1) * x) \ 0) (at_right 0)" proof (rule Lim_transform_eventually) show "((\x::real. sqrt ((4 - x) * x)) \ 0) (at_right 0)" by (force intro: tendsto_eq_intros) show "\\<^sub>F x in at_right 0. sqrt ((4 - x) * x) = sqrt (4 / x - 1) * x" proof (rule eventually_at_rightI) fix x :: real assume x: "x \ {0<..<1}" from x have "(4 - x) * x = (4/x - 1) * x^2" by (simp add: field_simps power2_eq_square) also from x have "sqrt \ = sqrt (4/x - 1) * x" by (subst real_sqrt_mult) simp_all finally show "sqrt ((4 - x) * x) = sqrt (4/x - 1) * x" . qed simp_all qed text \ Now we prove the integral by explicitly constructing the indefinite integral. \ lemma catalan_aux_integral: "((\x::real. sqrt ((4 - x) / x)) has_integral 2 * pi) {0..4}" proof - define F where "F \ \x. sqrt (4/x - 1) * x - 2 * arctan ((sqrt (4/x - 1) * (x - 2)) / (x - 4))" \ \The nice part of the indefinite integral. The endpoints are removable singularities.\ define G where "G \ \x. if x = 4 then pi else if x = 0 then -pi else F x" \ \The actual indefinite integral including the endpoints.\ \ \We now prove that the indefinite integral indeed tends to @{term "pi"} resp. @{term "-pi"} at the edges of the integration interval.\ have "filterlim (\x. (x - 2) / (x - 4) * sqrt (-1 + 4/x)) at_top (at_right 0)" by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros filterlim_tendsto_add_at_top filterlim_compose[OF sqrt_at_top] LIM_at_top_divide eventually_at_rightI[of 0 1] | simp)+ hence *: "filterlim (\x. (sqrt (4/x - 1) * (x - 2)) / (x - 4)) at_top (at_right 0)" by (simp add: field_simps) have "(F \ 0 - 2 * (pi/2)) (at_right 0)" unfolding F_def by (intro tendsto_intros limit2 filterlim_compose[OF tendsto_arctan_at_top] *) hence "(F \ -pi) (at_right 0)" by simp hence G_0: "(G \ -pi) (at_right 0)" unfolding G_def by (rule Lim_transform_eventually) (auto intro!: eventually_at_rightI[of 0 1]) have *: "((\x. sqrt (4 / x - 1) * x) \ 0) (at_left 4)" by (force intro: tendsto_eq_intros) have "filterlim (\x. (x - 2) * (sqrt (4 / x - 1) / (x - 4))) at_bot (at_left 4)" by (rule filterlim_tendsto_pos_mult_at_bot tendsto_intros limit1 | simp)+ hence **: "LIM x at_left 4. sqrt (4 / x - 1) * (x - 2) / (x - 4) :> at_bot" by (simp add: mult_ac) have "(F \ 0 - 2 * -(pi/2)) (at_left 4)" unfolding F_def by (intro tendsto_intros * ** filterlim_compose[OF tendsto_arctan_at_bot]) hence "(F \ pi) (at_left 4)" by simp hence G_4: "(G \ pi) (at_left 4)" unfolding G_def by (rule Lim_transform_eventually) (auto intro!: eventually_at_leftI[of 1]) \ \The derivative of @{term G} is indeed the integrand in the interior of the integration interval.\ have deriv_G: "(G has_field_derivative sqrt ((4 - x) / x)) (at x)" if x: "x \ {0<..<4}" for x proof - from x have "eventually (\y. y \ {0<..<4}) (nhds x)" by (intro eventually_nhds_in_open) simp_all hence eq: "eventually (\x. F x = G x) (nhds x)" by eventually_elim (simp add: G_def) define T where "T \ \x::real. 4 / (sqrt (4/x - 1) * (x - 4) * x^2)" have *: "((\x. (sqrt (4/x - 1) * (x - 2)) / (x - 4)) has_field_derivative T x) (at x)" by (insert x, rule derivative_eq_intros refl | simp)+ (simp add: divide_simps T_def, simp add: field_simps power2_eq_square) have "((\x. 2 * arctan ((sqrt (4/x - 1) * (x - 2)) / (x - 4))) has_field_derivative 2 * T x / (1 + (sqrt (4 / x - 1) * (x - 2) / (x - 4))\<^sup>2)) (at x)" by (rule * derivative_eq_intros refl | simp)+ (simp add: field_simps) also from x have "(sqrt (4 / x - 1) * (x - 2) / (x - 4))\<^sup>2 = (4/x - 1) * (x-2)^2 / (x - 4)^2" by (simp add: power_mult_distrib power_divide) also from x have "1 + \ = -4 / ((x - 4) * x)" by (simp add: divide_simps power2_eq_square mult_ac) (simp add: algebra_simps)? also from x have "2 * T x / \ = - (2 / (x * sqrt (4 / x - 1)))" by (simp add: T_def power2_eq_square) finally have *: "((\x. 2 * arctan (sqrt (4 / x - 1) * (x - 2) / (x - 4))) has_real_derivative - (2 / (x * sqrt (4 / x - 1)))) (at x)" . have "(F has_field_derivative sqrt (4 / x - 1)) (at x)" unfolding F_def by (insert x, (rule * derivative_eq_intros refl | simp)+) (simp add: field_simps) thus ?thesis by (subst (asm) DERIV_cong_ev[OF refl eq refl]) (insert x, simp add: field_simps) qed \ \It is now obvious that @{term G} is continuous over the entire integration interval.\ have cont_G: "continuous_on {0..4} G" unfolding continuous_on proof safe fix x :: real assume "x \ {0..4}" then consider "x = 0" | "x = 4" | "x \ {0<..<4}" by force thus "(G \ G x) (at x within {0..4})" proof cases assume "x = 0" have *: "at (0::real) within {0..4} \ at_right 0" unfolding at_within_def by (rule inf_mono) auto from G_0 have "(G \ -pi) (at x within {0..4})" by (rule filterlim_mono) (simp_all add: * \x = 0\) also have "-pi = G x" by (simp add: G_def \x = 0\) finally show ?thesis . next assume "x = 4" have *: "at (4::real) within {0..4} \ at_left 4" unfolding at_within_def by (rule inf_mono) auto from G_4 have "(G \ pi) (at x within {0..4})" by (rule filterlim_mono) (simp_all add: * \x = 4\) also have "pi = G x" by (simp add: G_def \x = 4\) finally show ?thesis . next assume "x \ {0<..<4}" from deriv_G[OF this] have "isCont G x" by (rule DERIV_isCont) thus ?thesis unfolding isCont_def by (rule filterlim_mono) (auto simp: at_le) qed qed \ \Finally, we can apply the Fundamental Theorem of Calculus.\ have "((\x. sqrt ((4 - x) / x)) has_integral G 4 - G 0) {0..4}" using cont_G deriv_G by (intro fundamental_theorem_of_calculus_interior) - (auto simp: has_field_derivative_iff_has_vector_derivative) + (auto simp: has_real_derivative_iff_has_vector_derivative) also have "G 4 - G 0 = 2 * pi" by (simp add: G_def) finally show ?thesis . qed end end diff --git a/thys/Catalan_Numbers/Catalan_Numbers.thy b/thys/Catalan_Numbers/Catalan_Numbers.thy --- a/thys/Catalan_Numbers/Catalan_Numbers.thy +++ b/thys/Catalan_Numbers/Catalan_Numbers.thy @@ -1,484 +1,484 @@ (* File: Catalan_Numbers.thy Author: Manuel Eberl (TUM) The recursive definition of Catalan numbers with a proof of several closed form expressions for them using generating functions. Also contains reasonably efficient code generation and a proof of their asymptotic growth. *) theory Catalan_Numbers imports Complex_Main Catalan_Auxiliary_Integral "HOL-Analysis.Analysis" "HOL-Computational_Algebra.Formal_Power_Series" "HOL-Library.Landau_Symbols" Landau_Symbols.Landau_More begin subsection \ Other auxiliary lemmas\ lemma mult_eq_imp_eq_div: assumes "a * b = c" "(a :: 'a :: semidom_divide) \ 0" shows "b = c div a" by (simp add: assms(2) assms(1) [symmetric]) lemma Gamma_minus_one_half_real: "Gamma (-(1/2) :: real) = - 2 * sqrt pi" using rGamma_plus1[of "-1/2 :: real"] by (simp add: rGamma_inverse_Gamma divide_simps Gamma_one_half_real split: if_split_asm) lemma gbinomial_asymptotic': assumes "z \ \" shows "(\n. z gchoose (n + k)) \[at_top] (\n. (-1)^(n+k) / (Gamma (-z) * of_nat n powr (z + 1)) :: real)" proof - from assms have [simp]: "Gamma (-z) \ 0" by (simp_all add: Gamma_eq_zero_iff uminus_in_nonpos_Ints_iff) have "filterlim (\n. n + k) at_top at_top" by (intro filterlim_subseq strict_mono_add) from asymp_equivI'_const[OF gbinomial_asymptotic[of z]] assms have "(\n. z gchoose n) \[at_top] (\n. (-1)^n / (Gamma (-z) * exp ((z+1) * ln (real n))))" by (simp add: Gamma_eq_zero_iff uminus_in_nonpos_Ints_iff field_simps) also have "eventually (\n. exp ((z+1) * ln (real n)) = real n powr (z+1)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (simp add: powr_def) finally have "(\x. z gchoose (x + k)) \[at_top] (\x. (- 1) ^ (x + k) / (Gamma (- z) * real (x + k) powr (z + 1)))" by (rule asymp_equiv_compose') (simp add: filterlim_subseq strict_mono_add) also have "(\x. real x + real k) \[at_top] real" by (subst asymp_equiv_add_right) auto hence "(\x. real (x + k) powr (z + 1)) \[at_top] (\x. real x powr (z + 1))" by (intro asymp_equiv_powr_real) auto finally show ?thesis by - (simp_all add: asymp_equiv_intros) qed subsection \Definition\ text \ We define Catalan numbers by their well-known recursive definition. We shall later derive a few more equivalent definitions from this one. \ (*<*) context notes [fundef_cong] = sum.cong begin (*>*) fun catalan :: "nat \ nat" where "catalan 0 = 1" | "catalan (Suc n) = (\i\n. catalan i * catalan (n - i))" (*<*) end declare catalan.simps(2) [simp del] lemmas catalan_0 = catalan.simps(1) lemmas catalan_Suc = catalan.simps(2) lemma catalan_1 [simp]: "catalan (Suc 0) = 1" by (simp add: catalan_Suc) (*>*) text \ The easiest proof of the more profound properties of the Catalan numbers (such as their closed-form equation and their asymptotic growth) uses their ordinary generating function (OGF). This proof is almost mechanical in the sense that it does not require `guessing' the closed form; one can read it directly from the generating function. We therefore define the OGF of the Catalan numbers ($\sum_{n=0}^\infty C_n z^n$ in standard mathematical notation): \ definition "fps_catalan = Abs_fps (of_nat \ catalan)" lemma fps_catalan_nth [simp]: "fps_nth fps_catalan n = of_nat (catalan n)" by (simp add: fps_catalan_def) text \ Given their recursive definition, it is easy to see that the OGF of the Catalan numbers satisfies the following recursive equation: \ lemma fps_catalan_recurrence: "fps_catalan = 1 + fps_X * fps_catalan^2" proof (rule fps_ext) fix n :: nat show "fps_nth fps_catalan n = fps_nth (1 + fps_X * fps_catalan^2) n" by (cases n) (simp_all add: fps_square_nth catalan_Suc) qed text \ We can now easily solve this equation for @{term fps_catalan}: if we denote the unknown OGF as $F(z)$, we get $F(z) = \frac{1}{2}(1 - \sqrt{1 - 4z})$. Note that we do not actually use the square root as defined on real or complex numbers. Any $(1 + cz)^\alpha$ can be expressed using the formal power series whose coefficients are the generalised binomial coefficients, and thus we can do all of these transformations in a purely algebraic way: $\sqrt{1-4z} = (1+z)^{\frac{1}{2}} \circ (-4z)$ (where ${\circ}$ denotes composition of formal power series) and $(1+z)^\alpha$ has the well-known expansion $\sum_{n=0}^\infty {\alpha \choose n} z^n$. \ lemma fps_catalan_fps_binomial: "fps_catalan = (1/2 * (1 - (fps_binomial (1/2) oo (-4*fps_X)))) / fps_X" proof (rule mult_eq_imp_eq_div) let ?F = "fps_catalan :: 'a fps" have "fps_X * (1 + fps_X * ?F^2) = fps_X * ?F" by (simp only: fps_catalan_recurrence [symmetric]) hence "(1 / 2 - fps_X * ?F)\<^sup>2 = - fps_X + 1 / 4" by (simp add: algebra_simps power2_eq_square fps_numeral_simps) also have "\ = (1/2 * (fps_binomial (1/2) oo (-4*fps_X)))^2" by (simp add: power_mult_distrib div_power fps_binomial_1 fps_binomial_power fps_compose_power fps_compose_add_distrib ring_distribs) finally have "1/2 - fps_X * ?F = 1/2 * (fps_binomial (1/2) oo (-4*fps_X))" by (rule fps_power_eqD) simp_all thus "fps_X*?F = 1/2 * (1 - (fps_binomial (1/2) oo (-4*fps_X)))" by algebra qed simp_all subsection \Closed-form formulae and more recurrences\ text \ We can now read a closed-form formula for the Catalan numbers directly from the generating function $\frac{1}{2z}(1 - (1+z)^{\frac{1}{2}} \circ (-4z))$. \ theorem catalan_closed_form_gbinomial: "real (catalan n) = 2 * (- 4) ^ n * (1/2 gchoose Suc n)" proof - have "(catalan n :: real) = fps_nth fps_catalan n" by simp also have "\ = 2 * (- 4) ^ n * (1/2 gchoose Suc n)" by (subst fps_catalan_fps_binomial) (simp add: fps_div_fps_X_nth numeral_fps_const fps_compose_linear) finally show ?thesis . qed text \ This closed-form formula can easily be rewritten to the form $C_n = \frac{1}{n+1} {2n \choose n}$, which contains only `normal' binomial coefficients and not the generalised ones: \ lemma catalan_closed_form_aux: "catalan n * Suc n = (2*n) choose n" proof - have "real ((2*n) choose n) = fact (2*n) / (fact n)^2" by (simp add: binomial_fact power2_eq_square) also have "(fact (2*n) :: real) = 4^n * pochhammer (1 / 2) n * fact n" by (simp add: fact_double power_mult) also have "\ / (fact n)^2 / real (n+1) = real (catalan n)" by (simp add: catalan_closed_form_gbinomial gbinomial_pochhammer pochhammer_rec field_simps power2_eq_square power_mult_distrib [symmetric] del: of_nat_Suc) finally have "real (catalan n * Suc n) = real ((2*n) choose n)" by (simp add: field_simps) thus ?thesis by (simp only: of_nat_eq_iff) qed theorem of_nat_catalan_closed_form: "of_nat (catalan n) = (of_nat ((2*n) choose n) / of_nat (Suc n) :: 'a :: field_char_0)" proof - have "of_nat (catalan n * Suc n) = of_nat ((2*n) choose n)" by (subst catalan_closed_form_aux) (rule refl) also have "of_nat (catalan n * Suc n) = of_nat (catalan n) * of_nat (Suc n)" by (simp only: of_nat_mult) finally show ?thesis by (simp add: divide_simps del: of_nat_Suc) qed theorem catalan_closed_form: "catalan n = ((2*n) choose n) div Suc n" by (subst catalan_closed_form_aux [symmetric]) (simp del: mult_Suc_right) text \ The following is another nice closed-form formula for the Catalan numbers, which directly follows from the previous one: \ corollary catalan_closed_form': "catalan n = ((2*n) choose n) - ((2*n) choose (Suc n))" proof (cases n) case (Suc m) have "real ((2*n) choose n) - real ((2*n) choose (Suc n)) = fact (2*m+2) / (fact (m+1))^2 - fact (2*m+2) / (real (m+2) * fact (m+1) * fact m)" by (subst (1 2) binomial_fact) (simp_all add: Suc power2_eq_square) also have "\ = fact (2*m+2) / ((fact (m+1))^2 * real (m+2))" by (simp add: divide_simps power2_eq_square) (simp_all add: algebra_simps) also have "\ = real (catalan n)" by (subst of_nat_catalan_closed_form, subst binomial_fact) (simp_all add: Suc power2_eq_square) finally show ?thesis by linarith qed simp_all text \ We can now easily show that the Catalan numbers also satisfy another, simpler recurrence, namely $C_{n+1} = \frac{2(2n+1)}{n+2} C_n$. We will later use this to prove code equations to compute the Catalan numbers more efficiently. \ lemma catalan_Suc_aux: "(n + 2) * catalan (Suc n) = 2 * (2 * n + 1) * catalan n" proof - have "real (catalan (Suc n)) * real (n + 2) = real (catalan n) * 2 * real (2 * n + 1)" proof (cases n) case (Suc n) thus ?thesis by (subst (1 2) of_nat_catalan_closed_form, subst (1 2) binomial_fact) (simp_all add: divide_simps) qed simp_all hence "real ((n + 2) * catalan (Suc n)) = real (2 * (2 * n + 1) * catalan n)" by (simp only: mult_ac of_nat_mult) thus ?thesis by (simp only: of_nat_eq_iff) qed theorem of_nat_catalan_Suc': "of_nat (catalan (Suc n)) = (of_nat (2*(2*n+1)) / of_nat (n+2) * of_nat (catalan n) :: 'a :: field_char_0)" proof - have "(of_nat (2*(2*n+1)) / of_nat (n+2) * of_nat (catalan n) :: 'a) = of_nat (2*(2*n + 1) * catalan n) / of_nat (n+2)" by (simp add: divide_simps mult_ac del: mult_Suc mult_Suc_right) also note catalan_Suc_aux[of n, symmetric] also have "of_nat ((n + 2) * catalan (Suc n)) / of_nat (n + 2) = (of_nat (catalan (Suc n)) :: 'a)" by (simp del: of_nat_Suc mult_Suc_right mult_Suc) finally show ?thesis .. qed theorem catalan_Suc': "catalan (Suc n) = (catalan n * (2*(2*n+1))) div (n+2)" proof - from catalan_Suc_aux[of n] have "catalan n * (2*(2*n+1)) = catalan (Suc n) * (n+2)" by (simp add: algebra_simps) also have "\ div (n+2) = catalan (Suc n)" by (simp del: mult_Suc mult_Suc_right) finally show ?thesis .. qed subsection \Integral formula\ text \ The recursive formula we have just proven allows us to derive an integral formula for the Catalan numbers. The proof was adapted from a textbook proof by Steven Roman.~\cite{catalan} \ context begin private definition I :: "nat \ real" where "I n = integral {0..4} (\x. x powr (of_nat n - 1/2) * sqrt (4 - x))" private lemma has_integral_I0: "((\x. x powr (-(1/2)) * sqrt (4 - x)) has_integral 2*pi) {0..4}" proof - have "\x. x\{0..4}-{} \ x powr (-(1/2)) * sqrt (4 - x) = sqrt ((4 - x) / x)" by (auto simp: powr_minus field_simps powr_half_sqrt real_sqrt_divide) thus ?thesis by (rule has_integral_spike[OF negligible_empty _ catalan_aux_integral]) qed private lemma integrable_I: "(\x. x powr (of_nat n - 1/2) * sqrt (4 - x)) integrable_on {0..4}" proof (cases "n = 0") case True with has_integral_I0 show ?thesis by (simp add: has_integral_integrable) next case False thus ?thesis by (intro integrable_continuous_real continuous_on_mult continuous_on_powr') (auto intro!: continuous_intros) qed private lemma I_Suc: "I (Suc n) = real (2 * (2*n + 1)) / real (n + 2) * I n" proof - define u' u v v' where "u' = (\x. sqrt (4 - x :: real))" and "u = (\x. -2/3 * (4 - x) powr (3/2 :: real))" and "v = (\x. x powr (real n + 1/2))" and "v' = (\x. (real n + 1/2) * x powr (real n - 1/2))" define c where "c = -2/3 * (real n + 1/2)" define i where "i = (\n x. x powr (real n - 1/2) * sqrt (4 - x) :: real)" have "I (Suc n) = integral {0..4} (\x. u' x * v x)" unfolding I_def by (simp add: algebra_simps u'_def v_def) have "((\x. u' x * v x) has_integral - c * (4 * I n - I (Suc n))) {0..4}" proof (rule integration_by_parts_interior[OF bounded_bilinear_mult]) show "continuous_on {0..4} u" unfolding u_def by (intro continuous_on_powr' continuous_on_mult) (auto intro!: continuous_intros) show "continuous_on {0..4} v" unfolding v_def by (intro continuous_on_powr' continuous_on_mult) (auto intro!: continuous_intros) fix x :: real assume x: "x \ {0<..<4}" from x show "(u has_vector_derivative u' x) (at x)" - unfolding has_field_derivative_iff_has_vector_derivative [symmetric] u_def u'_def + unfolding has_real_derivative_iff_has_vector_derivative [symmetric] u_def u'_def by (auto intro!: derivative_eq_intros simp: field_simps powr_half_sqrt) from x show "(v has_vector_derivative v' x) (at x)" - unfolding has_field_derivative_iff_has_vector_derivative [symmetric] v_def v'_def + unfolding has_real_derivative_iff_has_vector_derivative [symmetric] v_def v'_def by (auto intro!: derivative_eq_intros simp: field_simps) next show "((\x. u x * v' x) has_integral u 4 * v 4 - u 0 * v 0 - - c * (4 * I n - I (Suc n))) {0..4}" proof (rule has_integral_spike; (intro ballI)?) fix x :: real assume x: "x \ {0..4}-{0}" have "u x * v' x = c * ((4 - x) powr (1 + 1/2) * x powr (real n - 1/2))" by (simp add: u_def v'_def c_def) also from x have "(4 - x) powr (1 + 1/2) = (4 - x) * sqrt (4 - x)" by (subst powr_add) (simp_all add: powr_half_sqrt) also have "\ * x powr (real n - 1/2) = 4 * sqrt (4 - x) * x powr (real n - 1/2) - sqrt (4 - x) * x powr (real n - 1/2 + 1)" by (subst powr_add) (insert x, simp add: field_simps) also have "real n - 1/2 + 1 = real (Suc n) - 1/2" by simp finally show "u x * v' x = c * (4 * i n x - i (Suc n) x)" by (simp add: i_def) next have "((\x. c * (4 * i n x - i (Suc n) x)) has_integral c * (4 * I n - I (Suc n))) {0..4}" unfolding i_def I_def by (intro has_integral_mult_right has_integral_diff integrable_integral integrable_I) thus "((\x. c * (4 * i n x - i (Suc n) x)) has_integral u 4 * v 4 - u 0 * v 0 - - c * (4 * I n - I (Suc n))) {0..4}" by (simp add: u_def v_def) qed simp_all qed simp_all also have "(\x. u' x * v x) = i (Suc n)" by (rule ext) (simp add: u'_def v_def i_def algebra_simps) finally have "I (Suc n) = - c * (4 * I n - I (Suc n))" unfolding I_def i_def by blast hence "(1 - c) * I (Suc n) = -4 * c * I n" by algebra hence "I (Suc n) = (-4 * c) / (1 - c) * I n" by (simp add: field_simps c_def) also have "(-4 * c) / (1 - c) = real (2*(2*n + 1)) / real (n + 2)" by (simp add: c_def field_simps) finally show ?thesis . qed private lemma catalan_eq_I: "real (catalan n) = I n / (2 * pi)" proof (induction n) case 0 thus ?case using has_integral_I0 by (simp add: I_def integral_unique) next case (Suc n) show ?case by (simp add: of_nat_catalan_Suc' Suc.IH I_Suc) qed theorem catalan_integral_form: "((\x. x powr (real n - 1 / 2) * sqrt (4 - x) / (2*pi)) has_integral real (catalan n)) {0..4}" proof - have "((\x. x powr (real n - 1 / 2) * sqrt (4 - x) * inverse (2*pi)) has_integral I n * inverse (2 * pi)) {0..4}" unfolding I_def by (intro has_integral_mult_left integrable_integral integrable_I) thus ?thesis by (simp add: catalan_eq_I field_simps) qed end subsection \Asymptotics\ text \ Using the closed form $C_n = 2 \cdot (-4)^n {\frac{1}{2} \choose n+1}$ and the fact that ${\alpha \choose n} \sim \frac{(-1)^n}{\Gamma(-\alpha) n^{\alpha + 1}}$ for any $\alpha \notin \mathbb{N}$, wwe can now easily analyse the asymptotic behaviour of the Catalan numbers: \ theorem catalan_asymptotics: "catalan \[at_top] (\n. 4 ^ n / (sqrt pi * n powr (3/2)))" proof - have "catalan \[at_top] (\n. 2 * (- 4) ^ n * (1/2 gchoose (n+1)))" by (subst catalan_closed_form_gbinomial) simp_all also have "(\n. 1/2 gchoose (n+1)) \[at_top] (\n. (-1)^(n+1) / (Gamma (-(1/2)) * real n powr (1/2 + 1)))" using fraction_not_in_nats[of 2 1] by (intro asymp_equiv_intros gbinomial_asymptotic') simp_all also have "(\n. 2 * (- 4) ^ n * \ n) = (\n. 4 ^ n / (sqrt pi * n powr (3/2)))" by (intro ext) (simp add: Gamma_minus_one_half_real power_mult_distrib [symmetric]) finally show ?thesis by - (simp_all add: asymp_equiv_intros) qed subsection \Relation to binary trees\ (*<*) context begin (*>*) text \ It is well-known that the Catalan number $C_n$ is the number of rooted binary trees with $n$ internal nodes (where internal nodes are those with two children and external nodes are those with no children). We will briefly show this here to show that the above asymptotic formula also describes the number of binary trees of a given size. \ qualified datatype tree = Leaf | Node tree tree qualified primrec count_nodes :: "tree \ nat" where "count_nodes Leaf = 0" | "count_nodes (Node l r) = 1 + count_nodes l + count_nodes r" qualified definition trees_of_size :: "nat \ tree set" where "trees_of_size n = {t. count_nodes t = n}" lemma count_nodes_eq_0_iff [simp]: "count_nodes t = 0 \ t = Leaf" by (cases t) simp_all lemma trees_of_size_0 [simp]: "trees_of_size 0 = {Leaf}" by (simp add: trees_of_size_def) lemma trees_of_size_Suc: "trees_of_size (Suc n) = (\(l,r). Node l r) ` (\k\n. trees_of_size k \ trees_of_size (n - k))" (is "?lhs = ?rhs") proof (rule set_eqI) fix t show "t \ ?lhs \ t \ ?rhs" by (cases t) (auto simp: trees_of_size_def) qed lemma finite_trees_of_size [simp,intro]: "finite (trees_of_size n)" by (induction n rule: catalan.induct) (auto simp: trees_of_size_Suc intro!: finite_imageI finite_cartesian_product) lemma trees_of_size_nonempty: "trees_of_size n \ {}" by (induction n rule: catalan.induct) (auto simp: trees_of_size_Suc) lemma trees_of_size_disjoint: assumes "m \ n" shows "trees_of_size m \ trees_of_size n = {}" using assms by (auto simp: trees_of_size_def) theorem card_trees_of_size: "card (trees_of_size n) = catalan n" by (induction n rule: catalan.induct) (simp_all add: catalan_Suc trees_of_size_Suc card_image inj_on_def trees_of_size_disjoint Times_Int_Times catalan_Suc card_UN_disjoint) (*<*) end (*>*) subsection \Efficient computation\ (*<*) context begin (*>*) text \ We shall now prove code equations that allow more efficient computation of Catalan numbers. In order to do this, we define a tail-recursive function that uses the recurrence @{thm catalan_Suc'[no_vars]}: \ qualified function catalan_aux where [simp del]: "catalan_aux n k acc = (if k \ n then acc else catalan_aux n (Suc k) ((acc * (2*(2*k+1))) div (k+2)))" by auto termination by (relation "Wellfounded.measure (\(a,b,_). a - b)") simp_all qualified lemma catalan_aux_simps: "k \ n \ catalan_aux n k acc = acc" "k < n \ catalan_aux n k acc = catalan_aux n (Suc k) ((acc * (2*(2*k+1))) div (k+2))" by (subst catalan_aux.simps, simp)+ qualified lemma catalan_aux_correct: assumes "k \ n" shows "catalan_aux n k (catalan k) = catalan n" using assms proof (induction n k "catalan k" rule: catalan_aux.induct) case (1 n k) show ?case proof (cases "k < n") case True hence "catalan_aux n k (catalan k) = catalan_aux n (Suc k) (catalan (Suc k))" by (subst catalan_Suc') (simp_all add: catalan_aux_simps) with 1 True show ?thesis by (simp add: catalan_Suc') qed (insert "1.prems", simp_all add: catalan_aux_simps) qed lemma catalan_code [code]: "catalan n = catalan_aux n 0 1" using catalan_aux_correct[of 0 n] by simp (*<*) end (*>*) end diff --git a/thys/Dirichlet_L/Dirichlet_L_Functions.thy b/thys/Dirichlet_L/Dirichlet_L_Functions.thy --- a/thys/Dirichlet_L/Dirichlet_L_Functions.thy +++ b/thys/Dirichlet_L/Dirichlet_L_Functions.thy @@ -1,1267 +1,1267 @@ (* File: Dirichlet_L_Functions.thy Author: Manuel Eberl, TU München *) section \Dirichlet $L$-functions\ theory Dirichlet_L_Functions imports Dirichlet_Characters "HOL-Library.Landau_Symbols" "Zeta_Function.Zeta_Function" begin text \ We can now define the Dirichlet $L$-functions. These are essentially the functions in the complex plane that the Dirichlet series $\sum_{k=1}^\infty \chi(k) k^{-s}$ converge to, for some fixed Dirichlet character $\chi$. First of all, we need to take care of a syntactical problem: The notation for vectors uses $\chi$ as syntax, which causes some annoyance to us, so we disable it locally. \ (*<*) bundle vec_lambda_notation begin notation vec_lambda (binder "\" 10) end bundle no_vec_lambda_notation begin no_notation vec_lambda (binder "\" 10) end (*>*) subsection \Definition and basic properties\ (*<*) context includes no_vec_lambda_notation begin (*>*) text \ We now define Dirichlet $L$ functions as a finite linear combination of Hurwitz $\zeta$ functions. This has the advantage that we directly get the analytic continuation over the full domain and only need to prove that the series really converges to this definition whenever it does converge, which is not hard to do. \ definition Dirichlet_L :: "nat \ (nat \ complex) \ complex \ complex" where "Dirichlet_L m \ s = (if s = 1 then if \ = principal_dchar m then 0 else eval_fds (fds \) 1 else of_nat m powr - s * (\k = 1..m. \ k * hurwitz_zeta (real k / real m) s))" lemma Dirichlet_L_conv_hurwitz_zeta_nonprincipal: assumes "s \ 1" shows "Dirichlet_L n \ s = of_nat n powr -s * (\k = 1..n. \ k * hurwitz_zeta (real k / real n) s)" using assms by (simp add: Dirichlet_L_def) text \ Analyticity everywhere except $1$ is trivial by the above definition, since the Hurwitz $\zeta$ function is analytic everywhere except $1$. For $L$ functions of non principal characters, we will have to show the analyticity at $1$ separately later. \ lemma holomorphic_Dirichlet_L_weak: assumes "m > 0" "1 \ A" shows "Dirichlet_L m \ holomorphic_on A" proof - have "(\s. of_nat m powr - s * (\k = 1..m. \ k * hurwitz_zeta (real k / real m) s)) holomorphic_on A" using assms unfolding Dirichlet_L_def by (intro holomorphic_intros) auto also have "?this \ ?thesis" using assms by (intro holomorphic_cong refl) (auto simp: Dirichlet_L_def) finally show ?thesis . qed (*<*) end (*>*) context dcharacter begin (*<*) context includes no_vec_lambda_notation dcharacter_syntax begin (*>*) text \ For a real value greater than 1, the formal Dirichlet series of an $L$ function for some character $\chi$ converges to the L function. \ lemma fixes s :: complex assumes s: "Re s > 1" shows abs_summable_Dirichlet_L: "summable (\n. norm (\ n * of_nat n powr -s))" and summable_Dirichlet_L: "summable (\n. \ n * of_nat n powr -s)" and sums_Dirichlet_L: "(\n. \ n * n powr -s) sums Dirichlet_L n \ s" and Dirichlet_L_conv_eval_fds_weak: "Dirichlet_L n \ s = eval_fds (fds \) s" proof - define L where "L = (\n. \ n * of_nat n powr -s)" show "summable (\n. norm (\ n * of_nat n powr -s))" by (subst summable_Suc_iff [symmetric], rule summable_comparison_test [OF _ summable_zeta_real[of "Re s"]]) (insert s norm, auto intro!: exI[of _ 0] simp: norm_mult norm_powr_real_powr) thus summable: "summable (\n. \ n * of_nat n powr -s)" by (rule summable_norm_cancel) hence "(\n. \ n * of_nat n powr -s) sums L" by (simp add: L_def sums_iff) from this have "(\m. \k = m * n.. k * of_nat k powr - s) sums L" by (rule sums_group) (use n in auto) also have "(\m. \k = m * n.. k * of_nat k powr - s) = (\m. of_nat n powr -s * (\k = 1..n. \ k * (of_nat m + of_nat k / of_nat n) powr - s))" proof (rule ext, goal_cases) case (1 m) have "(\k = m * n.. k * of_nat k powr - s) = (\k=0.. (k + m * n) * of_nat (m * n + k) powr - s)" by (intro sum.reindex_bij_witness[of _ "\k. k + m * n" "\k. k - m * n"]) auto also have "\ = (\k=0.. k * of_nat (m * n + k) powr - s)" by (simp add: periodic_mult) also have "\ = (\k=0.. k * (of_nat m + of_nat k / of_nat n) powr - s * of_nat n powr -s)" proof (intro sum.cong refl, goal_cases) case (1 k) have "of_nat (m * n + k) = (of_nat m + of_nat k / of_nat n :: complex) * of_nat n" using n by (simp add: divide_simps del: div_mult_self1 div_mult_self2 div_mult_self3 div_mult_self4) also have "\ powr -s = (of_nat m + of_nat k / of_nat n) powr -s * of_nat n powr -s" by (rule powr_times_real) auto finally show ?case by simp qed also have "\ = of_nat n powr -s * (\k=0.. k * (of_nat m + of_nat k / of_nat n) powr - s)" by (subst sum_distrib_left) (simp_all add: mult_ac) also have "(\k = 0.. k * (of_nat m + of_nat k / of_nat n) powr - s) = (\k = 1.. k * (of_nat m + of_nat k / of_nat n) powr - s)" by (intro sum.mono_neutral_right) (auto simp: Suc_le_eq) also have "\ = (\k = 1..n. \ k * (of_nat m + of_nat k / of_nat n) powr - s)" using periodic_mult[of 0 1] by (intro sum.mono_neutral_left) auto finally show ?case . qed finally have "\ sums L" . moreover have "(\m. of_nat n powr - s * (\k=1..n. \ k * (of_nat m + of_real (of_nat k / of_nat n)) powr - s)) sums (of_nat n powr - s * (\k=1..n. \ k * hurwitz_zeta (of_nat k / of_nat n) s))" using s by (intro sums_sum sums_mult sums_hurwitz_zeta) auto ultimately have "L = \" by (simp add: sums_iff) also have "\ = Dirichlet_L n \ s" using assms by (auto simp: Dirichlet_L_def) finally have "Dirichlet_L n \ s = (\n. \ n * of_nat n powr -s)" by (simp add: L_def) with summable show "(\n. \ n * n powr -s) sums Dirichlet_L n \ s" by (simp add: sums_iff L_def) thus "Dirichlet_L n \ s = eval_fds (fds \) s" by (simp add: eval_fds_def sums_iff powr_minus field_simps fds_nth_fds') qed lemma fds_abs_converges_weak: "Re s > 1 \ fds_abs_converges (fds \) s" using abs_summable_Dirichlet_L[of s] by (simp add: fds_abs_converges_def powr_minus divide_simps fds_nth_fds') lemma abs_conv_abscissa_weak: "abs_conv_abscissa (fds \) \ 1" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) thus ?case by (intro exI[of _ "of_real c"] conjI fds_abs_converges_weak) auto qed text \ Dirichlet $L$ functions have the Euler product expansion \[L(\chi, s) = \prod_p \left(1 - \frac{\chi(p)}{p^{-s}}\right)\] for all $s$ with $\mathfrak{R}(s) > 1$. \ lemma fixes s :: complex assumes s: "Re s > 1" shows Dirichlet_L_euler_product_LIMSEQ: "(\n. \p\n. if prime p then inverse (1 - \ p / nat_power p s) else 1) \ Dirichlet_L n \ s" (is ?th1) and Dirichlet_L_abs_convergent_euler_product: "abs_convergent_prod (\p. if prime p then inverse (1 - \ p / p powr s) else 1)" (is ?th2) proof - have mult: "completely_multiplicative_function (fds_nth (fds \))" using mult.completely_multiplicative_function_axioms by (simp add: fds_nth_fds') have conv: "fds_abs_converges (fds \) s" using abs_summable_Dirichlet_L[OF s] by (simp add: fds_abs_converges_def fds_nth_fds' powr_minus divide_simps) have "(\n. \p\n. if prime p then inverse (1 - \ p / nat_power p s) else 1) \ eval_fds (fds \) s" using fds_euler_product_LIMSEQ' [OF mult conv] by (simp add: fds_nth_fds' cong: if_cong) also have "eval_fds (fds \) s = Dirichlet_L n \ s" using sums_Dirichlet_L[OF s] unfolding eval_fds_def by (simp add: sums_iff fds_nth_fds' powr_minus divide_simps) finally show ?th1 . from fds_abs_convergent_euler_product' [OF mult conv] show ?th2 by (simp add: fds_nth_fds cong: if_cong) qed lemma Dirichlet_L_Re_gt_1_nonzero: assumes "Re s > 1" shows "Dirichlet_L n \ s \ 0" proof - have "completely_multiplicative_function (fds_nth (fds \))" by (simp add: fds_nth_fds' mult.completely_multiplicative_function_axioms) moreover have "fds_abs_converges (fds \) s" using abs_summable_Dirichlet_L[OF assms] by (simp add: fds_abs_converges_def fds_nth_fds' powr_minus divide_simps) ultimately have "(eval_fds (fds \) s = 0) \ (\p. prime p \ fds_nth (fds \) p = nat_power p s)" by (rule fds_abs_convergent_zero_iff) also have "eval_fds (fds \) s = Dirichlet_L n \ s" using Dirichlet_L_conv_eval_fds_weak[OF assms] by simp also have "\(\p. prime p \ fds_nth (fds \) p = nat_power p s)" proof safe fix p :: nat assume p: "prime p" "fds_nth (fds \) p = nat_power p s" from p have "real 1 < real p" by (subst of_nat_less_iff) (auto simp: prime_gt_Suc_0_nat) also have "\ = real p powr 1" by simp also from p and assms have "real p powr 1 \ real p powr Re s" by (intro powr_mono) (auto simp: real_of_nat_ge_one_iff prime_ge_Suc_0_nat) also have "\ = norm (nat_power p s)" by (simp add: norm_nat_power norm_powr_real_powr) also have "nat_power p s = fds_nth (fds \) p" using p by simp also have "norm \ \ 1" by (auto simp: fds_nth_fds' norm) finally show False by simp qed finally show ?thesis . qed lemma sum_dcharacter_antimono_bound: fixes x0 a b :: real and f f' :: "real \ real" assumes nonprincipal: "\ \ \\<^sub>0" assumes x0: "x0 \ 0" and ab: "x0 \ a" "a < b" assumes f': "\x. x \ x0 \ (f has_field_derivative f' x) (at x)" assumes f_nonneg: "\x. x \ x0 \ f x \ 0" assumes f'_nonpos: "\x. x \ x0 \ f' x \ 0" shows "norm (\n\real -` {a<..b}. \ n * (f (real n))) \ 2 * real (totient n) * f a" proof - note deriv = has_field_derivative_at_within [OF f'] let ?A = "sum_upto \" have cont: "continuous_on {a..b} f" by (rule DERIV_continuous_on[OF deriv]) (use ab in auto) have I': "(f' has_integral (f b - f a)) {a..b}" using ab deriv by (intro fundamental_theorem_of_calculus) - (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric]) + (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric]) define I where "I = integral {a..b} (\t. ?A t * of_real (f' t))" define C where "C = real (totient n)" have C_nonneg: "C \ 0" by (simp add: C_def) have C: "norm (?A x) \ C" for x proof - have "?A x = (\k\nat \x\. \ k)" unfolding sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C" unfolding C_def using nonprincipal by (rule sum_dcharacter_atMost_le) finally show ?thesis . qed have I: "((\t. ?A t * f' t) has_integral ?A b * f b - ?A a * f a - (\n\real -` {a<..b}. \ n * f (real n))) {a..b}" using ab x0 cont f' by (intro partial_summation_strong[of "{}"] has_vector_derivative_of_real) auto hence "(\n\real -` {a<..b}. \ n * f (real n)) = ?A b * f b - ?A a * f a - I" by (simp add: has_integral_iff I_def) also have "norm \ \ norm (?A b) * norm (f b) + norm (?A a) * norm (f a) + norm I" by (rule order.trans[OF norm_triangle_ineq4] add_mono)+ (simp_all add: norm_mult) also have "norm I \ integral {a..b} (\t. of_real (-C) * of_real (f' t))" unfolding I_def using I I' f'_nonpos ab C by (intro integral_norm_bound_integral integrable_on_cmult_left) (simp_all add: has_integral_iff norm_mult mult_right_mono_neg) also have "\ = - (C * (f b - f a))" using integral_linear[OF _ bounded_linear_of_real, of f' "{a..b}"] I' by (simp add: has_integral_iff o_def ) also have "\ = C * (f a - f b)" by (simp add: algebra_simps) also have "norm (sum_upto \ b) \ C" by (rule C) also have "norm (sum_upto \ a) \ C" by (rule C) also have "C * norm (f b) + C * norm (f a) + C * (f a - f b) = 2 * C * f a" using f_nonneg[of a] f_nonneg[of b] ab by (simp add: algebra_simps) finally show ?thesis by (simp add: mult_right_mono C_def) qed lemma summable_dcharacter_antimono: fixes x0 a b :: real and f f' :: "real \ real" assumes nonprincipal: "\ \ \\<^sub>0" assumes f': "\x. x \ x0 \ (f has_field_derivative f' x) (at x)" assumes f_nonneg: "\x. x \ x0 \ f x \ 0" assumes f'_nonpos: "\x. x \ x0 \ f' x \ 0" assumes lim: "(f \ 0) at_top" shows "summable (\n. \ n * f n)" proof (rule summable_bounded_partials [where ?g = "\x. 2 * real (totient n) * f x"], goal_cases) case 1 from eventually_ge_at_top[of "nat \x0\"] show ?case proof eventually_elim case (elim x) show ?case proof (safe, goal_cases) case (1 a b) with elim have *: "max 0 x0 \ 0" "max 0 x0 \ a" "real a < real b" by (simp_all add: nat_le_iff ceiling_le_iff) have "(\n\{a<..b}. \ n * complex_of_real (f (real n))) = (\n\real -` {real a<..real b}. \ n * complex_of_real (f (real n)))" by (intro sum.cong refl) auto also have "norm \ \ 2 * real (totient n) * f a" using nonprincipal * f' f_nonneg f'_nonpos by (rule sum_dcharacter_antimono_bound) simp_all finally show ?case . qed qed qed (auto intro!: tendsto_mult_right_zero filterlim_compose[OF lim] filterlim_real_sequentially) lemma conv_abscissa_le_0: fixes s :: real assumes nonprincipal: "\ \ \\<^sub>0" shows "conv_abscissa (fds \) \ 0" proof (rule conv_abscissa_leI) fix s assume s: "0 < ereal s" have "summable (\n. \ n * of_real (n powr -s))" proof (rule summable_dcharacter_antimono[of 1]) fix x :: real assume "x \ 1" thus "((\x. x powr -s) has_field_derivative (-s * x powr (-s-1))) (at x)" by (auto intro!: derivative_eq_intros) qed (insert s assms, auto intro!: tendsto_neg_powr filterlim_ident) thus "\s'::complex. s' \ 1 = s \ fds_converges (fds \) s'" using s by (intro exI[of _ "of_real s"]) (auto simp: fds_converges_def powr_minus divide_simps powr_of_real [symmetric] fds_nth_fds') qed lemma summable_Dirichlet_L': assumes nonprincipal: "\ \ \\<^sub>0" assumes s: "Re s > 0" shows "summable (\n. \ n * of_nat n powr -s)" proof - from assms have "fds_converges (fds \) s" by (intro fds_converges le_less_trans[OF conv_abscissa_le_0]) auto thus ?thesis by (simp add: fds_converges_def powr_minus divide_simps fds_nth_fds') qed lemma assumes "\ \ \\<^sub>0" shows Dirichlet_L_conv_eval_fds: "\s. Re s > 0 \ Dirichlet_L n \ s = eval_fds (fds \) s" and holomorphic_Dirichlet_L: "Dirichlet_L n \ holomorphic_on A" proof - show eq: "Dirichlet_L n \ s = eval_fds (fds \) s" (is "?f s = ?g s") if "Re s > 0" for s proof (cases "s = 1") case False show ?thesis proof (rule analytic_continuation_open[where ?f = ?f and ?g = ?g]) show "{s. Re s > 1} \ {s. Re s > 0} - {1}" by auto show "connected ({s. 0 < Re s} - {1})" using aff_dim_halfspace_gt[of 0 "1::complex"] by (intro connected_punctured_convex convex_halfspace_Re_gt) auto qed (insert that n assms False, auto intro!: convex_halfspace_Re_gt open_halfspace_Re_gt exI[of _ 2] holomorphic_intros holomorphic_Dirichlet_L_weak Dirichlet_L_conv_eval_fds_weak le_less_trans[OF conv_abscissa_le_0]) qed (insert assms, simp_all add: Dirichlet_L_def) have "Dirichlet_L n \ holomorphic_on UNIV" proof (rule no_isolated_singularity') from n show "Dirichlet_L n \ holomorphic_on (UNIV - {1})" by (intro holomorphic_Dirichlet_L_weak) auto next fix s :: complex assume s: "s \ {1}" show "Dirichlet_L n \ \s\ Dirichlet_L n \ s" proof (rule Lim_transform_eventually) from assms have "continuous_on {s. Re s > 0} (eval_fds (fds \))" by (intro holomorphic_fds_eval holomorphic_on_imp_continuous_on) (auto intro: le_less_trans[OF conv_abscissa_le_0]) hence "eval_fds (fds \) \s\ eval_fds (fds \) s" using s by (subst (asm) continuous_on_eq_continuous_at) (auto simp: open_halfspace_Re_gt isCont_def) also have "eval_fds (fds \) s = Dirichlet_L n \ s" using assms s by (simp add: Dirichlet_L_def) finally show "eval_fds (fds \) \s\ Dirichlet_L n \ s" . next have "eventually (\z. z \ {z. Re z > 0}) (nhds s)" using s by (intro eventually_nhds_in_open) (auto simp: open_halfspace_Re_gt) hence "eventually (\z. z \ {z. Re z > 0}) (at s)" unfolding eventually_at_filter by eventually_elim auto then show "eventually (\z. eval_fds (fds \) z = Dirichlet_L n \ z) (at s)" by eventually_elim (auto intro!: eq [symmetric]) qed qed auto thus "Dirichlet_L n \ holomorphic_on A" by (rule holomorphic_on_subset) auto qed lemma cnj_Dirichlet_L: "cnj (Dirichlet_L n \ s) = Dirichlet_L n (inv_character \) (cnj s)" proof - { assume *: "\ \ \\<^sub>0" "s = 1" with summable_Dirichlet_L'[of 1] have "(\n. \ n / n) sums eval_fds (fds \) 1" by (simp add: eval_fds_def fds_nth_fds' powr_minus sums_iff divide_simps) hence "(\n. inv_character \ n / n) sums cnj (eval_fds (fds \) 1)" by (subst (asm) sums_cnj [symmetric]) (simp add: inv_character_def) hence "eval_fds (fds (inv_character \)) 1 = cnj (eval_fds (fds \) 1)" by (simp add: eval_fds_def fds_nth_fds' inv_character_def sums_iff) } thus ?thesis by (auto simp add: Dirichlet_L_def cnj_powr eval_inv_character) qed (*<*) end (*>*) end (*<*) context includes no_vec_lambda_notation begin (*>*) lemma holomorphic_Dirichlet_L [holomorphic_intros]: assumes "n > 1" "\ \ principal_dchar n \ dcharacter n \ \ \ = principal_dchar n \ 1 \ A" shows "Dirichlet_L n \ holomorphic_on A" using assms(2) proof assume "\ = principal_dchar n \ 1 \ A" with holomorphic_Dirichlet_L_weak[of n A "principal_dchar n"] assms(1) show ?thesis by auto qed (insert dcharacter.holomorphic_Dirichlet_L[of n \ A], auto) lemma holomorphic_Dirichlet_L' [holomorphic_intros]: assumes "n > 1" "f holomorphic_on A" "\ \ principal_dchar n \ dcharacter n \ \ \ = principal_dchar n \ (\x\A. f x \ 1)" shows "(\s. Dirichlet_L n \ (f s)) holomorphic_on A" using holomorphic_on_compose[OF assms(2) holomorphic_Dirichlet_L[OF assms(1), of \]] assms by (auto simp: o_def image_iff) lemma continuous_on_Dirichlet_L: assumes "n > 1" "\ \ principal_dchar n \ dcharacter n \ \ \ = principal_dchar n \ 1 \ A" shows "continuous_on A (Dirichlet_L n \)" using assms by (intro holomorphic_on_imp_continuous_on holomorphic_intros) lemma continuous_on_Dirichlet_L' [continuous_intros]: assumes "continuous_on A f" "n > 1" and "\ \ principal_dchar n \ dcharacter n \ \ \ = principal_dchar n \ (\x\A. f x \ 1)" shows "continuous_on A (\x. Dirichlet_L n \ (f x))" using continuous_on_compose2[OF continuous_on_Dirichlet_L[of n \ "f ` A"] assms(1)] assms by (auto simp: image_iff) corollary continuous_Dirichlet_L [continuous_intros]: "n > 1 \ \ \ principal_dchar n \ dcharacter n \ \ \ = principal_dchar n \ s \ 1 \ continuous (at s within A) (Dirichlet_L n \)" by (rule continuous_within_subset[of _ UNIV]) (insert continuous_on_Dirichlet_L[of n \ "(if \ = principal_dchar n then -{1} else UNIV)"], auto simp: continuous_on_eq_continuous_at open_Compl) corollary continuous_Dirichlet_L' [continuous_intros]: "n > 1 \ continuous (at s within A) f \ \ \ principal_dchar n \ dcharacter n \ \ \ = principal_dchar n \ f s \ 1 \ continuous (at s within A) (\x. Dirichlet_L n \ (f x))" by (rule continuous_within_compose3[OF continuous_Dirichlet_L]) auto (*<*) end (*>*) context residues_nat begin (*<*) context includes no_vec_lambda_notation dcharacter_syntax begin (*>*) text \ Applying the above to the $L(\chi_0,s)$, the $L$ function of the principal character, we find that it differs from the Riemann $\zeta$ function only by multiplication with a constant that depends only on the modulus $n$. They therefore have the same analytic properties as the $\zeta$ function itself. \ lemma Dirichlet_L_principal: fixes s :: complex shows "Dirichlet_L n \\<^sub>0 s = (\p | prime p \ p dvd n. (1 - 1 / p powr s)) * zeta s" (is "?f s = ?g s") proof (cases "s = 1") case False show ?thesis proof (rule analytic_continuation_open[where ?f = ?f and ?g = ?g]) show "{s. Re s > 1} \ - {1}" by auto show "?f s = ?g s" if "s \ {s. Re s > 1}" for s proof - from that have s: "Re s > 1" by simp let ?P = "(\p | prime p \ p dvd n. (1 - 1 / p powr s))" have "(\n. \p\n. if prime p then inverse (1 - \\<^sub>0 p / nat_power p s) else 1) \ Dirichlet_L n \\<^sub>0 s" using s by (rule principal.Dirichlet_L_euler_product_LIMSEQ) also have "?this \ (\n. ?P * (\p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1)) \ Dirichlet_L n \\<^sub>0 s" (is "_ = filterlim ?g _ _") proof (intro tendsto_cong eventually_mono [OF eventually_ge_at_top, of n], goal_cases) case (1 m) let ?f = "\p. inverse (1 - 1 / p powr s)" have "(\p\m. if prime p then inverse (1 - \\<^sub>0 p / nat_power p s) else 1) = (\p | p \ m \ prime p \ coprime p n. ?f p)" (is "_ = prod _ ?A") by (intro prod.mono_neutral_cong_right) (auto simp: principal_dchar_def) also have "?A = {p. p \ m \ prime p} - {p. prime p \ p dvd n}" (is "_ = ?B - ?C") using n by (auto dest: prime_imp_coprime simp: coprime_absorb_left) also { have *: "(\p\?B. ?f p) = (\p\?B - ?C. ?f p) * (\p\?C. ?f p)" using 1 n by (intro prod.subset_diff) (auto dest: dvd_imp_le) have "(\p\?B. ?f p) * ?P = (\p\?B - ?C. ?f p) * ((\p\?C. ?f p) * ?P)" by (subst *) (simp add: mult_ac) also have "(\p\?C. ?f p) * ?P = (\p\?C. 1)" by (subst prod.distrib [symmetric], rule prod.cong) (insert s, auto simp: divide_simps powr_def exp_eq_1) also have "\ = 1" by simp finally have "(\p\?B - ?C. ?f p) = (\p\?B. ?f p) * ?P" by simp } also have "(\p\?B. ?f p) = (\p\m. if prime p then ?f p else 1)" by (intro prod.mono_neutral_cong_left) auto finally show ?case by (simp only: mult_ac) qed finally have "?g \ Dirichlet_L n \\<^sub>0 s" . moreover have "?g \ ?P * zeta s" by (intro tendsto_mult tendsto_const euler_product_zeta s) ultimately show "Dirichlet_L n \\<^sub>0 s = ?P * zeta s" by (rule LIMSEQ_unique) qed qed (insert \s \ 1\ n, auto intro!: holomorphic_intros holomorphic_Dirichlet_L_weak open_halfspace_Re_gt exI[of _ 2] connected_punctured_universe) qed (simp_all add: Dirichlet_L_def zeta_1) (*<*) end (*>*) end subsection \The non-vanishing for $\mathfrak{R}(s)\geq 1$\ lemma coprime_prime_exists: assumes "n > (0 :: nat)" obtains p where "prime p" "coprime p n" proof - from bigger_prime[of n] obtain p where p: "prime p" "p > n" by auto with assms have "\p dvd n" by (auto dest: dvd_imp_le) with p have "coprime p n" by (intro prime_imp_coprime) with that[of p] and p show ?thesis by auto qed text \ The case of the principal character is trivial, since it differs from the Riemann $\zeta(s)$ only in a multiplicative factor that is clearly non-zero for $\mathfrak{R}(s) \geq 1$. \ theorem (in residues_nat) Dirichlet_L_Re_ge_1_nonzero_principal: assumes "Re s \ 1" "s \ 1" shows "Dirichlet_L n (principal_dchar n) s \ 0" proof - have "(\p | prime p \ p dvd n. 1 - 1 / p powr s) \ (0 :: complex)" proof (subst prod_zero_iff) from n show "finite {p. prime p \ p dvd n}" by (intro finite_prime_divisors) auto show "\(\p\{p. prime p \ p dvd n}. 1 - 1 / p powr s = 0)" proof safe fix p assume p: "prime p" "p dvd n" and "1 - 1 / p powr s = 0" hence "norm (p powr s) = 1" by simp also have "norm (p powr s) = real p powr Re s" by (simp add: norm_powr_real_powr) finally show False using p assms by (simp add: powr_def prime_gt_0_nat) qed qed with zeta_Re_ge_1_nonzero[OF assms] show ?thesis by (simp add: Dirichlet_L_principal) qed text \ The proof for non-principal character is quite involved and is typically very complicated and technical in most textbooks. For instance, Apostol~\cite{apostol1976analytic} proves the result separately for real and non-real characters, where the non-real case is relatively short and nice, but the real case involves a number of complicated asymptotic estimates. The following proof, on the other hand -- like our proof of the analogous result for the Riemann $\zeta$ function -- is based on Newman's book~\cite{newman1998analytic}. Newman gives a very short, concise, and high-level sketch that we aim to reproduce faithfully here. \ context dcharacter begin (*<*) context includes no_vec_lambda_notation dcharacter_syntax begin (*>*) theorem Dirichlet_L_Re_ge_1_nonzero_nonprincipal: assumes "\ \ \\<^sub>0" and "Re u \ 1" shows "Dirichlet_L n \ u \ 0" proof (cases "Re u > 1") include dcharacter_syntax case False define a where "a = -Im u" from False and assms have "Re u = 1" by simp hence [simp]: "u = 1 - \ * a" by (simp add: a_def complex_eq_iff) show ?thesis proof assume "Dirichlet_L n \ u = 0" hence zero: "Dirichlet_L n \ (1 - \ * a) = 0" by simp define \' where [simp]: "\' = inv_character \" \ \We define the function $Z(s)$, which is the product of all the Dirichlet $L$ functions, and its Dirichlet series. Then, similarly to the proof of the non-vanishing of the Riemann $\zeta$ function for $\mathfrak{R}(s) \geq 1$, we define $Q(s) = Z(s) Z(s + ia) Z(s - ia)$. Our objective is to show that the Dirichlet series of this function $Q$ converges everywhere.\ define Z where "Z = (\s. \\\dcharacters n. Dirichlet_L n \ s)" define Z_fds where "Z_fds = (\\\dcharacters n. fds \)" define Q where "Q = (\s. Z s ^ 2 * Z (s + \ * a) * Z (s - \ * a))" define Q_fds where "Q_fds = Z_fds ^ 2 * fds_shift (\ * a) Z_fds * fds_shift (-\ * a) Z_fds" let ?sings = "{1, 1 + \ * a, 1 - \ * a}" \ \Some preliminary auxiliary facts\ define P where "P = (\s. (\x\{p. prime p \ p dvd n}. 1 - 1 / of_nat x powr s :: complex))" have \\<^sub>0: "\\<^sub>0 \ dcharacters n" by (auto simp: principal.dcharacter_axioms dcharacters_def) have [continuous_intros]: "continuous_on A P" for A unfolding P_def by (intro continuous_intros) (auto simp: prime_gt_0_nat) from this[of UNIV] have [continuous_intros]: "isCont P s" for s by (auto simp: continuous_on_eq_continuous_at) have \: "\ \ dcharacters n" "\' \ dcharacters n" using dcharacter_axioms by (auto simp add: dcharacters_def dcharacter.dcharacter_inv_character) from zero dcharacter.cnj_Dirichlet_L[of n \ "1 - \ * a"] dcharacter_axioms have zero': "Dirichlet_L n \' (1 + \ * a) = 0" by simp have "Z = (\s. Dirichlet_L n \\<^sub>0 s * (\\\dcharacters n - {\\<^sub>0}. Dirichlet_L n \ s))" unfolding Z_def using \\<^sub>0 by (intro ext prod.remove) auto also have "\ = (\s. P s * zeta s * (\\\dcharacters n - {\\<^sub>0}. Dirichlet_L n \ s))" by (simp add: Dirichlet_L_principal P_def) finally have Z_eq: "Z = (\s. P s * zeta s * (\\\dcharacters n - {\\<^sub>0}. Dirichlet_L n \ s))" . have Z_eq': "Z = (\s. P s * zeta s * Dirichlet_L n \ s * (\\\dcharacters n - {\\<^sub>0} - {\}. Dirichlet_L n \ s))" if "\ \ dcharacters n" "\ \ \\<^sub>0" for \ proof (rule ext, goal_cases) case (1 s) from that have \: "\ \ dcharacters n" by (simp add: dcharacters_def) have "Z s = P s * zeta s * (\\\dcharacters n - {\\<^sub>0}. Dirichlet_L n \ s)" by (simp add: Z_eq) also have "(\\\dcharacters n - {\\<^sub>0}. Dirichlet_L n \ s) = Dirichlet_L n \ s * (\\\dcharacters n - {\\<^sub>0} - {\}. Dirichlet_L n \ s)" using assms \ that by (intro prod.remove) auto finally show ?case by (simp add: mult_ac) qed \ \We again show that @{term Q} is locally bounded everywhere by showing that every singularity is cancelled by some zero. Since now, @{term a} can be zero, we do a case distinction here to make things a bit easier.\ have Q_bigo_1: "Q \ O[at s](\_. 1)" for s proof (cases "a = 0") case True have "(\s. Dirichlet_L n \ s - Dirichlet_L n \ 1) \ O[at 1](\s. s - 1)" using \ assms n by (intro taylor_bigo_linear holomorphic_on_imp_differentiable_at[of _ UNIV] holomorphic_intros) (auto simp: dcharacters_def) hence *: "Dirichlet_L n \ \ O[at 1](\s. s - 1)" using zero True by simp have "Z = (\s. P s * zeta s * Dirichlet_L n \ s * (\\\dcharacters n - {\\<^sub>0} - {\}. Dirichlet_L n \ s))" using \ assms by (intro Z_eq') auto also have "\ \ O[at 1](\s. 1 * (1 / (s - 1)) * (s - 1) * 1)" using n \ by (intro landau_o.big.mult continuous_imp_bigo_1 zeta_bigo_at_1 continuous_intros *) (auto simp: dcharacters_def) also have "(\s::complex. 1 * (1 / (s - 1)) * (s - 1) * 1) \ \[at 1](\_. 1)" by (intro bigthetaI_cong) (auto simp: eventually_at_filter) finally have Z_at_1: "Z \ O[at 1](\_. 1)" . have "Z \ O[at s](\_. 1)" proof (cases "s = 1") case False thus ?thesis unfolding Z_def using n \ by (intro continuous_imp_bigo_1 continuous_intros) (auto simp: dcharacters_def) qed (insert Z_at_1, auto) from \a = 0\ have "Q = (\s. Z s * Z s * Z s * Z s)" by (simp add: Q_def power2_eq_square) also have "\ \ O[at s](\_. 1 * 1 * 1 * 1)" by (intro landau_o.big.mult) fact+ finally show ?thesis by simp next case False have bigo1: "(\s. Z s * Z (s - \ * a)) \ O[at 1](\_. 1)" if "Dirichlet_L n \ (1 - \ * a) = 0" "a \ 0" "\ \ dcharacters n" "\ \ \\<^sub>0" for a :: real and \ proof - have "(\s. Dirichlet_L n \ (s - \ * a) - Dirichlet_L n \ (1 - \ * a)) \ O[at 1](\s. s - 1)" using assms n that by (intro taylor_bigo_linear holomorphic_on_imp_differentiable_at[of _ UNIV] holomorphic_intros) (auto simp: dcharacters_def) hence *: "(\s. Dirichlet_L n \ (s - \ * a)) \ O[at 1](\s. s - 1)" using that by simp have "(\s. Z (s - \*a)) = (\s. P (s - \*a) * zeta (s - \*a) * Dirichlet_L n \ (s - \*a) * (\\\dcharacters n - {\\<^sub>0} - {\}. Dirichlet_L n \ (s - \*a)))" using that by (subst Z_eq'[of \]) auto also have "\ \ O[at 1](\s. 1 * 1 * (s - 1) * 1)" unfolding P_def using that n by (intro landau_o.big.mult continuous_imp_bigo_1 continuous_intros *) (auto simp: prime_gt_0_nat dcharacters_def) finally have "(\s. Z (s - \ * a)) \ O[at 1](\s. s - 1)" by simp moreover have "Z \ O[at 1](\s. 1 * (1 / (s - 1)) * 1)" unfolding Z_eq using n that by (intro landau_o.big.mult zeta_bigo_at_1 continuous_imp_bigo_1 continuous_intros) (auto simp: dcharacters_def) hence "Z \ O[at 1](\s. 1 / (s - 1))" by simp ultimately have "(\s. Z s * Z (s - \ * a)) \ O[at 1](\s. 1 / (s - 1) * (s - 1))" by (intro landau_o.big.mult) also have "(\s. 1 / (s - 1) * (s - 1)) \ \[at 1](\_. 1)" by (intro bigthetaI_cong) (auto simp add: eventually_at_filter) finally show ?thesis . qed have bigo1': "(\s. Z s * Z (s + \ * a)) \ O[at 1](\_. 1)" if "Dirichlet_L n \ (1 - \ * a) = 0" "a \ 0" "\ \ dcharacters n" "\ \ \\<^sub>0" for a :: real and \ proof - from that interpret dcharacter n G \ by (simp_all add: dcharacters_def G_def) from bigo1[of "inv_character \" "-a"] that cnj_Dirichlet_L[of "1 - \ * a"] show ?thesis by (simp add: dcharacters_def dcharacter_inv_character) qed have bigo2: "(\s. Z s * Z (s - \ * a)) \ O[at (1 + \ * a)](\_. 1)" if "Dirichlet_L n \ (1 - \ * a) = 0" "a \ 0" "\ \ dcharacters n" "\ \ \\<^sub>0" for a :: real and \ proof - have "(\s. Z s * Z (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 have bigo2': "(\s. Z s * Z (s + \ * a)) \ O[at (1 - \ * a)](\_. 1)" if "Dirichlet_L n \ (1 - \ * a) = 0" "a \ 0" "\ \ dcharacters n" "\ \ \\<^sub>0" for a :: real and \ proof - from that interpret dcharacter n G \ by (simp_all add: dcharacters_def G_def) from bigo2[of "inv_character \" "-a"] that cnj_Dirichlet_L[of "1 - \ * a"] show ?thesis by (simp add: dcharacters_def dcharacter_inv_character) qed have Q_eq: "Q = (\s. (Z s * Z (s + \ * a)) * (Z s * Z (s - \ * a)))" by (simp add: Q_def power2_eq_square mult_ac) consider "s = 1" | "s = 1 + \ * a" | "s = 1 - \ * a" | "s \ ?sings" by blast thus ?thesis proof cases case 1 have "Q \ O[at 1](\_. 1 * 1)" unfolding Q_eq using assms zero zero' False \ by (intro landau_o.big.mult bigo1[of \ a] bigo1'[of \ a]; simp)+ with 1 show ?thesis by simp next case 2 have "Q \ O[at (1 + \ * a)](\_. 1 * 1)" unfolding Q_eq using assms zero zero' False \ n by (intro landau_o.big.mult bigo2[of \ a] continuous_imp_bigo_1) (auto simp: Z_def dcharacters_def intro!: continuous_intros) with 2 show ?thesis by simp next case 3 have "Q \ O[at (1 - \ * a)](\_. 1 * 1)" unfolding Q_eq using assms zero zero' False \ n by (intro landau_o.big.mult bigo2'[of \ a] continuous_imp_bigo_1) (auto simp: Z_def dcharacters_def intro!: continuous_intros) with 3 show ?thesis by simp next case 4 thus ?thesis unfolding Q_def Z_def using n by (intro continuous_imp_bigo_1 continuous_intros) (auto simp: dcharacters_def complex_eq_iff) qed qed \ \Again, 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)" using n by (intro removable_singularities Q_bigo_1) (auto simp: Q_def Z_def dcharacters_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 [simp]: "dcharacter n \" if "\ \ dcharacters n" for \ using that by (simp add: dcharacters_def) from that have "abs_conv_abscissa (fds \) < ereal (Re s)" if "\ \ dcharacters n" for \ using that by (intro le_less_trans[OF dcharacter.abs_conv_abscissa_weak[of n \]]) auto hence "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 eval_fds_prod fds_abs_converges_prod dcharacter.Dirichlet_L_conv_eval_fds_weak fds_abs_converges_power eval_fds_zeta Z_fds_def Z_def fds_abs_converges) also from that have "\ = Q' s" by (subst Q') auto finally show ?thesis . qed \ \Since the characters are completely multiplicative, the series for this logarithm can be rewritten like this:\ define I where "I = (\k. if [k = 1] (mod n) then totient n else 0 :: real)" have ln_Q_fds_eq: "fds_ln 0 Q_fds = fds (\k. of_real (2 * I k * mangoldt k / ln k * (1 + cos (a * ln k))))" proof - have nz: "\ (Suc 0) = 1" if "\ \ dcharacters n" for \ using dcharacter.Suc_0[of n \] that by (simp add: dcharacters_def) 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 = (\\\dcharacters n. 2 * fds_ln 0 (fds \) + fds_shift (\ * a) (fds_ln 0 (fds \)) + fds_shift (-\ * a) (fds_ln 0 (fds \)))" by (auto simp: Q_fds_def Z_fds_def simps nz sum.distrib sum_distrib_left) also have "\ = (\\\dcharacters n. fds (\k. \ k * of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k)))))" (is "(\\\_. ?l \) = _") proof (intro sum.cong refl, goal_cases) case (1 \) then interpret dcharacter n G \ by (simp_all add: dcharacters_def G_def) have mult: "completely_multiplicative_function (fds_nth (fds \))" by (simp add: fds_nth_fds' mult.completely_multiplicative_function_axioms) have *: "fds_ln 0 (fds \) = fds (\n. \ n * mangoldt n /\<^sub>R ln (real n))" by (simp add: fds_ln_completely_multiplicative[OF mult] fds_nth_fds' fds_eq_iff) have "?l \ = fds (\k. \ k * mangoldt k /\<^sub>R ln k * (2 + k powr (\ * a) + k powr (-\ * a)))" by (unfold *, rule fds_eqI) (simp add: algebra_simps scaleR_conv_of_real numeral_fds) also have "\ = fds (\k. \ k * 2 * mangoldt k /\<^sub>R ln k * (1 + cos (of_real(a * ln k))))" unfolding cos_exp_eq by (intro fds_eqI) (simp add: powr_def algebra_simps) also have "\ = fds (\k. \ k * of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k))))" unfolding cos_of_real by (simp add: field_simps scaleR_conv_of_real) finally show ?case . qed also have "\ = fds (\k. (\\\dcharacters n. \ k) * of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k))))" by (simp add: sum_distrib_right sum_divide_distrib scaleR_conv_of_real sum_distrib_left) also have "\ = fds (\k. of_real (2 * I k * mangoldt k / ln k * (1 + cos (a * ln k))))" by (intro fds_eqI, subst sum_dcharacters) (simp_all add: I_def algebra_simps) finally show ?thesis . qed \ \The coefficients of that logarithm series are clearly nonnegative:\ have "nonneg_dirichlet_series (fds_ln 0 Q_fds)" proof show "fds_nth (fds_ln 0 Q_fds) k \ \\<^sub>\\<^sub>0" for k proof (cases "k < 2") case False have cos: "1 + cos x \ 0" for x :: real using cos_ge_minus_one[of x] by linarith have "fds_nth (fds_ln 0 Q_fds) k = of_real (2 * I k * mangoldt k / ln k * (1 + cos (a * ln k)))" by (auto simp: fds_nth_fds' ln_Q_fds_eq) also have "\ \ \\<^sub>\\<^sub>0" using False unfolding I_def by (subst nonneg_Reals_of_real_iff) (intro mult_nonneg_nonneg divide_nonneg_pos cos mangoldt_nonneg, auto) finally show ?thesis . qed (cases k; auto simp: ln_Q_fds_eq) qed \ \Therefore @{term Q_fds} also has non-negative coefficients.\ hence nonneg: "nonneg_dirichlet_series Q_fds" proof (rule nonneg_dirichlet_series_lnD) have "(\x\dcharacters n. x (Suc 0)) = 1" by (intro prod.neutral) (auto simp: dcharacters_def dcharacter.Suc_0) thus "exp 0 = fds_nth Q_fds (Suc 0)" by (simp add: Q_fds_def Z_fds_def) qed \ \And by Pringsheim--Landau, we get that the Dirichlet series of @{term Q} converges everywhere.\ have "abs_conv_abscissa Q_fds \ 1" unfolding Q_fds_def Z_fds_def fds_shift_prod by (intro abs_conv_abscissa_power_leI abs_conv_abscissa_mult_leI abs_conv_abscissa_prod_le) (auto simp: dcharacters_def dcharacter.abs_conv_abscissa_weak) 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 g = Q' and c = 1]) (auto simp: one_ereal_def) \ \Again, similarly to the proof for $\zeta$, we select a subseries of @{term Q}. This time we cannot simply pick powers of 2, since 2 might not be coprime to @{term n}, in which case the subseries would simply be 1 everywhere, which is not helpful. However, it is clear that there \emph{is} always some prime $p$ that is coprime to @{term n}, so we just use the subseries @{term Q} that corresponds to powers of $p$.\ from n obtain p where p: "prime p" "coprime p n" using coprime_prime_exists[of n] by auto define R_fds where "R_fds = fds_primepow_subseries p 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 \ \The following function $g(a,s)$ is the denominator in the Euler product expansion of the subseries of $Z(s + ia)$. It is clear that it is entire and non-zero for $\mathfrak{R}(s) > 0$ and all real $a$.\ define g :: "real \ complex \ complex" where "g = (\a s. (\\\dcharacters n. (1 - \ p * p powr (-s + \ * of_real a))))" have g_nz: "g a s \ 0" if "Re s > 0" for s a unfolding g_def proof (subst prod_zero_iff[OF finite_dcharacters], safe) fix \ assume "\ \ dcharacters n" and *: "1 - \ p * p powr (-s + \*a) = 0" then interpret dcharacter n G \ by (simp_all add: dcharacters_def G_def) from p have "real p > real 1" by (subst of_nat_less_iff) (auto simp: prime_gt_Suc_0_nat) hence "real p powr - Re s < real p powr 0" using p that by (intro powr_less_mono) auto hence "0 < norm (1 :: complex) - norm (\ p * p powr (-s + \*a))" using p by (simp add: norm_mult norm norm_powr_real_powr) also have "\ \ norm (1 - \ p * p powr (-s + \*a))" by (rule norm_triangle_ineq2) finally show False by (subst (asm) *) simp_all qed have [holomorphic_intros]: "g a holomorphic_on A" for a A unfolding g_def using p by (intro holomorphic_intros) \ \By taking Euler product expansions of every factor, we get \[R(s) = \frac{1}{g(0,s)^2 g(a,s) g(-a,s)} = (1 - 2^{-s})^{-2} (1 - 2^{-s+ia})^{-1} (1 - 2^{-s-ia})^{-1}\] for every $s$ with $\mathfrak{R}(s) > 1$, and by analytic continuation also for $\mathfrak{R}(s) > 0$.\ have eval_R: "eval_fds R_fds s = 1 / (g 0 s ^ 2 * g a s * g (-a) s)" (is "_ = ?f s") if "Re s > 0" for s :: complex proof - show ?thesis proof (rule analytic_continuation_open[where f = "eval_fds R_fds"]) show "?f holomorphic_on {s. Re s > 0}" using p g_nz[of 0] g_nz[of a] g_nz[of "-a"] by (intro holomorphic_intros) (auto simp: g_nz) next fix z assume z: "z \ {s. Re s > 1}" have [simp]: "completely_multiplicative_function \" "fds_nth (fds \) = \" if "\ \ dcharacters n" for \ proof - from that interpret dcharacter n G \ by (simp_all add: G_def dcharacters_def) show "completely_multiplicative_function \" "fds_nth (fds \) = \" by (simp_all add: fds_nth_fds' mult.completely_multiplicative_function_axioms) qed have [simp]: "dcharacter n \" if "\ \ dcharacters n" for \ using that by (simp add: dcharacters_def) from that have "abs_conv_abscissa (fds \) < ereal (Re z)" if "\ \ dcharacters n" for \ using that z by (intro le_less_trans[OF dcharacter.abs_conv_abscissa_weak[of n \]]) auto thus "eval_fds R_fds z = ?f z" using z p by (simp add: R_fds_def Q_fds_def Z_fds_def eval_fds_mult eval_fds_prod eval_fds_power fds_abs_converges_mult fds_abs_converges_power fds_abs_converges_prod g_def mult_ac fds_primepow_subseries_euler_product_cm powr_minus powr_diff powr_add prod_dividef fds_abs_summable_zeta g_nz fds_abs_converges power_one_over divide_inverse [symmetric]) qed (insert that abscissa', auto intro!: exI[of _ 2] convex_connected open_halfspace_Re_gt convex_halfspace_Re_gt holomorphic_intros) qed \ \We again have our contradiction: $R(s)$ is entire, but the right-hand side has a pole at 0 since $g(0,0) = 0$.\ show False proof (rule not_tendsto_and_filterlim_at_infinity) have g_limit: "(g a \ g a 0) (at 0 within {s. Re s > 0})" for a proof - have "continuous_on UNIV (g a)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) hence "isCont (g a) 0" by (rule continuous_on_interior) auto hence "continuous (at 0 within {s. Re s > 0}) (g a)" by (rule continuous_within_subset) auto thus ?thesis by (auto simp: continuous_within) qed have "((\s. g 0 s ^ 2 * g a s * g (-a) s) \ g 0 0 ^ 2 * g a 0 * g (-a) 0) (at 0 within {s. Re s > 0})" by (intro tendsto_intros g_limit) also have "g 0 0 = 0" unfolding g_def proof (rule prod_zero) from p and \\<^sub>0 show "\\\dcharacters n. 1 - \ p * of_nat p powr (- 0 + \ * of_real 0) = 0" by (intro bexI[of _ \\<^sub>0]) (auto simp: principal_dchar_def) qed 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. g 0 s ^ 2 * g a s * g (-a) s \ 0) (at 0 within {s. Re s > 0})" by eventually_elim (auto simp: g_nz) ultimately have "filterlim (\s. g 0 s ^ 2 * g a s * g (-a) s) (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 Dirichlet_L_Re_gt_1_nonzero) subsection \Asymptotic bounds on partial sums of Dirichlet $L$ functions\ text \ The following are some bounds on partial sums of the $L$-function of a character that are useful for asymptotic reasoning, particularly for Dirichlet's Theorem. \ lemma sum_upto_dcharacter_le: assumes "\ \ \\<^sub>0" shows "norm (sum_upto \ x) \ totient n" proof - have "sum_upto \ x = (\k\nat \x\. \ k)" unfolding sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ totient n" by (rule sum_dcharacter_atMost_le) fact finally show ?thesis . qed lemma Dirichlet_L_minus_partial_sum_bound: fixes s :: complex and x :: real assumes "\ \ \\<^sub>0" and "Re s > 0" and "x > 0" defines "\ \ Re s" shows "norm (sum_upto (\n. \ n * n powr -s) x - Dirichlet_L n \ s) \ real (totient n) * (2 + cmod s / \) / x powr \" proof (rule Lim_norm_ubound) from assms have "summable (\n. \ n * of_nat n powr -s)" by (intro summable_Dirichlet_L') with assms have "(\n. \ n * of_nat n powr -s) sums Dirichlet_L n \ s" using Dirichlet_L_conv_eval_fds[OF assms(1,2)] by (simp add: sums_iff eval_fds_def powr_minus divide_simps fds_nth_fds') hence "(\m. \k\m. \ k * of_nat k powr -s) \ Dirichlet_L n \ s" by (simp add: sums_def' atLeast0AtMost) thus "(\m. sum_upto (\k. \ k * of_nat k powr -s) x - (\k\m. \ k * of_nat k powr -s)) \ sum_upto (\k. \ k * of_nat k powr -s) x - Dirichlet_L n \ s" by (intro tendsto_intros) next define M where "M = sum_upto \" have le: "norm (\n\real-`{x<..y}. \ n * of_nat n powr - s) \ real (totient n) * (2 + cmod s / \) / x powr \" if xy: "0 < x" "x < y" for x y proof - from xy have I: "((\t. M t * (-s * t powr (-s-1))) has_integral M y * of_real y powr - s - M x * of_real x powr - s - (\n\real-`{x<..y}. \ n * of_real (real n) powr -s)) {x..y}" unfolding M_def by (intro partial_summation_strong [of "{}"]) (auto intro!: has_vector_derivative_real_field derivative_eq_intros continuous_intros) hence "(\n\real-`{x<..y}. \ n * real n powr -s) = M y * of_real y powr - s - M x * of_real x powr - s - integral {x..y} (\t. M t * (-s * t powr (-s-1)))" by (simp add: has_integral_iff) also have "norm \ \ norm (M y * of_real y powr -s) + norm (M x * of_real x powr -s) + norm (integral {x..y} (\t. M t * (-s * t powr (-s-1))))" by (intro order.trans[OF norm_triangle_ineq4] add_mono order.refl) also have "norm (M y * of_real y powr -s) \ totient n * y powr -\" using xy assms unfolding norm_mult M_def \_def by (intro mult_mono sum_upto_dcharacter_le) (auto simp: norm_powr_real_powr) also have "\ \ totient n * x powr -\" using assms xy by (intro mult_left_mono powr_mono2') (auto simp: \_def) also have "norm (M x * of_real x powr -s) \ totient n * x powr -\" using xy assms unfolding norm_mult M_def \_def by (intro mult_mono sum_upto_dcharacter_le) (auto simp: norm_powr_real_powr) also have "norm (integral {x..y} (\t. M t * (- s * of_real t powr (-s-1)))) \ integral {x..y} (\t. real (totient n) * norm s * t powr (-\-1))" proof (rule integral_norm_bound_integral integrable_on_cmult_left) show "(\t. real (totient n) * norm s * t powr (- \ - 1)) integrable_on {x..y}" using xy by (intro integrable_continuous_real continuous_intros) auto next fix t assume t: "t \ {x..y}" have "norm (M t * (-s * of_real t powr (-s-1))) \ real (totient n) * (norm s * t powr (-\-1))" unfolding norm_mult M_def \_def using xy t assms by (intro mult_mono sum_upto_dcharacter_le) (auto simp: norm_mult norm_powr_real_powr) thus "norm (M t * (-s * of_real t powr (-s-1))) \ real (totient n) * norm s * t powr (-\-1)" by (simp add: algebra_simps) qed (insert I, auto simp: has_integral_iff) also have "\ = real (totient n) * norm s * integral {x..y} (\t. t powr (-\-1))" by simp also have "((\t. t powr (-\-1)) has_integral (y powr -\ / (-\) - x powr -\ / (-\))) {x..y}" using xy assms by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros - simp: has_field_derivative_iff_has_vector_derivative [symmetric] \_def) + simp: has_real_derivative_iff_has_vector_derivative [symmetric] \_def) hence "integral {x..y} (\t. t powr (-\-1)) = y powr -\ / (-\) - x powr -\ / (-\)" by (simp add: has_integral_iff) also from assms have "\ \ x powr -\ / \" by (simp add: \_def) also have "real (totient n) * x powr -\ + real (totient n) * x powr -\ + real (totient n) * norm s * (x powr -\ / \) = real (totient n) * (2 + norm s / \) / x powr \" using xy by (simp add: field_simps powr_minus) finally show ?thesis by (simp add: mult_left_mono) qed show "eventually (\m. norm (sum_upto (\k. \ k * of_nat k powr - s) x - (\k\m. \ k * of_nat k powr - s)) \ real (totient n) * (2 + cmod s / \) / x powr \) at_top" using eventually_gt_at_top[of "nat \x\"] proof eventually_elim case (elim m) have "(\k\m. \ k * of_nat k powr - s) - sum_upto (\k. \ k * of_nat k powr - s) x = (\k\{..m} - {k. 0 < k \ real k \ x}. \ k * of_nat k powr -s)" unfolding sum_upto_def using elim \x > 0\ by (intro Groups_Big.sum_diff [symmetric]) (auto simp: nat_less_iff floor_less_iff) also have "\ = (\k\{..m} - {k. real k \ x}. \ k * of_nat k powr -s)" by (intro sum.mono_neutral_right) auto also have "{..m} - {k. real k \ x} = real -` {x<..real m}" using elim \x > 0\ by (auto simp: nat_less_iff floor_less_iff not_less) also have "norm (\k\\. \ k * of_nat k powr -s) \ real (totient n) * (2 + cmod s / \) / x powr \" using elim \x > 0\ by (intro le) (auto simp: nat_less_iff floor_less_iff) finally show ?case by (simp add: norm_minus_commute) qed qed auto lemma partial_Dirichlet_L_sum_bigo: fixes s :: complex and x :: real assumes "\ \ \\<^sub>0" "Re s > 0" shows "(\x. sum_upto (\n. \ n * n powr -s) x - Dirichlet_L n \ s) \ O(\x. x powr -s)" proof (rule bigoI) show "eventually (\x. norm (sum_upto (\n. \ n * of_nat n powr -s) x - Dirichlet_L n \ s) \ real (totient n) * (2 + norm s / Re s) * norm (of_real x powr - s)) at_top" using eventually_gt_at_top[of 0] proof eventually_elim case (elim x) have "norm (sum_upto (\n. \ n * of_nat n powr -s) x - Dirichlet_L n \ s) \ real (totient n) * (2 + norm s / Re s) / x powr Re s" using elim assms by (intro Dirichlet_L_minus_partial_sum_bound) auto thus ?case using elim assms by (simp add: norm_powr_real_powr powr_minus divide_simps norm_divide del: div_mult_self1 div_mult_self2 div_mult_self3 div_mult_self4) qed qed (*<*) end (*>*) end subsection \Evaluation of $L(\chi, 0)$\ context residues_nat begin (*<*) context includes no_vec_lambda_notation dcharacter_syntax begin (*>*) lemma Dirichlet_L_0_principal [simp]: "Dirichlet_L n \\<^sub>0 0 = 0" proof - have "Dirichlet_L n \\<^sub>0 0 = -1/2 * (\p | prime p \ p dvd n. 1 - 1 / p powr 0)" by (simp add: Dirichlet_L_principal prime_gt_0_nat) also have "(\p | prime p \ p dvd n. 1 - 1 / p powr 0) = (\p | prime p \ p dvd n. 0 :: complex)" by (intro prod.cong) (auto simp: prime_gt_0_nat) also have "(\p | prime p \ p dvd n. 0 :: complex) = 0" using prime_divisor_exists[of n] n by (auto simp: card_gt_0_iff) finally show ?thesis by simp qed end (*<*) end (*>*) context dcharacter begin (*<*) context includes no_vec_lambda_notation dcharacter_syntax begin (*>*) lemma Dirichlet_L_0_nonprincipal: assumes nonprincipal: "\ \ \\<^sub>0" shows "Dirichlet_L n \ 0 = -(\k=1.. k) / of_nat n" proof - have "Dirichlet_L n \ 0 = (\k=1..n. \ k * (1 / 2 - of_nat k / of_nat n))" using assms n by (simp add: Dirichlet_L_conv_hurwitz_zeta_nonprincipal) also have "\ = -1/n * (\k=1..n. of_nat k * \ k)" using assms by (simp add: algebra_simps sum_subtractf sum_dcharacter_block' sum_divide_distrib [symmetric]) also have "(\k=1..n. of_nat k * \ k) = (\k=1.. k)" using n by (intro sum.mono_neutral_right) (auto simp: eq_zero_iff) finally show eq: "Dirichlet_L n \ 0 = -(\k=1.. k) / of_nat n" by simp qed lemma Dirichlet_L_0_even [simp]: assumes "\ (n - 1) = 1" shows "Dirichlet_L n \ 0 = 0" proof (cases "\ = \\<^sub>0") case False hence "Dirichlet_L n \ 0 = -(\k=Suc 0.. k) / of_nat n" by (simp add: Dirichlet_L_0_nonprincipal) also have "\ = 0" using assms False by (subst even_dcharacter_linear_sum_eq_0) auto finally show "Dirichlet_L n \ 0 = 0" . qed auto lemma Dirichlet_L_0: "Dirichlet_L n \ 0 = (if \ (n - 1) = 1 then 0 else -(\k=1.. k) / of_nat n)" by (cases "\ = \\<^sub>0") (auto simp: Dirichlet_L_0_nonprincipal) (*<*) end (*>*) end subsection \Properties of $L(\chi, s)$ for real $\chi$\ (*<*) unbundle no_vec_lambda_notation (*>*) locale real_dcharacter = dcharacter + assumes real: "\ k \ \" begin lemma Im_eq_0 [simp]: "Im (\ k) = 0" using real[of k] by (auto elim!: Reals_cases) lemma of_real_Re [simp]: "of_real (Re (\ k)) = \ k" by (simp add: complex_eq_iff) lemma char_cases: "\ k \ {-1, 0, 1}" proof - from norm[of k] have "Re (\ k) \ {-1,0,1}" by (auto simp: cmod_def split: if_splits) hence "of_real (Re (\ k)) \ {-1, 0, 1}" by auto also have "of_real (Re (\ k)) = \ k" by (simp add: complex_eq_iff) finally show ?thesis . qed lemma cnj [simp]: "cnj (\ k) = \ k" by (simp add: complex_eq_iff) lemma inv_character_id [simp]: "inv_character \ = \" by (simp add: inv_character_def fun_eq_iff) lemma Dirichlet_L_in_Reals: assumes "s \ \" shows "Dirichlet_L n \ s \ \" proof - have "cnj (Dirichlet_L n \ s) = Dirichlet_L n \ s" using assms by (subst cnj_Dirichlet_L) (auto elim!: Reals_cases) thus ?thesis using Reals_cnj_iff by blast qed text \ The following property of real characters is used by Apostol to show the non-vanishing of $L(\chi, 1)$. We have already shown this in a much easier way, but this particular result is still of general interest. \ lemma assumes k: "k > 0" shows sum_char_divisors_ge: "Re (\d | d dvd k. \ d) \ 0" (is "Re (?A k) \ 0") and sum_char_divisors_square_ge: "is_square k \ Re (\d | d dvd k. \ d) \ 1" proof - interpret sum: multiplicative_function ?A by (fact mult.multiplicative_sum_divisors) have A: "?A k \ \" for k by (intro sum_in_Reals real) hence [simp]: "Im (?A k) = 0" for k by (auto elim!: Reals_cases) have *: "Re (?A (p ^ m)) \ (if even m then 1 else 0)" if p: "prime p" for p m proof - have sum_neg1: "(\i\m. (-1) ^ i) = (if even m then 1 else (0::real))" by (induction m) auto from p have "?A (p ^ m) = (\k\m. \ (p ^ k))" by (intro sum.reindex_bij_betw [symmetric] bij_betw_prime_power_divisors) also have "Re \ = (\k\m. Re (\ p) ^ k)" by (simp add: mult.power) also have "\ \ (if even m then 1 else 0)" using sum_neg1 char_cases[of p] by (auto simp: power_0_left) finally show ?thesis . qed have *: "Re (?A (p ^ m)) \ 0" "even m \ Re (?A (p ^ m)) \ 1" if "prime p" for p m using *[of p m] that by (auto split: if_splits) have eq: "Re (?A k) = (\p\prime_factors k. Re (?A (p ^ multiplicity p k)))" using k A by (subst sum.prod_prime_factors) (auto simp: Re_prod_Reals) show "Re (\d | d dvd k. \ d) \ 0" by (subst eq, intro prod_nonneg ballI *) auto assume "is_square k" then obtain m where m: "k = m ^ 2" by (auto elim!: is_nth_powerE) have "even (multiplicity p k)" if "prime p" for p using k that unfolding m by (subst prime_elem_multiplicity_power_distrib) (auto intro!: Nat.gr0I) thus "Re (\d | d dvd k. \ d) \ 1" by (subst eq, intro prod_ge_1 ballI *) auto qed end (*<*) unbundle vec_lambda_notation (*>*) end diff --git a/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy b/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy --- a/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy +++ b/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy @@ -1,3090 +1,3090 @@ (* File: Dirichlet_Series_Analysis.thy Author: Manuel Eberl, TU München *) section \Analytic properties of Dirichlet series\ theory Dirichlet_Series_Analysis imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Library.Going_To_Filter" "HOL-Real_Asymp.Real_Asymp" Dirichlet_Series Moebius_Mu Partial_Summation Euler_Products begin (* TODO Move *) lemma frequently_eventually_frequently: "frequently P F \ eventually Q F \ frequently (\x. P x \ Q x) F" by (erule frequently_rev_mp, erule eventually_mono) auto text \ The following illustrates a concept we will need later on: A property holds for \f\ going to \F\ if we can find e.\,g.\ a sequence that tends to \F\ and whose elements eventually satisfy \P\. \ lemma frequently_going_toI: assumes "filterlim (\n. f (g n)) F G" assumes "eventually (\n. P (g n)) G" assumes "eventually (\n. g n \ A) G" assumes "G \ bot" shows "frequently P (f going_to F within A)" unfolding frequently_def proof assume "eventually (\x. \P x) (f going_to F within A)" hence "eventually (\x. \P x) (inf (filtercomap f F) (principal A))" by (simp add: going_to_within_def) moreover have "filterlim (\n. g n) (inf (filtercomap f F) (principal A)) G" using assms unfolding filterlim_inf filterlim_principal by (auto simp add: filterlim_iff_le_filtercomap filtercomap_filtercomap) ultimately have "eventually (\n. \P (g n)) G" by (rule eventually_compose_filterlim) with assms(2) have "eventually (\_. False) G" by eventually_elim auto with assms(4) show False by simp qed lemma frequently_filtercomapI: assumes "filterlim (\n. f (g n)) F G" assumes "eventually (\n. P (g n)) G" assumes "G \ bot" shows "frequently P (filtercomap f F)" using frequently_going_toI[of f g F G P UNIV] assms by (simp add: going_to_def) lemma frequently_going_to_at_topE: fixes f :: "'a \ real" assumes "frequently P (f going_to at_top)" obtains g where "\n. P (g n)" and "filterlim (\n. f (g n)) at_top sequentially" proof - from assms have "\k. \x. f x \ real k \ P x" by (auto simp: frequently_def eventually_going_to_at_top_linorder) hence "\g. \k. f (g k) \ real k \ P (g k)" by metis then obtain g where g: "\k. f (g k) \ real k" "\k. P (g k)" by blast have "filterlim (\n. f (g n)) at_top sequentially" by (rule filterlim_at_top_mono[OF filterlim_real_sequentially]) (use g in auto) from g(2) and this show ?thesis using that[of g] by blast qed text \ Apostol often uses statements like `$P(s_k)$ for all $k$ in an infinite sequence $s_k$ such that $\mathfrak{R}(s_k)\longrightarrow\infty$ as $k\to\infty$'. Instead, we write @{prop "frequently P (Re going_to at_top)"}. This lemma shows that our statement is equivalent to his. \ lemma frequently_going_to_at_top_iff: "frequently P (f going_to (at_top :: real filter)) \ (\g. \n. P (g n) \ filterlim (\n. f (g n)) at_top sequentially)" by (auto intro: frequently_going_toI elim!: frequently_going_to_at_topE) (* END TODO *) lemma surj_bullet_1: "surj (\s::'a::{real_normed_algebra_1, real_inner}. s \ 1)" proof (rule surjI) fix x :: real show "(x *\<^sub>R 1) \ (1 :: 'a) = x" by (simp add: dot_square_norm) qed lemma bullet_1_going_to_at_top_neq_bot [simp]: "((\s::'a::{real_normed_algebra_1, real_inner}. s \ 1) going_to at_top) \ bot" unfolding going_to_def by (rule filtercomap_neq_bot_surj[OF _ surj_bullet_1]) auto lemma fds_abs_converges_altdef: "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on {1..}" by (auto simp add: fds_abs_converges_def abs_summable_on_nat_iff intro!: summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) lemma fds_abs_converges_altdef': "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on UNIV" by (subst fds_abs_converges_altdef, rule abs_summable_on_cong_neutral) (auto simp: Suc_le_eq) lemma eval_fds_altdef: assumes "fds_abs_converges f s" shows "eval_fds f s = (\\<^sub>an. fds_nth f n / nat_power n s)" proof - have "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on UNIV" unfolding fds_abs_converges_altdef by (intro abs_summable_on_cong_neutral) (auto simp: Suc_le_eq) with assms show ?thesis unfolding eval_fds_def fds_abs_converges_altdef by (intro infsetsum_nat' [symmetric]) simp_all qed lemma multiplicative_function_divide_nat_power: fixes f :: "nat \ 'a :: {nat_power, field}" assumes "multiplicative_function f" shows "multiplicative_function (\n. f n / nat_power n s)" proof interpret f: multiplicative_function f by fact show "f 0 / nat_power 0 s = 0" "f 1 / nat_power 1 s = 1" by simp_all fix a b :: nat assume "a > 1" "b > 1" "coprime a b" thus "f (a * b) / nat_power (a * b) s = f a / nat_power a s * (f b / nat_power b s)" by (simp_all add: f.mult_coprime nat_power_mult_distrib) qed lemma completely_multiplicative_function_divide_nat_power: fixes f :: "nat \ 'a :: {nat_power, field}" assumes "completely_multiplicative_function f" shows "completely_multiplicative_function (\n. f n / nat_power n s)" proof interpret f: completely_multiplicative_function f by fact show "f 0 / nat_power 0 s = 0" "f (Suc 0) / nat_power (Suc 0) s = 1" by simp_all fix a b :: nat assume "a > 1" "b > 1" thus "f (a * b) / nat_power (a * b) s = f a / nat_power a s * (f b / nat_power b s)" by (simp_all add: f.mult nat_power_mult_distrib) qed subsection \Convergence and absolute convergence\ class nat_power_normed_field = nat_power_field + real_normed_field + real_inner + real_algebra_1 + fixes real_power :: "real \ 'a \ 'a" assumes real_power_nat_power: "n > 0 \ real_power (real n) c = nat_power n c" assumes real_power_1_right_aux: "d > 0 \ real_power d 1 = d *\<^sub>R 1" assumes real_power_add: "d > 0 \ real_power d (a + b) = real_power d a * real_power d b" assumes real_power_nonzero [simp]: "d > 0 \ real_power d a \ 0" assumes norm_real_power: "x > 0 \ norm (real_power x c) = x powr (c \ 1)" assumes nat_power_of_real_aux: "nat_power n (x *\<^sub>R 1) = ((real n powr x) *\<^sub>R 1)" assumes has_field_derivative_nat_power_aux: "\x::'a. n > 0 \ LIM y inf_class.inf (Inf (principal ` {S. open S \ x \ S})) (principal (UNIV - {x})). (nat_power n y - nat_power n x - ln (real n) *\<^sub>R nat_power n x * (y - x)) /\<^sub>R norm (y - x) :> Inf (principal ` {S. open S \ 0 \ S})" assumes has_vector_derivative_real_power_aux: "x > 0 \ filterlim (\y. (real_power y c - real_power x (c :: 'a) - (y - x) *\<^sub>R (c * real_power x (c - 1))) /\<^sub>R norm (y - x)) (INF S\{S. open S \ 0 \ S}. principal S) (at x)" assumes norm_nat_power: "n > 0 \ norm (nat_power n y) = real n powr (y \ 1)" begin lemma real_power_diff: "d > 0 \ real_power d (a - b) = real_power d a / real_power d b" using real_power_add[of d b "a - b"] by (simp add: field_simps) end lemma real_power_1_right [simp]: "d > 0 \ real_power d 1 = of_real d" using real_power_1_right_aux[of d] by (simp add: scaleR_conv_of_real) lemma has_vector_derivative_real_power [derivative_intros]: "x > 0 \ ((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x within A)" by (rule has_vector_derivative_at_within) (insert has_vector_derivative_real_power_aux[of x c], simp add: has_vector_derivative_def has_derivative_def nhds_def bounded_linear_scaleR_left) lemma has_field_derivative_nat_power [derivative_intros]: "n > 0 \ ((\y. nat_power n y) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at (x :: 'a :: nat_power_normed_field) within A)" by (rule has_field_derivative_at_within) (insert has_field_derivative_nat_power_aux[of n x], simp only: has_field_derivative_def has_derivative_def netlimit_at, simp add: nhds_def at_within_def bounded_linear_mult_right) lemma continuous_on_real_power [continuous_intros]: "A \ {0<..} \ continuous_on A (\x. real_power x s)" by (rule continuous_on_vector_derivative has_vector_derivative_real_power)+ auto instantiation real :: nat_power_normed_field begin definition real_power_real :: "real \ real \ real" where [simp]: "real_power_real = (powr)" instance proof (standard, goal_cases) case (7 n x) hence "((\x. nat_power n x) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at x)" by (auto intro!: derivative_eq_intros simp: powr_def) thus ?case unfolding has_field_derivative_def netlimit_at has_derivative_def by (simp add: nhds_def at_within_def) next case (8 x c) hence "((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x)" by (auto intro!: derivative_eq_intros - simp: has_field_derivative_iff_has_vector_derivative [symmetric]) + simp: has_real_derivative_iff_has_vector_derivative [symmetric]) thus ?case by (simp add: has_vector_derivative_def has_derivative_def nhds_def) qed (simp_all add: powr_add) end instantiation complex :: nat_power_normed_field begin definition nat_power_complex :: "nat \ complex \ complex" where [simp]: "nat_power_complex n z = of_nat n powr z" definition real_power_complex :: "real \ complex \ complex" where [simp]: "real_power_complex = (\x y. of_real x powr y)" instance proof fix m n :: nat and z :: complex assume "m > 0" "n > 0" thus "nat_power (m * n) z = nat_power m z * nat_power n z" unfolding nat_power_complex_def of_nat_mult by (subst powr_times_real) simp_all next fix n :: nat and z :: complex assume "n > 0" show "norm (nat_power n z) = real n powr (z \ 1)" unfolding nat_power_complex_def using norm_powr_real_powr[of "of_nat n" z] by simp next fix n :: nat and x :: complex assume n: "n > 0" hence "((\x. nat_power n x) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at x)" by (auto intro!: derivative_eq_intros simp: powr_def scaleR_conv_of_real mult_ac) thus "LIM y inf_class.inf (Inf (principal ` {S. open S \ x \ S})) (principal (UNIV - {x})). (nat_power n y - nat_power n x - ln (real n) *\<^sub>R nat_power n x * (y - x)) /\<^sub>R cmod (y - x) :> (Inf (principal ` {S. open S \ 0 \ S}))" unfolding has_field_derivative_def netlimit_at has_derivative_def by (simp add: nhds_def at_within_def) next fix x :: real and c :: complex assume "x > 0" hence "((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x)" by (auto intro!: derivative_eq_intros has_vector_derivative_real_field) thus "LIM y at x. (real_power y c - real_power x c - (y - x) *\<^sub>R (c * real_power x (c - 1))) /\<^sub>R norm (y - x) :> INF S\{S. open S \ 0 \ S}. principal S" by (simp add: has_vector_derivative_def has_derivative_def nhds_def) next fix n :: nat and x :: real show "nat_power n (x *\<^sub>R 1 :: complex) = (real n powr x) *\<^sub>R 1" by (simp add: powr_Reals_eq scaleR_conv_of_real) qed (auto simp: powr_def exp_add exp_of_nat_mult [symmetric] algebra_simps scaleR_conv_of_real simp del: Ln_of_nat) end lemma nat_power_of_real [simp]: "nat_power n (of_real x :: 'a :: nat_power_normed_field) = of_real (real n powr x)" using nat_power_of_real_aux[of n x] by (simp add: scaleR_conv_of_real) lemma fds_abs_converges_of_real [simp]: "fds_abs_converges (fds_of_real f) (of_real s :: 'a :: {nat_power_normed_field,banach}) \ fds_abs_converges f s" unfolding fds_abs_converges_def by (subst (1 2) summable_Suc_iff [symmetric]) (simp add: norm_divide norm_nat_power) lemma eval_fds_of_real [simp]: assumes "fds_converges f s" shows "eval_fds (fds_of_real f) (of_real s :: 'a :: {nat_power_normed_field,banach}) = of_real (eval_fds f s)" using assms unfolding eval_fds_def by (auto simp: fds_converges_def suminf_of_real) lemma fds_abs_summable_zeta_iff [simp]: fixes s :: "'a :: {banach, nat_power_normed_field}" shows "fds_abs_converges fds_zeta s \ s \ 1 > (1 :: real)" proof - have "fds_abs_converges fds_zeta s \ summable (\n. real n powr -(s \ 1))" unfolding fds_abs_converges_def by (intro summable_cong always_eventually) (auto simp: norm_divide fds_nth_zeta powr_minus norm_nat_power divide_simps) also have "\ \ s \ 1 > 1" by (simp add: summable_real_powr_iff) finally show ?thesis . qed lemma fds_abs_summable_zeta: "(s :: 'a :: {banach, nat_power_normed_field}) \ 1 > 1 \ fds_abs_converges fds_zeta s" by simp lemma fds_abs_converges_moebius_mu: fixes s :: "'a :: {banach,nat_power_normed_field}" assumes "s \ 1 > 1" shows "fds_abs_converges (fds moebius_mu) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (norm (fds_nth (fds moebius_mu) n / nat_power n s)) \ real n powr (-s \ 1)" by (auto simp: powr_minus divide_simps abs_moebius_mu_le norm_nat_power norm_divide moebius_mu_def norm_power) next from assms show "summable (\n. real n powr (-s \ 1))" by (simp add: summable_real_powr_iff) qed definition conv_abscissa :: "'a :: {nat_power,banach,real_normed_field, real_inner} fds \ ereal" where "conv_abscissa f = (INF s\{s. fds_converges f s}. ereal (s \ 1))" definition abs_conv_abscissa :: "'a :: {nat_power,banach,real_normed_field, real_inner} fds \ ereal" where "abs_conv_abscissa f = (INF s\{s. fds_abs_converges f s}. ereal (s \ 1))" lemma conv_abscissa_mono: assumes "\s. fds_converges g s \ fds_converges f s" shows "conv_abscissa f \ conv_abscissa g" unfolding conv_abscissa_def by (rule INF_mono) (use assms in auto) lemma abs_conv_abscissa_mono: assumes "\s. fds_abs_converges g s \ fds_abs_converges f s" shows "abs_conv_abscissa f \ abs_conv_abscissa g" unfolding abs_conv_abscissa_def by (rule INF_mono) (use assms in auto) class dirichlet_series = euclidean_space + real_normed_field + nat_power_normed_field + assumes one_in_Basis: "1 \ Basis" instance real :: dirichlet_series by standard simp_all instance complex :: dirichlet_series by standard (simp_all add: Basis_complex_def) context assumes "SORT_CONSTRAINT('a :: dirichlet_series)" begin lemma fds_abs_converges_Re_le: fixes f :: "'a fds" assumes "fds_abs_converges f z" "z \ 1 \ z' \ 1" shows "fds_abs_converges f z'" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat assume n: "n \ 1" thus "norm (norm (fds_nth f n / nat_power n z')) \ norm (fds_nth f n / nat_power n z)" using assms(2) by (simp add: norm_divide norm_nat_power divide_simps powr_mono mult_left_mono) qed (insert assms(1), simp add: fds_abs_converges_def) lemma fds_abs_converges: assumes "s \ 1 > abs_conv_abscissa (f :: 'a fds)" shows "fds_abs_converges f s" proof - from assms obtain s0 where "fds_abs_converges f s0" "s0 \ 1 < s \ 1" by (auto simp: INF_less_iff abs_conv_abscissa_def) with fds_abs_converges_Re_le[OF this(1), of s] this(2) show ?thesis by simp qed lemma fds_abs_diverges: assumes "s \ 1 < abs_conv_abscissa (f :: 'a fds)" shows "\fds_abs_converges f s" proof assume "fds_abs_converges f s" hence "abs_conv_abscissa f \ s \ 1" unfolding abs_conv_abscissa_def by (intro INF_lower) auto with assms show False by simp qed lemma uniformly_Cauchy_eval_fds_aux: fixes s0 :: "'a :: dirichlet_series" assumes bounded: "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_Cauchy_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (cases "B = {}") case False show ?thesis proof (rule uniformly_Cauchy_onI', goal_cases) case (1 \) define \ where "\ = Inf ((\s. s \ 1) ` B)" have \_le: "s \ 1 \ \" if "s \ B" for s unfolding \_def using that by (intro cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded B) auto have "\ \ ((\s. s \ 1) ` B)" unfolding \_def using B \B \ {}\ by (intro closed_contains_Inf bounded_inner_imp_bdd_below compact_imp_bounded B compact_imp_closed compact_continuous_image continuous_intros) auto with B(2) have \_gt: "\ > s0 \ 1" by auto define \ where "\ = \ - s0 \ 1" have "bounded B" by (rule compact_imp_bounded) fact then obtain norm_B_aux where norm_B_aux: "\s. s \ B \ norm s \ norm_B_aux" by (auto simp: bounded_iff) define norm_B where "norm_B = norm_B_aux + norm s0" from norm_B_aux have norm_B: "norm (s - s0) \ norm_B" if "s \ B" for s using norm_triangle_ineq4[of s s0] norm_B_aux[OF that] by (simp add: norm_B_def) then have "0 \ norm_B" by (meson \\ \ (\s. s \ 1) ` B\ imageE norm_ge_zero order.trans) define A where "A = sum_upto (\k. fds_nth f k / nat_power k s0)" from bounded obtain C_aux where C_aux: "\n. norm (\k\n. fds_nth f k / nat_power k s0) \ C_aux" by (auto simp: Bseq_def) define C where "C = max C_aux 1" have C_pos: "C > 0" by (simp add: C_def) have C: "norm (A x) \ C" for x proof - have "A x = (\k\nat \x\. fds_nth f k / nat_power k s0)" unfolding A_def sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C_aux" by (rule C_aux) also have "\ \ C" by (simp add: C_def) finally show ?thesis . qed have "(\m. 2 * C * (1 + norm_B / \) * real m powr (-\)) \ 0" unfolding \_def using \_gt by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially) simp_all from order_tendstoD(2)[OF this \\ > 0\] obtain M where M: "\m. m \ M \ 2 * C * (1 + norm_B / \) * real m powr - \ < \" by (auto simp: eventually_at_top_linorder) show ?case proof (intro exI[of _ "max M 1"] ballI allI impI, goal_cases) case (1 s m n) from 1 have s: "s \ 1 > s0 \ 1" using B(2)[of s] by simp have mn: "m \ M" "m < n" "m > 0" "n > 0" using 1 by (simp_all add: ) have "dist (\n\m. fds_nth f n / nat_power n s) (\n\n. fds_nth f n / nat_power n s) = dist (\n\n. fds_nth f n / nat_power n s) (\n\m. fds_nth f n / nat_power n s)" by (simp add: dist_commute) also from 1 have "\ = norm (\k\{..n}-{..m}. fds_nth f k / nat_power k s)" by (subst Groups_Big.sum_diff) (simp_all add: dist_norm) also from 1 have "{..n} - {..m} = real -` {real m<..real n}" by auto also have "(\k\\. fds_nth f k / nat_power k s) = (\k\\. fds_nth f k / nat_power k s0 * real_power (real k) (s0 - s))" (is "_ = ?S") by (intro sum.cong refl) (simp_all add: nat_power_diff real_power_nat_power) also have *: "((\t. A t * ((s0 - s) * real_power t (s0 - s - 1))) has_integral (A (real n) * real_power n (s0 - s) - A (real m) * real_power m (s0 - s) - ?S)) {real m..real n}" (is "(?h has_integral _) _") unfolding A_def using mn by (intro partial_summation_strong[of "{}"]) (auto intro!: derivative_eq_intros continuous_intros) hence "?S = A (real n) * nat_power n (s0 - s) - A (real m) * nat_power m (s0 - s) - integral {real m..real n} ?h" using mn by (simp add: has_integral_iff real_power_nat_power) also have "norm \ \ norm (A (real n) * nat_power n (s0 - s)) + norm (A (real m) * nat_power m (s0 - s)) + norm (integral {real m..real n} ?h)" by (intro order.trans[OF norm_triangle_ineq4] add_right_mono order.refl) also have "norm (A (real n) * nat_power n (s0 - s)) \ C * nat_power m ((s0 - s) \ 1)" using mn \s \ B\ C_pos s by (auto simp: norm_mult norm_nat_power algebra_simps intro!: mult_mono C powr_mono2') also have "norm (A (real m) * nat_power m (s0 - s)) \ C * nat_power m ((s0 - s) \ 1)" using mn by (auto simp: norm_mult norm_nat_power intro!: mult_mono C) also have "norm (integral {real m..real n} ?h) \ integral {real m..real n} (\t. C * (norm (s0 - s) * t powr ((s0 - s) \ 1 - 1)))" proof (intro integral_norm_bound_integral ballI, goal_cases) case 1 with * show ?case by (simp add: has_integral_iff) next case 2 from mn show ?case by (auto intro!: integrable_continuous_real continuous_intros) next case (3 t) thus ?case unfolding norm_mult using C_pos mn by (intro mult_mono C) (auto simp: norm_real_power dot_square_norm algebra_simps) qed also have "\ = C * norm (s0 - s) * integral {real m..real n} (\t. t powr ((s0 - s) \ 1 - 1))" by (simp add: algebra_simps dot_square_norm) also { have "((\t. t powr ((s0 - s) \ 1 - 1)) has_integral (real n powr ((s0 - s) \ 1) / ((s0 - s) \ 1) - real m powr ((s0 - s) \ 1) / ((s0 - s) \ 1))) {m..n}" (is "(?l has_integral ?I) _") using mn s by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros - simp: has_field_derivative_iff_has_vector_derivative [symmetric] inner_diff_left) + simp: has_real_derivative_iff_has_vector_derivative [symmetric] inner_diff_left) hence "integral {real m..real n} ?l = ?I" by (simp add: has_integral_iff) also have "\ \ -(real m powr ((s0 - s) \ 1) / ((s0 - s) \ 1))" using s mn by (simp add: divide_simps inner_diff_left) also have "\ = 1 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1))" using s by (simp add: field_simps inner_diff_left) also have "\ \ 2 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1))" using mn s by (intro mult_right_mono divide_nonneg_pos) (simp_all add: inner_diff_left) finally have "integral {m..n} ?l \ \" . } hence "C * norm (s0 - s) * integral {real m..real n} (\t. t powr ((s0 - s) \ 1 - 1)) \ C * norm (s0 - s) * (2 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1)))" using C_pos mn by (intro mult_mono mult_nonneg_nonneg integral_nonneg integrable_continuous_real continuous_intros) auto also have "C * nat_power m ((s0 - s) \ 1) + C * nat_power m ((s0 - s) \ 1) + \ = 2 * C * nat_power m ((s0 - s) \ 1) * (1 + norm (s - s0) / ((s - s0) \ 1))" by (simp add: algebra_simps norm_minus_commute) also have "\ \ 2 * C * nat_power m (-\) * (1 + norm_B / \)" using C_pos s mn \_le[of s] \s \ B\ \_gt \0 \ norm_B\ unfolding nat_power_real_def \_def by (intro mult_mono powr_mono frac_le add_mono norm_B; simp add: inner_diff_left) also have "\ = 2 * C * (1 + norm_B / \) * real m powr (-\)" by simp also from \m \ M\ have "\ < \" by (rule M) finally show ?case by - simp_all qed qed qed (auto simp: uniformly_Cauchy_on_def) lemma uniformly_convergent_eval_fds_aux: assumes "Bseq (\n. \k\n. fds_nth f k / nat_power k (s0 :: 'a))" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" by (rule Cauchy_uniformly_convergent uniformly_Cauchy_eval_fds_aux assms)+ lemma uniformly_convergent_eval_fds_aux': assumes conv: "fds_converges f (s0 :: 'a)" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (rule uniformly_convergent_eval_fds_aux) from conv have "convergent (\n. \k\n. fds_nth f k / nat_power k s0)" by (simp add: fds_converges_def summable_iff_convergent') thus "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" by (rule convergent_imp_Bseq) qed (insert assms, auto) lemma bounded_partial_sums_imp_fps_converges: fixes s0 :: "'a :: dirichlet_series" assumes "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" and "s \ 1 > s0 \ 1" shows "fds_converges f s" proof - have "uniformly_convergent_on {s} (\N z. \n\N. fds_nth f n / nat_power n z)" using assms(2) by (intro uniformly_convergent_eval_fds_aux[OF assms(1)]) auto thus ?thesis by (auto simp: fds_converges_def summable_iff_convergent' dest: uniformly_convergent_imp_convergent) qed theorem fds_converges_Re_le: assumes "fds_converges f (s0 :: 'a)" "s \ 1 > s0 \ 1" shows "fds_converges f s" proof - have "uniformly_convergent_on {s} (\N z. \n\N. fds_nth f n / nat_power n z)" by (rule uniformly_convergent_eval_fds_aux' assms)+ (insert assms(2), auto) then obtain l where "uniform_limit {s} (\N z. \n\N. fds_nth f n / nat_power n z) l at_top" by (auto simp: uniformly_convergent_on_def) from tendsto_uniform_limitI[OF this, of s] have "(\n. fds_nth f n / nat_power n s) sums l s" unfolding sums_def' by (simp add: atLeast0AtMost) thus ?thesis by (simp add: fds_converges_def sums_iff) qed lemma fds_converges: assumes "s \ 1 > conv_abscissa (f :: 'a fds)" shows "fds_converges f s" proof - from assms obtain s0 where "fds_converges f s0" "s0 \ 1 < s \ 1" by (auto simp: INF_less_iff conv_abscissa_def) with fds_converges_Re_le[OF this(1), of s] this(2) show ?thesis by simp qed lemma fds_diverges: assumes "s \ 1 < conv_abscissa (f :: 'a fds)" shows "\fds_converges f s" proof assume "fds_converges f s" hence "conv_abscissa f \ s \ 1" unfolding conv_abscissa_def by (intro INF_lower) auto with assms show False by simp qed theorem fds_converges_imp_abs_converges: assumes "fds_converges (f :: 'a fds) s" "s' \ 1 > s \ 1 + 1" shows "fds_abs_converges f s'" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from assms(2) show "summable (\n. real n powr ((s - s') \ 1))" by (subst summable_real_powr_iff) (simp_all add: inner_diff_left) next from assms(1) have "(\n. fds_nth f n / nat_power n s) \ 0" unfolding fds_converges_def by (rule summable_LIMSEQ_zero) from tendsto_norm[OF this] have "(\n. norm (fds_nth f n / nat_power n s)) \ 0" by simp hence "eventually (\n. norm (fds_nth f n / nat_power n s) < 1) at_top" by (rule order_tendstoD) simp_all thus "eventually (\n. norm (norm (fds_nth f n / nat_power n s')) \ real n powr ((s - s') \ 1)) at_top" proof eventually_elim case (elim n) thus ?case proof (cases "n = 0") case False have "norm (fds_nth f n / nat_power n s') = norm (fds_nth f n) / real n powr (s' \ 1)" using False by (simp add: norm_divide norm_nat_power) also have "\ = norm (fds_nth f n / nat_power n s) / real n powr ((s' - s) \ 1)" using False by (simp add: norm_divide norm_nat_power inner_diff_left powr_diff) also have "\ \ 1 / real n powr ((s' - s) \ 1)" using elim by (intro divide_right_mono elim) simp_all also have "\ = real n powr ((s - s') \ 1)" using False by (simp add: field_simps inner_diff_left powr_diff) finally show ?thesis by simp qed simp_all qed qed lemma conv_le_abs_conv_abscissa: "conv_abscissa f \ abs_conv_abscissa f" unfolding conv_abscissa_def abs_conv_abscissa_def by (intro INF_superset_mono) auto lemma conv_abscissa_PInf_iff: "conv_abscissa f = \ \ (\s. \fds_converges f s)" unfolding conv_abscissa_def by (subst Inf_eq_PInfty) auto lemma conv_abscissa_PInfI [intro]: "(\s. \fds_converges f s) \ conv_abscissa f = \" by (subst conv_abscissa_PInf_iff) auto lemma conv_abscissa_MInf_iff: "conv_abscissa (f :: 'a fds) = -\ \ (\s. fds_converges f s)" proof safe assume *: "\s. fds_converges f s" have "conv_abscissa f \ B" for B :: real using spec[OF *, of "of_real B"] fds_diverges[of "of_real B" f] by (cases "conv_abscissa f \ B") simp_all thus "conv_abscissa f = -\" by (rule ereal_bot) qed (auto intro: fds_converges) lemma conv_abscissa_MInfI [intro]: "(\s. fds_converges (f::'a fds) s) \ conv_abscissa f = -\" by (subst conv_abscissa_MInf_iff) auto lemma abs_conv_abscissa_PInf_iff: "abs_conv_abscissa f = \ \ (\s. \fds_abs_converges f s)" unfolding abs_conv_abscissa_def by (subst Inf_eq_PInfty) auto lemma abs_conv_abscissa_PInfI [intro]: "(\s. \fds_converges f s) \ abs_conv_abscissa f = \" by (subst abs_conv_abscissa_PInf_iff) auto lemma abs_conv_abscissa_MInf_iff: "abs_conv_abscissa (f :: 'a fds) = -\ \ (\s. fds_abs_converges f s)" proof safe assume *: "\s. fds_abs_converges f s" have "abs_conv_abscissa f \ B" for B :: real using spec[OF *, of "of_real B"] fds_abs_diverges[of "of_real B" f] by (cases "abs_conv_abscissa f \ B") simp_all thus "abs_conv_abscissa f = -\" by (rule ereal_bot) qed (auto intro: fds_abs_converges) lemma abs_conv_abscissa_MInfI [intro]: "(\s. fds_abs_converges (f::'a fds) s) \ abs_conv_abscissa f = -\" by (subst abs_conv_abscissa_MInf_iff) auto lemma conv_abscissa_geI: assumes "\c'. ereal c' < c \ \s. s \ 1 = c' \ \fds_converges f s" shows "conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\conv_abscissa f \ c" hence "c > conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c > ereal c'" "c' > conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "\fds_converges f s" by blast ultimately show False using fds_converges[of f s] by auto qed lemma conv_abscissa_leI: assumes "\c'. ereal c' > c \ \s. s \ 1 = c' \ fds_converges f s" shows "conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\conv_abscissa f \ c" hence "c < conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c < ereal c'" "c' < conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "fds_converges f s" by blast ultimately show False using fds_diverges[of s f] by auto qed lemma abs_conv_abscissa_geI: assumes "\c'. ereal c' < c \ \s. s \ 1 = c' \ \fds_abs_converges f s" shows "abs_conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\abs_conv_abscissa f \ c" hence "c > abs_conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c > ereal c'" "c' > abs_conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "\fds_abs_converges f s" by blast ultimately show False using fds_abs_converges[of f s] by auto qed lemma abs_conv_abscissa_leI: assumes "\c'. ereal c' > c \ \s. s \ 1 = c' \ fds_abs_converges f s" shows "abs_conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\abs_conv_abscissa f \ c" hence "c < abs_conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c < ereal c'" "c' < abs_conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "fds_abs_converges f s" by blast ultimately show False using fds_abs_diverges[of s f] by auto qed lemma conv_abscissa_leI_weak: assumes "\x. ereal x > d \ fds_converges f (of_real x)" shows "conv_abscissa (f :: 'a fds) \ d" proof (rule conv_abscissa_leI) fix x assume "d < ereal x" from assms[OF this] show "\s. s \ 1 = x \ fds_converges f s" by (intro exI[of _ "of_real x"]) auto qed lemma abs_conv_abscissa_leI_weak: assumes "\x. ereal x > d \ fds_abs_converges f (of_real x)" shows "abs_conv_abscissa (f :: 'a fds) \ d" proof (rule abs_conv_abscissa_leI) fix x assume "d < ereal x" from assms[OF this] show "\s. s \ 1 = x \ fds_abs_converges f s" by (intro exI[of _ "of_real x"]) auto qed lemma conv_abscissa_truncate [simp]: "conv_abscissa (fds_truncate m (f :: 'a fds)) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_truncate [simp]: "abs_conv_abscissa (fds_truncate m (f :: 'a fds)) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) theorem abs_conv_le_conv_abscissa_plus_1: "abs_conv_abscissa (f :: 'a fds) \ conv_abscissa f + 1" proof (rule abs_conv_abscissa_leI) fix c assume less: "conv_abscissa f + 1 < ereal c" define c' where "c' = (if conv_abscissa f = -\ then c - 2 else (c - 1 + real_of_ereal (conv_abscissa f)) / 2)" from less have c': "conv_abscissa f < ereal c' \ c' < c - 1" by (cases "conv_abscissa f") (simp_all add: c'_def field_simps) from c' have "fds_converges f (of_real c')" by (intro fds_converges) (simp_all add: inner_diff_left dot_square_norm) hence "fds_abs_converges f (of_real c)" by (rule fds_converges_imp_abs_converges) (insert c', simp_all) thus "\s. s \ 1 = c \ fds_abs_converges f s" by (intro exI[of _ "of_real c"]) auto qed lemma uniformly_convergent_eval_fds: assumes B: "compact B" "\z. z \ B \ z \ 1 > conv_abscissa (f :: 'a fds)" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (cases "B = {}") case False define \ where "\ = Inf ((\s. s \ 1) ` B)" have \_le: "s \ 1 \ \" if "s \ B" for s unfolding \_def using that by (intro cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded B) auto have "\ \ ((\s. s \ 1) ` B)" unfolding \_def using B \B \ {}\ by (intro closed_contains_Inf bounded_inner_imp_bdd_below compact_imp_bounded B compact_imp_closed compact_continuous_image continuous_intros) auto with B(2) have \_gt: "\ > conv_abscissa f" by auto define s where "s = (if conv_abscissa f = -\ then \ - 1 else (\ + real_of_ereal (conv_abscissa f)) / 2)" from \_gt have s: "conv_abscissa f < s \ s < \" by (cases "conv_abscissa f") (auto simp: s_def) show ?thesis using s \compact B\ by (intro uniformly_convergent_eval_fds_aux'[of f "of_real s"] fds_converges) (auto dest: \_le) qed auto corollary uniformly_convergent_eval_fds': assumes B: "compact B" "\z. z \ B \ z \ 1 > conv_abscissa (f :: 'a fds)" shows "uniformly_convergent_on B (\N z. \nN z. \n\N. fds_nth f n / nat_power n z) l at_top" by (auto simp: uniformly_convergent_on_def) also have "(\N z. \n\N. fds_nth f n / nat_power n z) = (\N z. \nN z. \nDerivative of a Dirichlet series\ lemma fds_converges_deriv_aux: assumes conv: "fds_converges f (s0 :: 'a)" and gt: "s \ 1 > s0 \ 1" shows "fds_converges (fds_deriv f) s" proof - have "Cauchy (\n. \k\n. (-ln (real k) *\<^sub>R fds_nth f k) / nat_power k s)" proof (rule CauchyI', goal_cases) case (1 \) define \ where "\ = s \ 1 - s0 \ 1" define \' where "\' = \ / 2" from gt have \_pos: "\ > 0" by (simp add: \_def) define A where "A = sum_upto (\k. fds_nth f k / nat_power k s0)" from conv have "convergent (\n. \k\n. fds_nth f k / nat_power k s0)" by (simp add: fds_converges_def summable_iff_convergent') hence "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" by (rule convergent_imp_Bseq) then obtain C_aux where C_aux: "\n. norm (\k\n. fds_nth f k / nat_power k s0) \ C_aux" by (auto simp: Bseq_def) define C where "C = max C_aux 1" have C_pos: "C > 0" by (simp add: C_def) have C: "norm (A x) \ C" for x proof - have "A x = (\k\nat \x\. fds_nth f k / nat_power k s0)" unfolding A_def sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C_aux" by (rule C_aux) also have "\ \ C" by (simp add: C_def) finally show ?thesis . qed define C' where "C' = 2 * C + C * (norm (s0 - s) * (1 + 1 / \) + 1) / \" have "(\m. C' * real m powr (-\')) \ 0" unfolding \'_def using gt \_pos by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially) simp_all from order_tendstoD(2)[OF this \\ > 0\] obtain M1 where M1: "\m. m \ M1 \ C' * real m powr - \' < \" by (auto simp: eventually_at_top_linorder) have "((\x. ln (real x) / real x powr \') \ 0) at_top" using \_pos by (intro lim_ln_over_power) (simp_all add: \'_def) from order_tendstoD(2)[OF this zero_less_one] eventually_gt_at_top[of "1::nat"] have "eventually (\n. ln (real n) \ n powr \') at_top" by eventually_elim simp_all then obtain M2 where M2: "\n. n \ M2 \ ln (real n) \ n powr \'" by (auto simp: eventually_at_top_linorder) let ?f' = "\k. -ln (real k) *\<^sub>R fds_nth f k" show ?case proof (intro exI[of _ "max (max M1 M2) 1"] allI impI, goal_cases) case (1 m n) hence mn: "m \ M1" "m \ M2" "m > 0" "m < n" by simp_all define g :: "real \ 'a" where "g = (\t. real_power t (s0 - s) * of_real (ln t))" define g' :: "real \ 'a" where "g' = (\t. real_power t (s0 - s - 1) * ((s0 - s) * of_real (ln t) + 1))" define norm_g' :: "real \ real" where "norm_g' = (\t. t powr (-\ - 1) * (norm (s0 - s) * ln t + 1))" define norm_g :: "real \ real" where "norm_g = (\t. -(t powr -\) * (norm (s0 - s) * (\ * ln t + 1) + \) / \^2)" have g_g': "(g has_vector_derivative g' t) (at t)" if "t \ {real m..real n}" for t using mn that by (auto simp: g_def g'_def real_power_diff field_simps real_power_add intro!: derivative_eq_intros) have [continuous_intros]: "continuous_on {real m..real n} g" using mn by (auto simp: g_def intro!: continuous_intros) let ?S = "\k\real -` {real m<..real n}. fds_nth f k / nat_power k s0 * g k" have "dist (\k\m. ?f' k / nat_power k s) (\k\n. ?f' k / nat_power k s) = norm (\k\{..n} - {..m}. fds_nth f k / nat_power k s * of_real (ln (real k)))" using mn by (subst sum_diff) (simp_all add: dist_norm norm_minus_commute sum_negf scaleR_conv_of_real mult_ac) also have "{..n} - {..m} = real -` {real m<..real n}" by auto also have "(\k\\. fds_nth f k / nat_power k s * of_real (ln (real k))) = (\k\\. fds_nth f k / nat_power k s0 * g k)" using mn unfolding g_def by (intro sum.cong refl) (auto simp: real_power_nat_power field_simps nat_power_diff) also have *: "((\t. A t * g' t) has_integral (A (real n) * g n - A (real m) * g m - ?S)) {real m..real n}" (is "(?h has_integral _) _") unfolding A_def using mn by (intro partial_summation_strong[of "{}"]) (auto intro!: g_g' simp: field_simps continuous_intros) hence "?S = A (real n) * g n - A (real m) * g m - integral {real m..real n} ?h" using mn by (simp add: has_integral_iff field_simps) also have "norm \ \ norm (A (real n) * g n) + norm (A (real m) * g m) + norm (integral {real m..real n} ?h)" by (intro order.trans[OF norm_triangle_ineq4] add_right_mono order.refl) also have "norm (A (real n) * g n) \ C * norm (g n)" unfolding norm_mult using mn C_pos by (intro mult_mono C) auto also have "norm (g n) \ n powr -\ * n powr \'" using mn M2[of n] by (simp add: g_def norm_real_power norm_mult \_def inner_diff_left) also have "\ = n powr -\'" using mn by (simp add: \'_def powr_minus field_simps powr_add [symmetric]) also have "norm (A (real m) * g m) \ C * norm (g m)" unfolding norm_mult using mn C_pos by (intro mult_mono C) auto also have "norm (g m) \ m powr -\ * m powr \'" using mn M2[of m] by (simp add: g_def norm_real_power norm_mult \_def inner_diff_left) also have "\ = m powr -\'" using mn by (simp add: \'_def powr_minus field_simps powr_add [symmetric]) also have "C * real n powr - \' \ C * real m powr - \'" using \_pos mn C_pos by (intro mult_left_mono powr_mono2') (simp_all add: \'_def) also have "\ + \ = 2 * \" by simp also have "norm (integral {m..n} ?h) \ integral {m..n} (\t. C * norm_g' t)" proof (intro integral_norm_bound_integral ballI, goal_cases) case 1 with * show ?case by (simp add: has_integral_iff) next case 2 from mn show ?case by (auto intro!: integrable_continuous_real continuous_intros simp: norm_g'_def) next case (3 t) have "norm (g' t) \ norm_g' t" unfolding g'_def norm_g'_def using 3 mn unfolding norm_mult by (intro mult_mono order.trans[OF norm_triangle_ineq]) (auto simp: norm_real_power inner_diff_left dot_square_norm norm_mult \_def intro!: mult_left_mono) thus ?case unfolding norm_mult using C_pos mn by (intro mult_mono C) simp_all qed also have "\ = C * integral {m..n} norm_g'" unfolding norm_g'_def by (simp add: norm_g'_def \_def inner_diff_left) also { have "(norm_g' has_integral (norm_g n - norm_g m)) {m..n}" unfolding norm_g'_def norm_g_def power2_eq_square using mn \_pos by (intro fundamental_theorem_of_calculus) - (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] + (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric] field_simps powr_diff intro!: derivative_eq_intros) hence "integral {m..n} norm_g' = norm_g n - norm_g m" by (simp add: has_integral_iff) also have "norm_g n \ 0" unfolding norm_g_def using \_pos mn by (intro divide_nonpos_pos mult_nonpos_nonneg add_nonneg_nonneg mult_nonneg_nonneg) simp_all hence "norm_g n - norm_g m \ -norm_g m" by simp also have "\ = real m powr -\ * ln (real m) * (norm (s0 - s)) / \ + real m powr -\ * ((norm (s0 - s) / \ + 1) / \)" using \_pos by (simp add: field_simps norm_g_def power2_eq_square) also { have "ln (real m) \ real m powr \'" using M2[of m] mn by simp also have "real m powr -\ * \ = real m powr -\'" by (simp add: powr_add [symmetric] \'_def) finally have "real m powr -\ * ln (real m) * (norm (s0 - s)) / \ \ \ * (norm (s0 - s)) / \" using \_pos by (intro divide_right_mono mult_right_mono) (simp_all add: mult_left_mono) } also have "real m powr -\ * ((norm (s0 - s) / \ + 1) / \) \ real m powr -\' * ((norm (s0 - s) / \ + 1) / \)" using mn \_pos by (intro mult_right_mono powr_mono) (simp_all add: \'_def) also have "real m powr - \' * norm (s0 - s) / \ + \ = real m powr -\' * (norm (s0 - s) * (1 + 1 / \) + 1) / \" using \_pos by (simp add: field_simps power2_eq_square) finally have "integral {real m..real n} norm_g' \ real m powr - \' * (norm (s0 - s) * (1 + 1 / \) + 1) / \" by - simp_all } also have "2 * (C * m powr - \') + C * (m powr - \' * (norm (s0 - s) * (1 + 1 / \) + 1) / \) = C' * m powr -\'" by (simp add: algebra_simps C'_def) also have "\ < \" using M1[of m] mn by simp finally show ?case using C_pos by - simp_all qed qed from Cauchy_convergent[OF this] show ?thesis by (simp add: summable_iff_convergent' fds_converges_def fds_nth_deriv) qed theorem assumes "s \ 1 > conv_abscissa (f :: 'a fds)" shows fds_converges_deriv: "fds_converges (fds_deriv f) s" and has_field_derivative_eval_fds [derivative_intros]: "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s within A)" proof - define s1 :: real where "s1 = (if conv_abscissa f = -\ then s \ 1 - 2 else (s \ 1 * 1 / 3 + real_of_ereal (conv_abscissa f) * 2 / 3))" define s2 :: real where "s2 = (if conv_abscissa f = -\ then s \ 1 - 1 else (s \ 1 * 2 / 3 + real_of_ereal (conv_abscissa f) * 1 / 3))" from assms have s: "conv_abscissa f < s1 \ s1 < s2 \ s2 < s \ 1" by (cases "conv_abscissa f") (auto simp: s1_def s2_def field_simps) from s have *: "fds_converges f (of_real s1)" by (intro fds_converges) simp_all thus conv': "fds_converges (fds_deriv f) s" by (rule fds_converges_deriv_aux) (insert s, simp_all) from * have conv: "fds_converges (fds_deriv f) (of_real s2)" by (rule fds_converges_deriv_aux) (insert s, simp_all) define \ :: real where "\ = (s \ 1 - s2) / 2" from s have \_pos: "\ > 0" by (simp add: \_def) have "uniformly_convergent_on (cball s \) (\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s)" proof (intro uniformly_convergent_eval_fds_aux'[OF conv]) fix s'' :: 'a assume s'': "s'' \ cball s \" have "dist (s \ 1) (s'' \ 1) \ dist s s''" by (intro Euclidean_dist_upper) (simp_all add: one_in_Basis) also from s'' have "\ \ \" by simp finally show "s'' \ 1 > (of_real s2 :: 'a) \ 1" using s by (auto simp: \_def dist_real_def abs_if split: if_splits) qed (insert \_pos, auto) then obtain l where "uniform_limit (cball s \) (\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s) l at_top" by (auto simp: uniformly_convergent_on_def) also have "(\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s) = (\n s. \k) (\n s. \k) (\n s. \k cball s \" "s \ interior (cball s \)" using s by (simp_all add: \_def) show "summable (\n. fds_nth f n / nat_power n s)" using assms fds_converges[of f s] by (simp add: fds_converges_def) next fix s' :: 'a and n :: nat show "((\s. fds_nth f n / nat_power n s) has_field_derivative fds_nth (fds_deriv f) n / nat_power n s') (at s' within cball s \)" by (cases "n = 0") (simp, auto intro!: derivative_eq_intros simp: fds_nth_deriv field_simps) qed (auto simp: fds_nth_deriv intro!: derivative_eq_intros) thus "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s within A)" by (rule has_field_derivative_at_within) qed lemmas has_field_derivative_eval_fds' [derivative_intros] = DERIV_chain2[OF has_field_derivative_eval_fds] lemma continuous_eval_fds [continuous_intros]: assumes "s \ 1 > conv_abscissa f" shows "continuous (at s within A) (eval_fds (f :: 'a :: dirichlet_series fds))" proof - have "isCont (eval_fds f) s" by (rule has_field_derivative_eval_fds DERIV_isCont assms)+ thus ?thesis by (rule continuous_within_subset) auto qed lemma continuous_eval_fds' [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "continuous (at s within A) g" "g s \ 1 > conv_abscissa f" shows "continuous (at s within A) (\x. eval_fds f (g x))" by (rule continuous_within_compose3[OF _ assms(1)] continuous_intros assms)+ lemma continuous_on_eval_fds [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "A \ {s. s \ 1 > conv_abscissa f}" shows "continuous_on A (eval_fds f)" by (rule DERIV_continuous_on derivative_intros)+ (insert assms, auto) lemma continuous_on_eval_fds' [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "continuous_on A g" "g ` A \ {s. s \ 1 > conv_abscissa f}" shows "continuous_on A (\x. eval_fds f (g x))" by (rule continuous_on_compose2[OF continuous_on_eval_fds assms(1)]) (insert assms, auto simp: image_iff) lemma conv_abscissa_deriv_le: fixes f :: "'a fds" shows "conv_abscissa (fds_deriv f) \ conv_abscissa f" proof (rule conv_abscissa_leI) fix c' :: real assume "ereal c' > conv_abscissa f" thus "\s. s \ 1 = c' \ fds_converges (fds_deriv f) s" by (intro exI[of _ "of_real c'"]) (auto simp: fds_converges_deriv) qed lemma abs_conv_abscissa_integral: fixes f :: "'a fds" shows "abs_conv_abscissa (fds_integral a f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa (fds_integral a f) \ abs_conv_abscissa f" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) have "fds_abs_converges (fds_integral a f) (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from 1 have "fds_abs_converges f (of_real c)" by (intro fds_abs_converges) auto thus "summable (\n. norm (fds_nth f n / nat_power n (of_real c)))" by (simp add: fds_abs_converges_def) next show "\\<^sub>F n in sequentially. norm (norm (fds_nth (fds_integral a f) n / nat_power n (of_real c))) \ norm (fds_nth f n / nat_power n (of_real c))" using eventually_gt_at_top[of 3] proof eventually_elim case (elim n) from elim and exp_le have "ln (exp 1) \ ln (real n)" by (subst ln_le_cancel_iff) auto hence "1 * norm (fds_nth f n) \ ln (real n) * norm (fds_nth f n)" by (intro mult_right_mono) auto with elim show ?case by (simp add: norm_divide norm_nat_power fds_integral_def field_simps) qed qed thus ?case by (intro exI[of _ "of_real c"]) auto qed next show "abs_conv_abscissa f \ abs_conv_abscissa (fds_integral a f)" (is "_ \ ?s0") proof (cases "abs_conv_abscissa (fds_integral a f) = \") case False show ?thesis proof (rule abs_conv_abscissa_leI) fix c :: real define \ where "\ = (if ?s0 = -\ then 1 else (c - real_of_ereal ?s0) / 2)" assume "ereal c > ?s0" with False have \: "\ > 0" "c - \ > ?s0" by (cases ?s0; force simp: \_def field_simps)+ have "fds_abs_converges f (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from \ have "fds_abs_converges (fds_integral a f) (of_real (c - \))" by (intro fds_abs_converges) (auto simp: algebra_simps) thus "summable (\n. norm (fds_nth (fds_integral a f) n / nat_power n (of_real (c - \))))" by (simp add: fds_abs_converges_def) next have "\\<^sub>F n in at_top. ln (real n) / real n powr \ < 1" by (rule order_tendstoD lim_ln_over_power \\ > 0\ zero_less_one)+ thus "\\<^sub>F n in sequentially. norm (norm (fds_nth f n / nat_power n (of_real c))) \ norm (fds_nth (fds_integral a f) n / nat_power n (of_real (c - \)))" using eventually_gt_at_top[of 1] proof eventually_elim case (elim n) hence "ln (real n) * norm (fds_nth f n) \ real n powr \ * norm (fds_nth f n)" by (intro mult_right_mono) auto with elim show ?case by (simp add: norm_divide norm_nat_power field_simps powr_diff inner_diff_left fds_integral_def) qed qed thus "\s. s \ 1 = c \ fds_abs_converges f s" by (intro exI[of _ "of_real c"]) auto qed qed auto qed lemma abs_conv_abscissa_ln: "abs_conv_abscissa (fds_ln l (f :: 'a :: dirichlet_series fds)) = abs_conv_abscissa (fds_deriv f / f)" by (simp add: fds_ln_def abs_conv_abscissa_integral) lemma abs_conv_abscissa_deriv: fixes f :: "'a fds" shows "abs_conv_abscissa (fds_deriv f) = abs_conv_abscissa f" proof - have "abs_conv_abscissa (fds_deriv f) = abs_conv_abscissa (fds_integral (fds_nth f 1) (fds_deriv f))" by (rule abs_conv_abscissa_integral [symmetric]) also have "fds_integral (fds_nth f 1) (fds_deriv f) = f" by (rule fds_integral_fds_deriv) finally show ?thesis . qed lemma abs_conv_abscissa_higher_deriv: "abs_conv_abscissa ((fds_deriv ^^ n) f) = abs_conv_abscissa (f :: 'a :: dirichlet_series fds)" by (induction n) (simp_all add: abs_conv_abscissa_deriv) lemma conv_abscissa_higher_deriv_le: "conv_abscissa ((fds_deriv ^^ n) f) \ conv_abscissa (f :: 'a :: dirichlet_series fds)" by (induction n) (auto intro: order.trans[OF conv_abscissa_deriv_le]) lemma abs_conv_abscissa_restrict: "abs_conv_abscissa (fds_subseries P f) \ abs_conv_abscissa f" by (rule abs_conv_abscissa_mono) auto lemma eval_fds_deriv: fixes f :: "'a fds" assumes "s \ 1 > conv_abscissa f" shows "eval_fds (fds_deriv f) s = deriv (eval_fds f) s" by (intro DERIV_imp_deriv [symmetric] derivative_intros assms) lemma eval_fds_higher_deriv: assumes "(s :: 'a :: dirichlet_series) \ 1 > conv_abscissa f" shows "eval_fds ((fds_deriv ^^ n) f) s = (deriv ^^ n) (eval_fds f) s" using assms proof (induction n arbitrary: f s) case (Suc n f s) have ev: "eventually (\s. s \ {s. s \ 1 > conv_abscissa f}) (nhds s)" using Suc.prems open_halfspace_gt[of _ "1::'a"] by (intro eventually_nhds_in_open, cases "conv_abscissa f") (auto simp: open_halfspace_gt inner_commute) have "eval_fds ((fds_deriv ^^ Suc n) f) s = eval_fds ((fds_deriv ^^ n) (fds_deriv f)) s" by (subst funpow_Suc_right) simp also have "\ = (deriv ^^ n) (eval_fds (fds_deriv f)) s" by (intro Suc.IH le_less_trans[OF conv_abscissa_deriv_le] Suc.prems) also have "\ = (deriv ^^ n) (deriv (eval_fds f)) s" by (intro higher_deriv_cong_ev refl eventually_mono[OF ev] eval_fds_deriv) auto also have "\ = (deriv ^^ Suc n) (eval_fds f) s" by (subst funpow_Suc_right) simp finally show ?case . qed auto end subsection \Multiplication of two series\ lemma fixes f g :: "nat \ 'a :: {banach, real_normed_field, second_countable_topology, nat_power}" fixes s :: 'a assumes [simp]: "f 0 = 0" "g 0 = 0" assumes summable: "summable (\n. norm (f n / nat_power n s))" "summable (\n. norm (g n / nat_power n s))" shows summable_dirichlet_prod: "summable (\n. norm (dirichlet_prod f g n / nat_power n s))" and suminf_dirichlet_prod: "(\n. dirichlet_prod f g n / nat_power n s) = (\n. f n / nat_power n s) * (\n. g n / nat_power n s)" proof - have summable': "(\n. f n / nat_power n s) abs_summable_on A" "(\n. g n / nat_power n s) abs_summable_on A" for A by ((rule abs_summable_on_subset[OF _ subset_UNIV], insert summable, simp add: abs_summable_on_nat_iff'); fail)+ have f_g: "f a / nat_power a s * (g b / nat_power b s) = f a * g b / nat_power (a * b) s" for a b by (cases "a * b = 0") (auto simp: nat_power_mult_distrib) have eq: "(\\<^sub>a(m, n)\{(m, n). m * n = x}. f m * g n / nat_power x s) = dirichlet_prod f g x / nat_power x s" for x :: nat proof (cases "x > 0") case False hence "(\\<^sub>a(m,n) | m * n = x. f m * g n / nat_power x s) = (\\<^sub>a(m,n) | m * n = x. 0)" by (intro infsetsum_cong) auto with False show ?thesis by simp next case True from finite_divisors_nat'[OF this] show ?thesis by (simp add: dirichlet_prod_altdef2 case_prod_unfold sum_divide_distrib) qed have "(\(m,n). (f m / nat_power m s) * (g n / nat_power n s)) abs_summable_on UNIV \ UNIV" using summable' by (intro abs_summable_on_product) auto also have "?this \ (\(m,n). f m * g n / nat_power (m*n) s) abs_summable_on UNIV" using f_g by (intro abs_summable_on_cong) auto also have "\ \ (\(x,(m,n)). f m * g n / nat_power (m*n) s) abs_summable_on (SIGMA x:UNIV. {(m,n). m * n = x})" unfolding case_prod_unfold by (rule abs_summable_on_reindex_bij_betw [symmetric]) (auto simp: bij_betw_def inj_on_def image_iff) also have "\ \ (\(x,(m,n)). f m * g n / nat_power x s) abs_summable_on (SIGMA x:UNIV. {(m,n). m * n = x})" by (intro abs_summable_on_cong) auto finally have summable'': \ . from abs_summable_on_Sigma_project1'[OF this] show summable''': "summable (\n. norm (dirichlet_prod f g n / nat_power n s))" by (simp add: eq abs_summable_on_nat_iff') have "(\n. f n / nat_power n s) * (\n. g n / nat_power n s) = (\\<^sub>an. f n / nat_power n s) * (\\<^sub>an. g n / nat_power n s)" using summable' by (simp add: infsetsum_nat') also have "\ = (\\<^sub>a(m,n). (f m / nat_power m s) * (g n / nat_power n s))" using summable' by (subst infsetsum_product [symmetric]) simp_all also have "\ = (\\<^sub>a(m,n). f m * g n / nat_power (m * n) s)" using f_g by (intro infsetsum_cong refl) auto also have "\ = (\\<^sub>a(x,(m,n))\(SIGMA x:UNIV. {(m,n). m * n = x}). f m * g n / nat_power (m * n) s)" unfolding case_prod_unfold by (rule infsetsum_reindex_bij_betw [symmetric]) (auto simp: bij_betw_def inj_on_def image_iff) also have "\ = (\\<^sub>a(x,(m,n))\(SIGMA x:UNIV. {(m,n). m * n = x}). f m * g n / nat_power x s)" by (intro infsetsum_cong refl) (auto simp: case_prod_unfold) also have "\ = (\\<^sub>ax. dirichlet_prod f g x / nat_power x s)" (is "_ = infsetsum ?T _") using summable'' by (subst infsetsum_Sigma) (auto simp: eq) also have "\ = (\x. dirichlet_prod f g x / nat_power x s)" using summable''' by (intro infsetsum_nat') (simp_all add: abs_summable_on_nat_iff') finally show "\ = (\n. f n / nat_power n s) * (\n. g n / nat_power n s)" .. qed lemma fixes f g :: "nat \ real" fixes s :: real assumes "f 0 = 0" "g 0 = 0" assumes summable: "summable (\n. norm (f n / real n powr s))" "summable (\n. norm (g n / real n powr s))" shows summable_dirichlet_prod_real: "summable (\n. norm (dirichlet_prod f g n / real n powr s))" and suminf_dirichlet_prod_real: "(\n. dirichlet_prod f g n / real n powr s) = (\n. f n / nat_power n s) * (\n. g n / real n powr s)" using summable_dirichlet_prod[of f g s] suminf_dirichlet_prod[of f g s] assms by simp_all lemma fds_abs_converges_mult: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges g s" shows "fds_abs_converges (f * g) s" using summable_dirichlet_prod[OF _ _ assms[unfolded fds_abs_converges_def]] by (simp add: fds_abs_converges_def fds_nth_mult) lemma fds_abs_converges_power: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" shows "fds_abs_converges f s \ fds_abs_converges (f ^ n) s" by (induction n) (auto intro!: fds_abs_converges_mult) lemma fds_abs_converges_prod: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" shows "(\x. x \ A \ fds_abs_converges (f x) s) \ fds_abs_converges (prod f A) s" by (induction A rule: infinite_finite_induct) (auto intro!: fds_abs_converges_mult) lemma abs_conv_abscissa_mult_le: "abs_conv_abscissa (f * g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c') thus ?case by (auto intro!: exI[of _ "of_real c'"] fds_abs_converges_mult intro: fds_abs_converges) qed lemma abs_conv_abscissa_mult_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_mult_le[of f g] by (auto simp add: le_max_iff_disj) lemma abs_conv_abscissa_shift [simp]: "abs_conv_abscissa (fds_shift c f) = abs_conv_abscissa (f :: 'a :: dirichlet_series fds) + c \ 1" proof - have "abs_conv_abscissa (fds_shift c f) \ abs_conv_abscissa f + c \ 1" for c :: 'a and f proof (rule abs_conv_abscissa_leI) fix d assume "abs_conv_abscissa f + c \ 1 < ereal d" hence "abs_conv_abscissa f < ereal (d - c \ 1)" by (cases "abs_conv_abscissa f") auto hence "fds_abs_converges (fds_shift c f) (of_real d)" by (auto intro!: fds_abs_converges_shift fds_abs_converges simp: algebra_simps) thus "\s. s \ 1 = d \ fds_abs_converges (fds_shift c f) s" by (auto intro!: exI[of _ "of_real d"]) qed note * = this[of c f] this[of "-c" "fds_shift c f"] show ?thesis by (cases "abs_conv_abscissa (fds_shift c f)"; cases "abs_conv_abscissa f") (insert *, auto intro!: antisym) qed lemma eval_fds_mult: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges g s" shows "eval_fds (f * g) s = eval_fds f s * eval_fds g s" using suminf_dirichlet_prod[OF _ _ assms[unfolded fds_abs_converges_def]] by (simp_all add: eval_fds_def fds_nth_mult) lemma eval_fds_power: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" shows "eval_fds (f ^ n) s = eval_fds f s ^ n" using assms by (induction n) (simp_all add: eval_fds_mult fds_abs_converges_power) lemma eval_fds_prod: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "(\x. x \ A \ fds_abs_converges (f x) s)" shows "eval_fds (prod f A) s = (\x\A. eval_fds (f x) s)" using assms by (induction A rule: infinite_finite_induct) (auto simp: eval_fds_mult fds_abs_converges_prod) lemma eval_fds_inverse: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges (inverse f) s" "fds_nth f 1 \ 0" shows "eval_fds (inverse f) s = inverse (eval_fds f s)" proof - have "eval_fds (inverse f * f) s = eval_fds (inverse f) s * eval_fds f s" by (intro eval_fds_mult assms) also have "inverse f * f = 1" by (intro fds_left_inverse assms) also have "eval_fds 1 s = 1" by simp finally show ?thesis by (auto simp: divide_simps) qed lemma eval_fds_integral_has_field_derivative: fixes s :: "'a :: dirichlet_series" assumes "ereal (s \ 1) > abs_conv_abscissa f" assumes "fds_nth f 1 = 0" shows "(eval_fds (fds_integral c f) has_field_derivative eval_fds f s) (at s)" proof - have "conv_abscissa (fds_integral c f) \ abs_conv_abscissa (fds_integral c f)" by (rule conv_le_abs_conv_abscissa) also from assms have "\ < ereal (s \ 1)" by (simp add: abs_conv_abscissa_integral) finally have "(eval_fds (fds_integral c f) has_field_derivative eval_fds (fds_deriv (fds_integral c f)) s) (at s)" by (intro derivative_eq_intros) auto also from assms have "fds_deriv (fds_integral c f) = f" by simp finally show ?thesis . qed lemma holomorphic_fds_eval [holomorphic_intros]: "A \ {z. Re z > conv_abscissa f} \ eval_fds f holomorphic_on A" unfolding holomorphic_on_def field_differentiable_def by (rule ballI exI derivative_intros)+ auto lemma analytic_fds_eval [holomorphic_intros]: assumes "A \ {z. Re z > conv_abscissa f}" shows "eval_fds f analytic_on A" proof - have "eval_fds f analytic_on {z. Re z > conv_abscissa f}" proof (subst analytic_on_open) show "open {z. Re z > conv_abscissa f}" by (cases "conv_abscissa f") (simp_all add: open_halfspace_Re_gt) qed (intro holomorphic_intros, simp_all) from analytic_on_subset[OF this assms] show ?thesis . qed lemma conv_abscissa_0 [simp]: "conv_abscissa (0 :: 'a :: dirichlet_series fds) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_0 [simp]: "abs_conv_abscissa (0 :: 'a :: dirichlet_series fds) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_1 [simp]: "conv_abscissa (1 :: 'a :: dirichlet_series fds) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_1 [simp]: "abs_conv_abscissa (1 :: 'a :: dirichlet_series fds) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_const [simp]: "conv_abscissa (fds_const (c :: 'a :: dirichlet_series)) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_const [simp]: "abs_conv_abscissa (fds_const (c :: 'a :: dirichlet_series)) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_numeral [simp]: "conv_abscissa (numeral n :: 'a :: dirichlet_series fds) = -\" by (auto simp: numeral_fds) lemma abs_conv_abscissa_numeral [simp]: "abs_conv_abscissa (numeral n :: 'a :: dirichlet_series fds) = -\" by (auto simp: numeral_fds) lemma abs_conv_abscissa_power_le: "abs_conv_abscissa (f ^ n :: 'a :: dirichlet_series fds) \ abs_conv_abscissa f" by (induction n) (auto intro!: order.trans[OF abs_conv_abscissa_mult_le]) lemma abs_conv_abscissa_power_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa (f ^ n) \ d" by (rule order.trans[OF abs_conv_abscissa_power_le]) lemma abs_conv_abscissa_prod_le: assumes "\x. x \ A \ abs_conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "abs_conv_abscissa (prod f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: abs_conv_abscissa_mult_leI) lemma conv_abscissa_add_le: "conv_abscissa (f + g :: 'a :: dirichlet_series fds) \ max (conv_abscissa f) (conv_abscissa g)" by (rule conv_abscissa_leI_weak) (auto intro!: fds_converges_add intro: fds_converges) lemma conv_abscissa_add_leI: "conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ conv_abscissa g \ d \ conv_abscissa (f + g) \ d" using conv_abscissa_add_le[of f g] by (auto simp: le_max_iff_disj) lemma conv_abscissa_sum_leI: assumes "\x. x \ A \ conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "conv_abscissa (sum f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: conv_abscissa_add_leI) lemma abs_conv_abscissa_add_le: "abs_conv_abscissa (f + g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" by (rule abs_conv_abscissa_leI_weak) (auto intro!: fds_abs_converges_add intro: fds_abs_converges) lemma abs_conv_abscissa_add_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_add_le[of f g] by (auto simp: le_max_iff_disj) lemma abs_conv_abscissa_sum_leI: assumes "\x. x \ A \ abs_conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "abs_conv_abscissa (sum f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: abs_conv_abscissa_add_leI) lemma fds_converges_cmult_left [intro]: assumes "fds_converges f s" shows "fds_converges (fds_const c * f) s" proof - from assms have "summable (\n. c * (fds_nth f n / nat_power n s))" by (intro summable_mult) (auto simp: fds_converges_def) thus ?thesis by (simp add: fds_converges_def mult_ac) qed lemma fds_converges_cmult_right [intro]: assumes "fds_converges f s" shows "fds_converges (f * fds_const c) s" using fds_converges_cmult_left[OF assms] by (simp add: mult_ac) lemma conv_abscissa_cmult_left [simp]: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "conv_abscissa (fds_const c * f) = conv_abscissa f" proof - have "fds_converges (fds_const c * f) s \ fds_converges f s" for s proof assume "fds_converges (fds_const c * f) s" hence "fds_converges (fds_const (inverse c) * (fds_const c * f)) s" by (rule fds_converges_cmult_left) also have "fds_const (inverse c) * (fds_const c * f) = fds_const (inverse c * c) * f" by simp also have "inverse c * c = 1" using assms by simp finally show "fds_converges f s" by simp qed auto thus ?thesis by (simp add: conv_abscissa_def) qed lemma conv_abscissa_cmult_right [simp]: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "conv_abscissa (f * fds_const c) = conv_abscissa f" using assms by (subst mult.commute) auto lemma abs_conv_abscissa_cmult: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "abs_conv_abscissa (fds_const c * f) = abs_conv_abscissa f" proof (intro antisym) have "abs_conv_abscissa (fds_const (inverse c) * (fds_const c * f)) \ abs_conv_abscissa (fds_const c * f)" using abs_conv_abscissa_mult_le[of "fds_const (inverse c)" "fds_const c * f"] by (auto simp: max_def) also have "fds_const (inverse c) * (fds_const c * f) = fds_const (inverse c * c) * f" by (simp add: mult_ac) also have "inverse c * c = 1" using assms by simp finally show "abs_conv_abscissa f \ abs_conv_abscissa (fds_const c * f)" by simp qed (insert abs_conv_abscissa_mult_le[of "fds_const c" f], auto simp: max_def) lemma conv_abscissa_minus [simp]: fixes f :: "'a :: dirichlet_series fds" shows "conv_abscissa (-f) = conv_abscissa f" using conv_abscissa_cmult_left[of "-1" f] by simp lemma abs_conv_abscissa_minus [simp]: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (-f) = abs_conv_abscissa f" using abs_conv_abscissa_cmult[of "-1" f] by simp lemma conv_abscissa_diff_le: "conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (conv_abscissa f) (conv_abscissa g)" using conv_abscissa_add_le[of f "-g"] by simp lemma conv_abscissa_diff_leI: "conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ conv_abscissa g \ d \ conv_abscissa (f - g) \ d" using conv_abscissa_add_le[of f "-g"] by (auto simp: le_max_iff_disj) 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 simp 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_add_le[of f "-g"] by (auto simp: le_max_iff_disj) lemmas eval_fds_integral_has_field_derivative' [derivative_intros] = DERIV_chain'[OF _ eval_fds_integral_has_field_derivative] lemma abs_conv_abscissa_completely_multiplicative_log_deriv: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "abs_conv_abscissa (fds_deriv f / f) \ abs_conv_abscissa f" proof - have "fds_deriv f = - fds (\n. fds_nth f n * mangoldt n) * f" using assms by (subst completely_multiplicative_fds_deriv') simp_all also have "\ / f = - fds (\n. fds_nth f n * mangoldt n) * (f * inverse f)" by (simp add: divide_fds_def) also have "f * inverse f = 1" using assms by (intro fds_right_inverse) finally have "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n)" by simp also have "abs_conv_abscissa \ = abs_conv_abscissa (fds (\n. fds_nth f n * mangoldt n))" (is "_ = abs_conv_abscissa ?f") by (rule abs_conv_abscissa_minus) also have "\ \ abs_conv_abscissa f" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) have "fds_abs_converges ?f (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from 1 have "fds_abs_converges (fds_deriv f) (of_real c)" by (intro fds_abs_converges) (auto simp: abs_conv_abscissa_deriv) thus "summable (\n. \ln (real n)\ * norm (fds_nth f n) / norm (nat_power n (of_real c :: 'a)))" by (simp add: fds_abs_converges_def fds_deriv_def fds_nth_fds' scaleR_conv_of_real powr_minus norm_mult norm_divide norm_nat_power) next show "\\<^sub>F n in sequentially. norm (norm (fds_nth (fds (\n. fds_nth f n * mangoldt n)) n / nat_power n (of_real c))) \ \ln (real n)\ * norm (fds_nth f n) / norm (nat_power n (of_real c) :: 'a)" using eventually_gt_at_top[of 0] proof eventually_elim case (elim n) have "norm (norm (fds_nth (fds (\n. fds_nth f n * mangoldt n)) n / nat_power n (of_real c))) = norm (fds_nth f n) * mangoldt n / real n powr c" using elim by (simp add: fds_nth_fds' norm_mult norm_divide norm_nat_power abs_mult mangoldt_nonneg) also have "\ \ norm (fds_nth f n) * ln n / real n powr c" using elim by (intro mult_left_mono divide_right_mono mangoldt_le) (simp_all add: mangoldt_def) finally show ?case using elim by (simp add: norm_nat_power algebra_simps) qed qed thus ?case by (intro exI[of _ "of_real c"]) auto qed finally show ?thesis . qed subsection \Uniqueness\ context assumes "SORT_CONSTRAINT ('a :: dirichlet_series)" begin lemma norm_dirichlet_series_cutoff_le: assumes "fds_abs_converges f (s0 :: 'a)" "N > 0" "s \ 1 \ c" "c \ s0 \ 1" shows "summable (\n. fds_nth f (n + N) / nat_power (n + N) s)" "summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" and "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c) / nat_power N (s \ 1 - c)" proof - from assms have "fds_abs_converges f (of_real c)" using fds_abs_converges_Re_le[of f s0 "of_real c"] by auto hence "summable (\n. norm (fds_nth f (n + N) / nat_power (n + N) (of_real c)))" unfolding fds_abs_converges_def by (rule summable_ignore_initial_segment) also have "?this \ summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_divide norm_nat_power) finally show *: "summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" . from assms have "fds_abs_converges f s" using fds_abs_converges_Re_le[of f s0 s] by auto hence **: "summable (\n. norm (fds_nth f (n + N) / nat_power (n + N) s))" unfolding fds_abs_converges_def by (rule summable_ignore_initial_segment) thus "summable (\n. fds_nth f (n + N) / nat_power (n + N) s)" by (rule summable_norm_cancel) have "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ (\n. norm (fds_nth f (n + N) / nat_power (n + N) s))" by (intro summable_norm **) also have "\ \ (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c / nat_power N (s \ 1 - c))" proof (intro suminf_le * ** summable_divide allI) fix n :: nat have "real N powr (s \ 1 - c) \ real (n + N) powr (s \ 1 - c)" using assms by (intro powr_mono2) simp_all also have "real (n + N) powr c * \ = real (n + N) powr (s \ 1)" by (simp add: powr_diff) finally have "norm (fds_nth f (n + N)) / real (n + N) powr (s \ 1) \ norm (fds_nth f (n + N)) / (real (n + N) powr c * real N powr (s \ 1 - c))" using \N > 0\ by (intro divide_left_mono) (simp_all add: mult_left_mono) thus "norm (fds_nth f (n + N) / nat_power (n + N) s) \ norm (fds_nth f (n + N)) / nat_power (n + N) c / nat_power N (s \ 1 - c)" using \N > 0\ by (simp add: norm_divide norm_nat_power ) qed also have "\ = (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c) / nat_power N (s \ 1 - c)" using * by (rule suminf_divide) finally show "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ \" . qed lemma eval_fds_zeroD_aux: fixes h :: "'a fds" assumes conv: "fds_abs_converges h (s0 :: 'a)" assumes freq: "frequently (\s. eval_fds h s = 0) ((\s. s \ 1) going_to at_top)" shows "h = 0" proof (rule ccontr) assume "h \ 0" hence ex: "\n>0. fds_nth h n \ 0" by (auto simp: fds_eq_iff) define N :: nat where "N = (LEAST n. n > 0 \ fds_nth h n \ 0)" have N: "N > 0" "fds_nth h N \ 0" using LeastI_ex[OF ex, folded N_def] by auto have less_N: "fds_nth h n = 0" if "n < N" for n using Least_le[of "\n. n > 0 \ fds_nth h n \ 0" n, folded N_def] that by (cases "n = 0") (auto simp: not_less) define c where "c = s0 \ 1" define remainder where "remainder = (\s. (\n. fds_nth h (n + Suc N) / nat_power (n + Suc N) s))" define A where "A = (\n. norm (fds_nth h (n + Suc N)) / nat_power (n + Suc N) c) * nat_power (Suc N) c" have eq: "fds_nth h N = nat_power N s * eval_fds h s - nat_power N s * remainder s" if "s \ 1 \ c" for s :: 'a proof - from conv and that have conv': "fds_abs_converges h s" unfolding c_def by (rule fds_abs_converges_Re_le) hence conv'': "fds_converges h s" by blast from conv'' have "(\n. fds_nth h n / nat_power n s) sums eval_fds h s" by (simp add: fds_converges_iff) hence "(\n. fds_nth h (n + Suc N) / nat_power (n + Suc N) s) sums (eval_fds h s - (\nnn = fds_nth h N / nat_power N s" by (subst sum.delta) auto finally show ?thesis unfolding remainder_def using \N > 0\ by (auto simp: sums_iff field_simps) qed have remainder_bound: "norm (remainder s) \ A / real (Suc N) powr (s \ 1)" if "s \ 1 \ c" for s :: 'a proof - note * = norm_dirichlet_series_cutoff_le[of h s0 "Suc N" c s, folded remainder_def] have "norm (remainder s) \ (\n. norm (fds_nth h (n + Suc N)) / nat_power (n + Suc N) c) / nat_power (Suc N) (s \ 1 - c)" using that assms unfolding remainder_def by (intro *) (simp_all add: c_def) also have "\ = A / real (Suc N) powr (s \ 1)" by (simp add: A_def powr_diff) finally show ?thesis . qed from freq have "\c. \s. s \ 1 \ c \ eval_fds h s = 0" unfolding frequently_def by (auto simp: eventually_going_to_at_top_linorder) hence "\k. \s. s \ 1 \ real k \ eval_fds h s = 0" by blast then obtain S where S: "\k. S k \ 1 \ real k \ eval_fds h (S k) = 0" by metis have S_limit: "filterlim (\k. S k \ 1) at_top sequentially" by (rule filterlim_at_top_mono[OF filterlim_real_sequentially]) (use S in auto) have "eventually (\k. real k \ c) sequentially" by real_asymp hence "eventually (\k. norm (fds_nth h N) \ (real N / real (Suc N)) powr (S k \ 1) * A) sequentially" proof eventually_elim case (elim k) hence "norm (fds_nth h N) = real N powr (S k \ 1) * norm (remainder (S k))" (is "_ = _ * ?X") using \N > 0\ S[of k] eq[of "S k"] by (auto simp: norm_mult norm_nat_power c_def) also have "norm (remainder (S k)) \ A / real (Suc N) powr (S k \ 1)" using elim S[of k] by (intro remainder_bound) (simp_all add: c_def) finally show ?case using N by (simp add: mult_left_mono powr_divide field_simps del: of_nat_Suc) qed moreover have "((\k. (real N / real (Suc N)) powr (S k \ 1) * A) \ 0) sequentially" by (rule filterlim_compose[OF _ S_limit]) (use \N > 0\ in real_asymp) ultimately have "((\_. fds_nth h N) \ 0) sequentially" by (rule Lim_null_comparison) hence "fds_nth h N = 0" by (simp add: tendsto_const_iff) with \fds_nth h N \ 0\ show False by contradiction qed lemma eval_fds_zeroD: fixes h :: "'a fds" assumes conv: "conv_abscissa h < \" assumes freq: "frequently (\s. eval_fds h s = 0) ((\s. s \ 1) going_to at_top)" shows "h = 0" proof - have [simp]: "2 \ (1 :: 'a) = 2" using of_real_inner_1[of 2] unfolding of_real_numeral by simp from conv obtain s where "fds_converges h s" by auto hence "fds_abs_converges h (s + 2)" by (rule fds_converges_imp_abs_converges) (auto simp: algebra_simps) from this assms(2-) show ?thesis by (rule eval_fds_zeroD_aux) qed lemma eval_fds_eqD: fixes f g :: "'a fds" assumes conv: "conv_abscissa f < \" "conv_abscissa g < \" assumes eq: "frequently (\s. eval_fds f s = eval_fds g s) ((\s. s \ 1) going_to at_top)" shows "f = g" proof - have conv': "conv_abscissa (f - g) < \" using assms by (intro le_less_trans[OF conv_abscissa_diff_le]) (auto simp: max_def) have "max (conv_abscissa f) (conv_abscissa g) < \" using conv by (auto simp: max_def) from ereal_dense2[OF this] obtain c where c: "max (conv_abscissa f) (conv_abscissa g) < ereal c" by auto (* TODO: something like "frequently_elim" would be great here *) have "frequently (\s. eval_fds f s = eval_fds g s \ s \ 1 \ c) ((\s. s \ 1) going_to at_top)" using eq by (rule frequently_eventually_frequently) auto hence *: "frequently (\s. eval_fds (f - g) s = 0) ((\s. s \ 1) going_to at_top)" proof (rule frequently_mono [rotated], safe, goal_cases) case (1 s) thus ?case using c by (subst eval_fds_diff) (auto intro!: fds_converges intro: less_le_trans) qed have "f - g = 0" by (rule eval_fds_zeroD fds_abs_converges_diff assms * conv')+ thus ?thesis by simp qed end subsection \Limit at infinity\ lemma eval_fds_at_top_tail_bound: fixes f :: "'a :: dirichlet_series fds" assumes c: "ereal c > abs_conv_abscissa f" defines "B \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" assumes s: "s \ 1 \ c" shows "norm (eval_fds f s - fds_nth f 1) \ B / 2 powr (s \ 1)" proof - from c have "fds_abs_converges f (of_real c)" by (intro fds_abs_converges) simp_all also have "?this \ summable (\n. norm (fds_nth f n) / real n powr c)" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_divide norm_nat_power norm_powr_real_powr) finally have summable_c: \ . note c also from s have "ereal c \ ereal (s \ 1)" by simp finally have "fds_abs_converges f s" by (intro fds_abs_converges) auto hence summable: "summable (\n. norm (fds_nth f n / nat_power n s))" by (simp add: fds_abs_converges_def) from summable_norm_cancel[OF this] have "(\n. fds_nth f n / nat_power n s) sums eval_fds f s" by (simp add: eval_fds_def sums_iff) from sums_split_initial_segment[OF this, of "Suc (Suc 0)"] have "norm (eval_fds f s - fds_nth f 1) = norm (\n. fds_nth f (n+2) / nat_power (n+2) s)" by (auto simp: sums_iff) also have "\ \ (\n. norm (fds_nth f (n+2) / nat_power (n+2) s))" by (intro summable_norm summable_ignore_initial_segment summable) also have "\ \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c / 2 powr (s \ 1 - c))" proof (intro suminf_le allI) fix n :: nat have "norm (fds_nth f (n + 2) / nat_power (n + 2) s) = norm (fds_nth f (n + 2)) / real (n+2) powr c / real (n+2) powr (s \ 1 - c)" by (simp add: field_simps powr_diff norm_divide norm_nat_power) also have "\ \ norm (fds_nth f (n + 2)) / real (n+2) powr c / 2 powr (s \ 1 - c)" using s by (intro divide_left_mono divide_nonneg_pos powr_mono2 mult_pos_pos) simp_all finally show "norm (fds_nth f (n + 2) / nat_power (n + 2) s) \ \" . qed (intro summable_ignore_initial_segment summable summable_divide summable_c)+ also have "\ = (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) / 2 powr (s \ 1 - c)" by (intro suminf_divide summable_ignore_initial_segment summable_c) also have "\ = B / 2 powr (s \ 1)" by (simp add: B_def powr_diff) finally show ?thesis . qed lemma tendsto_eval_fds_Re_at_top: assumes "conv_abscissa (f :: 'a :: dirichlet_series fds) \ \" assumes lim: "filterlim (\x. S x \ 1) at_top F" shows "((\x. eval_fds f (S x)) \ fds_nth f 1) F" proof - from assms(1) have "abs_conv_abscissa f < \" using abs_conv_le_conv_abscissa_plus_1[of f] by auto from ereal_dense2[OF this] obtain c where c: "abs_conv_abscissa f < ereal c" by auto define B where "B = (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" have *: "norm (eval_fds f s - fds_nth f 1) \ B / 2 powr (s \ 1)" if s: "s \ 1 \ c" for s using eval_fds_at_top_tail_bound[of f c s] that c by (simp add: B_def) moreover from lim have "eventually (\x. S x \ 1 \ c) F" by (auto simp: filterlim_at_top) ultimately have "eventually (\x. norm (eval_fds f (S x) - fds_nth f 1) \ B / 2 powr (S x \ 1)) F" by (auto elim!: eventually_mono) moreover have "((\x. B / 2 powr (S x \ 1)) \ 0) F" using filterlim_tendsto_pos_mult_at_top[OF tendsto_const[of "ln 2"] _ lim] by (intro real_tendsto_divide_at_top[OF tendsto_const]) (auto simp: powr_def mult_ac intro!: filterlim_compose[OF exp_at_top]) ultimately have "((\x. eval_fds f (S x) - fds_nth f 1) \ 0) F" by (rule Lim_null_comparison) thus ?thesis by (subst (asm) Lim_null [symmetric]) qed lemma tendsto_eval_fds_Re_at_top': assumes "conv_abscissa (f :: complex fds) \ \" shows "uniform_limit UNIV (\\ t. eval_fds f (of_real \ + of_real t * \) ) (\_ .fds_nth f 1) at_top" proof - from assms(1) have "abs_conv_abscissa f < \" using abs_conv_le_conv_abscissa_plus_1[of f] by auto from ereal_dense2[OF this] obtain c where c: "abs_conv_abscissa f < ereal c" by auto define B where "B \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" show ?thesis unfolding uniform_limit_iff proof safe fix \ :: real assume "\ > 0" hence "eventually (\\. B / 2 powr \ < \) at_top" by real_asymp thus "eventually (\\. \t\UNIV. dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) < \) at_top" using eventually_ge_at_top[of c] proof eventually_elim case (elim \) show ?case proof fix t :: real have "dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) \ B / 2 powr \" using eval_fds_at_top_tail_bound[of f c "of_real \ + of_real t * \"] elim c by (simp add: dist_norm B_def) also have "\ < \" by fact finally show "dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) < \" . qed qed qed qed lemma tendsto_eval_fds_Re_going_to_at_top: assumes "conv_abscissa (f :: 'a :: dirichlet_series fds) \ \" shows "((\s. eval_fds f s) \ fds_nth f 1) ((\s. s \ 1) going_to at_top)" using assms by (rule tendsto_eval_fds_Re_at_top) auto lemma tendsto_eval_fds_Re_going_to_at_top': assumes "conv_abscissa (f :: complex fds) \ \" shows "((\s. eval_fds f s) \ fds_nth f 1) (Re going_to at_top)" using assms by (rule tendsto_eval_fds_Re_at_top) auto text \ Any Dirichlet series that is not identically zero and does not diverge everywhere has a half-plane in which it converges and is non-zero. \ theorem fds_nonzero_halfplane_exists: fixes f :: "'a :: dirichlet_series fds" assumes "conv_abscissa f < \" "f \ 0" shows "eventually (\s. fds_converges f s \ eval_fds f s \ 0) ((\s. s \ 1) going_to at_top)" proof - from ereal_dense2[OF assms(1)] obtain c where c: "conv_abscissa f < ereal c" by auto have "eventually (\s::'a. s \ 1 > c) ((\s. s \ 1) going_to at_top)" using eventually_gt_at_top[of c] by auto hence "eventually (\s. fds_converges f s) ((\s. s \ 1) going_to at_top)" by eventually_elim (use c in \auto intro!: fds_converges simp: less_le_trans\) moreover have "eventually (\s. eval_fds f s \ 0) ((\s. s \ 1) going_to at_top)" using eval_fds_zeroD[OF assms(1)] assms(2) by (auto simp: frequently_def) ultimately show ?thesis by (rule eventually_conj) qed subsection \Normed series\ lemma fds_converges_norm_iff [simp]: fixes s :: "'a :: {nat_power_normed_field,banach}" shows "fds_converges (fds_norm f) (s \ 1) \ fds_abs_converges f s" unfolding fds_converges_def fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_abs_converges_norm_iff [simp]: fixes s :: "'a :: {nat_power_normed_field,banach}" shows "fds_abs_converges (fds_norm f) (s \ 1) \ fds_abs_converges f s" unfolding fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_converges_norm_iff': fixes f :: "'a :: {nat_power_normed_field,banach} fds" shows "fds_converges (fds_norm f) s \ fds_abs_converges f (of_real s)" unfolding fds_converges_def fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_abs_converges_norm_iff': fixes f :: "'a :: {nat_power_normed_field,banach} fds" shows "fds_abs_converges (fds_norm f) s \ fds_abs_converges f (of_real s)" unfolding fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma abs_conv_abscissa_norm [simp]: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (fds_norm f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa f \ abs_conv_abscissa (fds_norm f)" proof (rule abs_conv_abscissa_leI_weak) fix x assume "abs_conv_abscissa (fds_norm f) < ereal x" hence "fds_abs_converges (fds_norm f) (of_real x)" by (intro fds_abs_converges) auto thus "fds_abs_converges f (of_real x)" by (simp add: fds_abs_converges_norm_iff') qed qed (auto intro!: abs_conv_abscissa_leI_weak simp: fds_abs_converges_norm_iff' fds_abs_converges) lemma conv_abscissa_norm [simp]: fixes f :: "'a :: dirichlet_series fds" shows "conv_abscissa (fds_norm f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa f \ conv_abscissa (fds_norm f)" proof (rule abs_conv_abscissa_leI_weak) fix x assume "conv_abscissa (fds_norm f) < ereal x" hence "fds_converges (fds_norm f) (of_real x)" by (intro fds_converges) auto thus "fds_abs_converges f (of_real x)" by (simp add: fds_converges_norm_iff') qed qed (auto intro!: conv_abscissa_leI_weak simp: fds_abs_converges) lemma fixes f g :: "'a :: dirichlet_series fds" assumes "fds_abs_converges (fds_norm f) s" "fds_abs_converges (fds_norm g) s" shows fds_abs_converges_norm_mult: "fds_abs_converges (fds_norm (f * g)) s" and eval_fds_norm_mult_le: "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" proof - show conv: "fds_abs_converges (fds_norm (f * g)) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) have "fds_abs_converges (fds_norm f * fds_norm g) s" by (rule fds_abs_converges_mult assms)+ thus "summable (\n. norm (fds_nth (fds_norm f * fds_norm g) n) / nat_power n s)" by (simp add: fds_abs_converges_def) qed (auto intro!: always_eventually divide_right_mono order.trans[OF fds_nth_norm_mult_le] simp: norm_divide) have conv': "fds_abs_converges (fds_norm f * fds_norm g) s" by (intro fds_abs_converges_mult assms) hence "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f * fds_norm g) s" using conv unfolding eval_fds_def fds_abs_converges_def norm_divide by (intro suminf_le allI divide_right_mono) (simp_all add: norm_mult fds_nth_norm_mult_le) also have "\ = eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" by (intro eval_fds_mult assms) finally show "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" . qed lemma eval_fds_norm_nonneg: assumes "fds_abs_converges (fds_norm f) s" shows "eval_fds (fds_norm f) s \ 0" using assms unfolding eval_fds_def fds_abs_converges_def by (intro suminf_nonneg) auto lemma fixes f :: "'a :: dirichlet_series fds" assumes "fds_abs_converges (fds_norm f) s" shows fds_abs_converges_norm_power: "fds_abs_converges (fds_norm (f ^ n)) s" and eval_fds_norm_power_le: "eval_fds (fds_norm (f ^ n)) s \ eval_fds (fds_norm f) s ^ n" proof - show *: "fds_abs_converges (fds_norm (f ^ n)) s" for n by (induction n) (auto intro!: fds_abs_converges_norm_mult assms) show "eval_fds (fds_norm (f ^ n)) s \ eval_fds (fds_norm f) s ^ n" by (induction n) (auto intro!: order.trans[OF eval_fds_norm_mult_le] assms * mult_left_mono eval_fds_norm_nonneg) qed subsection \Logarithms of Dirichlet series\ (* TODO: Move ? *) lemma eventually_gt_ereal_at_top: "c \ \ \ eventually (\x. ereal x > c) at_top" by (cases c) auto lemma eval_fds_log_deriv: fixes s :: "'a :: dirichlet_series" assumes "fds_nth f 1 \ 0" "s \ 1 > abs_conv_abscissa f" "s \ 1 > abs_conv_abscissa (fds_deriv f / f)" assumes "eval_fds f s \ 0" shows "eval_fds (fds_deriv f / f) s = eval_fds (fds_deriv f) s / eval_fds f s" proof - have "eval_fds (fds_deriv f / f * f) s = eval_fds (fds_deriv f / f) s * eval_fds f s" using assms by (intro eval_fds_mult fds_abs_converges) auto also have "fds_deriv f / f * f = fds_deriv f * (f * inverse f)" by (simp add: divide_fds_def algebra_simps) also have "f * inverse f = 1" using assms by (intro fds_right_inverse) finally show ?thesis using assms by simp qed text \ Given a sufficiently nice absolutely convergent Dirichlet series that converges to some function $f(s)$ and a holomorphic branch of $\ln f(s)$, we can construct a Dirichlet series that absolutely converges to that logarithm. \ lemma eval_fds_ln: fixes s0 :: ereal assumes nz: "\s. Re s > s0 \ eval_fds f s \ 0" "fds_nth f 1 \ 0" assumes l: "exp l = fds_nth f 1" "((g \ of_real) \ l) at_top" assumes g: "\s. Re s > s0 \ exp (g s) = eval_fds f s" assumes holo_g: "g holomorphic_on {s. Re s > s0}" assumes "ereal (Re s) > s0" assumes "s0 \ abs_conv_abscissa f" and "s0 \ abs_conv_abscissa (fds_deriv f / f)" shows "eval_fds (fds_ln l f) s = g s" proof - let ?s0 = "abs_conv_abscissa f" and ?s1 = "abs_conv_abscissa (inverse f)" let ?h = "\s. eval_fds (fds_ln l f) s - g s" let ?A = "{s. Re s > s0}" have open_A: "open ?A" by (cases s0) (auto simp: open_halfspace_Re_gt) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) moreover from assms have "\ \ \" by auto ultimately have "conv_abscissa f \ \" by auto have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) finally have "conv_abscissa (fds_ln l f) \ \" using assms by (auto simp: max_def abs_conv_abscissa_deriv split: if_splits) have deriv_g [derivative_intros]: "(g has_field_derivative eval_fds (fds_deriv f) s / eval_fds f s) (at s within B)" if s: "Re s > s0" for s B proof - have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ s0" using assms by simp also have "\ < Re s" by fact finally have s': "Re s > conv_abscissa f" . have deriv_g: "(g has_field_derivative deriv g s) (at s)" using holomorphic_derivI[OF holo_g open_A, of s] s by (auto simp: at_within_open[OF _ open_A]) have "((\s. exp (g s)) has_field_derivative eval_fds f s * deriv g s) (at s)" (is ?P) by (rule derivative_eq_intros deriv_g s)+ (insert s, simp_all add: g) also from s have ev: "eventually (\t. t \ ?A) (nhds s)" by (intro eventually_nhds_in_open open_A) auto have "?P \ (eval_fds f has_field_derivative eval_fds f s * deriv g s) (at s)" by (intro DERIV_cong_ev refl eventually_mono[OF ev]) (auto simp: g) finally have "(eval_fds f has_field_derivative eval_fds f s * deriv g s) (at s)" . moreover have "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s)" using s' assms by (intro derivative_intros) auto ultimately have "eval_fds f s * deriv g s = eval_fds (fds_deriv f) s" by (rule DERIV_unique) hence "deriv g s = eval_fds (fds_deriv f) s / eval_fds f s" using s nz by (simp add: field_simps) with deriv_g show ?thesis by (auto intro: has_field_derivative_at_within) qed have "\c. \z\{z. Re z > s0}. ?h z = c" proof (rule has_field_derivative_zero_constant, goal_cases) case 1 show ?case using convex_halfspace_gt[of _ "1::complex"] by (cases s0) auto next case (2 z) have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" by (simp add: abs_conv_abscissa_ln) also have "\ < Re z" using 2 assms by (auto simp: abs_conv_abscissa_deriv) finally have s1: "conv_abscissa (fds_ln l f) < ereal (Re z)" . have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ < Re z" using 2 assms by auto finally have s2: "conv_abscissa f < ereal (Re z)" . from l have "fds_nth f 1 \ 0" by auto with 2 assms have *: "eval_fds (fds_deriv f / f) z = eval_fds (fds_deriv f) z / (eval_fds f z)" by (auto simp: eval_fds_log_deriv) have "eval_fds f z \ 0" using 2 assms by auto show ?case using s1 s2 2 nz by (auto intro!: derivative_eq_intros simp: * field_simps) qed then obtain c where c: "\z. Re z > s0 \ ?h z = c" by blast have "(at_top :: real filter) \ bot" by simp moreover from assms have "s0 \ \" by auto have "eventually (\x. c = (?h \ of_real) x) at_top" using eventually_gt_ereal_at_top[OF \s0 \ \\] by eventually_elim (simp add: c) hence "((?h \ of_real) \ c) at_top" by (force intro: Lim_transform_eventually) moreover have "((?h \ of_real) \ fds_nth (fds_ln l f) 1 - l) at_top" using \conv_abscissa (fds_ln l f) \ \\ and l unfolding o_def by (intro tendsto_intros tendsto_eval_fds_Re_at_top) (auto simp: filterlim_ident) ultimately have "c = fds_nth (fds_ln l f) 1 - l" by (rule tendsto_unique) with c[OF \Re s > s0\] and l and nz show ?thesis by (simp add: exp_minus field_simps) qed text \ Less explicitly: For a sufficiently nice absolutely convergent Dirichlet series converging to a function $f(s)$, the formal logarithm absolutely converges to some logarithm of $f(s)$. \ lemma eval_fds_ln': fixes s0 :: ereal assumes "ereal (Re s) > s0" assumes "s0 \ abs_conv_abscissa f" and "s0 \ abs_conv_abscissa (fds_deriv f / f)" and nz: "\s. Re s > s0 \ eval_fds f s \ 0" "fds_nth f 1 \ 0" assumes l: "exp l = fds_nth f 1" shows "exp (eval_fds (fds_ln l f) s) = eval_fds f s" proof - let ?s0 = "abs_conv_abscissa f" and ?s1 = "abs_conv_abscissa (inverse f)" let ?h = "\s. eval_fds f s * exp (-eval_fds (fds_ln l f) s)" have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) moreover from assms have "\ \ \" by auto ultimately have "conv_abscissa f \ \" by auto have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) finally have "conv_abscissa (fds_ln l f) \ \" using assms by (auto simp: max_def abs_conv_abscissa_deriv split: if_splits) have "\c. \z\{z. Re z > s0}. ?h z = c" proof (rule has_field_derivative_zero_constant, goal_cases) case 1 show ?case using convex_halfspace_gt[of _ "1::complex"] by (cases s0) auto next case (2 z) have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) also have "\ < Re z" using 2 assms by (auto simp: abs_conv_abscissa_deriv) finally have s1: "conv_abscissa (fds_ln l f) < ereal (Re z)" . have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ < Re z" using 2 assms by auto finally have s2: "conv_abscissa f < ereal (Re z)" . from l have "fds_nth f 1 \ 0" by auto with 2 assms have *: "eval_fds (fds_deriv f / f) z = eval_fds (fds_deriv f) z / (eval_fds f z)" by (subst eval_fds_log_deriv) auto have "eval_fds f z \ 0" using 2 assms by auto thus ?case using s1 s2 by (auto intro!: derivative_eq_intros simp: *) qed then obtain c where c: "\z. Re z > s0 \ ?h z = c" by blast have "(at_top :: real filter) \ bot" by simp moreover from assms have "s0 \ \" by auto have "eventually (\x. c = (?h \ of_real) x) at_top" using eventually_gt_ereal_at_top[OF \s0 \ \\] by eventually_elim (simp add: c) hence "((?h \ of_real) \ c) at_top" by (force intro: Lim_transform_eventually) moreover have "((?h \ of_real) \ fds_nth f 1 * exp (-fds_nth (fds_ln l f) 1)) at_top" unfolding o_def using \conv_abscissa (fds_ln l f) \ \\ and \conv_abscissa f \ \\ by (intro tendsto_intros tendsto_eval_fds_Re_at_top) (auto simp: filterlim_ident) ultimately have "c = fds_nth f 1 * exp (-fds_nth (fds_ln l f) 1)" by (rule tendsto_unique) with c[OF \Re s > s0\] and l and nz show ?thesis by (simp add: exp_minus field_simps) qed lemma fds_ln_completely_multiplicative: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" assumes "fds_nth f 1 \ 0" shows "fds_ln l f = fds (\n. if n = 1 then l else fds_nth f n * mangoldt n /\<^sub>R ln n)" proof - have "fds_ln l f = fds_integral l (fds_deriv f / f)" by (simp add: fds_ln_def) also have "fds_deriv f = -fds (\n. fds_nth f n * mangoldt n) * f" by (intro completely_multiplicative_fds_deriv' assms) also have "\ / f = -fds (\n. fds_nth f n * mangoldt n) * (f * inverse f)" by (simp add: divide_fds_def) also from assms have "f * inverse f = 1" by (simp add: fds_right_inverse) also have "fds_integral l (- fds (\n. fds_nth f n * mangoldt n) * 1) = fds (\n. if n = 1 then l else fds_nth f n * mangoldt n /\<^sub>R ln n)" by (simp add: fds_integral_def cong: if_cong) finally show ?thesis . qed lemma eval_fds_ln_completely_multiplicative_strong: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" and g :: "nat \ 'a" defines "h \ fds (\n. fds_nth (fds_ln l f) n * g n)" assumes "fds_abs_converges h s" assumes "completely_multiplicative_function (fds_nth f)" and "fds_nth f 1 \ 0" shows "(\(p,k). (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k)) abs_summable_on ({p. prime p} \ UNIV)" (is ?th1) and "eval_fds h s = l * g 1 + (\\<^sub>a(p, k)\{p. prime p}\UNIV. (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k))" (is ?th2) proof - let ?P = "{p::nat. prime p}" interpret f: completely_multiplicative_function "fds_nth f" by fact from assms have *: "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV" by (auto simp: abs_summable_on_nat_iff' fds_abs_converges_def) have eq: "h = fds (\n. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n))" using fds_ln_completely_multiplicative [OF assms(3), of l] by (simp add: h_def fds_eq_iff) note * also have "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV \ (\x. if x = Suc 0 then l * g 1 else fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on {1} \ Collect primepow" using eq by (intro abs_summable_on_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) finally have sum1: "(\x. if x = Suc 0 then l * g 1 else fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on Collect primepow" by (rule abs_summable_on_subset) auto also have "?this \ (\x. fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on Collect primepow" by (intro abs_summable_on_cong) (insert primepow_gt_Suc_0, auto) also have "\ \ (\(p,k). fds_nth f (p ^ Suc k) * g (p ^ Suc k) * mangoldt (p ^ Suc k) /\<^sub>R ln (real (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 "\ \ ?th1" by (intro abs_summable_on_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps scaleR_conv_of_real simp del: power_Suc) 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 \ {1} \ {n. primepow n}. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" by (intro infsetsum_cong_neutral) (auto simp: eq fds_nth_fds mangoldt_def) also have "\ = l * g 1 + (\\<^sub>an | primepow n. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" (is "_ = _ + ?x") using sum1 primepow_gt_Suc_0 by (subst infsetsum_Un_disjoint) auto also have "?x = (\\<^sub>an\Collect primepow. fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" (is "_ = infsetsum ?f _") by (intro infsetsum_cong refl) (insert primepow_gt_Suc_0, auto) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). fds_nth f (p ^ Suc k) * g (p ^ Suc k) * mangoldt (p ^ Suc k) /\<^sub>R ln (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 * g (p ^ Suc k) / of_nat (Suc k))" 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 scaleR_conv_of_real simp del: power_Suc) finally show ?th2 . qed lemma eval_fds_ln_completely_multiplicative: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_nth f 1 \ 0" assumes "s \ 1 > abs_conv_abscissa (fds_deriv f / f)" shows "(\(p,k). (fds_nth f p / nat_power p s) ^ Suc k / of_nat (Suc k)) abs_summable_on ({p. prime p} \ UNIV)" (is ?th1) and "eval_fds (fds_ln l f) s = l + (\\<^sub>a(p, k)\{p. prime p}\UNIV. (fds_nth f p / nat_power p s) ^ Suc k / of_nat (Suc k))" (is ?th2) proof - from assms have "fds_abs_converges (fds_ln l f) s" by (intro fds_abs_converges_ln) (auto intro!: fds_abs_converges_mult intro: fds_abs_converges) hence "fds_abs_converges (fds (\n. fds_nth (fds_ln l f) n * 1)) s" by simp from eval_fds_ln_completely_multiplicative_strong [OF this assms(1,2)] show ?th1 ?th2 by simp_all qed subsection \Exponential and logarithm\ lemma summable_fds_exp_aux: assumes "fds_nth f' 1 = (0 :: 'a :: real_normed_algebra_1)" shows "summable (\k. fds_nth (f' ^ k) n /\<^sub>R fact k)" proof (rule summable_finite) fix k assume "k \ {..n}" hence "n < k" by simp also have "\ < 2 ^ k" by (rule less_exp) finally have "fds_nth (f' ^ k) n = 0" using assms by (intro fds_nth_power_eq_0) auto thus "fds_nth (f' ^ k) n /\<^sub>R fact k = 0" by simp qed auto lemma fixes f :: "'a :: dirichlet_series fds" assumes "fds_abs_converges f s" shows fds_abs_converges_exp: "fds_abs_converges (fds_exp f) s" and eval_fds_exp: "eval_fds (fds_exp f) s = exp (eval_fds f s)" proof - have conv: "fds_abs_converges (fds_exp f) s" and ev: "eval_fds (fds_exp f) s = exp (eval_fds f s)" if "fds_abs_converges f s" and [simp]: "fds_nth f (Suc 0) = 0" for f proof - have [simp]: "fds (\n. if n = Suc 0 then 0 else fds_nth f n) = f" by (intro fds_eqI) simp_all have "(\(k,n). fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on (UNIV \ {1..})" proof (subst abs_summable_on_Sigma_iff, safe, goal_cases) case (3 k) from that have "fds_abs_converges (f ^ k) s" by (intro fds_abs_converges_power) hence "(\n. fds_nth (f ^ k) n / nat_power n s * inverse (fact k)) abs_summable_on {1..}" unfolding fds_abs_converges_altdef by (intro abs_summable_on_cmult_left) thus ?case by (simp add: field_simps) next case 4 show ?case unfolding abs_summable_on_nat_iff' proof (rule summable_comparison_test_ev[OF always_eventually[OF allI]]) fix k :: nat from that have *: "fds_abs_converges (fds_norm (f ^ k)) (s \ 1)" by (auto simp: fds_abs_converges_power) have "(\\<^sub>an\{1..}. norm (fds_nth (f ^ k) n / fact k / nat_power n s)) = (\\<^sub>an\{1..}. fds_nth (fds_norm (f ^ k)) n / nat_power n (s \ 1) / fact k)" (is "?S = _") by (intro infsetsum_cong) (simp_all add: norm_divide norm_mult norm_nat_power) also have "\ = (\\<^sub>an\{1..}. fds_nth (fds_norm (f ^ k)) n / nat_power n (s \ 1)) /\<^sub>R fact k" (is "_ = ?S' /\<^sub>R _") using * unfolding fds_abs_converges_altdef by (subst infsetsum_cdiv) (auto simp: abs_summable_on_nat_iff scaleR_conv_of_real divide_simps) also have "?S' = eval_fds (fds_norm (f ^ k)) (s \ 1)" using * unfolding fds_abs_converges_altdef eval_fds_def by (subst infsetsum_nat) (auto intro!: suminf_cong) finally have eq: "?S = \ /\<^sub>R fact k" . note eq also have "?S \ 0" by (intro infsetsum_nonneg) auto hence "?S = norm (norm ?S)" by simp also have "eval_fds (fds_norm (f ^ k)) (s \ 1) \ eval_fds (fds_norm f) (s \ 1) ^ k" using that by (intro eval_fds_norm_power_le) auto finally show "norm (norm (\\<^sub>an\{1..}. norm (fds_nth (f ^ k) n / fact k / nat_power n s))) \ eval_fds (fds_norm f) (s \ 1) ^ k /\<^sub>R fact k" by (simp add: divide_right_mono) next from exp_converges[of "eval_fds (fds_norm f) (s \ 1)"] show "summable (\x. eval_fds (fds_norm f) (s \ 1) ^ x /\<^sub>R fact x)" by (simp add: sums_iff) qed qed auto hence summable: "(\(n,k). fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on {1..} \ UNIV" by (subst abs_summable_on_Times_swap) (simp add: case_prod_unfold) have summable': "(\k. fds_nth (f ^ k) n / fact k) abs_summable_on UNIV" for n using abs_summable_on_cmult_left[of "nat_power n s", OF abs_summable_on_Sigma_project2 [OF summable, of n]] by (cases "n = 0") simp_all have "(\n. \\<^sub>ak. fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on {1..}" using summable by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\n. (\k. fds_nth (f ^ k) n / fact k) * inverse (nat_power n s)) abs_summable_on {1..}" proof (intro abs_summable_on_cong refl, goal_cases) case (1 n) hence "(\\<^sub>ak. fds_nth (f ^ k) n / fact k / nat_power n s) = (\\<^sub>ak. fds_nth (f ^ k) n / fact k) * inverse (nat_power n s)" using summable'[of n] by (subst infsetsum_cmult_left [symmetric]) (auto simp: field_simps) also have "(\\<^sub>ak. fds_nth (f ^ k) n / fact k) = (\k. fds_nth (f ^ k) n / fact k)" using summable'[of n] 1 by (intro abs_summable_on_cong refl infsetsum_nat') auto finally show ?case . qed finally show "fds_abs_converges (fds_exp f) s" by (simp add: fds_exp_def fds_nth_fds' abs_summable_on_Sigma_iff scaleR_conv_of_real fds_abs_converges_altdef field_simps) have "eval_fds (fds_exp f) s = (\n. (\k. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" by (simp add: fds_exp_def eval_fds_def fds_nth_fds') also have "\ = (\n. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" proof (intro suminf_cong, goal_cases) case (1 n) show ?case proof (cases "n = 0") case False have "(\k. fds_nth (f ^ k) n /\<^sub>R fact k) = (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k)" using summable'[of n] False by (intro infsetsum_nat' [symmetric]) (auto simp: scaleR_conv_of_real field_simps) thus ?thesis by simp qed simp_all qed also have "\ = (\\<^sub>an. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" proof (intro infsetsum_nat' [symmetric], goal_cases) case 1 have *: "UNIV - {Suc 0..} = {0}" by auto have "(\x. \\<^sub>ay. fds_nth (f ^ y) x / fact y / nat_power x s) abs_summable_on {1..}" by (intro abs_summable_on_Sigma_project1'[OF summable]) auto also have "?this \ (\x. (\\<^sub>ay. fds_nth (f ^ y) x / fact y) * inverse (nat_power x s)) abs_summable_on {1..}" using summable' by (intro abs_summable_on_cong refl, subst infsetsum_cmult_left [symmetric]) (auto simp: field_simps) also have "\ \ (\x. (\\<^sub>ay. fds_nth (f ^ y) x /\<^sub>R fact y) / (nat_power x s)) abs_summable_on {1..}" by (simp add: field_simps scaleR_conv_of_real) finally show ?case by (rule abs_summable_on_finite_diff) (use * in auto) qed also have "\ = (\\<^sub>an. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k * inverse (nat_power n s)))" using summable' by (subst infsetsum_cmult_left) (auto simp: field_simps scaleR_conv_of_real) also have "\ = (\\<^sub>an\{1..}. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k * inverse (nat_power n s)))" by (intro infsetsum_cong_neutral) (auto simp: Suc_le_eq) also have "\ = (\\<^sub>ak. \\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s /\<^sub>R fact k)" using summable by (subst infsetsum_swap) (auto simp: field_simps scaleR_conv_of_real case_prod_unfold) also have "\ = (\\<^sub>ak. (\\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s) /\<^sub>R fact k)" by (subst infsetsum_scaleR_right) simp also have "\ = (\\<^sub>ak. eval_fds f s ^ k /\<^sub>R fact k)" proof (intro infsetsum_cong refl, goal_cases) case (1 k) have *: "fds_abs_converges (f ^ k) s" by (intro fds_abs_converges_power that) have "(\\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s) = (\\<^sub>an. fds_nth (f ^ k) n / nat_power n s)" by (intro infsetsum_cong_neutral) (auto simp: Suc_le_eq) also have "\ = eval_fds (f ^ k) s" using * unfolding eval_fds_def by (intro infsetsum_nat') (auto simp: fds_abs_converges_def abs_summable_on_nat_iff') also from that have "\ = eval_fds f s ^ k" by (simp add: eval_fds_power) finally show ?case by simp qed also have "\ = (\k. eval_fds f s ^ k /\<^sub>R fact k)" using exp_converges[of "norm (eval_fds f s)"] by (intro infsetsum_nat') (auto simp: abs_summable_on_nat_iff' sums_iff field_simps norm_power) also have "\ = exp (eval_fds f s)" by (simp add: exp_def) finally show "eval_fds (fds_exp f) s = exp (eval_fds f s)" . qed define f' where "f' = f - fds_const (fds_nth f 1)" have *: "fds_abs_converges (fds_exp f') s" by (auto simp: f'_def intro!: fds_abs_converges_diff conv assms) have "fds_abs_converges (fds_const (exp (fds_nth f 1)) * fds_exp f') s" unfolding f'_def by (intro fds_abs_converges_mult conv fds_abs_converges_diff assms) auto thus "fds_abs_converges (fds_exp f) s" unfolding f'_def by (simp add: fds_exp_times_fds_nth_0) have "eval_fds (fds_exp f) s = eval_fds (fds_const (exp (fds_nth f 1)) * fds_exp f') s" by (simp add: f'_def fds_exp_times_fds_nth_0) also have "\ = exp (fds_nth f (Suc 0)) * eval_fds (fds_exp f') s" using * using assms by (subst eval_fds_mult) (simp_all) also have "\ = exp (eval_fds f s)" using ev[of f'] assms unfolding f'_def by (auto simp: fds_abs_converges_diff eval_fds_diff fds_abs_converges_imp_converges exp_diff) finally show "eval_fds (fds_exp f) s = exp (eval_fds f s)" . qed lemma fds_exp_add: fixes f :: "'a :: dirichlet_series fds" shows "fds_exp (f + g) = fds_exp f * fds_exp g" proof (rule fds_eqI_truncate) fix m :: nat assume m: "m > 0" let ?T = "fds_truncate m" have "?T (fds_exp (f + g)) = ?T (fds_exp (?T f + ?T g))" by (simp add: fds_truncate_exp fds_truncate_add_strong [symmetric]) also have "fds_exp (?T f + ?T g) = fds_exp (?T f) * fds_exp (?T g)" proof (rule eval_fds_eqD) have "fds_abs_converges (fds_exp (?T f + ?T g)) 0" by (intro fds_abs_converges_exp fds_abs_converges_add) auto thus "conv_abscissa (fds_exp (?T f + ?T g)) < \" using conv_abscissa_PInf_iff by blast hence "fds_abs_converges (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) 0" by (intro fds_abs_converges_mult fds_abs_converges_exp) auto thus "conv_abscissa (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) < \" using conv_abscissa_PInf_iff by blast show "frequently (\s. eval_fds (fds_exp (fds_truncate m f + fds_truncate m g)) s = eval_fds (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) s) ((\s. s \ 1) going_to at_top)" by (auto simp: eval_fds_add eval_fds_mult eval_fds_exp fds_abs_converges_add fds_abs_converges_exp exp_add) qed also have "?T \ = ?T (fds_exp f * fds_exp g)" by (subst fds_truncate_mult [symmetric], subst (1 2) fds_truncate_exp) (simp add: fds_truncate_mult) finally show "?T (fds_exp (f + g)) = \" . qed lemma fds_exp_minus: fixes f :: "'a :: dirichlet_series fds" shows "fds_exp (-f) = inverse (fds_exp f)" proof (rule fds_right_inverse_unique) have "fds_exp f * fds_exp (- f) = fds_exp (f + (-f))" by (subst fds_exp_add) simp_all also have "f + (-f) = 0" by simp also have "fds_exp \ = 1" by simp finally show "fds_exp f * fds_exp (-f) = 1" . qed lemma abs_conv_abscissa_exp: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (fds_exp f) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono fds_abs_converges_exp) lemma fds_deriv_exp [simp]: fixes f :: "'a :: dirichlet_series fds" shows "fds_deriv (fds_exp f) = fds_exp f * fds_deriv f" proof (rule fds_eqI_truncate) fix m :: nat assume m: "m > 0" let ?T = "fds_truncate m" have "abs_conv_abscissa (fds_deriv (?T f)) = -\" by (simp add: abs_conv_abscissa_deriv) have "?T (fds_deriv (fds_exp f)) = ?T (fds_deriv (fds_exp (?T f)))" by (simp add: fds_truncate_deriv fds_truncate_exp) also have "fds_deriv (fds_exp (?T f)) = fds_exp (?T f) * fds_deriv (?T f)" proof (rule eval_fds_eqD) note abscissa = conv_le_abs_conv_abscissa abs_conv_abscissa_exp note abscissa' = abscissa[THEN le_less_trans] have "fds_abs_converges (fds_deriv (fds_exp (fds_truncate m f))) 0" by (intro fds_abs_converges ) (auto simp: abs_conv_abscissa_deriv intro: le_less_trans[OF abs_conv_abscissa_exp]) thus "conv_abscissa (fds_deriv (fds_exp (fds_truncate m f))) < \" using conv_abscissa_PInf_iff by blast have "fds_abs_converges (fds_exp (fds_truncate m f) * fds_deriv (fds_truncate m f)) 0" by (intro fds_abs_converges_mult fds_abs_converges_exp) (auto intro: fds_abs_converges simp add: fds_truncate_deriv [symmetric]) thus "conv_abscissa (fds_exp (fds_truncate m f) * fds_deriv (fds_truncate m f)) < \" using conv_abscissa_PInf_iff by blast show "\\<^sub>F s in (\s. s \ 1) going_to at_top. eval_fds (fds_deriv (fds_exp (?T f))) s = eval_fds (fds_exp (?T f) * fds_deriv (?T f)) s" proof (intro always_eventually eventually_frequently allI, goal_cases) case (2 s) have "eval_fds (fds_deriv (fds_exp (?T f))) s = deriv (eval_fds (fds_exp (?T f))) s" by (auto simp: eval_fds_exp eval_fds_mult fds_abs_converges_mult fds_abs_converges_exp fds_abs_converges eval_fds_deriv abscissa') also have "eval_fds (fds_exp (?T f)) = (\s. exp (eval_fds (?T f) s))" by (intro ext eval_fds_exp) auto also have "deriv \ = (\s. exp (eval_fds (?T f) s) * deriv (eval_fds (?T f)) s)" by (auto intro!: DERIV_imp_deriv derivative_eq_intros simp: eval_fds_deriv) also have "\ = eval_fds (fds_exp (?T f) * fds_deriv (?T f))" by (auto simp: eval_fds_exp eval_fds_mult fds_abs_converges_mult fds_abs_converges_exp fds_abs_converges eval_fds_deriv abs_conv_abscissa_deriv) finally show ?case . qed auto qed also have "?T \ = ?T (fds_exp f * fds_deriv f)" by (subst fds_truncate_mult [symmetric]) (simp add: fds_truncate_exp fds_truncate_deriv [symmetric], simp add: fds_truncate_mult) finally show "?T (fds_deriv (fds_exp f)) = \" . qed lemma fds_exp_ln_strong: fixes f :: "'a :: dirichlet_series fds" assumes "fds_nth f (Suc 0) \ 0" shows "fds_exp (fds_ln l f) = fds_const (exp l / fds_nth f (Suc 0)) * f" proof - let ?c = "exp l / fds_nth f (Suc 0)" have "f * fds_const ?c = f * (fds_exp (-fds_ln l f) * fds_exp (fds_ln l f)) * fds_const ?c" (is "_ = _ * (?g * ?h) * _") by (subst fds_exp_add [symmetric]) simp also have "\ = fds_const ?c * (f * ?g) * ?h" by (simp add: mult_ac) also have "f * ?g = fds_const (inverse ?c)" proof (rule fds_deriv_eq_imp_eq) have "fds_deriv (f * fds_exp (-fds_ln l f)) = fds_exp (- fds_ln l f) * fds_deriv f * (1 - f / f)" by (simp add: divide_fds_def algebra_simps) also from assms have "f / f = 1" by (simp add: divide_fds_def fds_right_inverse) finally show "fds_deriv (f * fds_exp (-fds_ln l f)) = fds_deriv (fds_const (inverse ?c))" by simp qed (insert assms, auto simp: exp_minus field_simps) also have "fds_const ?c * fds_const (inverse ?c) = 1" using assms by (subst fds_const_mult [symmetric]) (simp add: divide_simps) finally show ?thesis by (simp add: mult_ac) qed lemma fds_exp_ln [simp]: fixes f :: "'a :: dirichlet_series fds" assumes "exp l = fds_nth f (Suc 0)" shows "fds_exp (fds_ln l f) = f" using assms by (subst fds_exp_ln_strong) auto lemma fds_ln_exp [simp]: fixes f :: "'a :: dirichlet_series fds" assumes "l = fds_nth f (Suc 0)" shows "fds_ln l (fds_exp f) = f" proof (rule fds_deriv_eq_imp_eq) have "fds_deriv (fds_ln l (fds_exp f)) = fds_deriv f * (fds_exp f / fds_exp f)" by (simp add: algebra_simps divide_fds_def) also have "fds_exp f / fds_exp f = 1" by (simp add: divide_fds_def fds_right_inverse) finally show "fds_deriv (fds_ln l (fds_exp f)) = fds_deriv f" by simp qed (insert assms, auto simp: field_simps) subsection \Euler products\ lemma fds_euler_product_LIMSEQ: fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "(\n. \p\n. if prime p then \i. fds_nth f (p ^ i) / nat_power (p ^ i) s else 1) \ eval_fds f s" unfolding eval_fds_def proof (rule euler_product_LIMSEQ) show "multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_euler_product_LIMSEQ': fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "(\n. \p\n. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1) \ eval_fds f s" unfolding eval_fds_def proof (rule euler_product_LIMSEQ') show "completely_multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule completely_multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_euler_product: fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "abs_convergent_prod (\p. if prime p then \i. fds_nth f (p ^ i) / nat_power (p ^ i) s else 1)" unfolding eval_fds_def proof (rule abs_convergent_euler_product) show "multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_euler_product': fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "abs_convergent_prod (\p. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1)" unfolding eval_fds_def proof (rule abs_convergent_euler_product') show "completely_multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule completely_multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_zero_iff: fixes f :: "'a :: {nat_power_field, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" assumes "fds_abs_converges f s" shows "eval_fds f s = 0 \ (\p. prime p \ fds_nth f p = nat_power p s)" proof - let ?g = "\p. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1" have lim: "(\n. \p\n. ?g p) \ eval_fds f s" by (intro fds_euler_product_LIMSEQ' assms) have conv: "convergent_prod ?g" by (intro abs_convergent_prod_imp_convergent_prod fds_abs_convergent_euler_product' assms) { assume "eval_fds f s = 0" from convergent_prod_to_zero_iff[OF conv] and this and lim have "\p. prime p \ fds_nth f p = nat_power p s" by (auto split: if_splits) } moreover { assume "\p. prime p \ fds_nth f p = nat_power p s" then obtain p where "prime p" "fds_nth f p = nat_power p s" by blast moreover from this have "nat_power p s \ 0" by (intro nat_power_nonzero) (auto simp: prime_gt_0_nat) ultimately have "(\n. \p\n. ?g p) \ 0" using convergent_prod_to_zero_iff[OF conv] by (auto intro!: exI[of _ p] split: if_splits) from tendsto_unique[OF _ lim this] have "eval_fds f s = 0" by simp } ultimately show ?thesis by blast qed lemma fixes s :: "'a :: {nat_power_normed_field,banach,euclidean_space}" assumes "s \ 1 > 1" shows euler_product_fds_zeta: "(\n. \p\n. if prime p then inverse (1 - 1 / nat_power p s) else 1) \ eval_fds fds_zeta s" (is ?th1) and eval_fds_zeta_nonzero: "eval_fds fds_zeta s \ 0" proof - have *: "completely_multiplicative_function (fds_nth fds_zeta)" by standard auto have lim: "(\n. \p\n. if prime p then inverse (1 - fds_nth fds_zeta p / nat_power p s) else 1) \ eval_fds fds_zeta s" (is "filterlim ?g _ _") using assms by (intro fds_euler_product_LIMSEQ' * fds_abs_summable_zeta) also have "?g = (\n. \p\n. if prime p then inverse (1 - 1 / nat_power p s) else 1)" by (intro ext prod.cong refl) (auto simp: fds_zeta_def fds_nth_fds) finally show ?th1 . { fix p :: nat assume "prime p" from this have "p > 1" by (simp add: prime_gt_Suc_0_nat) hence "norm (nat_power p s) = real p powr (s \ 1)" by (simp add: norm_nat_power) also have "\ > real p powr 0" using assms and \p > 1\ by (intro powr_less_mono) auto finally have "nat_power p s \ 1" using \p > 1\ by auto } hence **: "\p. prime p \ fds_nth fds_zeta p = nat_power p s" by (auto simp: fds_zeta_def fds_nth_fds) show "eval_fds fds_zeta s \ 0" using assms * ** by (subst fds_abs_convergent_zero_iff) simp_all qed lemma fds_primepow_subseries_euler_product_cm: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" "prime p" assumes "s \ 1 > abs_conv_abscissa f" shows "eval_fds (fds_primepow_subseries p f) s = 1 / (1 - fds_nth f p / nat_power p s)" proof - let ?f = "(\n. \pa\n. if prime pa then inverse (1 - fds_nth (fds_primepow_subseries p f) pa / nat_power pa s) else 1)" have "sequentially \ bot" by simp moreover have "?f \ eval_fds (fds_primepow_subseries p f) s" by (intro fds_euler_product_LIMSEQ' completely_multiplicative_function_only_pows assms fds_abs_converges_subseries) (insert assms, auto intro!: fds_abs_converges) moreover have "eventually (\n. ?f n = 1 / (1 - fds_nth f p / nat_power p s)) at_top" using eventually_ge_at_top[of p] proof eventually_elim case (elim n) have "(\pa\n. if prime pa then inverse (1 - fds_nth (fds_primepow_subseries p f) pa / nat_power pa s) else 1) = (\q\n. if q = p then inverse (1 - fds_nth f p / nat_power p s) else 1)" using \prime p\ by (intro prod.cong) (auto simp: fds_nth_subseries prime_prime_factors) also have "\ = 1 / (1 - fds_nth f p / nat_power p s)" using elim by (subst prod.delta) (auto simp: divide_simps) finally show ?case . qed hence "?f \ 1 / (1 - fds_nth f p / nat_power p s)" by (rule tendsto_eventually) ultimately show ?thesis by (rule tendsto_unique) qed subsection \Non-negative Dirichlet series\ lemma nonneg_Reals_sum: "(\x. x \ A \ f x \ \\<^sub>\\<^sub>0) \ sum f A \ \\<^sub>\\<^sub>0" by (induction A rule: infinite_finite_induct) auto locale nonneg_dirichlet_series = fixes f :: "'a :: dirichlet_series fds" assumes nonneg_coeffs_aux: "n > 0 \ fds_nth f n \ \\<^sub>\\<^sub>0" begin lemma nonneg_coeffs: "fds_nth f n \ \\<^sub>\\<^sub>0" using nonneg_coeffs_aux[of n] by (cases "n = 0") auto end lemma nonneg_dirichlet_series_0 [simp,intro]: "nonneg_dirichlet_series 0" by standard (auto simp: zero_fds_def) lemma nonneg_dirichlet_series_1 [simp,intro]: "nonneg_dirichlet_series 1" by standard (auto simp: one_fds_def) lemma nonneg_dirichlet_series_const [simp,intro]: "c \ \\<^sub>\\<^sub>0 \ nonneg_dirichlet_series (fds_const c)" by standard (auto simp: fds_const_def) lemma nonneg_dirichlet_series_add [intro]: assumes "nonneg_dirichlet_series f" "nonneg_dirichlet_series g" shows "nonneg_dirichlet_series (f + g)" proof - interpret f: nonneg_dirichlet_series f by fact interpret g: nonneg_dirichlet_series g by fact show ?thesis by standard (auto intro!: nonneg_Reals_add_I f.nonneg_coeffs g.nonneg_coeffs) qed lemma nonneg_dirichlet_series_mult [intro]: assumes "nonneg_dirichlet_series f" "nonneg_dirichlet_series g" shows "nonneg_dirichlet_series (f * g)" proof - interpret f: nonneg_dirichlet_series f by fact interpret g: nonneg_dirichlet_series g by fact show ?thesis by standard (auto intro!: nonneg_Reals_sum nonneg_Reals_mult_I f.nonneg_coeffs g.nonneg_coeffs simp: fds_nth_mult dirichlet_prod_def) qed lemma nonneg_dirichlet_series_power [intro]: assumes "nonneg_dirichlet_series f" shows "nonneg_dirichlet_series (f ^ n)" using assms by (induction n) auto context nonneg_dirichlet_series begin lemma nonneg_exp [intro]: "nonneg_dirichlet_series (fds_exp f)" proof fix n :: nat assume "n > 0" define c where "c = exp (fds_nth f (Suc 0))" define f' where "f' = fds (\n. if n = Suc 0 then 0 else fds_nth f n)" from nonneg_coeffs[of 1] obtain c' where "fds_nth f (Suc 0) = of_real c'" by (auto elim!: nonneg_Reals_cases) hence "c = of_real (exp c')" by (simp add: c_def exp_of_real) hence c: "c \ \\<^sub>\\<^sub>0" by simp have less: "n < 2 ^ k" if "n < k" for k proof - have "n < k" by fact also have "\ < 2 ^ k" by (rule less_exp) finally show ?thesis . qed have nonneg_power: "fds_nth (f' ^ k) n \ \\<^sub>\\<^sub>0" for k proof - have "nonneg_dirichlet_series f'" by standard (insert nonneg_coeffs, auto simp: f'_def) interpret nonneg_dirichlet_series "f' ^ k" by (intro nonneg_dirichlet_series_power) fact+ from nonneg_coeffs[of n] show ?thesis . qed hence "fds_nth (fds_exp f) n = c * (\k. fds_nth (f' ^ k) n /\<^sub>R fact k)" by (simp add: fds_exp_def fds_nth_fds' f'_def c_def) also have "(\k. fds_nth (f' ^ k) n /\<^sub>R fact k) = (\k\n. fds_nth (f' ^ k) n /\<^sub>R fact k)" by (intro suminf_finite) (auto intro!: fds_nth_power_eq_0 less simp: f'_def not_le) also have "c * \ \ \\<^sub>\\<^sub>0" unfolding scaleR_conv_of_real by (intro nonneg_Reals_mult_I nonneg_Reals_sum nonneg_power, unfold nonneg_Reals_of_real_iff ) (auto simp: c) finally show "fds_nth (fds_exp f) n \ \\<^sub>\\<^sub>0" . qed end lemma nonneg_dirichlet_series_lnD: assumes "nonneg_dirichlet_series (fds_ln l f)" "exp l = fds_nth f (Suc 0)" shows "nonneg_dirichlet_series f" proof - from assms have "nonneg_dirichlet_series (fds_exp (fds_ln l f))" by (intro nonneg_dirichlet_series.nonneg_exp) thus ?thesis using assms by simp qed context nonneg_dirichlet_series begin lemma fds_of_real_norm: "fds_of_real (fds_norm f) = f" proof (rule fds_eqI) fix n :: nat assume n: "n > 0" show "fds_nth (fds_of_real (fds_norm f)) n = fds_nth f n" using nonneg_coeffs[of n] by (auto elim!: nonneg_Reals_cases) qed end (* TODO: Simplify proof *) lemma pringsheim_landau_aux: fixes c :: real and f :: "complex fds" assumes "nonneg_dirichlet_series f" assumes abscissa: "c \ abs_conv_abscissa f" assumes g: "\s. s \ A \ Re s > c \ g s = eval_fds f s" assumes "g holomorphic_on A" "open A" "c \ A" shows "\x. x < c \ fds_abs_converges f (of_real x)" proof - interpret nonneg_dirichlet_series f by fact define a where "a = 1 + c" define g' where "g' = (\s. if s \ {s. Re s > c} then eval_fds f s else g s)" \ \We can find some $\varepsilon > 0$ such that the Dirichlet series can be continued analytically in a ball of radius $1 + \varepsilon$ around $a$.\ from \open A\ \c \ A\ obtain \ where \: "\ > 0" "ball c \ \ A" by (auto simp: open_contains_ball) define \ where "\ = sqrt (1 + \^2) - 1" from \ have \: "\ > 0" by (simp add: \_def) have ball_a_subset: "ball a (1 + \) \ {s. Re s > c} \ A" proof (intro subsetI) fix s :: complex assume s: "s \ ball a (1 + \)" define x y where "x = Re s" and "y = Im s" have [simp]: "s = x + \ * y" by (simp add: complex_eq_iff x_def y_def) show "s \ {s. Re s > c} \ A" proof (cases "Re s \ c") case True hence "(c - x)\<^sup>2 + y\<^sup>2 \ (1 + c - x)\<^sup>2 + y\<^sup>2 - 1" by (simp add: power2_eq_square algebra_simps) also from s have "(1 + c - x)\<^sup>2 + y\<^sup>2 - 1 < \\<^sup>2" by (auto simp: dist_norm cmod_def a_def \_def) finally have "sqrt ((c - x)\<^sup>2 + y\<^sup>2) < \" using \ by (intro real_less_lsqrt) auto hence "s \ ball c \" by (auto simp: dist_norm cmod_def) also have "\ \ A" by fact finally show ?thesis .. qed auto qed have holo: "g' holomorphic_on ball a (1 + \)" unfolding g'_def proof (intro holomorphic_on_subset[OF _ ball_a_subset] holomorphic_on_If_Un) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ ereal c" by fact finally have*: "conv_abscissa f \ ereal c" . show "eval_fds f holomorphic_on {s. c < Re s}" by (intro holomorphic_intros) (auto intro: le_less_trans[OF *]) qed (insert assms, auto intro!: holomorphic_intros open_halfspace_Re_gt) define f' where "f' = fds_norm f" have f_f': "f = fds_of_real f'" by (simp add: f'_def fds_of_real_norm) have f'_nonneg: "fds_nth f' n \ 0" for n using nonneg_coeffs[of n] by (auto elim!: nonneg_Reals_cases simp: f'_def) have deriv: "(\n. (deriv ^^ n) g' a) = (\n. eval_fds ((fds_deriv ^^ n) f) a)" proof fix n :: nat have ev: "eventually (\s. s \ {s. Re s > c}) (nhds (complex_of_real a))" by (intro eventually_nhds_in_open open_halfspace_Re_gt) (auto simp: a_def) have "(deriv ^^ n) g' a = (deriv ^^ n) (eval_fds f) a" by (intro higher_deriv_cong_ev refl eventually_mono[OF ev]) (auto simp: g'_def) also have "\ = eval_fds ((fds_deriv ^^ n) f) a" proof (intro eval_fds_higher_deriv [symmetric]) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ ereal c" by (rule assms) also have "\ < a" by (simp add: a_def) finally show "conv_abscissa f < ereal (complex_of_real a \ 1)" by simp qed finally show "(deriv ^^ n) g' a = eval_fds ((fds_deriv ^^ n) f) a" . qed have nth_deriv_conv: "fds_abs_converges ((fds_deriv ^^ n) f) (of_real a)" for n by (intro fds_abs_converges) (auto simp: abs_conv_abscissa_higher_deriv a_def intro!: le_less_trans[OF abscissa]) have nth_deriv_eq: "(fds_deriv ^^ n) f = fds (\k. (-1) ^ n * fds_nth f k * ln (real k) ^ n)" for n proof - have "fds_nth ((fds_deriv ^^ n) f) k = (-1) ^ n * fds_nth f k * ln (real k) ^ n" for k by (induction n) (simp_all add: fds_deriv_def fds_eq_iff fds_nth_fds' scaleR_conv_of_real) thus ?thesis by (intro fds_eqI) simp_all qed have deriv': "(\n. eval_fds ((fds_deriv ^^ n) f) (complex_of_real a)) = (\n. (- 1) ^ n * complex_of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / real k powr a))" proof fix n have "eval_fds ((fds_deriv ^^ n) f) (of_real a) = (\\<^sub>ak. fds_nth ((fds_deriv ^^ n) f) k / of_nat k powr complex_of_real a)" using nth_deriv_conv by (subst eval_fds_altdef) auto hence "eval_fds ((fds_deriv ^^ n) f) (of_real a) = (\\<^sub>ak. (- 1) ^ n *\<^sub>R (fds_nth f k * ln (real k) ^ n / k powr a))" by (simp add: nth_deriv_eq fds_nth_fds' powr_Reals_eq scaleR_conv_of_real algebra_simps) also have "\ = (- 1) ^ n * (\\<^sub>ak. of_real (fds_nth f' k * ln (real k) ^ n / k powr a))" by (subst infsetsum_scaleR_right) (simp_all add: scaleR_conv_of_real f_f') also have "\ = (- 1) ^ n * of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / k powr a)" by (subst infsetsum_of_real) (rule refl) finally show "eval_fds ((fds_deriv ^^ n) f) (complex_of_real a) = (- 1) ^ n * complex_of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / real k powr a)" . qed define s :: complex where "s = c - \ / 2" have s: "Re s < c" using assms \ by (simp_all add: s_def \_def field_simps) have "s \ ball a (1 + \)" using s by (simp add: a_def dist_norm cmod_def s_def) from holomorphic_power_series[OF holo this] have sums: "(\n. (deriv ^^ n) g' a / fact n * (s - a) ^ n) sums g' s" by simp also note deriv also have "s - a = -of_real (1 + \ / 2)" by (simp add: s_def a_def) also have "(\n. \ ^ n) = (\n. of_real ((-1) ^ n * (1 + \ / 2) ^ n))" by (intro ext) (subst power_minus, auto) also have "(\n. eval_fds ((fds_deriv ^^ n) f) a / fact n * \ n) = (\n. of_real ((-1) ^ n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2) ^ n))" using nth_deriv_conv by (simp add: f_f' fds_abs_converges_imp_converges mult_ac) finally have "summable \" by (simp add: sums_iff) hence summable: "summable (\n. (-1)^n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2)^n)" by (subst (asm) summable_of_real_iff) have "(\(n,k). (-1)^n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s-a) ^ n / fact n)) abs_summable_on (UNIV \ UNIV)" proof (subst abs_summable_on_Sigma_iff, safe, goal_cases) case (3 n) from nth_deriv_conv[of n] show ?case unfolding fds_abs_converges_altdef' by (intro abs_summable_on_cmult_left) (simp add: nth_deriv_eq fds_nth_fds' powr_Reals_eq) next case 4 have nth_deriv_f_f': "(fds_deriv ^^ n) f = fds_of_real ((fds_deriv ^^ n) f')" for n by (induction n) (auto simp: f'_def fds_of_real_norm) have norm_nth_deriv_f: "norm (fds_nth ((fds_deriv ^^ n) f) k) = (-1) ^ n * of_real (fds_nth ((fds_deriv ^^ n) f') k)" for n k proof (induction n) case (Suc n) thus ?case by (cases k) (auto simp: f_f' fds_nth_deriv scaleR_conv_of_real norm_mult) qed (auto simp: f'_nonneg f_f') note summable also have "(\n. (-1)^n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2)^n) = (\n. \\<^sub>ak. norm ((- 1) ^ n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s - a) ^ n / fact n)))" (is "_ = ?h") proof (rule ext, goal_cases) case (1 n) have "(\\<^sub>ak. norm ((- 1) ^ n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s - a) ^ n / fact n))) = (norm ((s - a) ^ n / fact n) * (-1) ^ n) *\<^sub>R (\\<^sub>ak. (-1) ^ n * norm (fds_nth ((fds_deriv ^^ n) f) k / real k powr a))" (is "_ = _ *\<^sub>R ?S") by (subst infsetsum_scaleR_right [symmetric]) (auto simp: norm_mult norm_divide norm_power mult_ac nth_deriv_eq fds_nth_fds') also have "?S = (\\<^sub>ak. fds_nth ((fds_deriv ^^ n) f') k / real k powr a)" by (intro infsetsum_cong) (auto simp: norm_mult norm_divide norm_power norm_nth_deriv_f) also have "\ = eval_fds ((fds_deriv ^^ n) f') a" using nth_deriv_conv[of n] by (subst eval_fds_altdef) (auto simp: f'_def nth_deriv_f_f') also have "(norm ((s - a) ^ n / fact n) * (- 1) ^ n) *\<^sub>R eval_fds ((fds_deriv ^^ n) f') a = (-1) ^ n * eval_fds ((fds_deriv ^^ n) f') a / fact n * norm (s - a) ^ n" by (simp add: norm_divide norm_power) also have s_a: "s - a = -of_real (1 + \ / 2)" by (simp add: s_def a_def) have "norm (s - a) = 1 + \ / 2" unfolding s_a norm_minus_cancel norm_of_real using \ by simp finally show ?case .. qed also have "?h n \ 0" for n by (intro infsetsum_nonneg) auto hence "?h = (\n. norm (?h n))" by simp finally show ?case unfolding abs_summable_on_nat_iff' . qed auto hence "(\(k,n). (-1)^n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s-a) ^ n / fact n)) abs_summable_on (UNIV \ UNIV)" by (subst (asm) abs_summable_on_Times_swap) (simp add: case_prod_unfold) hence "(\k. \\<^sub>an. (- 1) ^ n * fds_nth f k * ln (real k) ^ n / (k powr a) * ((s - a) ^ n / fact n)) abs_summable_on UNIV" (is "?h abs_summable_on _") by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\k. fds_nth f k / nat_power k s) abs_summable_on UNIV" proof (intro abs_summable_on_cong refl, goal_cases) case (1 k) have "?h k = (fds_nth f' k / k powr a) *\<^sub>R (\\<^sub>an. (-ln (real k) * (s - a)) ^ n / fact n)" by (subst infsetsum_scaleR_right [symmetric], rule infsetsum_cong) (simp_all add: scaleR_conv_of_real f_f' power_minus' power_mult_distrib divide_simps) also have "(\\<^sub>an. (-ln (real k) * (s - a)) ^ n / fact n) = exp (-ln (real k) * (s - a))" using exp_converges[of "-ln k * (s - a)"] exp_converges[of "norm (-ln k * (s - a))"] by (subst infsetsum_nat') (auto simp: abs_summable_on_nat_iff' sums_iff scaleR_conv_of_real divide_simps norm_divide norm_mult norm_power) also have "(fds_nth f' k / k powr a) *\<^sub>R \ = fds_nth f k / nat_power k s" by (auto simp: scaleR_conv_of_real f_f' powr_def exp_minus field_simps exp_of_real [symmetric] exp_diff) finally show ?case . qed finally have "fds_abs_converges f s" by (simp add: fds_abs_converges_def abs_summable_on_nat_iff') thus ?thesis by (intro exI[of _ "(c - \ / 2)"]) (auto simp: s_def a_def \) qed theorem pringsheim_landau: fixes c :: real and f :: "complex fds" assumes "nonneg_dirichlet_series f" assumes abscissa: "abs_conv_abscissa f = c" assumes g: "\s. s \ A \ Re s > c \ g s = eval_fds f s" assumes "g holomorphic_on A" "open A" "c \ A" shows False proof - have "\x complex_of_real x \ 1" unfolding abs_conv_abscissa_def by (intro Inf_lower) (auto simp: image_iff intro!: exI[of _ "of_real x"]) also have "\ < abs_conv_abscissa f" using assms x by simp finally show False by simp qed corollary entire_continuation_imp_abs_conv_abscissa_MInfty: assumes "nonneg_dirichlet_series f" assumes c: "c \ abs_conv_abscissa f" assumes g: "\s. Re s > c \ g s = eval_fds f s" assumes holo: "g holomorphic_on UNIV" shows "abs_conv_abscissa f = -\" proof (rule ccontr) assume "abs_conv_abscissa f \ -\" with c obtain a where abscissa [simp]: "abs_conv_abscissa f = ereal a" by (cases "abs_conv_abscissa f") auto show False proof (rule pringsheim_landau[OF assms(1) abscissa _ holo]) fix s assume s: "Re s > a" show "g s = eval_fds f s" proof (rule sym, rule analytic_continuation_open[of _ _ _ g]) show "g holomorphic_on {s. Re s > a}" by (rule holomorphic_on_subset[OF holo]) auto from assms show "{s. Re s > c} \ {s. Re s > a}" by auto next have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ = ereal a" by simp finally show "eval_fds f holomorphic_on {s. Re s > a}" by (intro holomorphic_intros) (auto intro: le_less_trans) qed (insert assms s, auto intro!: exI[of _ "of_real (c + 1)"] open_halfspace_Re_gt convex_connected convex_halfspace_Re_gt) qed auto qed subsection \Convergence of the $\zeta$ and M\"{o}bius $\mu$ series\ lemma fds_abs_summable_zeta_real_iff [simp]: "fds_abs_converges fds_zeta s \ s > (1 :: real)" proof - have "fds_abs_converges fds_zeta s \ summable (\n. real n powr -s)" unfolding fds_abs_converges_def by (intro summable_cong always_eventually) (auto simp: fds_nth_zeta powr_minus divide_simps) also have "\ \ s > 1" by (simp add: summable_real_powr_iff) finally show ?thesis . qed lemma fds_abs_summable_zeta_real: "s > (1 :: real) \ fds_abs_converges fds_zeta s" by simp lemma fds_abs_converges_moebius_mu_real: assumes "s > (1 :: real)" shows "fds_abs_converges (fds moebius_mu) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (norm (fds_nth (fds moebius_mu) n / nat_power n s)) \ n powr (-s)" by (simp add: powr_minus divide_simps abs_moebius_mu_le) next from assms show "summable (\n. real n powr -s)" by (simp add: summable_real_powr_iff) qed subsection \Application to the M\"{o}bius $\mu$ function\ lemma inverse_squares_sums': "(\n. 1 / real n ^ 2) sums (pi ^ 2 / 6)" using inverse_squares_sums sums_Suc_iff[of "\n. 1 / real n ^ 2" "pi^2 / 6"] by simp lemma norm_summable_moebius_over_square: "summable (\n. norm (moebius_mu n / real n ^ 2))" proof (subst summable_Suc_iff [symmetric], rule summable_comparison_test) show "summable (\n. 1 / real (Suc n) ^ 2)" using inverse_squares_sums by (simp add: sums_iff) qed (auto simp del: of_nat_Suc simp: field_simps abs_moebius_mu_le) lemma summable_moebius_over_square: "summable (\n. moebius_mu n / real n ^ 2)" using norm_summable_moebius_over_square by (rule summable_norm_cancel) lemma moebius_over_square_sums: "(\n. moebius_mu n / n\<^sup>2) sums (6 / pi\<^sup>2)" proof - have "1 = eval_fds (1 :: real fds) 2" by simp also have "(1 :: real fds) = fds_zeta * fds moebius_mu" by (rule fds_zeta_times_moebius_mu [symmetric]) also have "eval_fds \ 2 = eval_fds fds_zeta 2 * eval_fds (fds moebius_mu) 2" by (intro eval_fds_mult fds_abs_converges_moebius_mu_real) simp_all also have "\ = pi ^ 2 / 6 * (\n. moebius_mu n / (real n)\<^sup>2)" using inverse_squares_sums' by (simp add: eval_fds_at_numeral suminf_fds_zeta_aux sums_iff) finally have "(\n. moebius_mu n / (real n)\<^sup>2) = 6 / pi ^ 2" by (simp add: field_simps) with summable_moebius_over_square show ?thesis by (simp add: sums_iff) qed end diff --git a/thys/Error_Function/Error_Function.thy b/thys/Error_Function/Error_Function.thy --- a/thys/Error_Function/Error_Function.thy +++ b/thys/Error_Function/Error_Function.thy @@ -1,657 +1,657 @@ (* File: Error_Function.thy Author: Manuel Eberl, TU München The complex and real error function (most results are on the real error function though) *) section \The complex and real error function\ theory Error_Function imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Library.Landau_Symbols" begin subsection \Auxiliary Facts\ lemma tendsto_sandwich_mono: assumes "(\n. f (real n)) \ (c::real)" assumes "eventually (\x. \y z. x \ y \ y \ z \ f y \ f z) at_top" shows "(f \ c) at_top" proof (rule tendsto_sandwich) from assms(2) obtain C where C: "\x y. C \ x \ x \ y \ f x \ f y" by (auto simp: eventually_at_top_linorder) show "eventually (\x. f (real (nat \x\)) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) show "eventually (\x. f (real (Suc (nat \x\))) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) qed (auto intro!: filterlim_compose[OF assms(1)] filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_Suc] filterlim_floor_sequentially simp del: of_nat_Suc) lemma tendsto_sandwich_antimono: assumes "(\n. f (real n)) \ (c::real)" assumes "eventually (\x. \y z. x \ y \ y \ z \ f y \ f z) at_top" shows "(f \ c) at_top" proof (rule tendsto_sandwich) from assms(2) obtain C where C: "\x y. C \ x \ x \ y \ f x \ f y" by (auto simp: eventually_at_top_linorder) show "eventually (\x. f (real (nat \x\)) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) show "eventually (\x. f (real (Suc (nat \x\))) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) qed (auto intro!: filterlim_compose[OF assms(1)] filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_Suc] filterlim_floor_sequentially simp del: of_nat_Suc) lemma has_bochner_integral_completion [intro]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "has_bochner_integral M f I \ has_bochner_integral (completion M) f I" by (auto simp: has_bochner_integral_iff integrable_completion integral_completion borel_measurable_integrable) lemma has_bochner_integral_imp_has_integral: "has_bochner_integral lebesgue (\x. indicator S x *\<^sub>R f x) I \ (f has_integral (I :: 'b :: euclidean_space)) S" using has_integral_set_lebesgue[of S f] by (simp add: has_bochner_integral_iff set_integrable_def set_lebesgue_integral_def) lemma has_bochner_integral_imp_has_integral': "has_bochner_integral lborel (\x. indicator S x *\<^sub>R f x) I \ (f has_integral (I :: 'b :: euclidean_space)) S" by (intro has_bochner_integral_imp_has_integral has_bochner_integral_completion) lemma has_bochner_integral_erf_aux: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2)) (sqrt pi / 2)" proof - let ?pI = "\f. (\\<^sup>+s. f (s::real) * indicator {0..} s \lborel)" let ?gauss = "\x. exp (- x\<^sup>2)" let ?I = "indicator {0<..} :: real \ real" let ?ff= "\x s. x * exp (- x\<^sup>2 * (1 + s\<^sup>2)) :: real" have *: "?pI ?gauss = (\\<^sup>+x. ?gauss x * ?I x \lborel)" by (intro nn_integral_cong_AE AE_I[where N="{0}"]) (auto split: split_indicator) have "?pI ?gauss * ?pI ?gauss = (\\<^sup>+x. \\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel \lborel)" by (auto simp: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] * ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x. \\<^sup>+s. ?ff x s * ?I s * ?I x \lborel \lborel)" proof (rule nn_integral_cong, cases) fix x :: real assume "x \ 0" then show "(\\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel) = (\\<^sup>+s. ?ff x s * ?I s * ?I x \lborel)" by (subst nn_integral_real_affine[where t="0" and c="x"]) (auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) qed simp also have "... = \\<^sup>+s. \\<^sup>+x. ?ff x s * ?I s * ?I x \lborel \lborel" by (rule lborel_pair.Fubini'[symmetric]) auto also have "... = ?pI (\s. ?pI (\x. ?ff x s))" by (rule nn_integral_cong_AE) (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] split: split_indicator_asm) also have "\ = ?pI (\s. ennreal (1 / (2 * (1 + s\<^sup>2))))" proof (intro nn_integral_cong ennreal_mult_right_cong) fix s :: real show "?pI (\x. ?ff x s) = ennreal (1 / (2 * (1 + s\<^sup>2)))" proof (subst nn_integral_FTC_atLeast) have "((\a. - (exp (- ((1 + s\<^sup>2) * a\<^sup>2)) / (2 + 2 * s\<^sup>2))) \ (- (0 / (2 + 2 * s\<^sup>2)))) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top] filterlim_tendsto_pos_mult_at_top) (auto intro!: filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff) then show "((\a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) \ 0) at_top" by (simp add: field_simps) qed (auto intro!: derivative_eq_intros simp: field_simps add_nonneg_eq_0_iff) qed also have "... = ennreal (pi / 4)" proof (subst nn_integral_FTC_atLeast) show "((\a. arctan a / 2) \ (pi / 2) / 2 ) at_top" by (intro tendsto_intros) (simp_all add: tendsto_arctan_at_top) qed (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps power2_eq_square) finally have "?pI ?gauss^2 = pi / 4" by (simp add: power2_eq_square) then have "?pI ?gauss = sqrt (pi / 4)" using power_eq_iff_eq_base[of 2 "enn2real (?pI ?gauss)" "sqrt (pi / 4)"] by (cases "?pI ?gauss") (auto simp: power2_eq_square ennreal_mult[symmetric] ennreal_top_mult) also have "?pI ?gauss = (\\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "sqrt (pi / 4) = sqrt pi / 2" by (simp add: real_sqrt_divide) finally show ?thesis by (rule has_bochner_integral_nn_integral[rotated 3]) auto qed lemma has_integral_erf_aux: "((\t::real. exp (-t\<^sup>2)) has_integral (sqrt pi / 2)) {0..}" by (intro has_bochner_integral_imp_has_integral' has_bochner_integral_erf_aux) lemma contour_integrable_on_linepath_neg_exp_squared [simp, intro]: "(\t. exp (-(t^2))) contour_integrable_on linepath 0 z" by (auto intro!: contour_integrable_continuous_linepath continuous_intros) lemma holomorphic_on_chain: "g holomorphic_on t \ f holomorphic_on s \ f ` s \ t \ (\x. g (f x)) holomorphic_on s" using holomorphic_on_compose_gen[of f s g t] by (simp add: o_def) lemma holomorphic_on_chain_UNIV: "g holomorphic_on UNIV \ f holomorphic_on s\ (\x. g (f x)) holomorphic_on s" using holomorphic_on_compose_gen[of f s g UNIV] by (simp add: o_def) lemmas holomorphic_on_exp' [holomorphic_intros] = holomorphic_on_exp [THEN holomorphic_on_chain_UNIV] lemma leibniz_rule_field_derivative_real: fixes f::"'a::{real_normed_field, banach} \ real \ 'a" assumes fx: "\x t. x \ U \ t \ {a..b} \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on {a..b}" assumes cont_fx: "continuous_on (U \ {a..b}) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral {a..b} (f x)) has_field_derivative integral {a..b} (fx x0)) (at x0 within U)" using leibniz_rule_field_derivative[of U a b f fx x0] assms by simp lemma has_vector_derivative_linepath_within [derivative_intros]: assumes [derivative_intros]: "(f has_vector_derivative f') (at x within S)" "(g has_vector_derivative g') (at x within S)" "(h has_real_derivative h') (at x within S)" shows "((\x. linepath (f x) (g x) (h x)) has_vector_derivative (1 - h x) *\<^sub>R f' + h x *\<^sub>R g' - h' *\<^sub>R (f x - g x)) (at x within S)" unfolding linepath_def [abs_def] by (auto intro!: derivative_eq_intros simp: field_simps scaleR_diff_right) lemma has_field_derivative_linepath_within [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative f') (at x within S)" "(g has_field_derivative g') (at x within S)" "(h has_real_derivative h') (at x within S)" shows "((\x. linepath (f x) (g x) (h x)) has_field_derivative (1 - h x) *\<^sub>R f' + h x *\<^sub>R g' - h' *\<^sub>R (f x - g x)) (at x within S)" unfolding linepath_def [abs_def] by (auto intro!: derivative_eq_intros simp: field_simps scaleR_diff_right) lemma continuous_on_linepath' [continuous_intros]: assumes [continuous_intros]: "continuous_on A f" "continuous_on A g" "continuous_on A h" shows "continuous_on A (\x. linepath (f x) (g x) (h x))" using assms unfolding linepath_def by (auto intro!: continuous_intros) lemma contour_integral_has_field_derivative: assumes A: "open A" "convex A" "a \ A" "z \ A" assumes integrable: "\z. z \ A \ f contour_integrable_on linepath a z" assumes holo: "f holomorphic_on A" shows "((\z. contour_integral (linepath a z) f) has_field_derivative f z) (at z within B)" proof - have "(f has_field_derivative deriv f z) (at z)" if "z \ A" for z using that assms by (auto intro!: holomorphic_derivI) note [derivative_intros] = DERIV_chain2[OF this] note [continuous_intros] = continuous_on_compose2[OF holomorphic_on_imp_continuous_on [OF holo]] continuous_on_compose2[OF holomorphic_on_imp_continuous_on [OF holomorphic_deriv[OF holo]]] have [derivative_intros]: "((\x. linepath a x t) has_field_derivative of_real t) (at x within A)" for t x by (auto simp: linepath_def scaleR_conv_of_real intro!: derivative_eq_intros) have *: "linepath a b t \ A" if "a \ A" "b \ A" "t \ {0..1}" for a b t using that linepath_in_convex_hull[of a A b t] \convex A\ by (simp add: hull_same) have "((\z. integral {0..1} (\x. f (linepath a z x)) * (z - a)) has_field_derivative integral {0..1} (\t. deriv f (linepath a z t) * of_real t) * (z - a) + integral {0..1} (\x. f (linepath a z x))) (at z within A)" (is "(_ has_field_derivative ?I) _") by (rule derivative_eq_intros leibniz_rule_field_derivative_real)+ (insert assms, auto intro!: derivative_eq_intros leibniz_rule_field_derivative_real integrable_continuous_real continuous_intros simp: split_beta scaleR_conv_of_real *) also have "(\z. integral {0..1} (\x. f (linepath a z x)) * (z - a)) = (\z. contour_integral (linepath a z) f)" by (simp add: contour_integral_integral) also have "?I = integral {0..1} (\x. deriv f (linepath a z x) * of_real x * (z - a) + f (linepath a z x))" (is "_ = integral _ ?g") by (subst integral_mult_left [symmetric], subst integral_add [symmetric]) (insert assms, auto intro!: integrable_continuous_real continuous_intros simp: *) also have "(?g has_integral of_real 1 * f (linepath a z 1) - of_real 0 * f (linepath a z 0)) {0..1}" using * A by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: linepath_def scaleR_conv_of_real) hence "integral {0..1} ?g = f (linepath a z 1)" by (simp add: has_integral_iff) also have "linepath a z 1 = z" by (simp add: linepath_def) also from \z \ A\ and \open A\ have "at z within A = at z" by (rule at_within_open) finally show ?thesis by (rule DERIV_subset) simp_all qed subsection \Definition of the error function\ definition erf_coeffs :: "nat \ real" where "erf_coeffs n = (if odd n then 2 / sqrt pi * (-1) ^ (n div 2) / (of_nat n * fact (n div 2)) else 0)" lemma summable_erf: fixes z :: "'a :: {real_normed_div_algebra, banach}" shows "summable (\n. of_real (erf_coeffs n) * z ^ n)" proof - define b where "b = (\n. 2 / sqrt pi * (if odd n then inverse (fact (n div 2)) else 0))" show ?thesis proof (rule summable_comparison_test[OF exI[of _ 1]], clarify) fix n :: nat assume n: "n \ 1" hence "norm (of_real (erf_coeffs n) * z ^ n) \ b n * norm z ^ n" unfolding norm_mult norm_power erf_coeffs_def b_def by (intro mult_right_mono) (auto simp: field_simps norm_divide abs_mult) thus "norm (of_real (erf_coeffs n) * z ^ n) \ b n * norm z ^ n" by (simp add: mult_ac) next have "summable (\n. (norm z * 2 / sqrt pi) * (inverse (fact n) * norm z ^ (2*n)))" (is "summable ?c") unfolding power_mult by (intro summable_mult summable_exp) also have "?c = (\n. b (2*n+1) * norm z ^ (2*n+1))" unfolding b_def by (auto simp: fun_eq_iff b_def) also have "summable \ \ summable (\n. b n * norm z ^ n)" using summable_mono_reindex [of "\n. 2*n+1"] by (intro summable_mono_reindex [of "\n. 2*n+1"]) (auto elim!: oddE simp: strict_mono_def b_def) finally show \ . qed qed definition erf :: "('a :: {real_normed_field, banach}) \ 'a" where "erf z = (\n. of_real (erf_coeffs n) * z ^ n)" lemma erf_converges: "(\n. of_real (erf_coeffs n) * z ^ n) sums erf z" using summable_erf by (simp add: sums_iff erf_def) lemma erf_0 [simp]: "erf 0 = 0" unfolding erf_def powser_zero by (simp add: erf_coeffs_def) lemma erf_minus [simp]: "erf (-z) = - erf z" unfolding erf_def by (subst suminf_minus [OF summable_erf, symmetric], rule suminf_cong) (simp_all add: erf_coeffs_def) lemma erf_of_real [simp]: "erf (of_real x) = of_real (erf x)" unfolding erf_def using summable_erf[of x] by (subst suminf_of_real) (simp_all add: summable_erf) lemma of_real_erf_numeral [simp]: "of_real (erf (numeral n)) = erf (numeral n)" by (simp only: erf_of_real [symmetric] of_real_numeral) lemma of_real_erf_1 [simp]: "of_real (erf 1) = erf 1" by (simp only: erf_of_real [symmetric] of_real_1) lemma erf_has_field_derivative: "(erf has_field_derivative of_real (2 / sqrt pi) * exp (-(z^2))) (at z within A)" proof - define a' where "a' = (\n. 2 / sqrt pi * (if even n then (- 1) ^ (n div 2) / fact (n div 2) else 0))" have "(erf has_field_derivative (\n. diffs (\n. of_real (erf_coeffs n)) n * z ^ n)) (at z)" using summable_erf unfolding erf_def by (rule termdiffs_strong_converges_everywhere) also have "diffs (\n. of_real (erf_coeffs n)) = (\n. of_real (a' n) :: 'a)" by (simp add: erf_coeffs_def a'_def diffs_def fun_eq_iff del: of_nat_Suc) hence "(\n. diffs (\n. of_real (erf_coeffs n)) n * z ^ n) = (\n. of_real (a' n) * z ^ n)" by simp also have "\ = (\n. of_real (a' (2*n)) * z ^ (2*n))" by (intro suminf_mono_reindex [symmetric]) (auto simp: strict_mono_def a'_def elim!: evenE) also have "(\n. of_real (a' (2*n)) * z ^ (2*n)) = (\n. of_real (2 / sqrt pi) * (inverse (fact n) * (-(z^2))^n))" by (simp add: fun_eq_iff power_mult [symmetric] a'_def field_simps power_minus') also have "suminf \ = of_real (2 / sqrt pi) * exp (-(z^2))" by (subst suminf_mult, intro summable_exp) (auto simp: field_simps scaleR_conv_of_real exp_def) finally show ?thesis by (rule DERIV_subset) simp_all qed lemmas erf_has_field_derivative' [derivative_intros] = erf_has_field_derivative [THEN DERIV_chain2] lemma erf_continuous_on: "continuous_on A erf" by (rule DERIV_continuous_on erf_has_field_derivative)+ lemma continuous_on_compose2_UNIV: "continuous_on UNIV g \ continuous_on s f \ continuous_on s (\x. g (f x))" by (rule continuous_on_compose2[of UNIV g s f]) simp_all lemmas erf_continuous_on' [continuous_intros] = erf_continuous_on [THEN continuous_on_compose2_UNIV] lemma erf_continuous [continuous_intros]: "continuous (at x within A) erf" by (rule continuous_within_subset[OF _ subset_UNIV]) (insert erf_continuous_on[of UNIV], auto simp: continuous_on_eq_continuous_at) lemmas erf_continuous' [continuous_intros] = continuous_within_compose2[OF _ erf_continuous] lemmas tendsto_erf [tendsto_intros] = isCont_tendsto_compose[OF erf_continuous] lemma erf_cnj [simp]: "erf (cnj z) = cnj (erf z)" proof - interpret bounded_linear cnj by (rule bounded_linear_cnj) from suminf[OF summable_erf] show ?thesis by (simp add: erf_def erf_coeffs_def) qed lemma integral_exp_minus_squared_real: assumes "a \ b" shows "((\t. exp (-(t^2))) has_integral (sqrt pi / 2 * (erf b - erf a))) {a..b}" proof - have "((\t. exp (-(t^2))) has_integral (sqrt pi / 2 * erf b - sqrt pi / 2 * erf a)) {a..b}" using assms by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros - simp: has_field_derivative_iff_has_vector_derivative [symmetric]) + simp: has_real_derivative_iff_has_vector_derivative [symmetric]) thus ?thesis by (simp add: field_simps) qed lemma erf_real_altdef_nonneg: "x \ 0 \ erf (x::real) = 2 / sqrt pi * integral {0..x} (\t. exp (-(t^2)))" using integral_exp_minus_squared_real[of 0 x] by (simp add: has_integral_iff field_simps) lemma erf_real_altdef_nonpos: "x \ 0 \ erf (x::real) = -2 / sqrt pi * integral {0..-x} (\t. exp (-(t^2)))" using erf_real_altdef_nonneg[of "-x"] by simp lemma less_imp_erf_real_less: assumes "a < (b::real)" shows "erf a < erf b" proof - from assms have "\z. z > a \ z < b \ erf b - erf a = (b - a) * (2 / sqrt pi * exp (- z\<^sup>2))" by (intro MVT2) (auto intro!: derivative_eq_intros simp: field_simps) then obtain z where z: "a < z" "z < b" and erf: "erf b - erf a = (b - a) * (2 / sqrt pi * exp (- z\<^sup>2))" by blast note erf also from assms have "(b - a) * (2 / sqrt pi * exp (- z\<^sup>2)) > 0" by (intro mult_pos_pos divide_pos_pos) simp_all finally show ?thesis by simp qed lemma le_imp_erf_real_le: "a \ (b::real) \ erf a \ erf b" by (cases "a < b") (auto dest: less_imp_erf_real_less) lemma erf_real_less_cancel [simp]: "(erf (a :: real) < erf b) \ a < b" using less_imp_erf_real_less[of a b] less_imp_erf_real_less[of b a] by (cases a b rule: linorder_cases) simp_all lemma erf_real_eq_iff [simp]: "erf (a::real) = erf b \ a = b" by (cases a b rule: linorder_cases) (auto dest: less_imp_erf_real_less) lemma erf_real_le_cancel [simp]: "(erf (a :: real) \ erf b) \ a \ b" by (cases a b rule: linorder_cases) (auto simp: less_eq_real_def) lemma inj_on_erf_real [intro]: "inj_on (erf :: real \ real) A" by (auto simp: inj_on_def) lemma strict_mono_erf_real [intro]: "strict_mono (erf :: real \ real)" by (auto simp: strict_mono_def) lemma mono_erf_real [intro]: "mono (erf :: real \ real)" by (auto simp: mono_def) lemma erf_real_ge_0_iff [simp]: "erf (x::real) \ 0 \ x \ 0" using erf_real_le_cancel[of 0 x] unfolding erf_0 . lemma erf_real_le_0_iff [simp]: "erf (x::real) \ 0 \ x \ 0" using erf_real_le_cancel[of x 0] unfolding erf_0 . lemma erf_real_gt_0_iff [simp]: "erf (x::real) > 0 \ x > 0" using erf_real_less_cancel[of 0 x] unfolding erf_0 . lemma erf_real_less_0_iff [simp]: "erf (x::real) < 0 \ x < 0" using erf_real_less_cancel[of x 0] unfolding erf_0 . lemma erf_at_top [tendsto_intros]: "((erf :: real \ real) \ 1) at_top" proof - have *: "(\n. {0..real n}) = {0..}" by (auto intro!: real_nat_ceiling_ge) let ?f = "\t::real. exp (-t\<^sup>2)" have "(\n. set_lebesgue_integral lborel {0..real n} ?f) \ set_lebesgue_integral lborel (\n. {0..real n}) ?f" using has_bochner_integral_erf_aux by (intro set_integral_cont_up ) (insert *, auto simp: incseq_def has_bochner_integral_iff set_integrable_def) also note * also have "(\n. set_lebesgue_integral lborel {0..real n} ?f) = (\n. integral {0..real n} ?f)" proof - have "\n. set_integrable lborel {0..real n} (\x. exp (- x\<^sup>2))" unfolding set_integrable_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?thesis by (intro set_borel_integral_eq_integral ext) qed also have "\ = (\n. sqrt pi / 2 * erf (real n))" by (simp add: erf_real_altdef_nonneg) also have "set_lebesgue_integral lborel {0..} ?f = sqrt pi / 2" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_lebesgue_integral_def) finally have "(\n. 2 / sqrt pi * (sqrt pi / 2 * erf (real n))) \ (2 / sqrt pi) * (sqrt pi / 2)" by (intro tendsto_intros) hence "(\n. erf (real n)) \ 1" by simp thus ?thesis by (rule tendsto_sandwich_mono) auto qed lemma erf_at_bot [tendsto_intros]: "((erf :: real \ real) \ -1) at_bot" by (simp add: filterlim_at_bot_mirror tendsto_minus_cancel_left erf_at_top) lemmas tendsto_erf_at_top [tendsto_intros] = filterlim_compose[OF erf_at_top] lemmas tendsto_erf_at_bot [tendsto_intros] = filterlim_compose[OF erf_at_bot] subsection \The complimentary error function\ definition erfc where "erfc z = 1 - erf z" lemma erf_conv_erfc: "erf z = 1 - erfc z" by (simp add: erfc_def) lemma erfc_0 [simp]: "erfc 0 = 1" by (simp add: erfc_def) lemma erfc_minus: "erfc (-z) = 2 - erfc z" by (simp add: erfc_def) lemma erfc_of_real [simp]: "erfc (of_real x) = of_real (erfc x)" by (simp add: erfc_def) lemma of_real_erfc_numeral [simp]: "of_real (erfc (numeral n)) = erfc (numeral n)" by (simp add: erfc_def) lemma of_real_erfc_1 [simp]: "of_real (erfc 1) = erfc 1" by (simp add: erfc_def) lemma less_imp_erfc_real_less: "a < (b::real) \ erfc a > erfc b" by (simp add: erfc_def) lemma le_imp_erfc_real_le: "a \ (b::real) \ erfc a \ erfc b" by (simp add: erfc_def) lemma erfc_real_less_cancel [simp]: "(erfc (a :: real) < erfc b) \ a > b" by (simp add: erfc_def) lemma erfc_real_eq_iff [simp]: "erfc (a::real) = erfc b \ a = b" by (simp add: erfc_def) lemma erfc_real_le_cancel [simp]: "(erfc (a :: real) \ erfc b) \ a \ b" by (simp add: erfc_def) lemma inj_on_erfc_real [intro]: "inj_on (erfc :: real \ real) A" by (auto simp: inj_on_def) lemma antimono_erfc_real [intro]: "antimono (erfc :: real \ real)" by (auto simp: antimono_def) lemma erfc_real_ge_0_iff [simp]: "erfc (x::real) \ 1 \ x \ 0" by (simp add: erfc_def) lemma erfc_real_le_0_iff [simp]: "erfc (x::real) \ 1 \ x \ 0" by (simp add: erfc_def) lemma erfc_real_gt_0_iff [simp]: "erfc (x::real) > 1 \ x < 0" by (simp add: erfc_def) lemma erfc_real_less_0_iff [simp]: "erfc (x::real) < 1 \ x > 0" by (simp add: erfc_def) lemma erfc_has_field_derivative: "(erfc has_field_derivative -of_real (2 / sqrt pi) * exp (-(z^2))) (at z within A)" unfolding erfc_def [abs_def] by (auto intro!: derivative_eq_intros) lemmas erfc_has_field_derivative' [derivative_intros] = erfc_has_field_derivative [THEN DERIV_chain2] lemma erfc_continuous_on: "continuous_on A erfc" by (rule DERIV_continuous_on erfc_has_field_derivative)+ lemmas erfc_continuous_on' [continuous_intros] = erfc_continuous_on [THEN continuous_on_compose2_UNIV] lemma erfc_continuous [continuous_intros]: "continuous (at x within A) erfc" by (rule continuous_within_subset[OF _ subset_UNIV]) (insert erfc_continuous_on[of UNIV], auto simp: continuous_on_eq_continuous_at) lemmas erfc_continuous' [continuous_intros] = continuous_within_compose2[OF _ erfc_continuous] lemmas tendsto_erfc [tendsto_intros] = isCont_tendsto_compose[OF erfc_continuous] lemma erfc_at_top [tendsto_intros]: "((erfc :: real \ real) \ 0) at_top" unfolding erfc_def [abs_def] by (auto intro!: tendsto_eq_intros) lemma erfc_at_bot [tendsto_intros]: "((erfc :: real \ real) \ 2) at_bot" unfolding erfc_def [abs_def] by (auto intro!: tendsto_eq_intros) lemmas tendsto_erfc_at_top [tendsto_intros] = filterlim_compose[OF erfc_at_top] lemmas tendsto_erfc_at_bot [tendsto_intros] = filterlim_compose[OF erfc_at_bot] lemma integrable_exp_minus_squared: assumes "A \ {0..}" "A \ sets lborel" shows "set_integrable lborel A (\t::real. exp (-t\<^sup>2))" (is ?thesis1) and "(\t::real. exp (-t\<^sup>2)) integrable_on A" (is ?thesis2) proof - show ?thesis1 by (rule set_integrable_subset[of _ "{0..}"]) (insert assms has_bochner_integral_erf_aux, auto simp: has_bochner_integral_iff set_integrable_def) thus ?thesis2 by (rule set_borel_integral_eq_integral) qed lemma assumes "x \ 0" shows erfc_real_altdef_nonneg: "erfc x = 2 / sqrt pi * integral {x..} (\t. exp (-t\<^sup>2))" and has_integral_erfc: "((\t. exp (-t\<^sup>2)) has_integral (sqrt pi / 2 * erfc x)) {x..}" proof - let ?f = "\t::real. exp (-t\<^sup>2)" have int: "set_integrable lborel {0..} ?f" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_integrable_def) from assms have *: "{(0::real)..} = {0..x} \ {x..}" by auto have "set_lebesgue_integral lborel ({0..x} \ {x..}) ?f = set_lebesgue_integral lborel {0..x} ?f + set_lebesgue_integral lborel {x..} ?f" by (subst set_integral_Un_AE; (rule set_integrable_subset[OF int])?) (insert assms AE_lborel_singleton[of x], auto elim!: eventually_mono) also note * [symmetric] also have "set_lebesgue_integral lborel {0..} ?f = sqrt pi / 2" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_lebesgue_integral_def) also have "set_lebesgue_integral lborel {0..x} ?f = sqrt pi / 2 * erf x" by (subst set_borel_integral_eq_integral(2)[OF set_integrable_subset[OF int]]) (insert assms, auto simp: erf_real_altdef_nonneg) also have "set_lebesgue_integral lborel {x..} ?f = integral {x..} ?f" by (subst set_borel_integral_eq_integral(2)[OF set_integrable_subset[OF int]]) (insert assms, auto) finally show "erfc x = 2 / sqrt pi * integral {x..} ?f" by (simp add: field_simps erfc_def) with integrable_exp_minus_squared(2)[of "{x..}"] assms show "(?f has_integral (sqrt pi / 2 * erfc x)) {x..}" by (simp add: has_integral_iff) qed lemma erfc_real_gt_0 [simp, intro]: "erfc (x::real) > 0" proof (cases "x \ 0") case True have "0 < integral {x..x+1} (\t. exp (-(x+1)\<^sup>2))" by simp also from True have "\ \ integral {x..x+1} (\t. exp (-t\<^sup>2))" by (intro integral_le) (auto intro!: integrable_continuous_real continuous_intros power_mono) also have "\ \ sqrt pi / 2 * erfc x" by (rule has_integral_subset_le[OF _ integrable_integral has_integral_erfc]) (auto intro!: integrable_continuous_real continuous_intros True) finally have "sqrt pi / 2 * erfc x > 0" . hence "\ * (2 / sqrt pi) > 0" by (rule mult_pos_pos) simp_all thus "erfc x > 0" by simp next case False have "0 \ (1::real)" by simp also from False have "\ < erfc x" by simp finally show ?thesis . qed lemma erfc_real_less_2 [intro]: "erfc (x::real) < 2" using erfc_real_gt_0[of "-x"] unfolding erfc_minus by simp lemma erf_real_gt_neg1 [intro]: "erf (x::real) > -1" using erfc_real_less_2[of x] unfolding erfc_def by simp lemma erf_real_less_1 [intro]: "erf (x::real) < 1" using erfc_real_gt_0[of x] unfolding erfc_def by simp lemma erfc_cnj [simp]: "erfc (cnj z) = cnj (erfc z)" by (simp add: erfc_def) subsection \Specific facts about the complex case\ lemma erf_complex_altdef: "erf z = of_real (2 / sqrt pi) * contour_integral (linepath 0 z) (\t. exp (-(t^2)))" proof - define A where "A = (\z. contour_integral (linepath 0 z) (\t. exp (-(t^2))))" have [derivative_intros]: "(A has_field_derivative exp (-(z^2))) (at z)" for z :: complex unfolding A_def by (rule contour_integral_has_field_derivative[where A = UNIV]) (auto intro!: holomorphic_intros) have "erf z - 2 / sqrt pi * A z = erf 0 - 2 / sqrt pi * A 0" by (rule has_derivative_zero_unique [where f = "\z. erf z - 2 / sqrt pi * A z" and s = UNIV]) (auto intro!: has_field_derivative_imp_has_derivative derivative_eq_intros) also have "A 0 = 0" by (simp only: A_def contour_integral_trivial) finally show ?thesis unfolding A_def by (simp add: algebra_simps) qed lemma erf_holomorphic_on: "erf holomorphic_on A" by (auto simp: holomorphic_on_def field_differentiable_def intro!: erf_has_field_derivative) lemmas erf_holomorphic_on' [holomorphic_intros] = erf_holomorphic_on [THEN holomorphic_on_chain_UNIV] lemma erf_analytic_on: "erf analytic_on A" by (auto simp: analytic_on_def) (auto intro!: exI[of _ 1] holomorphic_intros) lemma erf_analytic_on' [analytic_intros]: assumes "f analytic_on A" shows "(\x. erf (f x)) analytic_on A" proof - from assms and erf_analytic_on have "erf \ f analytic_on A" by (rule analytic_on_compose_gen) auto thus ?thesis by (simp add: o_def) qed lemma erfc_holomorphic_on: "erfc holomorphic_on A" by (auto simp: holomorphic_on_def field_differentiable_def intro!: erfc_has_field_derivative) lemmas erfc_holomorphic_on' [holomorphic_intros] = erfc_holomorphic_on [THEN holomorphic_on_chain_UNIV] lemma erfc_analytic_on: "erfc analytic_on A" by (auto simp: analytic_on_def) (auto intro!: exI[of _ 1] holomorphic_intros) lemma erfc_analytic_on' [analytic_intros]: assumes "f analytic_on A" shows "(\x. erfc (f x)) analytic_on A" proof - from assms and erfc_analytic_on have "erfc \ f analytic_on A" by (rule analytic_on_compose_gen) auto thus ?thesis by (simp add: o_def) qed end diff --git a/thys/Error_Function/Error_Function_Asymptotics.thy b/thys/Error_Function/Error_Function_Asymptotics.thy --- a/thys/Error_Function/Error_Function_Asymptotics.thy +++ b/thys/Error_Function/Error_Function_Asymptotics.thy @@ -1,275 +1,275 @@ (* File: Error_Function_Asymptotics.thy Author: Manuel Eberl, TU München The asymptotics of the real error function *) subsection \Asymptotics\ theory Error_Function_Asymptotics imports Error_Function Landau_Symbols.Landau_More begin lemma real_powr_eq_powerI: "x > 0 \ y = real y' \ x powr y = x ^ y'" by (simp add: powr_realpow) definition erf_remainder_integral where "erf_remainder_integral n x = lim (\m. integral {x..x + real m} (\t. exp (-(t^2)) / t ^ (2*n)))" text \ The following is the remainder term in the asymptotic expansion of @{term erfc}. \ definition erf_remainder where "erf_remainder n x = ((-1)^n * 2 * fact (2*n)) / (sqrt pi * 4 ^ n * fact n) * erf_remainder_integral n x" lemma erf_remainder_integral_aux_nonneg: "x > 0 \ integral {x..x + real m} (\t. exp (-(t^2)) / t ^ (2*n)) \ 0" by (intro integral_nonneg integrable_continuous_real) (auto intro!: continuous_intros) lemma erf_remainder_integral_aux_bound: assumes "x > 0" shows "norm (integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n))) \ exp (-x\<^sup>2) / x ^ (2*n+1)" and "integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n)) \ exp (-x\<^sup>2) / x ^ (2*n+1)" proof - have "norm (integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n))) \ integral {x..x + real m} (\t. exp (-x*t) / x ^ (2*n))" proof (intro integral_norm_bound_integral ballI) fix t assume t: "t \ {x..x + real m}" from t have "norm (exp (-t\<^sup>2) / t ^ (2*n)) = exp (-t\<^sup>2) / t ^ (2*n)" by simp also have "\ \ exp (-x*t) / x ^ (2*n)" using t assms by (intro frac_le) (simp_all add: self_le_power power2_eq_square power_mono) finally show "norm (exp (-t\<^sup>2) / t ^ (2*n)) \ \" by simp qed (insert assms, auto intro!: continuous_intros integrable_continuous_real) also have "\ = -exp (-x*(x + real m)) / x ^ (2*n+1) - (-exp (-x*x) / x ^ (2*n+1))" using assms by (intro integral_unique fundamental_theorem_of_calculus) - (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] + (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric] intro!: derivative_eq_intros) also have "\ \ exp (-x\<^sup>2) / x ^ (2*n+1)" using assms by (simp add: power2_eq_square) finally show *: "norm (integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n))) \ exp (-x\<^sup>2) / x ^ (2*n+1)" . have "integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n)) \ norm (integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n)))" by simp also note * finally show "integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n)) \ exp (-x\<^sup>2) / x ^ (2*n+1)" . qed lemma convergent_erf_remainder_integral: assumes "x > 0" shows "convergent (\m. integral {x..x + real m} (\t. exp (-(t^2)) / t ^ (2*n)))" proof (intro Bseq_mono_convergent BseqI'; clarify?) fix m :: nat show "norm (integral {x..x + real m} (\t. exp (-t\<^sup>2) / t ^ (2*n))) \ exp (-x\<^sup>2) / x ^ (2*n+1)" using assms by (rule erf_remainder_integral_aux_bound) qed (insert assms, auto intro!: integral_subset_le integrable_continuous_real continuous_intros) lemma LIMSEQ_erf_remainder_integral: "x > 0 \ (\m. integral {x..x + real m} (\t. exp (-(t^2)) / t ^ (2*n))) \ erf_remainder_integral n x" using convergent_erf_remainder_integral[of x] by (simp add: convergent_LIMSEQ_iff erf_remainder_integral_def) text \ We show some bounds on the remainder term. \ lemma assumes "x > 0" shows erf_remainder_integral_nonneg: "erf_remainder_integral n x \ 0" and erf_remainder_integral_bound: "erf_remainder_integral n x \ exp (-x\<^sup>2) / x ^ (2*n+1)" proof - note * = LIMSEQ_erf_remainder_integral[OF assms] show "erf_remainder_integral n x \ 0" by (intro tendsto_le[OF _ * tendsto_const] always_eventually erf_remainder_integral_aux_nonneg allI assms sequentially_bot) show "erf_remainder_integral n x \ exp (-x\<^sup>2) / x ^ (2*n+1)" by (intro tendsto_le[OF _ tendsto_const *] always_eventually erf_remainder_integral_aux_bound allI assms sequentially_bot) qed lemma erf_remainder_integral_bigo: "erf_remainder_integral n \ O(\x. exp (-x\<^sup>2) / x ^ (2*n+1))" using erf_remainder_integral_nonneg erf_remainder_integral_bound by (auto intro!: bigoI[of _ 1] eventually_mono [OF eventually_gt_at_top[of "0::real"]]) theorem erf_remainder_bigo: "erf_remainder n \ O(\x. exp (-x\<^sup>2) / x ^ (2*n+1))" using erf_remainder_integral_bigo[of n] by (simp add: erf_remainder_def [abs_def]) text \ Next, we unroll the remainder term to develop the asymptotic expansion. \ lemma erf_remainder_integral_0_conv_erfc: assumes "(x::real) > 0" shows "erf_remainder_integral 0 x = sqrt pi / 2 * erfc x" proof - have "(\m. sqrt pi / 2 * (erf (x + real m) - erf x)) \ sqrt pi / 2 * erfc x" (is "filterlim ?f _ _") unfolding erfc_def by (intro tendsto_intros filterlim_tendsto_add_at_top[OF tendsto_const filterlim_real_sequentially]) also have "?f = (\m. integral {x..x+real m} (\t. exp (-t\<^sup>2)))" by (auto simp: fun_eq_iff integral_unique[OF integral_exp_minus_squared_real]) finally have "(\m. integral {x..x + real m} (\t. exp (- t\<^sup>2))) \ sqrt pi / 2 * erfc x" . moreover have "(\m. integral {x..x + real m} (\t. exp (- t\<^sup>2))) \ erf_remainder_integral 0 x" using LIMSEQ_erf_remainder_integral[of x 0] assms by simp ultimately show "erf_remainder_integral 0 x = sqrt pi / 2 * erfc x" by (intro LIMSEQ_unique) qed text \ The first remainder is the @{term erfc} function itself. \ lemma erf_remainder_0_conv_erfc: "x > 0 \ erf_remainder 0 x = erfc x" by (simp add: erf_remainder_def erf_remainder_integral_0_conv_erfc) text \ Also, the following recurrence allows us to get the next term of the asymptotic expansion. \ lemma erf_remainder_integral_conv_Suc: assumes "x > 0" shows "erf_remainder_integral n x = exp (-x\<^sup>2) / (2*x^(2*n+1)) - real (2*n+1) / 2 * erf_remainder_integral (Suc n) x" proof - let ?A = "\m. {x..x + real m}" let ?J = "\m n. integral {x..x+real m} (\t. exp(-t\<^sup>2) / t ^ (2*n))" define I where "I = (\m. exp (- (x + real m)\<^sup>2) / (- 2 * (x + real m) ^ (2 * n + 1)) - exp (- x\<^sup>2) * inverse (- 2 * x ^ (2 * n + 1)) - real (2*n+1)/2 * ?J m (Suc n))" have I_eq: "I = (\m. integral (?A m) (\t. exp (- t\<^sup>2) / t ^ (2 * n)))" proof fix m :: nat have "((\t. (-2*t*exp (-(t^2))) * inverse (-2*t^(2*n+1))) has_integral I m) (?A m)" proof (rule integration_by_parts[OF bounded_bilinear_mult]) fix t assume t: "t \ ?A m" with assms show "((\t. exp (-(t^2))) has_vector_derivative -2*t*exp (-(t^2))) (at t)" - by (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] + by (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric] field_simps intro!: derivative_eq_intros) from assms t have "((\t. -(1/2) * t powr (-2*n-1)) has_field_derivative (2*n+1)/2 * t powr (-2*n-2)) (at t)" by (auto intro!: derivative_eq_intros simp: field_simps powr_numeral power2_eq_square powr_minus powr_divide [symmetric] powr_add) also have "?this \ ((\t. inverse (-2*t^(2*n+1))) has_field_derivative (2*n+1)/2 / t ^ (2*Suc n)) (at t)" using t using eventually_nhds_in_open[of "{0<..}" t] assms by (intro DERIV_cong_ev refl) (auto elim!: eventually_mono simp: powr_minus field_simps powr_diff powr_realpow power2_eq_square intro!: real_powr_eq_powerI) finally show "((\t. inverse (-2*t^(2*n+1))) has_vector_derivative (2*n+1)/2 / t ^ (2*Suc n)) (at t)" - by (simp add: has_field_derivative_iff_has_vector_derivative) + by (simp add: has_real_derivative_iff_has_vector_derivative) next have "((\t. real (2*n+1)/2 * (exp (- t\<^sup>2) / t ^ (2 * Suc n))) has_integral real (2*n+1)/2 * ?J m (Suc n)) (?A m)" (is "(?f has_integral ?a) _") using assms by (intro has_integral_mult_right integrable_integral integrable_continuous_real) (auto intro!: continuous_intros) also have "?f = (\t. exp (- t\<^sup>2) * (real (2 * n + 1) / 2 / t ^ (2 * Suc n)))" by (simp add: fun_eq_iff field_simps) also have "?a = exp (- (x + real m)\<^sup>2) * inverse (- 2 * (x + real m) ^ (2 * n + 1)) - exp (- x\<^sup>2) * inverse (- 2 * x ^ (2 * n + 1)) - I m" using assms by (simp add: I_def algebra_simps inverse_eq_divide) finally show "((\t. exp (- t\<^sup>2) * (real (2 * n + 1) / 2 / t ^ (2 * Suc n))) has_integral \) {x..x + real m}" . qed (insert assms, auto intro!: continuous_intros) hence "I m = integral {x..x + real m} (\t. - 2*t*exp (- t\<^sup>2) * inverse (-2*t^(2 * n + 1)))" by (simp add: has_integral_iff) also have "\ = integral {x..x + real m} (\t. exp (- t\<^sup>2) / t ^ (2*n))" using assms by (intro integral_cong) (simp_all add: field_simps) finally show "I m = \" . qed have "filterlim (\m. (-exp (- (x + real m)\<^sup>2)) / (2 * (x + real m) ^ (2 * n + 1))) (nhds 0) at_top" by (rule real_tendsto_divide_at_top filterlim_real_sequentially tendsto_minus filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top] filterlim_pow_at_top filterlim_tendsto_add_at_top tendsto_const filterlim_ident filterlim_tendsto_pos_mult_at_top | simp)+ hence *: "filterlim (\m. (exp (- (x + real m)\<^sup>2)) / (-2 * (x + real m) ^ (2 * n + 1))) (nhds 0) at_top" by (simp add: add_ac) have "I \ 0 - exp (- x\<^sup>2) * inverse (- 2 * x ^ (2 * n + 1)) - real (2 * n + 1) / 2 * erf_remainder_integral (Suc n) x" unfolding I_def by (intro tendsto_diff * tendsto_const tendsto_mult LIMSEQ_erf_remainder_integral assms) moreover from LIMSEQ_erf_remainder_integral[OF assms, of n] I_eq have "I \ erf_remainder_integral n x" by simp ultimately have "0 - exp (- x\<^sup>2) * inverse (- 2 * x ^ (2 * n + 1)) - real (2 * n + 1) / 2 * erf_remainder_integral (Suc n) x = erf_remainder_integral n x" by (rule LIMSEQ_unique) thus ?thesis by (simp add: field_simps) qed lemma erf_remainder_conv_Suc: assumes "x > 0" shows "erf_remainder n x = (- 1) ^ n * fact (2 * n) / (sqrt pi * 4 ^ n * fact n) * exp (- x\<^sup>2) / (x ^ (2 * n + 1)) + erf_remainder (Suc n) x" proof - have "erf_remainder n x = (- 1) ^ n * 2 * fact (2 * n) / (sqrt pi * 4 ^ n * fact n) * exp (- x\<^sup>2) / (2 * x ^ (2 * n + 1)) + -( (- 1) ^ n * 2 * fact (2 * n) / (sqrt pi * 4 ^ n * fact n) * real (2 * n + 1) / 2 * erf_remainder_integral (Suc n) x)" (is "_ = ?A + ?B") unfolding erf_remainder_def using assms by (subst erf_remainder_integral_conv_Suc) (auto simp: assms algebra_simps simp del: power_Suc) also have "?B = erf_remainder (Suc n) x" by (simp add: divide_simps erf_remainder_def) also have "?A = (- 1) ^ n * fact (2 * n) / (sqrt pi * 4 ^ n * fact n) * exp (- x\<^sup>2) / (x ^ (2 * n + 1))" by (simp add: divide_simps) finally show ?thesis . qed text \ Finally, this gives us the full asymptotic expansion for @{term erfc}: \ theorem erfc_unroll: assumes "x > 0" shows "erfc x = exp (-x\<^sup>2) / sqrt pi * (\i2) / sqrt pi * (\i2) / x^(2*n+1) + erf_remainder (Suc n) x) = exp (- x\<^sup>2) / sqrt pi * (\i For convenience, we define another auxiliary function that is more suitable for use in an automated expansion framework, since it has a simple asymptotic expansion in powers of $x$. \ definition erfc_aux where "erfc_aux x = exp (x\<^sup>2) * sqrt pi * erfc x" definition erf_remainder' where "erf_remainder' n x = exp (x\<^sup>2) * sqrt pi * erf_remainder n x" lemma erfc_aux_unroll: "x > 0 \ erfc_aux x = (\i O(\x. 1 / x ^ (2*n+1))" proof - have "(\x. exp (x\<^sup>2) * erf_remainder n x) \ O(\x. exp (x\<^sup>2) * (exp (-x\<^sup>2) / x ^ (2*n+1)))" by (intro landau_o.big.mult erf_remainder_bigo) simp_all thus ?thesis by (simp add: exp_minus erf_remainder'_def [abs_def]) qed lemma has_field_derivative_erfc_aux: "(erfc_aux has_field_derivative (2 * x * erfc_aux x - 2)) (at x)" by (auto simp: erfc_aux_def [abs_def] exp_minus field_simps intro!: derivative_eq_intros) 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,1617 +1,1617 @@ section \The Euler--MacLaurin summation formula\ theory Euler_MacLaurin imports "HOL-Complex_Analysis.Complex_Analysis" 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 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 obtain C where C: "\x. norm (pbernpoly n x) \ C" by auto 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) + simp: has_real_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) + (auto simp: has_real_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 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]) + simp: has_real_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] + simp: fs_def has_real_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]) + simp: fs_def has_real_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] + simp: fs_def has_real_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) + simp: fs_def has_real_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/Fourier/Fourier.thy b/thys/Fourier/Fourier.thy --- a/thys/Fourier/Fourier.thy +++ b/thys/Fourier/Fourier.thy @@ -1,2742 +1,2742 @@ section\The basics of Fourier series\ text\Ported from HOL Light; thanks to Manuel Eberl for help with the real asymp proof methods\ theory "Fourier" imports Periodic Square_Integrable "HOL-Real_Asymp.Real_Asymp" Confine Fourier_Aux2 begin subsection\Orthonormal system of L2 functions and their Fourier coefficients\ definition orthonormal_system :: "'a::euclidean_space set \ ('b \ 'a \ real) \ bool" where "orthonormal_system S w \ \m n. l2product S (w m) (w n) = (if m = n then 1 else 0)" definition orthonormal_coeff :: "'a::euclidean_space set \ (nat \ 'a \ real) \ ('a \ real) \ nat \ real" where "orthonormal_coeff S w f n = l2product S (w n) f" lemma orthonormal_system_eq: "orthonormal_system S w \ l2product S (w m) (w n) = (if m = n then 1 else 0)" by (simp add: orthonormal_system_def) lemma orthonormal_system_l2norm: "orthonormal_system S w \ l2norm S (w i) = 1" by (simp add: l2norm_def orthonormal_system_def) lemma orthonormal_partial_sum_diff: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x - (\i\I. a i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 + (\i\I. (a i)\<^sup>2) - 2 * (\i\I. a i * orthonormal_coeff S w f i)" proof - have "S \ sets lebesgue" using f square_integrable_def by blast then have "(\x. \i\I. a i * w i x) square_integrable S" by (intro square_integrable_sum square_integrable_lmult w \finite I\) with assms show ?thesis apply (simp add: l2norm_pow_2 orthonormal_coeff_def orthonormal_system_def) apply (simp add: l2product_rdiff l2product_sym l2product_rsum l2product_rmult algebra_simps power2_eq_square if_distrib [of "\x. _ * x"]) done qed lemma orthonormal_optimal_partial_sum: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "l2norm S (\x. f x - (\i\I. orthonormal_coeff S w f i * w i x)) \ l2norm S (\x. f x - (\i\I. a i * w i x))" proof - have "2 * (a i * l2product S f(w i)) \ (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2" for i using sum_squares_bound [of "a i" "l2product S f(w i)"] by (force simp: algebra_simps) then have *: "2 * (\i\I. a i * l2product S f(w i)) \ (\i\I. (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2)" by (simp add: sum_distrib_left sum_mono) have S: "S \ sets lebesgue" using assms square_integrable_def by blast with assms * show ?thesis apply (simp add: l2norm_le square_integrable_sum square_integrable_diff square_integrable_lmult flip: l2norm_pow_2) apply (simp add: orthonormal_coeff_def orthonormal_partial_sum_diff) apply (simp add: l2product_sym power2_eq_square sum.distrib) done qed lemma Bessel_inequality: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "(\i\I. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" using orthonormal_partial_sum_diff [OF assms, of "orthonormal_coeff S w f"] apply (simp add: algebra_simps flip: power2_eq_square) by (metis (lifting) add_diff_cancel_left' diff_ge_0_iff_ge zero_le_power2) lemma Fourier_series_square_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" shows "summable (confine (\i. (orthonormal_coeff S w f i) ^ 2) I)" proof - have "incseq (\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2)" by (auto simp: incseq_def intro: sum_mono2) moreover have "\i. (\i\I \ {..i}. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" by (simp add: Bessel_inequality assms) ultimately obtain L where "(\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) \ L" using incseq_convergent by blast then have "\r>0. \no. \n\no. \(\i\{..n} \ I. (orthonormal_coeff S w f i)\<^sup>2) - L\ < r" by (simp add: LIMSEQ_iff Int_commute) then show ?thesis by (auto simp: summable_def sums_confine_le LIMSEQ_iff) qed lemma orthonormal_Fourier_partial_sum_diff_squared: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x -(\i\I. orthonormal_coeff S w f i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 - (\i\I. (orthonormal_coeff S w f i)\<^sup>2)" using orthonormal_partial_sum_diff [OF assms, where a = "orthonormal_coeff S w f"] by (simp add: power2_eq_square) lemma Fourier_series_l2_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?\ = "confine (\i. (orthonormal_coeff S w f i)\<^sup>2) I" have si: "?\ A square_integrable S" if "finite A" for A by (force intro: square_integrable_lmult w square_integrable_sum S that) have "\N. \m\N. \n\N. l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" if "e > 0" for e proof - have "summable ?\" by (rule Fourier_series_square_summable [OF os w f]) then have "Cauchy (\n. sum ?\ {..n})" by (simp add: summable_def sums_def_le convergent_eq_Cauchy) then obtain M where M: "\m n. \m\M; n\M\ \ dist (sum ?\ {..m}) (sum ?\ {..n}) < e\<^sup>2" using \e > 0\ unfolding Cauchy_def by (drule_tac x="e^2" in spec) auto have "\m\M; n\M\ \ l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" for m n proof (induct m n rule: linorder_class.linorder_wlog) case (le m n) have sum_diff: "sum f(I \ {..n}) - sum f(I \ {..m}) = sum f(I \ {Suc m..n})" for f :: "nat \ real" proof - have "(I \ {..n}) = (I \ {..m} \ I \ {Suc m..n})" "(I \ {..m}) \ (I \ {Suc m..n}) = {}" using le by auto then show ?thesis by (simp add: algebra_simps flip: sum.union_disjoint) qed have "(l2norm S (\x. ?\ {..n} x - ?\ {..m} x))^2 \ \(\i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) - (\i\I \ {..m}. (orthonormal_coeff S w f i)\<^sup>2)\" proof (simp add: sum_diff) have "(l2norm S (?\ {Suc m..n}))\<^sup>2 = (\i\I \ {Suc m..n}. l2product S (\x. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * w j x) (\x. orthonormal_coeff S w f i * w i x))" by (simp add: l2norm_pow_2 l2product_rsum si w) also have "\ = (\i\I \ {Suc m..n}. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * orthonormal_coeff S w f i * l2product S (w j) (w i))" by (simp add: l2product_lsum l2product_lmult l2product_rmult si w flip: mult.assoc) also have "\ \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" by (simp add: orthonormal_system_eq [OF os] power2_eq_square if_distrib [of "(*)_"] cong: if_cong) finally show "(l2norm S (?\ {Suc m..n}))\<^sup>2 \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" . qed also have "\ < e\<^sup>2" using M [OF \n\M\ \m\M\] by (simp add: dist_real_def sum_confine_eq_Int Int_commute) finally have "(l2norm S (\x. ?\ {..m} x - ?\ {..n} x))^2 < e^2" using l2norm_diff [OF si si] by simp with \e > 0\ show ?case by (simp add: power2_less_imp_less) next case (sym a b) then show ?case by (subst l2norm_diff) (auto simp: si) qed then show ?thesis by blast qed with that show thesis by (blast intro: si [of "{.._}"] l2_complete [of "\n. ?\ {..n}"]) qed lemma Fourier_series_l2_summable_strong: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "\i. i \ I \ orthonormal_coeff S w (\x. f x - g x) i = 0" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast obtain g where g: "g square_integrable S" and teng: "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" using Fourier_series_l2_summable [OF assms] . show thesis proof show "orthonormal_coeff S w (\x. f x - g x) i = 0" if "i \ I" for i proof (rule tendsto_unique) let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?f = "\n. l2product S (w i) (\x. (f x - ?\ {..n} x) + (?\ {..n} x - g x))" show "?f \ orthonormal_coeff S w (\x. f x - g x) i" by (simp add: orthonormal_coeff_def) have 1: "(\n. l2product S (w i) (\x. f x - ?\ {..n} x)) \ 0" proof (rule tendsto_eventually) have "l2product S (w i) (\x. f x - ?\ {..j} x) = 0" if "j \ i" for j using that \i \ I\ apply (simp add: l2product_rdiff l2product_rsum l2product_rmult orthonormal_coeff_def w f S) apply (simp add: orthonormal_system_eq [OF os] if_distrib [of "(*)_"] cong: if_cong) done then show "\\<^sub>F n in sequentially. l2product S (w i) (\x. f x - ?\ {..n} x) = 0" using eventually_at_top_linorder by blast qed have 2: "(\n. l2product S (w i) (\x. ?\ {..n} x - g x)) \ 0" proof (intro Lim_null_comparison [OF _ teng] eventuallyI) show "norm (l2product S (w i) (\x. ?\ {..n} x - g x)) \ l2norm S (\x. ?\ {..n} x - g x)" for n using Schwartz_inequality_abs [of "w i" S "(\x. ?\ {..n} x - g x)"] by (simp add: S g f w orthonormal_system_l2norm [OF os]) qed show "?f \ 0" using that tendsto_add [OF 1 2] by (subst l2product_radd) (simp_all add: l2product_radd w f g S) qed auto qed (auto simp: g teng) qed subsection\Actual trigonometric orthogonality relations\ lemma integrable_sin_cx: "integrable (lebesgue_on {-pi..pi}) (\x. sin(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integrable_cos_cx: "integrable (lebesgue_on {-pi..pi}) (\x. cos(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integral_cos_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(n * x)) = (if n = 0 then 2 * pi else 0)" by (metis assms integral_cos_Z mult_commute_abs) lemma integral_sin_and_cos: fixes m n::int shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\m \ 0; n \ 0\ \ integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" by (simp_all add: abs_if sin_times_sin cos_times_sin sin_times_cos cos_times_cos integrable_sin_cx integrable_cos_cx mult_ac flip: distrib_left distrib_right left_diff_distrib right_diff_distrib) lemma integral_sin_and_cos_Z [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" using assms unfolding Ints_def apply safe unfolding integral_sin_and_cos apply auto done lemma integral_sin_and_cos_N [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" using assms unfolding Nats_altdef1 by (auto simp: integral_sin_and_cos) lemma integrable_sin_and_cos: fixes m n::int shows "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * sin(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * sin(x * n))" by (intro continuous_imp_integrable_real continuous_intros)+ lemma sqrt_pi_ge1: "sqrt pi \ 1" using pi_gt3 by auto definition trigonometric_set :: "nat \ real \ real" where "trigonometric_set n \ if n = 0 then \x. 1 / sqrt(2 * pi) else if odd n then \x. sin(real(Suc (n div 2)) * x) / sqrt(pi) else (\x. cos((n div 2) * x) / sqrt pi)" lemma trigonometric_set: "trigonometric_set 0 x = 1 / sqrt(2 * pi)" "trigonometric_set (Suc (2 * n)) x = sin(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (2 * n + 2) x = cos(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (Suc (Suc (2 * n))) x = cos(real(Suc n) * x) / sqrt(pi)" by (simp_all add: trigonometric_set_def algebra_simps add_divide_distrib) lemma trigonometric_set_even: "trigonometric_set(2*k) = (if k = 0 then (\x. 1 / sqrt(2 * pi)) else (\x. cos(k * x) / sqrt pi))" by (induction k) (auto simp: trigonometric_set_def add_divide_distrib split: if_split_asm) lemma orthonormal_system_trigonometric_set: "orthonormal_system {-pi..pi} trigonometric_set" proof - have "l2product {-pi..pi} (trigonometric_set m) (trigonometric_set n) = (if m = n then 1 else 0)" for m n proof (induction m rule: odd_even_cases) case 0 show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def measure_restrict_space) next case (odd m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def double_not_eq_Suc_double) next case (even m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def Suc_double_not_eq_double) qed then show ?thesis unfolding orthonormal_system_def by auto qed lemma square_integrable_trigonometric_set: "(trigonometric_set i) square_integrable {-pi..pi}" proof - have "continuous_on {-pi..pi} (\x. sin ((1 + real n) * x) / sqrt pi)" for n by (intro continuous_intros) auto moreover have "continuous_on {-pi..pi} (\x. cos (real i * x / 2) / sqrt pi) " by (intro continuous_intros) auto ultimately show ?thesis by (simp add: trigonometric_set_def) qed subsection\Weierstrass for trigonometric polynomials\ lemma Weierstrass_trig_1: fixes g :: "real \ real" assumes contf: "continuous_on UNIV g" and periodic: "\x. g(x + 2 * pi) = g x" and 1: "norm z = 1" shows "continuous (at z within (sphere 0 1)) (g \ Im \ Ln)" proof (cases "Re z < 0") let ?f = "complex_of_real \ g \ (\x. x + pi) \ Im \ Ln \ uminus" case True have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" proof (rule continuous_transform_within) have [simp]: "z \ \\<^sub>\\<^sub>0" using True complex_nonneg_Reals_iff by auto show "continuous (at z within (sphere 0 1)) ?f" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto show "?f x' = (complex_of_real \ g \ Im \ Ln) x'" if "x' \ (sphere 0 1)" and "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto with that show ?thesis by (auto simp: Ln_minus add.commute periodic) qed qed (use 1 in auto) then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) next case False let ?f = "complex_of_real \ g \ Im \ Ln \ uminus" have "z \ 0" using 1 by auto with False have "z \ \\<^sub>\\<^sub>0" by (auto simp: not_less nonpos_Reals_def) then have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) qed inductive_set cx_poly :: "(complex \ real) set" where Re: "Re \ cx_poly" | Im: "Im \ cx_poly" | const: "(\x. c) \ cx_poly" | add: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x + g x) \ cx_poly" | mult: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x * g x) \ cx_poly" declare cx_poly.intros [intro] lemma Weierstrass_trig_polynomial: assumes contf: "continuous_on {-pi..pi} f" and fpi: "f(-pi) = f pi" and "0 < e" obtains n::nat and a b where "\x::real. x \ {-pi..pi} \ \f x - (\k\n. a k * sin (k * x) + b k * cos (k * x))\ < e" proof - interpret CR: function_ring_on "cx_poly" "sphere 0 1" proof show "continuous_on (sphere 0 1) f" if "f \ cx_poly" for f using that by induction (assumption | intro continuous_intros)+ fix x y::complex assume "x \ sphere 0 1" and "y \ sphere 0 1" and "x \ y" then consider "Re x \ Re y" | "Im x \ Im y" using complex_eqI by blast then show "\f\cx_poly. f x \ f y" by (metis cx_poly.Im cx_poly.Re) qed (auto simp: cx_poly.intros) have "continuous (at z within (sphere 0 1)) (f \ Im \ Ln)" if "norm z = 1" for z proof - obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ {-pi..pi} \ g x = f x" and periodic: "\x. g(x + 2*pi) = g x" using Tietze_periodic_interval [OF contf fpi] by auto let ?f = "(g \ Im \ Ln)" show ?thesis proof (rule continuous_transform_within) show "continuous (at z within (sphere 0 1)) ?f" using Weierstrass_trig_1 [OF contg periodic that] by (simp add: sphere_def) show "?f x' = (f \ Im \ Ln) x'" if "x' \ sphere 0 1" "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto then have "Im (Ln x') \ {-pi..pi}" using Im_Ln_le_pi [of x'] mpi_less_Im_Ln [of x'] by simp then show ?thesis using gf by simp qed qed (use that in auto) qed then have "continuous_on (sphere 0 1) (f \ Im \ Ln)" using continuous_on_eq_continuous_within mem_sphere_0 by blast then obtain g where "g \ cx_poly" and g: "\x. x \ sphere 0 1 \ \(f \ Im \ Ln) x - g x\ < e" using CR.Stone_Weierstrass_basic [of "f \ Im \ Ln"] \e > 0\ by meson define trigpoly where "trigpoly \ \f. \n a b. f = (\x. (\k\n. a k * sin(real k * x) + b k * cos(real k * x)))" have tp_const: "trigpoly(\x. c)" for c unfolding trigpoly_def by (rule_tac x=0 in exI) auto have tp_add: "trigpoly(\x. f x + g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast let ?a = "\n. (if n \ n1 then a1 n else 0) + (if n \ n2 then a2 n else 0)" let ?b = "\n. (if n \ n1 then b1 n else 0) + (if n \ n2 then b2 n else 0)" have eq: "{k. k \ max a b \ k \ a} = {..a}" "{k. k \ max a b \ k \ b} = {..b}" for a b::nat by auto have "(\x. f x + g x) = (\x. \k \ max n1 n2. ?a k * sin (real k * x) + ?b k * cos (real k * x))" by (simp add: feq geq algebra_simps eq sum.distrib if_distrib [of "\u. _ * u"] cong: if_cong flip: sum.inter_filter) then show ?thesis unfolding trigpoly_def by meson qed have tp_sum: "trigpoly(\x. \i\S. f i x)" if "finite S" "\i. i \ S \ trigpoly(f i)" for f and S :: "nat set" using that by induction (auto simp: tp_const tp_add) have tp_cmul: "trigpoly(\x. c * f x)" if "trigpoly f" for f c proof - obtain n a b where feq: "f = (\x. \k\n. a k * sin (real k * x) + b k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast have "(\x. c * f x) = (\x. \k \ n. (c * a k) * sin (real k * x) + (c * b k) * cos (real k * x))" by (simp add: feq algebra_simps sum_distrib_left) then show ?thesis unfolding trigpoly_def by meson qed have tp_cdiv: "trigpoly(\x. f x / c)" if "trigpoly f" for f c using tp_cmul [OF \trigpoly f\, of "inverse c"] by (simp add: divide_inverse_commute) have tp_diff: "trigpoly(\x. f x - g x)" if "trigpoly f" "trigpoly g" for f g using tp_add [OF \trigpoly f\ tp_cmul [OF \trigpoly g\, of "-1"]] by auto have tp_sin: "trigpoly(\x. sin (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. sin (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (rule_tac x="\i. 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_cos: "trigpoly(\x. cos (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. cos (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. 0" in exI) apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_sincos: "trigpoly(\x. sin(i * x) * sin(j * x)) \ trigpoly(\x. sin(i * x) * cos(j * x)) \ trigpoly(\x. cos(i * x) * sin(j * x)) \ trigpoly(\x. cos(i * x) * cos(j * x))" (is "?P i j") for i j::nat proof (rule linorder_wlog [of _ j i]) show "?P j i" if "i \ j" for j i using that by (simp add: sin_times_sin cos_times_cos sin_times_cos cos_times_sin diff_divide_distrib tp_add tp_diff tp_cdiv tp_cos tp_sin flip: left_diff_distrib distrib_right) qed (simp add: mult_ac) have tp_mult: "trigpoly(\x. f x * g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast then show ?thesis unfolding feq geq by (simp add: algebra_simps sum_product tp_sum tp_add tp_cmul tp_sincos del: mult.commute) qed have *: "trigpoly(\x. f(exp(\ * of_real x)))" if "f \ cx_poly" for f using that proof induction case Re then show ?case using tp_cos [of 1] by (auto simp: Re_exp) next case Im then show ?case using tp_sin [of 1] by (auto simp: Im_exp) qed (auto simp: tp_const tp_add tp_mult) obtain n a b where eq: "(g (iexp x)) = (\k\n. a k * sin (real k * x) + b k * cos (real k * x))" for x using * [OF \g \ cx_poly\] trigpoly_def by meson show thesis proof show "\f \ - (\k\n. a k * sin (real k * \) + b k * cos (real k * \))\ < e" if "\ \ {-pi..pi}" for \ proof - have "f \ - g (iexp \) = (f \ Im \ Ln) (iexp \) - g (iexp \)" proof (cases "\ = -pi") case True then show ?thesis by (simp add: exp_minus fpi) next case False then show ?thesis using that by auto qed then show ?thesis using g [of "exp(\ * of_real \)"] by (simp flip: eq) qed qed qed subsection\A bit of extra hacking round so that the ends of a function are OK\ lemma integral_tweak_ends: fixes a b :: real assumes "a < b" "e > 0" obtains f where "continuous_on {a..b} f" "f a = d" "f b = 0" "l2norm {a..b} f < e" proof - have cont: "continuous_on {a..b} (\x. if x \ a+1 / real(Suc n) then ((Suc n) * d) * ((a+1 / (Suc n)) - x) else 0)" for n proof (cases "a+1/(Suc n) \ b") case True have *: "1 / (1 + real n) > 0" by auto have abeq: "{a..b} = {a..a+1/(Suc n)} \ {a+1/(Suc n)..b}" using \a < b\ True apply auto using "*" by linarith show ?thesis unfolding abeq proof (rule continuous_on_cases) show "continuous_on {a..a+1 / real (Suc n)} (\x. real (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed auto next case False show ?thesis proof (rule continuous_on_eq [where f = "\x. ((Suc n) * d) * ((a+1/(Suc n)) - x)"]) show " continuous_on {a..b} (\x. (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed (use False in auto) qed let ?f = "\k x. (if x \ a+1 / (Suc k) then (Suc k) * d * (a+1 / (Suc k) - x) else 0)\<^sup>2" let ?g = "\x. if x = a then d\<^sup>2 else 0" have bmg: "?g \ borel_measurable (lebesgue_on {a..b})" apply (rule measurable_restrict_space1) using borel_measurable_if_I [of _ "{a}", OF borel_measurable_const] by auto have bmf: "?f k \ borel_measurable (lebesgue_on {a..b})" for k proof - have bm: "(\x. (Suc k) * d * (a+1 / real (Suc k) - x)) \ borel_measurable (lebesgue_on {..a+1 / (Suc k)})" by (intro measurable) (auto simp: measurable_completion measurable_restrict_space1) show ?thesis apply (intro borel_measurable_power measurable_restrict_space1) using borel_measurable_if_I [of _ "{.. a+1 / (Suc k)}", OF bm] apply auto done qed have int_d2: "integrable (lebesgue_on {a..b}) (\x. d\<^sup>2)" by (intro continuous_imp_integrable_real continuous_intros) have "(\k. ?f k x) \ ?g x" if x: "x \ {a..b}" for x proof (cases "x = a") case False then have "x > a" using x by auto with x obtain N where "N > 0" and N: "1 / real N < x-a" using real_arch_invD [of "x-a"] by (force simp: inverse_eq_divide) then have "x > a+1 / (1 + real k)" if "k \ N" for k proof - have "a+1 / (1 + real k) < a+1 / real N" using that \0 < N\ by (simp add: field_simps) also have "\ < x" using N by linarith finally show ?thesis . qed then show ?thesis apply (intro tendsto_eventually eventually_sequentiallyI [where c=N]) by (fastforce simp: False) qed auto then have tends: "AE x in (lebesgue_on {a..b}). (\k. ?f k x) \ ?g x" by force have le_d2: "\k. AE x in (lebesgue_on {a..b}). norm (?f k x) \ d\<^sup>2" proof show "norm ((if x \ a+1 / real (Suc k) then real (Suc k) * d * (a+1 / real (Suc k) - x) else 0)\<^sup>2) \ d\<^sup>2" if "x \ space (lebesgue_on {a..b})" for k x using that apply (simp add: abs_mult divide_simps flip: abs_le_square_iff) apply (auto simp: abs_if zero_less_mult_iff mult_left_le) done qed have integ: "integrable (lebesgue_on {a..b}) ?g" using integrable_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto have int: "(\k. integral\<^sup>L (lebesgue_on {a..b}) (?f k)) \ integral\<^sup>L (lebesgue_on {a..b}) ?g" using integral_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto then obtain N where N: "\k. k \ N \ \integral\<^sup>L (lebesgue_on {a..b}) (?f k) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" apply (simp add: lim_sequentially dist_real_def) apply (drule_tac x="e^2" in spec) using \e > 0\ by auto obtain M where "M > 0" and M: "1 / real M < b-a" using real_arch_invD [of "b-a"] by (metis \a < b\ diff_gt_0_iff_gt inverse_eq_divide neq0_conv) have *: "\integral\<^sup>L (lebesgue_on {a..b}) (?f (max M N)) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" using N by force let ?\ = "\x. if x \ a+1/(Suc (max M N)) then ((Suc (max M N)) * d) * ((a+1/(Suc (max M N))) - x) else 0" show ?thesis proof show "continuous_on {a..b} ?\" by (rule cont) have "1 / (1 + real (max M N)) \ 1 / (real M)" by (simp add: \0 < M\ frac_le) then have "\ (b \ a+1 / (1 + real (max M N)))" using M \a < b\ \M > 0\ max.cobounded1 [of M N] by linarith then show "?\ b = 0" by simp have null_a: "{a} \ null_sets (lebesgue_on {a..b})" using \a < b\ by (simp add: null_sets_restrict_space) have "LINT x|lebesgue_on {a..b}. ?g x = 0" by (force intro: integral_eq_zero_AE AE_I' [OF null_a]) then have "l2norm {a..b} ?\ < sqrt (e^2)" unfolding l2norm_def l2product_def power2_eq_square [symmetric] apply (intro real_sqrt_less_mono) using "*" by linarith then show "l2norm {a..b} ?\ < e" using \e > 0\ by auto qed auto qed lemma square_integrable_approximate_continuous_ends: assumes f: "f square_integrable {a..b}" and "a < b" "0 < e" obtains g where "continuous_on {a..b} g" "g b = g a" "g square_integrable {a..b}" "l2norm {a..b} (\x. f x - g x) < e" proof - obtain g where contg: "continuous_on UNIV g" and "g square_integrable {a..b}" and lg: "l2norm {a..b} (\x. f x - g x) < e/2" using f \e > 0\ square_integrable_approximate_continuous by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) obtain h where conth: "continuous_on {a..b} h" and "h a = g b - g a" "h b = 0" and lh: "l2norm {a..b} h < e/2" using integral_tweak_ends \e > 0\ by (metis \a < b\ zero_less_divide_iff zero_less_numeral) have "h square_integrable {a..b}" using \continuous_on {a..b} h\ continuous_imp_square_integrable by blast show thesis proof show "continuous_on {a..b} (\x. g x + h x)" by (blast intro: continuous_on_subset [OF contg] conth continuous_intros) then show "(\x. g x + h x) square_integrable {a..b}" using continuous_imp_square_integrable by blast show "g b + h b = g a + h a" by (simp add: \h a = g b - g a\ \h b = 0\) have "l2norm {a..b} (\x. (f x - g x) + - h x) < e" proof (rule le_less_trans [OF l2norm_triangle [of "\x. f x - g x" "{a..b}" "\x. - (h x)"]]) show "(\x. f x - g x) square_integrable {a..b}" using \g square_integrable {a..b}\ f square_integrable_diff by blast show "(\x. - h x) square_integrable {a..b}" by (simp add: \h square_integrable {a..b}\) show "l2norm {a..b} (\x. f x - g x) + l2norm {a..b} (\x. - h x) < e" using \h square_integrable {a..b}\ l2norm_neg lg lh by auto qed then show "l2norm {a..b} (\x. f x - (g x + h x)) < e" by (simp add: field_simps) qed qed subsection\Hence the main approximation result\ lemma Weierstrass_l2_trig_polynomial: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a b where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" proof - let ?\ = "\n a b x. \k\n. a k * sin(real k * x) + b k * cos(real k * x)" obtain g where contg: "continuous_on {-pi..pi} g" and geq: "g(-pi) = g pi" and g: "g square_integrable {-pi..pi}" and norm_fg: "l2norm {-pi..pi} (\x. f x - g x) < e/2" using \e > 0\ by (auto intro: square_integrable_approximate_continuous_ends [OF f, of "e/2"]) then obtain n a b where g_phi_less: "\x. x \ {-pi..pi} \ \g x - (?\ n a b x)\ < e/6" using \e > 0\ Weierstrass_trig_polynomial [OF contg geq, of "e/6"] by (meson zero_less_divide_iff zero_less_numeral) show thesis proof have si: "(?\ n u v) square_integrable {-pi..pi}" for n::nat and u v proof (intro square_integrable_sum continuous_imp_square_integrable) show "continuous_on {-pi..pi} (\x. u k * sin (real k * x) + v k * cos (real k * x))" if "k \ {..n}" for k using that by (intro continuous_intros) qed auto have "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) = l2norm {-pi..pi} (\x. (f x - g x) + (g x - (?\ n a b x)))" by simp also have "\ \ l2norm {-pi..pi} (\x. f x - g x) + l2norm {-pi..pi} (\x. g x - (?\ n a b x))" proof (rule l2norm_triangle) show "(\x. f x - g x) square_integrable {-pi..pi}" using f g square_integrable_diff by blast show "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast qed also have "\ < e" proof - have g_phi: "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast have "l2norm {-pi..pi} (\x. g x - (?\ n a b x)) \ e/2" unfolding l2norm_def l2product_def power2_eq_square [symmetric] proof (rule real_le_lsqrt) show "0 \ LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2" by (rule integral_nonneg_AE) auto have "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ LINT x|lebesgue_on {-pi..pi}. (e / 6) ^ 2" proof (rule integral_mono) show "integrable (lebesgue_on {-pi..pi}) (\x. (g x - (?\ n a b x))\<^sup>2)" using g_phi square_integrable_def by auto show "integrable (lebesgue_on {-pi..pi}) (\x. (e / 6)\<^sup>2)" by (intro continuous_intros continuous_imp_integrable_real) show "(g x - (?\ n a b x))\<^sup>2 \ (e / 6)\<^sup>2" if "x \ space (lebesgue_on {-pi..pi})" for x using \e > 0\ g_phi_less that apply (simp add: abs_le_square_iff [symmetric]) using less_eq_real_def by blast qed also have "\ \ (e / 2)\<^sup>2" using \e > 0\ pi_less_4 by (auto simp: power2_eq_square measure_restrict_space) finally show "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ (e / 2)\<^sup>2" . qed (use \e > 0\ in auto) with norm_fg show ?thesis by auto qed finally show "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) < e" . qed qed proposition Weierstrass_l2_trigonometric_set: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * trigonometric_set k x)) < e" proof - obtain n a b where lee: "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" using Weierstrass_l2_trig_polynomial [OF assms] . let ?a = "\k. if k = 0 then sqrt(2 * pi) * b 0 else if even k then sqrt pi * b(k div 2) else if k \ 2 * n then sqrt pi * a((Suc k) div 2) else 0" show thesis proof have [simp]: "Suc (i * 2) \ n * 2 \ i {..k\n. b k * cos (real k * x)) = (\i\n. if i = 0 then b 0 else b i * cos (real i * x))" for x by (rule sum.cong) auto moreover have "(\k\n. a k * sin (real k * x)) = (\i\n. (if Suc (2 * i) \ 2 * n then sqrt pi * a (Suc i) * sin ((1 + real i) * x) else 0) / sqrt pi)" (is "?lhs = ?rhs") for x proof (cases "n=0") case False then obtain n' where n': "n = Suc n'" using not0_implies_Suc by blast have "?lhs = (\k = Suc 0..n. a k * sin (real k * x))" by (simp add: atMost_atLeast0 sum_shift_lb_Suc0_0) also have "\ = (\ij\n'. a (Suc j) * sin (real (Suc j) * x)) = (\i = ?rhs" by (simp add: field_simps if_distrib [of "\x. x/_"] sum.inter_restrict [where B = "{..k\n. a k * sin(real k * x) + b k * cos(real k * x)) = (\k \ Suc(2*n). ?a k * trigonometric_set k x)" for x unfolding sum.in_pairs_0 trigonometric_set_def by (simp add: sum.distrib if_distrib [of "\x. x*_"] cong: if_cong) with lee show "l2norm {-pi..pi} (\x. f x - (\k \ Suc(2*n). ?a k * trigonometric_set k x)) < e" by auto qed qed subsection\Convergence wrt the L2 norm of trigonometric Fourier series\ definition Fourier_coefficient where "Fourier_coefficient \ orthonormal_coeff {-pi..pi} trigonometric_set" lemma Fourier_series_l2: assumes "f square_integrable {-pi..pi}" shows "(\n. l2norm {-pi..pi} (\x. f x - (\i\n. Fourier_coefficient f i * trigonometric_set i x))) \ 0" proof (clarsimp simp add: lim_sequentially dist_real_def Fourier_coefficient_def) let ?h = "\n x. (\i\n. orthonormal_coeff {-pi..pi} trigonometric_set f i * trigonometric_set i x)" show "\N. \n\N. \l2norm {-pi..pi} (\x. f x - ?h n x)\ < e" if "0 < e" for e proof - obtain N a where lte: "l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x)) < e" using Weierstrass_l2_trigonometric_set by (meson \0 < e\ assms) show ?thesis proof (intro exI allI impI) show "\l2norm {-pi..pi} (\x. f x - ?h m x)\ < e" if "N \ m" for m proof - have "\l2norm {-pi..pi} (\x. f x - ?h m x)\ = l2norm {-pi..pi} (\x. f x - ?h m x)" proof (rule abs_of_nonneg) show "0 \ l2norm {-pi..pi} (\x. f x - ?h m x)" apply (intro l2norm_pos_le square_integrable_diff square_integrable_sum square_integrable_lmult square_integrable_trigonometric_set assms, auto) done qed also have "\ \ l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x))" proof - have "(\i\m. (if i \ N then a i else 0) * trigonometric_set i x) = (\i\N. a i * trigonometric_set i x)" for x using sum.inter_restrict [where A = "{..m}" and B = "{..N}", symmetric] that by (force simp: if_distrib [of "\x. x * _"] min_absorb2 cong: if_cong) moreover have "l2norm {-pi..pi} (\x. f x - ?h m x) \ l2norm {-pi..pi} (\x. f x - (\i\m. (if i \ N then a i else 0) * trigonometric_set i x))" using orthonormal_optimal_partial_sum [OF orthonormal_system_trigonometric_set square_integrable_trigonometric_set assms] by simp ultimately show ?thesis by simp qed also have "\ < e" by (rule lte) finally show ?thesis . qed qed qed qed subsection\Fourier coefficients go to 0 (weak form of Riemann-Lebesgue)\ lemma trigonometric_set_mul_absolutely_integrable: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. trigonometric_set n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "trigonometric_set n \ borel_measurable (lebesgue_on {-pi..pi})" using square_integrable_def square_integrable_trigonometric_set by blast show "bounded (trigonometric_set n ` {-pi..pi})" unfolding bounded_iff using pi_gt3 sqrt_pi_ge1 by (rule_tac x=1 in exI) (auto simp: trigonometric_set_def dist_real_def intro: order_trans [OF abs_sin_le_one] order_trans [OF abs_cos_le_one]) qed (auto simp: assms) lemma trigonometric_set_mul_integrable: "f absolutely_integrable_on {-pi..pi} \ integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using trigonometric_set_mul_absolutely_integrable by (simp add: integrable_restrict_space set_integrable_def) lemma trigonometric_set_integrable [simp]: "integrable (lebesgue_on {-pi..pi}) (trigonometric_set n)" using trigonometric_set_mul_integrable [where f = id] by simp (metis absolutely_integrable_imp_integrable fmeasurableD interval_cbox lmeasurable_cbox square_integrable_imp_absolutely_integrable square_integrable_trigonometric_set) lemma absolutely_integrable_sin_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. sin(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_sin_cx mult_commute_abs) show "bounded ((\x. sin (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_sin_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma absolutely_integrable_cos_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. cos(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_cos_cx mult_commute_abs) show "bounded ((\x. cos (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_cos_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma assumes "f absolutely_integrable_on {-pi..pi}" shows Fourier_products_integrable_cos: "integrable (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x)" and Fourier_products_integrable_sin: "integrable (lebesgue_on {-pi..pi}) (\x. sin(k * x) * f x)" using absolutely_integrable_cos_product absolutely_integrable_sin_product assms by (auto simp: integrable_restrict_space set_integrable_def) lemma Riemann_lebesgue_square_integrable: assumes "orthonormal_system S w" "\i. w i square_integrable S" "f square_integrable S" shows "orthonormal_coeff S w f \ 0" using Fourier_series_square_summable [OF assms, of UNIV] summable_LIMSEQ_zero by force proposition Riemann_lebesgue: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient f \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain g where "continuous_on UNIV g" and gabs: "g absolutely_integrable_on {-pi..pi}" and fg_e2: "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. \f x - g x\) < e/2" using absolutely_integrable_approximate_continuous [OF assms, of "e/2"] by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) have "g square_integrable {-pi..pi}" using \continuous_on UNIV g\ continuous_imp_square_integrable continuous_on_subset by blast then have "orthonormal_coeff {-pi..pi} trigonometric_set g \ 0" using Riemann_lebesgue_square_integrable orthonormal_system_trigonometric_set square_integrable_trigonometric_set by blast with \e > 0\ obtain N where N: "\n. n \ N \ \orthonormal_coeff {-pi..pi} trigonometric_set g n\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def) have "\Fourier_coefficient f n\ < e" if "n \ N" for n proof - have "\LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x\ < e/2" using N [OF \n \ N\] by (auto simp: orthonormal_coeff_def l2product_def) have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * g x)" using gabs trigonometric_set_mul_integrable by blast moreover have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using assms trigonometric_set_mul_integrable by blast ultimately have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ = \(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * (g x - f x))\" by (simp add: algebra_simps flip: Bochner_Integration.integral_diff) also have "\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * (g x - f x))" by (simp add: gabs assms trigonometric_set_mul_integrable) have "(\x. f x - g x) absolutely_integrable_on {-pi..pi}" using gabs assms by blast then show "integrable (lebesgue_on {-pi..pi}) (\x. \f x - g x\)" by (simp add: absolutely_integrable_imp_integrable) fix x assume "x \ space (lebesgue_on {-pi..pi})" then have "-pi \ x" "x \ pi" by auto have "\trigonometric_set n x\ \ 1" using pi_ge_two apply (simp add: trigonometric_set_def) using sqrt_pi_ge1 abs_sin_le_one order_trans abs_cos_le_one by metis then show "\trigonometric_set n x * (g x - f x)\ \ \f x - g x\" using abs_ge_zero mult_right_mono by (fastforce simp add: abs_mult abs_minus_commute) qed finally have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" . then show ?thesis using N [OF \n \ N\] fg_e2 unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by linarith qed then show "\no. \n\no. dist (Fourier_coefficient f n) 0 < e" by auto qed lemma Riemann_lebesgue_sin: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have "\Fourier_coefficient f(Suc (2 * n))\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. sin ((1 + real n) * x) * f x\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by simp qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_cos: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have eq: "(x * 2 + x * (real n * 2)) / 2 = x + x * (real n)" for x by simp have "\Fourier_coefficient f(2*n + 2)\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. f x * cos (x + x * (real n))\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps eq) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by (simp add: field_simps) qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_sin_half: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. LINT x|lebesgue_on {-pi..pi}. sin ((real n + 1/2) * x) * f x) \ 0" proof (simp add: algebra_simps sin_add) let ?INT = "integral\<^sup>L (lebesgue_on {-pi..pi})" let ?f = "(\n. ?INT (\x. sin(n * x) * cos(1/2 * x) * f x) + ?INT (\x. cos(n * x) * sin(1/2 * x) * f x))" show "(\n. ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))) \ 0" proof (rule Lim_transform_eventually) have sin: "(\x. sin (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_sin_product assms) have cos: "(\x. cos (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_cos_product assms) show "\\<^sub>F n in sequentially. ?f n = ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))" unfolding mult.assoc apply (rule eventuallyI) apply (subst Bochner_Integration.integral_add [symmetric]) apply (safe intro!: cos absolutely_integrable_sin_product sin absolutely_integrable_cos_product absolutely_integrable_imp_integrable) apply (auto simp: algebra_simps) done have "?f \ 0 + 0" unfolding mult.assoc by (intro tendsto_add Riemann_lebesgue_sin Riemann_lebesgue_cos sin cos) then show "?f \ (0::real)" by simp qed qed lemma Fourier_sum_limit_pair: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. \k\2 * n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e::real assume "e > 0" then obtain N1 where N1: "\n. n \ N1 \ \Fourier_coefficient f n\ < e/2" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) obtain N2 where N2: "\n. n \ N2 \ \(\k\2 * n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using L unfolding lim_sequentially dist_real_def by (meson \0 < e\ half_gt_zero_iff) show "\no. \n\no. \(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof (intro exI allI impI) fix n assume n: "N1 + 2*N2 + 1 \ n" then have "n \ N1" "n \ N2" "n div 2 \ N2" by linarith+ consider "n = 2 * (n div 2)" | "n = Suc(2 * (n div 2))" by linarith then show "\(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof cases case 1 show ?thesis apply (subst 1) using N2 [OF \n div 2 \ N2\] by linarith next case 2 have "\(\k\2 * (n div 2). Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using N2 [OF \n div 2 \ N2\] by linarith moreover have "\Fourier_coefficient f(Suc (2 * (n div 2))) * trigonometric_set (Suc (2 * (n div 2))) t\ < (e/2) * 1" proof - have "\trigonometric_set (Suc (2 * (n div 2))) t\ \ 1" apply (simp add: trigonometric_set) using abs_sin_le_one sqrt_pi_ge1 by (blast intro: order_trans) moreover have "\Fourier_coefficient f(Suc (2 * (n div 2)))\ < e/2" using N1 \N1 \ n\ by auto ultimately show ?thesis unfolding abs_mult by (meson abs_ge_zero le_less_trans mult_left_mono mult_less_iff1 zero_less_one) qed ultimately show ?thesis apply (subst 2) unfolding sum.atMost_Suc by linarith qed qed qed next assume ?rhs then show ?lhs unfolding lim_sequentially by (auto simp: elim!: imp_forward ex_forward) qed subsection\Express Fourier sum in terms of the special expansion at the origin\ lemma Fourier_sum_0: "(\k \ n. Fourier_coefficient f k * trigonometric_set k 0) = (\k \ n div 2. Fourier_coefficient f(2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k = 2 * 0.. Suc (2 * (n div 2)). Fourier_coefficient f k * trigonometric_set k 0)" proof (rule sum.mono_neutral_left) show "\i\{2 * 0..Suc (2 * (n div 2))} - {..n}. Fourier_coefficient f i * trigonometric_set i 0 = 0" proof clarsimp fix i assume "i \ Suc (2 * (n div 2))" "\ i \ n" then have "i = Suc n" "even n" by presburger+ moreover assume "trigonometric_set i 0 \ 0" ultimately show "Fourier_coefficient f i = 0" by (simp add: trigonometric_set_def) qed qed auto also have "\ = ?rhs" unfolding sum.in_pairs by (simp add: trigonometric_set atMost_atLeast0) finally show ?thesis . qed lemma Fourier_sum_0_explicit: "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (Fourier_coefficient f 0 / sqrt 2 + (\k = 1..n div 2. Fourier_coefficient f(2*k))) / sqrt pi" (is "?lhs = ?rhs") proof - have "(\k=0..n. Fourier_coefficient f k * trigonometric_set k 0) = Fourier_coefficient f 0 * trigonometric_set 0 0 + (\k = 1..n. Fourier_coefficient f k * trigonometric_set k 0)" by (simp add: Fourier_sum_0 sum.atLeast_Suc_atMost) also have "\ = ?rhs" proof - have "Fourier_coefficient f 0 * trigonometric_set 0 0 = Fourier_coefficient f 0 / (sqrt 2 * sqrt pi)" by (simp add: real_sqrt_mult trigonometric_set(1)) moreover have "(\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k 0) = (\k = Suc 0..n div 2. Fourier_coefficient f(2*k)) / sqrt pi" proof (induction n) case (Suc n) show ?case by (simp add: Suc) (auto simp: trigonometric_set_def field_simps) qed auto ultimately show ?thesis by (simp add: add_divide_distrib) qed finally show ?thesis by (simp add: atMost_atLeast0) qed lemma Fourier_sum_0_integrals: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x))) / pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / (sqrt (2 * pi) * sqrt 2 * sqrt pi) = integral\<^sup>L (lebesgue_on {-pi..pi}) f / (2 * pi)" by (simp add: algebra_simps real_sqrt_mult) moreover have "(\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. trigonometric_set (2*k) x * f x) / sqrt pi = (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. cos (k * x) * f x) / pi" by (simp add: trigonometric_set_def sum_divide_distrib) ultimately show ?thesis unfolding Fourier_sum_0_explicit by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set add_divide_distrib) qed lemma Fourier_sum_0_integral: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. (1/2 + (\k = Suc 0..n div 2. cos(k * x))) * f x) / pi" proof - note * = Fourier_products_integrable_cos [OF assms] have "integrable (lebesgue_on {-pi..pi}) (\x. \n = Suc 0..n div 2. f x * cos (x * real n))" using * by (force simp: mult_ac) moreover have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. f x * cos (x * real k)) = (LINT x|lebesgue_on {-pi..pi}. f x / 2) + (LINT x|lebesgue_on {-pi..pi}. (\n = Suc 0..n div 2. f x * cos (x * real n)))" proof (subst Bochner_Integration.integral_sum) show "integrable (lebesgue_on {-pi..pi}) (\x. f x * cos (x * real i))" if "i \ {Suc 0..n div 2}" for i using that * [of i] by (auto simp: Bochner_Integration.integral_sum mult_ac) qed auto ultimately show ?thesis using Fourier_products_integrable_cos [OF assms] absolutely_integrable_imp_integrable [OF assms] by (simp add: Fourier_sum_0_integrals sum_distrib_left assms algebra_simps) qed subsection\How Fourier coefficients behave under addition etc\ lemma Fourier_coefficient_add: assumes "f absolutely_integrable_on {-pi..pi}" "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x + g x) i = Fourier_coefficient f i + Fourier_coefficient g i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def distrib_left) lemma Fourier_coefficient_minus: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. - f x) i = - Fourier_coefficient f i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def) lemma Fourier_coefficient_diff: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x - g x) i = Fourier_coefficient f i - Fourier_coefficient g i" proof - have mg: "(\x. - g x) absolutely_integrable_on {-pi..pi}" using g by (simp add: absolutely_integrable_measurable_real) show ?thesis using Fourier_coefficient_add [OF f mg] Fourier_coefficient_minus [OF g] by simp qed lemma Fourier_coefficient_const: "Fourier_coefficient (\x. c) i = (if i = 0 then c * sqrt(2 * pi) else 0)" by (auto simp: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set_def divide_simps measure_restrict_space) lemma Fourier_offset_term: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "Fourier_coefficient (\x. f(x+t)) (2 * n + 2) * trigonometric_set (2 * n + 2) 0 = Fourier_coefficient f(2 * n+1) * trigonometric_set (2 * n+1) t + Fourier_coefficient f(2 * n + 2) * trigonometric_set (2 * n + 2) t" proof - have eq: "(2 + 2 * real n) * x / 2 = (1 + real n) * x" for x by (simp add: divide_simps) have 1: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (cos (x + x * n) * cos (t + t * n)))" using Fourier_products_integrable_cos [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have 2: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (sin (x + x * n) * sin (t + t * n)))" using Fourier_products_integrable_sin [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * (x + t - t)) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" proof (rule has_integral_periodic_offset) have "(\x. cos (real (Suc n) * (x - t)) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (real (Suc n) * (x - t))) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. cos (real (Suc n) * (x - t))) ` {-pi..pi})" by (rule boundedI [where B=1]) auto qed (use f in auto) then show "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos ((Suc n) * (x - t)) * f x) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by (simp add: has_bochner_integral_integrable integrable_restrict_space set_integrable_def) next fix x show "cos (real (Suc n) * (x + (pi - - pi) - t)) * f(x + (pi - - pi)) = cos (real (Suc n) * (x - t)) * f x" using periodic cos.plus_of_nat [of "(Suc n) * (x - t)" "Suc n"] by (simp add: algebra_simps) qed then have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * x) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by simp then have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * (x - t)) * f x" using has_bochner_integral_integral_eq by blast also have "\ = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x - ((Suc n) * t)) * f x" by (simp add: algebra_simps) also have "\ = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" by (simp add: cos_diff algebra_simps 1 2 flip: integral_mult_left_zero integral_mult_right_zero) finally have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" . then show ?thesis unfolding Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def by (simp add: eq) (simp add: divide_simps algebra_simps l2product_def) qed lemma Fourier_sum_offset: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" (is "?lhs = ?rhs") proof - have "?lhs = Fourier_coefficient f 0 * trigonometric_set 0 t + (\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "(\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" proof (cases n) case (Suc n') have "(\k = Suc 0..2 * n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0.. Suc (Suc (2 * n')). Fourier_coefficient f k * trigonometric_set k t)" by (simp add: Suc) also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient f(Suc k) * trigonometric_set (Suc k) t)" unfolding atMost_atLeast0 using sum.shift_bounds_cl_Suc_ivl by blast also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient (\x. f(x+t)) (Suc k) * trigonometric_set (Suc k) 0)" unfolding sum.in_pairs_0 proof (rule sum.cong [OF refl]) show "Fourier_coefficient f(Suc (2 * k)) * trigonometric_set (Suc (2 * k)) t + Fourier_coefficient f(Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) t = Fourier_coefficient (\x. f(x + t)) (Suc (2 * k)) * trigonometric_set (Suc (2 * k)) 0 + Fourier_coefficient (\x. f(x + t)) (Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) 0" if "k \ {..n'}" for k using that Fourier_offset_term [OF assms, of t ] by (auto simp: trigonometric_set_def) qed also have "\ = (\k = Suc 0..Suc (Suc (2 * n')). Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp only: atMost_atLeast0 sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: Suc) finally show ?thesis . qed auto moreover have "?rhs = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0 + (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "Fourier_coefficient f 0 * trigonometric_set 0 t = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def integral_periodic_offset periodic) ultimately show ?thesis by simp qed lemma Fourier_sum_offset_unpaired: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0 + Fourier_coefficient (\x. f(x+t)) (Suc (2*k)) * trigonometric_set (Suc (2*k)) 0)" apply (simp add: assms Fourier_sum_offset) apply (subst sum.in_pairs_0 [symmetric]) by (simp add: trigonometric_set) also have "\ = ?rhs" by (auto simp: trigonometric_set) finally show ?thesis . qed subsection\Express partial sums using Dirichlet kernel\ definition Dirichlet_kernel where "Dirichlet_kernel \ \n x. if x = 0 then real n + 1/2 else sin((real n + 1/2) * x) / (2 * sin(x/2))" lemma Dirichlet_kernel_0 [simp]: "\x\ < 2 * pi \ Dirichlet_kernel 0 x = 1/2" by (auto simp: Dirichlet_kernel_def sin_zero_iff) lemma Dirichlet_kernel_minus [simp]: "Dirichlet_kernel n (-x) = Dirichlet_kernel n x" by (auto simp: Dirichlet_kernel_def) lemma Dirichlet_kernel_continuous_strong: "continuous_on {-(2 * pi)<..<2 * pi} (Dirichlet_kernel n)" proof - have "isCont (Dirichlet_kernel n) z" if "-(2 * pi) < z" "z < 2 * pi" for z proof (cases "z=0") case True have *: "(\x. sin ((n + 1/2) * x) / (2 * sin (x/2))) \0 \ real n + 1/2" by real_asymp show ?thesis unfolding Dirichlet_kernel_def continuous_at True apply (rule Lim_transform_eventually [where f = "\x. sin((n + 1/2) * x) / (2 * sin(x/2))"]) apply (auto simp: * eventually_at_filter) done next case False have "continuous (at z) (\x. sin((real n + 1/2) * x) / (2 * sin(x/2)))" using that False by (intro continuous_intros) (auto simp: sin_zero_iff) moreover have "\\<^sub>F x in nhds z. x \ 0" using False t1_space_nhds by blast ultimately show ?thesis using False by (force simp: Dirichlet_kernel_def continuous_at eventually_at_filter elim: Lim_transform_eventually) qed then show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma Dirichlet_kernel_continuous: "continuous_on {-pi..pi} (Dirichlet_kernel n)" apply (rule continuous_on_subset [OF Dirichlet_kernel_continuous_strong], clarsimp) using pi_gt_zero by linarith lemma absolutely_integrable_mult_Dirichlet_kernel: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Dirichlet_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Dirichlet_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) have "compact (Dirichlet_kernel n ` {-pi..pi})" by (auto simp: compact_continuous_image [OF Dirichlet_kernel_continuous]) then show "bounded (Dirichlet_kernel n ` {-pi..pi})" using compact_imp_bounded by blast qed (use assms in auto) lemma cosine_sum_lemma: "(1/2 + (\k = Suc 0..n. cos(real k * x))) * sin(x/2) = sin((real n + 1/2) * x) / 2" proof - consider "n=0" | "n \ 1" by linarith then show ?thesis proof cases case 2 then have "(\k = Suc 0..n. (sin (real k * x + x/2) * 2 - sin (real k * x - x/2) * 2) / 2) = sin (real n * x + x/2) - sin (x/2)" proof (induction n) case (Suc n) show ?case proof (cases "n=0") case False with Suc show ?thesis by (simp add: divide_simps algebra_simps) qed auto qed auto then show ?thesis by (simp add: distrib_right sum_distrib_right cos_times_sin) qed auto qed lemma Dirichlet_kernel_cosine_sum: assumes "\x\ < 2 * pi" shows "Dirichlet_kernel n x = 1/2 + (\k = Suc 0..n. cos(real k * x))" proof - have "sin ((real n + 1/2) * x) / (2 * sin (x/2)) = 1/2 + (\k = Suc 0..n. cos (real k * x))" if "x \ 0" proof - have "sin(x/2) \ 0" using assms that by (auto simp: sin_zero_iff) then show ?thesis using cosine_sum_lemma [of x n] by (simp add: divide_simps) qed then show ?thesis by (auto simp: Dirichlet_kernel_def) qed lemma integrable_Dirichlet_kernel: "integrable (lebesgue_on {-pi..pi}) (Dirichlet_kernel n)" using Dirichlet_kernel_continuous continuous_imp_integrable_real by blast lemma integral_Dirichlet_kernel [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = LINT x|lebesgue_on {-pi..pi}. 1/2 + (\k = Suc 0..n. cos (k * x))" using pi_ge_two by (force intro: Bochner_Integration.integral_cong [OF refl Dirichlet_kernel_cosine_sum]) also have "\ = (LINT x|lebesgue_on {-pi..pi}. 1/2) + (LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (k * x)))" proof (rule Bochner_Integration.integral_add) show "integrable (lebesgue_on {-pi..pi}) (\x. \k = Suc 0..n. cos (real k * x))" by (rule Bochner_Integration.integrable_sum) (metis integrable_cos_cx mult_commute_abs) qed auto also have "\ = pi" proof - have "\i. \Suc 0 \ i; i \ n\ \ integrable (lebesgue_on {-pi..pi}) (\x. cos (real i * x))" by (metis integrable_cos_cx mult_commute_abs) then have "LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (real k * x)) = 0" by (simp add: Bochner_Integration.integral_sum) then show ?thesis by (auto simp: measure_restrict_space) qed finally show ?thesis . qed lemma integral_Dirichlet_kernel_half [simp]: "integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi/2" proof - have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) + integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi" using integral_combine [OF integrable_Dirichlet_kernel] integral_Dirichlet_kernel by simp moreover have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) = integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using integral_reflect_real [of pi 0 "Dirichlet_kernel n"] by simp ultimately show ?thesis by simp qed lemma Fourier_sum_offset_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f(x+t)) / pi" (is "?lhs = ?rhs") proof - have ft: "(\x. f(x+t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp have "?lhs = (\k=0..n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" using Fourier_sum_offset_unpaired assms atMost_atLeast0 by auto also have "\ = Fourier_coefficient (\x. f(x+t)) 0 / sqrt (2 * pi) + (\k = Suc 0..n. Fourier_coefficient (\x. f(x+t)) (2*k) / sqrt pi)" by (simp add: sum.atLeast_Suc_atMost trigonometric_set_def) also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t)) / (2 * pi) + (\k = Suc 0..n. (LINT x|lebesgue_on {-pi..pi}. cos (real k * x) * f(x+t)) / pi)" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def) also have "\ = LINT x|lebesgue_on {-pi..pi}. f(x+t) / (2 * pi) + (\k = Suc 0..n. (cos (real k * x) * f(x+t)) / pi)" using Fourier_products_integrable_cos [OF ft] absolutely_integrable_imp_integrable [OF ft] by simp also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x+t))) / pi" by (simp add: divide_simps sum_distrib_right mult.assoc) also have "\ = ?rhs" proof - have "LINT x|lebesgue_on {-pi..pi}. f(x + t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)" proof - have eq: "f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = Dirichlet_kernel n x * f(x + t)" if "- pi \ x" "x \ pi" for x proof (cases "x = 0") case False then have "sin (x/2) \ 0" using that by (auto simp: sin_zero_iff) then have "f(x + t) * (1/2 + (\k = Suc 0..n. cos(real k * x))) = f(x + t) * sin((real n + 1/2) * x) / 2 / sin(x/2)" using cosine_sum_lemma [of x n] by (simp add: divide_simps) then show ?thesis by (simp add: Dirichlet_kernel_def False field_simps sum_distrib_left) qed (simp add: Dirichlet_kernel_def algebra_simps) show ?thesis by (rule Bochner_Integration.integral_cong [OF refl]) (simp add: eq) qed then show ?thesis by simp qed finally show ?thesis . qed lemma Fourier_sum_limit_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "((\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l) \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * l" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) / pi) \ l" by (simp add: Fourier_sum_limit_pair [OF f, symmetric] Fourier_sum_offset_Dirichlet_kernel assms) also have "\ \ (\k. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel k x * f(x + t)) * inverse pi) \ pi * l * inverse pi" by (auto simp: divide_simps) also have "\ \ ?rhs" using tendsto_mult_right_iff [of "inverse pi", symmetric] by auto finally show ?thesis . qed subsection\A directly deduced sufficient condition for convergence at a point\ lemma simple_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and ft: "(\x. (f(x+t) - f t) / sin(x/2)) absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - let ?\ = "\n. \k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t" let ?\ = "\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (f(x + t) - f t)" have fft: "(\x. f x - f t) absolutely_integrable_on {-pi..pi}" by (simp add: f absolutely_integrable_measurable_real) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have *: "?\ \ 0 \ ?\ \ 0" by (simp add: Fourier_sum_limit_Dirichlet_kernel fft periodic) moreover have "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t) \ 0" if "?\ \ 0" proof (rule Lim_transform_eventually [OF that eventually_sequentiallyI]) show "(\k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t" if "Suc 0 \ n" for n proof - have "(\k = Suc 0..n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k t)" proof (rule sum.cong [OF refl]) fix k assume k: "k \ {Suc 0..n}" then have [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (trigonometric_set k) = 0" by (auto simp: trigonometric_set_def) show "Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t = Fourier_coefficient f k * trigonometric_set k t" using k unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (simp add: right_diff_distrib f absolutely_integrable_measurable_real trigonometric_set_mul_integrable) qed moreover have "Fourier_coefficient (\x. f x - f t) 0 * trigonometric_set 0 t = Fourier_coefficient f 0 * trigonometric_set 0 t - f t" using f unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (auto simp: divide_simps trigonometric_set absolutely_integrable_imp_integrable measure_restrict_space) ultimately show ?thesis by (simp add: sum.atLeast_Suc_atMost atMost_atLeast0) qed qed moreover have "?\ \ 0" proof - have eq: "\n. ?\ n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n + 1/2) * x) * ((f(x+t) - f t) / sin(x/2) / 2))" unfolding Dirichlet_kernel_def by (rule Bochner_Integration.integral_cong [OF refl]) (auto simp: divide_simps sin_zero_iff) show ?thesis unfolding eq by (intro ft set_integrable_divide Riemann_lebesgue_sin_half) qed ultimately show ?thesis by (simp add: LIM_zero_cancel) qed subsection\A more natural sufficient Hoelder condition at a point\ lemma bounded_inverse_sin_half: assumes "d > 0" obtains B where "B>0" "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" proof - have "bounded ((\x. inverse (sin (x/2))) ` ({-pi..pi} - {- d<..x \ {-pi..pi}; x \ 0\ \ sin (x/2) \ 0" for x using \0 < d\ by (auto simp: sin_zero_iff) then show "continuous_on ({-pi..pi} - {-d<..x. inverse (sin (x/2)))" using \0 < d\ by (intro continuous_intros) auto show "compact ({-pi..pi} - {-d<.. 0" "a > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\ powr a" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof (rule simple_Fourier_convergence_periodic [OF f]) have half: "((\x. sin(x/2)) has_real_derivative 1/2) (at 0)" by (rule derivative_eq_intros refl | force)+ then obtain e0::real where "e0 > 0" and e0: "\x. \x \ 0; \x\ < e0\ \ \sin (x/2) / x - 1/2\ < 1/4" apply (simp add: DERIV_def Lim_at dist_real_def) apply (drule_tac x="1/4" in spec, auto) done obtain e where "e > 0" and e: "\x. \x\ < e \ \(f(x+t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a-1))" proof show "min d e0 > 0" using \d > 0\ \e0 > 0\ by auto next fix x assume x: "\x\ < min d e0" show "\(f(x + t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a - 1))" proof (cases "sin(x/2) = 0 \ x=0") case False have eq: "\(f(x + t) - f t) / sin (x/2)\ = \inverse (sin (x/2) / x)\ * (\f(x + t) - f t\ / \x\)" by simp show ?thesis unfolding eq proof (rule mult_mono) have "\sin (x/2) / x - 1/2\ < 1/4" using e0 [of x] x False by force then have "1/4 \ \sin (x/2) / x\" by (simp add: abs_if field_simps split: if_split_asm) then show "\inverse (sin (x/2) / x)\ \ 4" using False by (simp add: field_simps) have "\f(x + t) - f t\ / \x\ \ M * \x + t - t\ powr (a - 1)" using ft [of "x+t"] x by (auto simp: divide_simps algebra_simps Transcendental.powr_mult_base) also have "\ \ \M\ * \x\ powr (a - 1)" by (simp add: mult_mono) finally show "\f(x + t) - f t\ / \x\ \ \M\ * \x\ powr (a - 1)" . qed auto qed auto qed obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {- e<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \e > 0\] by metis let ?g = "\x. max (B * \f(x + t) - f t\) ((4 * \M\) * \x\ powr (a - 1))" show "(\x. (f(x + t) - f t) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto show "(\x. (f(x + t) - f t) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "(\x. f(x + t)) \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def fxt integrable_imp_measurable by blast show "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto qed auto have "(\x. f(x + t) - f t) absolutely_integrable_on {-pi..pi}" by (intro set_integral_diff fxt) (simp add: set_integrable_def) moreover have "(\x. \x\ powr (a - 1)) absolutely_integrable_on {-pi..pi}" proof - have "((\x. x powr (a - 1)) has_integral inverse a * pi powr a - inverse a * 0 powr a) {0..pi}" proof (rule fundamental_theorem_of_calculus_interior) show "continuous_on {0..pi} (\x. inverse a * x powr a)" using \a > 0\ by (intro continuous_on_powr' continuous_intros) auto show "((\b. inverse a * b powr a) has_vector_derivative x powr (a - 1)) (at x)" if "x \ {0<..a > 0\ - apply (simp flip: has_field_derivative_iff_has_vector_derivative) + apply (simp flip: has_real_derivative_iff_has_vector_derivative) apply (rule derivative_eq_intros | simp)+ done qed auto then have int: "(\x. x powr (a - 1)) integrable_on {0..pi}" by blast show ?thesis proof (rule nonnegative_absolutely_integrable_1) show "(\x. \x\ powr (a - 1)) integrable_on {-pi..pi}" proof (rule Henstock_Kurzweil_Integration.integrable_combine) show "(\x. \x\ powr (a - 1)) integrable_on {0..pi}" using int integrable_eq by fastforce then show "(\x. \x\ powr (a - 1)) integrable_on {- pi..0}" using Henstock_Kurzweil_Integration.integrable_reflect_real by fastforce qed auto qed auto qed ultimately show "?g integrable_on {-pi..pi}" apply (intro set_lebesgue_integral_eq_integral absolutely_integrable_max_1 absolutely_integrable_bounded_measurable_product_real set_integrable_abs measurable) apply (auto simp: image_constant_conv) done show "norm ((f(x + t) - f t) / sin (x/2)) \ ?g x" if "x \ {-pi..pi}" for x proof (cases "\x\ < e") case True then show ?thesis using e [OF True] by (simp add: max_def) next case False then have "\inverse (sin (x/2))\ \ B" using B that by force then have "\inverse (sin (x/2))\ * \f(x + t) - f t\ \ B * \f(x + t) - f t\" by (simp add: mult_right_mono) then have "\f(x + t) - f t\ / \sin (x/2)\ \ B * \f(x + t) - f t\" by (simp add: divide_simps split: if_split_asm) then show ?thesis by auto qed qed auto qed (auto simp: periodic) text\In particular, a Lipschitz condition at the point\ corollary Lipschitz_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and "d > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" using Hoelder_Fourier_convergence_periodic [OF f \d > 0\, of 1] assms by (fastforce simp add:) text\In particular, if left and right derivatives both exist\ proposition bi_differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and f_lt: "f differentiable at_left t" and f_gt: "f differentiable at_right t" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - obtain D_lt where D_lt: "\e. e > 0 \ \d>0. \x dist x t < d \ dist ((f x - f t) / (x - t)) D_lt < e" using f_lt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson lessThan_iff) then obtain d_lt where "d_lt > 0" and d_lt: "\x. \x < t; 0 < \x - t\; \x - t\ < d_lt\ \ \(f x - f t) / (x - t) - D_lt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) obtain D_gt where D_gt: "\e. e > 0 \ \d>0. \x>t. 0 < dist x t \ dist x t < d \ dist ((f x - f t) / (x - t)) D_gt < e" using f_gt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson greaterThan_iff) then obtain d_gt where "d_gt > 0" and d_gt: "\x. \x > t; 0 < \x - t\; \x - t\ < d_gt\ \ \(f x - f t) / (x - t) - D_gt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) show ?thesis proof (rule Lipschitz_Fourier_convergence_periodic [OF f]) fix x assume "\x - t\ < min d_lt d_gt" then have *: "\x - t\ < d_lt" "\x - t\ < d_gt" by auto consider "xt" by linarith then show "\f x - f t\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" proof cases case 1 then have "\(f x - f t) / (x - t) - D_lt\ < 1" using d_lt [OF 1] * by auto then have "\(f x - f t) / (x - t)\ - \D_lt\ < 1" by linarith then have "\f x - f t\ \ (\D_lt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . next case 3 then have "\(f x - f t) / (x - t) - D_gt\ < 1" using d_gt [OF 3] * by auto then have "\(f x - f t) / (x - t)\ - \D_gt\ < 1" by linarith then have "\f x - f t\ \ (\D_gt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . qed auto qed (auto simp: \0 < d_gt\ \0 < d_lt\ periodic) qed text\And in particular at points where the function is differentiable\ lemma differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and fdif: "f differentiable (at t)" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" by (rule bi_differentiable_Fourier_convergence_periodic) (auto simp: differentiable_at_withinI assms) text\Use reflection to halve the region of integration\ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" proof - show "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" apply (rule absolutely_integrable_mult_Dirichlet_kernel) using absolutely_integrable_periodic_offset [OF f, of t] periodic apply simp done then show "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" by (subst absolutely_integrable_reflect_real [symmetric]) simp show "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" by (simp add: absolutely_integrable_measurable_real borel_measurable_integrable integrable_Dirichlet_kernel) qed lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected [OF f periodic, of n] \d \ pi\ by (force intro: absolutely_integrable_on_subinterval)+ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * (f(t+x) + f(t-x))) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - c)) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part assms by (simp_all add: distrib_left right_diff_distrib) lemma integral_reflect_and_add: fixes f :: "real \ 'b::euclidean_space" assumes "integrable (lebesgue_on {-a..a}) f" shows "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" proof (cases "a \ 0") case True have f: "integrable (lebesgue_on {0..a}) f" "integrable (lebesgue_on {-a..0}) f" using integrable_subinterval assms by fastforce+ then have fm: "integrable (lebesgue_on {0..a}) (\x. f(-x))" using integrable_reflect_real by fastforce have "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {-a..0}) f + integral\<^sup>L (lebesgue_on {0..a}) f" by (simp add: True assms integral_combine) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f(-x)) + integral\<^sup>L (lebesgue_on {0..a}) f" by (metis (no_types) add.inverse_inverse add.inverse_neutral integral_reflect_real) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" by (simp add: fm f) finally show ?thesis . qed (use assms in auto) lemma Fourier_sum_offset_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) - l = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) / pi" proof - have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have int: "integrable (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using not_integrable_integral_eq by fastforce have "LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t) = LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * f(x + t) + Dirichlet_kernel n (- x) * f(- x + t)" by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel fxt) also have "\ = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) + pi * l" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part [OF f periodic order_refl, of n t] apply (simp add: algebra_simps absolutely_integrable_imp_integrable int) done finally show ?thesis by (simp add: Fourier_sum_offset_Dirichlet_kernel assms divide_simps) qed lemma Fourier_sum_limit_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l))) \ 0" apply (simp flip: Fourier_sum_limit_pair [OF f]) apply (simp add: Lim_null [where l=l] Fourier_sum_offset_Dirichlet_kernel_half assms) done subsection\Localization principle: convergence only depends on values "nearby"\ proposition Riemann_localization_integral: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and "d > 0" and d: "\x. \x\ < d \ f x = g x" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * g x)) \ 0" (is "?a \ 0") proof - let ?f = "\x. (if \x\ < d then 0 else f x - g x) / sin(x/2) / 2" have eq: "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" for n proof - have "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * (if \x\ < d then 0 else f x - g x))" apply (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel f g flip: Bochner_Integration.integral_diff) apply (auto simp: d algebra_simps intro: Bochner_Integration.integral_cong) done also have "\ = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" using \d > 0\ by (auto simp: Dirichlet_kernel_def intro: Bochner_Integration.integral_cong) finally show ?thesis . qed show ?thesis unfolding eq proof (rule Riemann_lebesgue_sin_half) obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \d > 0\] by metis have "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "f \ borel_measurable (lebesgue_on {-pi..pi})" "g \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def f g integrable_imp_measurable by blast+ show "(\x. x) \ borel_measurable (lebesgue_on {-pi..pi})" "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_intros continuous_imp_measurable_on_sets_lebesgue | force)+ qed auto have "(\x. B * \f x - g x\) absolutely_integrable_on {-pi..pi}" using \B > 0\ by (simp add: f g set_integrable_abs) then show "(\x. B * \f x - g x\) integrable_on {-pi..pi}" using absolutely_integrable_on_def by blast next fix x assume x: "x \ {-pi..pi}" then have [simp]: "sin (x/2) = 0 \ x=0" using \0 < d\ by (force simp: sin_zero_iff) show "norm ((if \x\ < d then 0 else f x - g x) / sin (x/2)) \ B * \f x - g x\" proof (cases "\x\ < d") case False then have "1 \ B * \sin (x/2)\" using \B > 0\ B [of x] x \d > 0\ by (auto simp: abs_less_iff divide_simps split: if_split_asm) then have "1 * \f x - g x\ \ B * \sin (x/2)\ * \f x - g x\" by (metis (full_types) abs_ge_zero mult.commute mult_left_mono) then show ?thesis using False by (auto simp: divide_simps algebra_simps) qed (simp add: d) qed auto then show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2) / 2) absolutely_integrable_on {-pi..pi}" using set_integrable_divide by blast qed qed lemma Riemann_localization_integral_range: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-d..d}) (\x. Dirichlet_kernel n x * f x)) \ 0" proof - have *: "(\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f x) - (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (if x \ {-d..d} then f x else 0))) \ 0" proof (rule Riemann_localization_integral [OF f _ \0 < d\]) have "f absolutely_integrable_on {-d..d}" using f absolutely_integrable_on_subinterval \d \ pi\ by fastforce moreover have "(\x. if - pi \ x \ x \ pi then if x \ {-d..d} then f x else 0 else 0) = (\x. if x \ {-d..d} then f x else 0)" using \d \ pi\ by force ultimately show "(\x. if x \ {-d..d} then f x else 0) absolutely_integrable_on {-pi..pi}" apply (subst absolutely_integrable_restrict_UNIV [symmetric]) apply (simp flip: absolutely_integrable_restrict_UNIV [of "{-d..d}" f]) done qed auto then show ?thesis using integral_restrict [of "{-d..d}" "{-pi..pi}" "\x. Dirichlet_kernel _ x * f x"] assms by (simp add: if_distrib cong: if_cong) qed lemma Riemann_localization: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and perf: "\x. f(x + 2*pi) = f x" and perg: "\x. g(x + 2*pi) = g x" and "d > 0" and d: "\x. \x-t\ < d \ f x = g x" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ c \ (\n. \k\n. Fourier_coefficient g k * trigonometric_set k t) \ c" proof - have "(\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * c \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * g (x + t)) \ pi * c" proof (intro Lim_transform_eq Riemann_localization_integral) show "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" "(\x. g (x + t)) absolutely_integrable_on {-pi..pi}" using assms by (auto intro: absolutely_integrable_periodic_offset) qed (use assms in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel assms) qed subsection\Localize the earlier integral\ lemma Riemann_localization_integral_range_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f x + f(-x))) - (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f x + f(-x)))) \ 0" proof - have *: "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-d..d}" for n by (metis (full_types) absolutely_integrable_mult_Dirichlet_kernel absolutely_integrable_on_subinterval \d \ pi\ atLeastatMost_subset_iff f neg_le_iff_le) show ?thesis using absolutely_integrable_mult_Dirichlet_kernel [OF f] using Riemann_localization_integral_range [OF assms] apply (simp add: "*" absolutely_integrable_imp_integrable integral_reflect_and_add) apply (simp add: algebra_simps) done qed lemma Fourier_sum_limit_Dirichlet_kernel_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l))) \ 0" proof - have "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp then have int: "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" by auto have "(\n. LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0 \ (\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0" by (rule Lim_transform_eq) (use Riemann_localization_integral_range_half [OF int d] in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel_half assms) qed subsection\Make a harmless simplifying tweak to the Dirichlet kernel\ lemma inte_Dirichlet_kernel_mul_expand: assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x = LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2))) \ (integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2))))" proof (cases "0 \ S") case True have *: "{x. x = 0 \ x \ S} \ sets (lebesgue_on S)" using True by (simp add: S sets_restrict_space_iff cong: conj_cong) have bm1: "(\x. Dirichlet_kernel n x * f x) \ borel_measurable (lebesgue_on S)" unfolding Dirichlet_kernel_def by (force intro: * assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) have bm2: "(\x. sin ((n + 1/2) * x) * f x / (2 * sin (x/2))) \ borel_measurable (lebesgue_on S)" by (intro assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto have 0: "{0} \ null_sets (lebesgue_on S)" using True by (simp add: S null_sets_restrict_space) show ?thesis apply (intro conjI integral_cong_AE integrable_cong_AE bm1 bm2 AE_I' [OF 0]) unfolding Dirichlet_kernel_def by auto next case False then show ?thesis unfolding Dirichlet_kernel_def by (auto intro!: Bochner_Integration.integral_cong Bochner_Integration.integrable_cong) qed lemma assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows integral_Dirichlet_kernel_mul_expand: "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x) = (LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th1") and integrable_Dirichlet_kernel_mul_expand: "integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th2") using inte_Dirichlet_kernel_mul_expand [OF assms] by auto proposition Fourier_sum_limit_sine_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. LINT x|lebesgue_on {0..d}. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)) \ 0" (is "?lhs \ ?\ \ 0") proof - have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have fbm: "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {-pi..pi}" by (force intro: ftmx ftx) let ?\ = "\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l)" have "?lhs \ ?\ \ 0" by (intro Fourier_sum_limit_Dirichlet_kernel_part assms) moreover have "?\ \ 0 \ ?\ \ 0" proof (rule Lim_transform_eq [OF Lim_transform_eventually]) let ?f = "\n. LINT x|lebesgue_on {0..d}. sin((real n + 1/2) * x) * (1 / (2 * sin(x/2)) - 1/x) * (f(t+x) + f(t-x) - 2*l)" have "?f n = (LINT x|lebesgue_on {-pi..pi}. sin ((n + 1/2) * x) * ((if x \ {0..d} then 1 / (2 * sin (x/2)) - 1/x else 0) * (f(t+x) + f(t-x) - 2*l)))" for n using d by (simp add: integral_restrict if_distrib [of "\u. _ * (u * _)"] mult.assoc flip: atLeastAtMost_iff cong: if_cong) moreover have "\ \ 0" proof (intro Riemann_lebesgue_sin_half absolutely_integrable_bounded_measurable_product_real) have "(\x. 1 / (2 * sin(x/2)) - 1/x) \ borel_measurable (lebesgue_on {0..d})" by (intro measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto then show "(\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) \ borel_measurable (lebesgue_on {-pi..pi})" using d by (simp add: borel_measurable_if_lebesgue_on flip: atLeastAtMost_iff) let ?g = "\x::real. 1 / (2 * sin(x/2)) - 1/x" have limg: "(?g \ ?g a) (at a within {0..d})" \\thanks to Manuel Eberl\ if a: "0 \ a" "a \ d" for a proof - have "(?g \ 0) (at_right 0)" and "(?g \ ?g d) (at_left d)" using d sin_gt_zero[of "d/2"] by (real_asymp simp: field_simps)+ moreover have "(?g \ ?g a) (at a)" if "a > 0" using a that d sin_gt_zero[of "a/2"] that by (real_asymp simp: field_simps) ultimately show ?thesis using that d by (cases "a = 0 \ a = d") (auto simp: at_within_Icc_at at_within_Icc_at_right at_within_Icc_at_left) qed have "((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi}) = ((\x. 1 / (2 * sin(x/2)) - 1/x) ` {0..d})" using d by (auto intro: image_eqI [where x=0]) moreover have "bounded \" by (intro compact_imp_bounded compact_continuous_image) (auto simp: continuous_on limg) ultimately show "bounded ((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi})" by simp qed (auto simp: fbm) ultimately show "?f \ (0::real)" by simp show "\\<^sub>F n in sequentially. ?f n = ?\ n - ?\ n" proof (rule eventually_sequentiallyI [where c = "Suc 0"]) fix n assume "n \ Suc 0" have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2)))" using d apply (subst integrable_Dirichlet_kernel_mul_expand [symmetric]) apply (intro assms absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF fbm] absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel_reflected_part2 | force)+ done moreover have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" proof (rule integrable_cong_AE_imp) let ?g = "\x. Dirichlet_kernel n x * (2 * sin(x/2) / x * (f(t+x) + f(t-x) - 2*l))" have *: "\2 * sin (x/2) / x\ \ 1" for x::real using abs_sin_x_le_abs_x [of "x/2"] by (simp add: divide_simps mult.commute abs_mult) have "bounded ((\x. 1 + (x/2)\<^sup>2) ` {-pi..pi})" by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then have bo: "bounded ((\x. 2 * sin (x/2) / x) ` {-pi..pi})" using * unfolding bounded_real by blast show "integrable (lebesgue_on {0..d}) ?g" using d apply (intro absolutely_integrable_imp_integrable absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Dirichlet_kernel] absolutely_integrable_bounded_measurable_product_real bo fbm measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros, auto) done show "(\x. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..d})" using d apply (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) done have "Dirichlet_kernel n x * (2 * sin(x/2)) / x = sin ((real n + 1/2) * x) / x" if "0 < x" "x \ d" for x using that d by (simp add: Dirichlet_kernel_def divide_simps sin_zero_pi_iff) then show "AE x in lebesgue_on {0..d}. ?g x = sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x" using d by (force intro: AE_I' [where N="{0}"]) qed ultimately have "?f n = (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2))) - (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" by (simp add: right_diff_distrib [of _ _ "1/_"] left_diff_distrib) also have "\ = ?\ n - ?\ n" using d by (simp add: measurable_restrict_mono [OF absolutely_integrable_imp_borel_measurable [OF fbm]] integral_Dirichlet_kernel_mul_expand) finally show "?f n = ?\ n - ?\ n" . qed qed ultimately show ?thesis by simp qed subsection\Dini's test for the convergence of a Fourier series\ proposition Fourier_Dini_test: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and int0d: "integrable (lebesgue_on {0..d}) (\x. \f(t+x) + f(t-x) - 2*l\ / x)" and "0 < d" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l" proof - define \ where "\ \ \n x. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)" have "(\n. LINT x|lebesgue_on {0..pi}. \ n x) \ 0" unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e :: real assume "e > 0" define \ where "\ \ \x. LINT x|lebesgue_on {0..x}. \f(t+x) + f(t-x) - 2*l\ / x" have [simp]: "\ 0 = 0" by (simp add: \_def integral_eq_zero_null_sets) have cont: "continuous_on {0..d} \" unfolding \_def using indefinite_integral_continuous_real int0d by blast with \d > 0\ have "\e>0. \da>0. \x'\{0..d}. \x' - 0\ < da \ \\ x' - \ 0\ < e" by (force simp: continuous_on_real_range dest: bspec [where x=0]) with \e > 0\ obtain k where "k > 0" and k: "\x'. \0 \ x'; x' < k; x' \ d\ \ \\ x'\ < e/2" by (drule_tac x="e/2" in spec) auto define \ where "\ \ min d (min (k/2) pi)" have e2: "\\ \\ < e/2" and "\ > 0" "\ \ pi" unfolding \_def using k \k > 0\ \d > 0\ by auto have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have 1: "\ n absolutely_integrable_on {0..\}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((real n + 1/2) * x)) \ borel_measurable (lebesgue_on {0..\})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((real n + 1/2) * x)) ` {0..\})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..\})" using \d > 0\ unfolding \_def by (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) moreover have "integrable (lebesgue_on {0..\}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" proof (rule integrable_subinterval [of 0 d]) show "integrable (lebesgue_on {0..d}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" using int0d by (subst integrable_cong) (auto simp: o_def) show "{0..\} \ {0..d}" using \d > 0\ by (auto simp: \_def) qed ultimately show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {0..\}" by (auto simp: absolutely_integrable_measurable) qed auto have 2: "\ n absolutely_integrable_on {\..pi}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((n + 1/2) * x)) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((n + 1/2) * x)) ` {\..pi})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. inverse x * (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) have "(\x. 1/x) \ borel_measurable (lebesgue_on {\..pi})" by (auto simp: measurable_completion measurable_restrict_space1) then show "inverse \ borel_measurable (lebesgue_on {\..pi})" by (metis (no_types, lifting) inverse_eq_divide measurable_lebesgue_cong) have "\x\{\..pi}. inverse \x\ \ inverse \" using \0 < \\ by auto then show "bounded (inverse ` {\..pi})" by (auto simp: bounded_iff) show "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_on_subinterval) show "(\x. (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {-pi..pi}" by (fast intro: ftx ftmx absolutely_integrable_on_const) show "{\..pi} \ {-pi..pi}" using \0 < \\ by auto qed qed auto then show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {\..pi}" by (metis (no_types, lifting) divide_inverse mult.commute set_integrable_cong) qed auto have "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" using ftx by auto moreover have "bounded ((\x. 0) ` {x. \x\ < \})" using bounded_def by blast moreover have "bounded (inverse ` {x. \ \x\ < \})" using \\ > 0\ by (auto simp: divide_simps intro: boundedI [where B = "1/\"]) ultimately have "(\x. (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)) absolutely_integrable_on UNIV" apply (intro absolutely_integrable_bounded_measurable_product_real measurable set_integral_diff) apply (auto simp: lebesgue_on_UNIV_eq measurable_completion simp flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) done moreover have "(if x \ {-pi..pi} then if \x\ < \ then 0 else (f(t+x) - l) / x else 0) = (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)" for x by (auto simp: divide_simps) ultimately have *: "(\x. if \x\ < \ then 0 else (f(t+x) - l) / x) absolutely_integrable_on {-pi..pi}" by (simp add: flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) have "(\n. LINT x|lebesgue_on {0..pi}. sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x)) \ 0" using Riemann_lebesgue_sin_half [OF *] * by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_sin_product cong: if_cong) moreover have "sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x) = (if \ \x\ < \ then \ n x else 0)" for x n by simp (auto simp: divide_simps algebra_simps \_def) ultimately have "(\n. LINT x|lebesgue_on {0..pi}. (if \ \x\ < \ then \ n x else 0)) \ 0" by simp moreover have "(if 0 \ x \ x \ pi then if \ \x\ < \ then \ n x else 0 else 0) = (if \ \ x \ x \ pi then \ n x else 0)" for x n using \\ > 0\ by auto ultimately have \: "(\n. LINT x|lebesgue_on {\..pi}. \ n x) \ 0" by (simp flip: Lebesgue_Measure.integral_restrict_UNIV) then obtain N::nat where N: "\n. n\N \ \LINT x|lebesgue_on {\..pi}. \ n x\ < e/2" unfolding lim_sequentially dist_real_def by (metis (full_types) \0 < e\ diff_zero half_gt_zero_iff) have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e" if "n \ N" for n::nat proof - have int: "integrable (lebesgue_on {0..pi}) (\ (real n))" by (intro integrable_combine [of concl: 0 pi] absolutely_integrable_imp_integrable) (use \\ > 0\ \\ \ pi\ 1 2 in auto) then have "integral\<^sup>L (lebesgue_on {0..pi}) (\ n) = integral\<^sup>L (lebesgue_on {0..\}) (\ n) + integral\<^sup>L (lebesgue_on {\..pi}) (\ n)" by (rule integral_combine) (use \0 < \\ \\ \ pi\ in auto) moreover have "\integral\<^sup>L (lebesgue_on {0..\}) (\ n)\ \ LINT x|lebesgue_on {0..\}. \f(t + x) + f(t - x) - 2 * l\ / x" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {0..\}) (\ (real n))" by (meson integrable_subinterval \\ \ pi\ int atLeastatMost_subset_iff order_refl) have "{0..\} \ {0..d}" by (auto simp: \_def) then show "integrable (lebesgue_on {0..\}) (\x. \f(t + x) + f(t - x) - 2 * l\ / x)" by (rule integrable_subinterval [OF int0d]) show "\\ (real n) x\ \ \f(t + x) + f(t - x) - 2 * l\ / x" if "x \ space (lebesgue_on {0..\})" for x using that apply (auto simp: \_def divide_simps abs_mult) by (simp add: mult.commute mult_left_le) qed ultimately have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e/2 + e/2" using N [OF that] e2 unfolding \_def by linarith then show ?thesis by simp qed then show "\N. \n\N. \integral\<^sup>L (lebesgue_on {0..pi}) (\ (real n)) - 0\ < e" by force qed then show ?thesis unfolding \_def using Fourier_sum_limit_sine_part assms pi_gt_zero by blast qed subsection\Cesaro summability of Fourier series using Fejér kernel\ definition Fejer_kernel :: "nat \ real \ real" where "Fejer_kernel \ \n x. if n = 0 then 0 else (\r sin(x/2) = 0") case True have "(\rrx. Fejer_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Fejer_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Fejer_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) show "bounded (Fejer_kernel n ` {-pi..pi})" using Fejer_kernel_continuous compact_continuous_image compact_imp_bounded by blast qed (use assms in auto) lemma absolutely_integrable_mult_Fejer_kernel_reflected1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {-pi..pi}" using assms by (force intro: absolutely_integrable_mult_Fejer_kernel absolutely_integrable_periodic_offset) lemma absolutely_integrable_mult_Fejer_kernel_reflected2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {-pi..pi}" proof - have "(\x. f(t - x)) absolutely_integrable_on {-pi..pi}" using assms apply (subst absolutely_integrable_reflect_real [symmetric]) apply (simp add: absolutely_integrable_periodic_offset) done then show ?thesis by (rule absolutely_integrable_mult_Fejer_kernel) qed lemma absolutely_integrable_mult_Fejer_kernel_reflected3: shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel by blast lemma absolutely_integrable_mult_Fejer_kernel_reflected_part1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected1]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part3: assumes "d \ pi" shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part4: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * (f(t + x) + f(t - x))) absolutely_integrable_on {0..d}" unfolding distrib_left by (intro set_integral_add absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part5: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - c)) absolutely_integrable_on {0..d}" unfolding distrib_left right_diff_distrib by (intro set_integral_add set_integral_diff absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms, auto) lemma Fourier_sum_offset_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "n > 0" shows "(\rk\2*r. Fourier_coefficient f k * trigonometric_set k t) / n - l = (LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi" proof - have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = (\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - l)" by (simp add: sum_subtractf) also have "\ = (\r = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" proof - have "integrable (lebesgue_on {0..pi}) (\x. Dirichlet_kernel i x * (f(t + x) + f(t - x) - 2 * l))" for i using absolutely_integrable_mult_Dirichlet_kernel_reflected_part2(2) f periodic by (force simp: intro!: absolutely_integrable_imp_integrable) then show ?thesis using \n > 0\ apply (simp add: Fejer_kernel_def flip: sum_divide_distrib) apply (simp add: sum_distrib_right flip: Bochner_Integration.integral_sum [symmetric]) done qed finally have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" . with \n > 0\ show ?thesis by (auto simp: mult.commute divide_simps) qed lemma Fourier_sum_limit_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) \ l \ ((\n. integral\<^sup>L (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - 2*l))) \ 0)" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n - l) \ 0" by (simp add: LIM_zero_iff) also have "\ \ (\n. ((((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) - l) * pi) \ 0" using tendsto_mult_right_iff [OF pi_neq_zero] by simp also have "\ \ ?rhs" apply (intro Lim_transform_eq [OF Lim_transform_eventually [of "\n. 0"]] eventually_sequentiallyI [of 1]) apply (simp_all add: Fourier_sum_offset_Fejer_kernel_half assms) done finally show ?thesis . qed lemma has_integral_Fejer_kernel: "has_bochner_integral (lebesgue_on {-pi..pi}) (Fejer_kernel n) (if n = 0 then 0 else pi)" proof - have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. (\rrx. (\rr 0" by (simp add: Fejer_kernel) theorem Fourier_Fejer_Cesaro_summable: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and fl: "(f \ l) (at t within atMost t)" and fr: "(f \ r) (at t within atLeast t)" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k t) / n) \ (l+r) / 2" proof - define h where "h \ \u. (f(t+u) + f(t-u)) - (l+r)" have "(\n. LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u) \ 0" proof - have h0: "(h \ 0) (at 0 within atLeast 0)" proof - have l0: "((\u. f(t-u) - l) \ 0) (at 0 within {0..})" using fl unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t-x" in bspec) apply (auto simp: dist_norm) done have r0: "((\u. f(t + u) - r) \ 0) (at 0 within {0..})" using fr unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t+x" in bspec) apply (auto simp: dist_norm) done show ?thesis using tendsto_add [OF l0 r0] by (simp add: h_def algebra_simps) qed show ?thesis unfolding lim_sequentially dist_real_def diff_0_right proof clarify fix e::real assume "e > 0" then obtain x' where "0 < x'" "\x. \0 < x; x < x'\ \ \h x\ < e / (2 * pi)" using h0 unfolding Lim_within dist_real_def by (auto simp: dest: spec [where x="e/2/pi"]) then obtain \ where "0 < \" "\ < pi" and \: "\x. 0 < x \ x \ \ \ \h x\ < e/2/pi" apply (intro that [where \="min x' pi/2"], auto) using m2pi_less_pi by linarith have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have h_aint: "h absolutely_integrable_on {-pi..pi}" unfolding h_def by (intro absolutely_integrable_on_const set_integral_diff set_integral_add, auto simp: ftx ftmx) have "(\n. LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ 0" proof (rule Lim_null_comparison) define \ where "\ \ \n. (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * sin(x/2) ^ 2)) / n" show "\\<^sub>F n in sequentially. norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" proof (rule eventually_sequentiallyI) fix n::nat assume "n \ 1" have hint: "(\x. h x / (2 * sin(x/2) ^ 2)) absolutely_integrable_on {\..pi}" unfolding divide_inverse mult.commute [of "h _"] proof (rule absolutely_integrable_bounded_measurable_product_real) have cont: "continuous_on {\..pi} (\x. inverse (2 * (sin (x * inverse 2))\<^sup>2))" using \0 < \\ by (intro continuous_intros) (auto simp: sin_zero_pi_iff) show "(\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) \ borel_measurable (lebesgue_on {\..pi})" using \0 < \\ by (intro cont continuous_imp_measurable_on_sets_lebesgue) auto show "bounded ((\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) ` {\..pi})" using cont by (simp add: compact_continuous_image compact_imp_bounded) show "h absolutely_integrable_on {\..pi}" using \0 < \\ \\ < pi\ by (auto intro: absolutely_integrable_on_subinterval [OF h_aint]) qed auto then have *: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2))" by (simp add: absolutely_integrable_measurable o_def) define \ where "\ \ \x. (if n = 0 then 0 else if x = 0 then n/2 else (sin (real n / 2 * x))\<^sup>2 / (2 * real n * (sin (x/2))\<^sup>2)) * h x" have "\LINT x|lebesgue_on {\..pi}. \ x\ \ (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" proof (rule integral_abs_bound_integral) show **: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" using Bochner_Integration.integrable_mult_left [OF *, of "1/n"] by (simp add: field_simps) show \: "\\ x\ \ \h x\ / (2 * (sin (x/2))\<^sup>2) / real n" if "x \ space (lebesgue_on {\..pi})" for x using that \0 < \\ apply (simp add: \_def divide_simps mult_less_0_iff abs_mult) apply (auto simp: square_le_1 mult_left_le_one_le) done show "integrable (lebesgue_on {\..pi}) \" proof (rule measurable_bounded_by_integrable_imp_lebesgue_integrable [OF _ **]) let ?g = "\x. \h x\ / (2 * sin(x/2) ^ 2) / n" have ***: "integrable (lebesgue_on {\..pi}) (\x. (sin (n/2 * x))\<^sup>2 * (inverse (2 * (sin (x/2))\<^sup>2) * h x))" proof (rule absolutely_integrable_imp_integrable [OF absolutely_integrable_bounded_measurable_product_real]) show "(\x. (sin (real n / 2 * x))\<^sup>2) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. (sin (real n / 2 * x))\<^sup>2) ` {\..pi})" by (force simp: square_le_1 intro: boundedI [where B=1]) show "(\x. inverse (2 * (sin (x/2))\<^sup>2) * h x) absolutely_integrable_on {\..pi}" using hint by (simp add: divide_simps) qed auto show "\ \ borel_measurable (lebesgue_on {\..pi})" apply (rule borel_measurable_integrable) apply (rule integrable_cong [where f = "\x. sin(n / 2 * x) ^ 2 / (2 * n * sin(x/2) ^ 2) * h x", OF refl, THEN iffD1]) using \0 < \\ ** apply (force simp: \_def divide_simps algebra_simps mult_less_0_iff abs_mult) using Bochner_Integration.integrable_mult_left [OF ***, of "1/n"] by (simp add: field_simps) show "norm (\ x) \ ?g x" if "x \ {\..pi}" for x using that \ by (simp add: \_def) qed auto qed then show "norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" by (simp add: Fejer_kernel \_def \_def flip: Bochner_Integration.integral_divide [OF *]) qed show "\ \ 0" unfolding \_def divide_inverse by (simp add: tendsto_mult_right_zero lim_inverse_n) qed then obtain N where N: "\n. n \ N \ \LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def \e > 0\) show "\N. \n\N. \(LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u)\ < e" proof (intro exI allI impI) fix n :: nat assume n: "n \ max 1 N" with N have 1: "\LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" by simp have "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ integral\<^sup>L (lebesgue_on {0..pi}) (Fejer_kernel n)" using \\ < pi\ has_bochner_integral_iff has_integral_Fejer_kernel_half by (force intro!: integral_mono_lebesgue_on_AE) also have "\ \ pi" using has_integral_Fejer_kernel_half by (simp add: has_bochner_integral_iff) finally have int_le_pi: "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ pi" . have 2: "\LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x\ \ (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * e/2/pi)" proof - have eq: "integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * h x) = integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" proof (rule integral_cong_AE) have \: "{u. u \ 0 \ P u} \ {0..\} = {0} \ Collect P \ {0..\}" "{u. u \ 0 \ P u} \ {0..\} = (Collect P \ {0..\}) - {0}" for P using \0 < \\ by auto have *: "- {0} \ A \ {0..\} = A \ {0..\} - {0}" for A by auto show "(\x. Fejer_kernel n x * h x) \ borel_measurable (lebesgue_on {0..\})" using \\ < pi\ by (intro absolutely_integrable_imp_borel_measurable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel], auto) then show "(\x. Fejer_kernel n x * (if x = 0 then 0 else h x)) \ borel_measurable (lebesgue_on {0..\})" apply (simp add: in_borel_measurable Ball_def vimage_def Collect_conj_eq Collect_imp_eq * flip: Collect_neg_eq) apply (elim all_forward imp_forward asm_rl) using \0 < \\ apply (auto simp: \ sets.insert_in_sets sets_restrict_space_iff cong: conj_cong) done have 0: "{0} \ null_sets (lebesgue_on {0..\})" using \0 < \\ by (simp add: null_sets_restrict_space) then show "AE x in lebesgue_on {0..\}. Fejer_kernel n x * h x = Fejer_kernel n x * (if x = 0 then 0 else h x)" by (auto intro: AE_I' [OF 0]) qed show ?thesis unfolding eq proof (rule integral_abs_bound_integral) have "(\x. if x = 0 then 0 else h x) absolutely_integrable_on {- pi..pi}" proof (rule absolutely_integrable_spike [OF h_aint]) show "negligible {0}" by auto qed auto with \0 < \\ \\ < pi\ show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" by (intro absolutely_integrable_imp_integrable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel]) auto show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * e / 2 / pi)" by (simp add: absolutely_integrable_imp_integrable \\ < pi\ absolutely_integrable_mult_Fejer_kernel_reflected_part3 less_eq_real_def) show "\Fejer_kernel n x * (if x = 0 then 0 else h x)\ \ Fejer_kernel n x * e / 2 / pi" if "x \ space (lebesgue_on {0..\})" for x using that \ [of x] \e > 0\ by (auto simp: abs_mult eq simp flip: times_divide_eq_right intro: mult_left_mono) qed qed have 3: "\ \ e/2" using int_le_pi \0 < e\ by (simp add: divide_simps mult.commute [of e]) have "integrable (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * h x)" unfolding h_def by (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Fejer_kernel_reflected_part5 assms) then have "LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * h x = (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x) + (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x)" by (rule integral_combine) (use \0 < \\ \\ < pi\ in auto) then show "\LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u\ < e" using 1 2 3 by linarith qed qed qed then show ?thesis unfolding h_def by (simp add: Fourier_sum_limit_Fejer_kernel_half assms add_divide_distrib) qed corollary Fourier_Fejer_Cesaro_summable_simple: assumes f: "continuous_on UNIV f" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ f x" proof - have "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ (f x + f x) / 2" proof (rule Fourier_Fejer_Cesaro_summable) show "f absolutely_integrable_on {- pi..pi}" using absolutely_integrable_continuous_real continuous_on_subset f by blast show "(f \ f x) (at x within {..x})" "(f \ f x) (at x within {x..})" using Lim_at_imp_Lim_at_within continuous_on_def f by blast+ qed (auto simp: periodic Lim_at_imp_Lim_at_within continuous_on_def f) then show ?thesis by simp qed end diff --git a/thys/Green/Derivs.thy b/thys/Green/Derivs.thy --- a/thys/Green/Derivs.thy +++ b/thys/Green/Derivs.thy @@ -1,651 +1,651 @@ theory Derivs imports General_Utils begin lemma field_simp_has_vector_derivative [derivative_intros]: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" - by (simp add: has_field_derivative_iff_has_vector_derivative) + by (simp add: has_real_derivative_iff_has_vector_derivative) lemma continuous_on_cases_empty [continuous_intros]: "\closed S; continuous_on S f; \x. \x \ S; \ P x\ \ f x = g x\ \ continuous_on S (\x. if P x then f x else g x)" using continuous_on_cases [of _ "{}"] by force lemma inj_on_cases: assumes "inj_on f (Collect P \ S)" "inj_on g (Collect (Not \ P) \ S)" "f ` (Collect P \ S) \ g ` (Collect (Not \ P) \ S) = {}" shows "inj_on (\x. if P x then f x else g x) S" using assms by (force simp: inj_on_def) lemma inj_on_arccos: "S \ {-1..1} \ inj_on arccos S" by (metis atLeastAtMost_iff cos_arccos inj_onI subsetCE) lemma has_vector_derivative_componentwise_within: "(f has_vector_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_vector_derivative (f' \ i)) (at a within S))" apply (simp add: has_vector_derivative_def) apply (subst has_derivative_componentwise_within) apply simp done lemma has_vector_derivative_pair_within: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "\u. u \ Basis \ ((\x. f x \ u) has_vector_derivative f' \ u) (at x within S)" "\u. u \ Basis \ ((\x. g x \ u) has_vector_derivative g' \ u) (at x within S)" shows "((\x. (f x, g x)) has_vector_derivative (f',g')) (at x within S)" apply (subst has_vector_derivative_componentwise_within) apply (auto simp: assms Basis_prod_def) done lemma piecewise_C1_differentiable_const: shows "(\x. c) piecewise_C1_differentiable_on s" using continuous_on_const by (auto simp add: piecewise_C1_differentiable_on_def) declare piecewise_C1_differentiable_const [simp, derivative_intros] declare piecewise_C1_differentiable_neg [simp, derivative_intros] declare piecewise_C1_differentiable_add [simp, derivative_intros] declare piecewise_C1_differentiable_diff [simp, derivative_intros] (*move to Derivative*) lemma piecewise_C1_differentiable_on_ident [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_vector" shows "(\x. x) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def using C1_differentiable_on_ident by (blast intro: continuous_on_id C1_differentiable_on_ident) lemma piecewise_C1_differentiable_on_mult [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_algebra" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. f x * g x) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def apply safe apply (blast intro: continuous_intros) apply (rename_tac A B) apply (rule_tac x="A \ B" in exI) apply (auto intro: C1_differentiable_on_mult C1_differentiable_on_subset) done lemma C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a :: real_normed_field" shows "f C1_differentiable_on S \ (\x. f x / c) C1_differentiable_on S" by (simp add: divide_inverse) lemma piecewise_C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_field" assumes "f piecewise_C1_differentiable_on S" shows "(\x. f x / c) piecewise_C1_differentiable_on S" by (simp add: divide_inverse piecewise_C1_differentiable_const piecewise_C1_differentiable_on_mult assms) lemma sqrt_C1_differentiable [simp, derivative_intros]: assumes f: "f C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) show ?thesis using assms - unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] + unfolding C1_differentiable_on_def has_real_derivative_iff_has_vector_derivative [symmetric] by (fastforce intro!: contf continuous_intros derivative_intros) qed lemma sqrt_piecewise_C1_differentiable [simp, derivative_intros]: assumes f: "f piecewise_C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (fastforce intro!: continuous_intros derivative_intros) lemma fixes f :: "real \ 'a::{banach,real_normed_field}" assumes f: "f C1_differentiable_on S" shows sin_C1_differentiable [simp, derivative_intros]: "(\x. sin (f x)) C1_differentiable_on S" and cos_C1_differentiable [simp, derivative_intros]: "(\x. cos (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df_sin = field_vector_diff_chain_at [where g=sin, unfolded o_def] note df_cos = field_vector_diff_chain_at [where g=cos, unfolded o_def] show "(\x. sin (f x)) C1_differentiable_on S" "(\x. cos (f x)) C1_differentiable_on S" using assms - unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] + unfolding C1_differentiable_on_def has_real_derivative_iff_has_vector_derivative [symmetric] apply auto by (rule contf continuous_intros df_sin df_cos derivative_intros exI conjI ballI | force)+ qed lemma has_derivative_abs: fixes a::real assumes "a \ 0" shows "(abs has_derivative ((*) (sgn a))) (at a)" proof - have [simp]: "norm = abs" using real_norm_def by force show ?thesis using has_derivative_norm [where 'a=real, simplified] assms by (simp add: mult_commute_abs) qed lemma abs_C1_differentiable [simp, derivative_intros]: fixes f :: "real \ real" assumes f: "f C1_differentiable_on S" and "0 \ f ` S" shows "(\x. abs (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df = DERIV_chain [where f=abs and g=f, unfolded o_def] show ?thesis using assms - unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] + unfolding C1_differentiable_on_def has_real_derivative_iff_has_vector_derivative [symmetric] apply clarify apply (rule df exI conjI ballI)+ apply (force simp: has_field_derivative_def intro: has_derivative_abs continuous_intros contf)+ done qed lemma C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f C1_differentiable_on S" "g C1_differentiable_on S" shows "(\x. (f x, g x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def apply safe apply (rename_tac A B) apply (intro exI ballI conjI) apply (rule_tac f'="A x" and g'="B x" in has_vector_derivative_pair_within) using has_vector_derivative_componentwise_within by (blast intro: continuous_on_Pair)+ lemma piecewise_C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (blast intro!: continuous_intros C1_differentiable_on_pair del: continuous_on_discrete intro: C1_differentiable_on_subset) lemma test2: assumes s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" and "x \ {0..1}" "x \ (\t. (v-u) *\<^sub>R t + u) -` s" shows "vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within{0..1})" proof - have i:"(g has_vector_derivative vector_derivative g (at ((v - u) * x + u))) (at ((v-u) * x + u))" using assms s [of "(v - u) * x + u"] uv mult_left_le [of x "v-u"] by (auto simp: vector_derivative_works) have ii:"((\x. g ((v - u) * x + u)) has_vector_derivative (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))) (at x)" by (intro vector_diff_chain_at [simplified o_def] derivative_eq_intros | simp add: i)+ have 0: "0 \ (v - u) * x + u" using assms uv by auto have 1: "(v - u) * x + u \ 1" using assms uv by simp (metis add.commute atLeastAtMost_iff atLeastatMost_empty_iff diff_ge_0_iff_ge empty_iff le_diff_eq mult_left_le) have iii: "vector_derivative g (at ((v - u) * x + u) within {0..1}) = vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF i, of "0" "1", OF 0 1] by auto have iv: "vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) = (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF ii, of "0" "1"] assms by auto show ?thesis using iii iv by auto qed lemma C1_differentiable_on_components: assumes "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s" shows "f C1_differentiable_on s" proof (clarsimp simp add: C1_differentiable_on_def has_vector_derivative_def) have *:"\f i x. x *\<^sub>R (f \ i) = (x *\<^sub>R f) \ i" by auto have "\f'. \i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" using assms lambda_skolem_euclidean[of "\i D. (\x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D"] apply (simp only: C1_differentiable_on_def has_vector_derivative_def *) using continuous_on_componentwise[of "s"] by metis then obtain f' where f': "\i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" by auto then have 0: "(\x\s. (f has_derivative (\z. z *\<^sub>R f' x)) (at x)) \ continuous_on s f'" using f' has_derivative_componentwise_within[of "f", where S= UNIV] by auto then show "\D. (\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D" by metis qed lemma piecewise_C1_differentiable_on_components: assumes "finite t" "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s - t" "\i. i \ Basis \ continuous_on s (\x. f x \ i)" shows "f piecewise_C1_differentiable_on s" using C1_differentiable_on_components assms continuous_on_componentwise piecewise_C1_differentiable_on_def by blast lemma all_components_smooth_one_pw_smooth_is_pw_smooth: assumes "\i. i \ Basis - {j} \ (\x. f x \ i) C1_differentiable_on s" assumes "(\x. f x \ j) piecewise_C1_differentiable_on s" shows "f piecewise_C1_differentiable_on s" proof - have is_cont: "\i\Basis. continuous_on s (\x. f x \ i)" using assms C1_differentiable_imp_continuous_on piecewise_C1_differentiable_on_def by fastforce obtain t where t:"(finite t \ (\x. f x \ j) C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "f"] using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t is_cont by fastforce qed lemma derivative_component_fun_component: fixes i::"'a::euclidean_space" assumes "f differentiable (at x)" shows "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" proof - have "((\x. f x \ i) has_vector_derivative vector_derivative f (at x) \ i) (at x)" using assms and bounded_linear.has_vector_derivative[of "(\x. x \ i)" "f" "(vector_derivative f (at x))" "(at x)"] and bounded_linear_inner_left[of "i"] and vector_derivative_works[of "f" "(at x)"] by blast then show "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" using vector_derivative_works[of "(\x. (f x) \ i)" "(at x)"] and differentiableI_vector[of "(\x. (f x) \ i)" "(vector_derivative f (at x) \ i)" "(at x)"] and Derivative.vector_derivative_at by force qed lemma gamma_deriv_at_within: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b}" and gamma_differentiable: "\x \ {a .. b}. \ differentiable at x" shows "vector_derivative \ (at x within {a..b}) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b by (auto simp add: vector_derivative_works) lemma islimpt_diff_finite: assumes "finite (t::'a::t1_space set)" shows "x islimpt s - t = x islimpt s" proof- have iii: "s - t = s - (t \ s)" by auto have "(t \ s) \ s" by auto have ii:"finite (t \ s)" using assms(1) by auto have i: "(t \ s) \ (s - (t \ s)) = ( s)" using assms by auto then have "x islimpt s - (t \ s) = x islimpt s" using islimpt_Un_finite[OF ii,where ?t = "s - (t \ s)"] i assms(1) by force then show ?thesis using iii by auto qed lemma ivl_limpt_diff: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x islimpt {a..b} - s" proof- have "x islimpt {a..b}" proof (cases "x \{a,b}") have i: "finite {a,b}" and ii: "{a, b} \ {a<..{a,b}" then show ?thesis using islimpt_Un_finite[OF i, where ?t= "{a<..{a,b}" then show "x islimpt {a..b}" using assms by auto qed then show "x islimpt {a..b} - s" using islimpt_diff_finite[OF assms(1)] assms by fastforce qed lemma ivl_closure_diff_del: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x \ closure (({a..b} - s) - {x})" using ivl_limpt_diff islimpt_in_closure assms by blast lemma ivl_not_trivial_limit_within: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "at x within {a..b} - s \ bot" using assms ivl_closure_diff_del not_trivial_limit_within by blast lemma vector_derivative_at_within_non_trivial_limit: "at x within s \ bot \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within s) = f'" using has_vector_derivative_at_within vector_derivative_within by fastforce lemma vector_derivative_at_within_ivl_diff: "finite s \ a < b \ (x::real) \ {a..b} - s \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within {a..b} - s) = f'" using vector_derivative_at_within_non_trivial_limit ivl_not_trivial_limit_within by fastforce lemma gamma_deriv_at_within_diff: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b} - s" and gamma_differentiable: "\x \ {a .. b} - s. \ differentiable at x" and s_subset: "s \ {a..b}" and finite_s: "finite s" shows "vector_derivative \ (at x within {a..b} - s) = vector_derivative \ (at x)" using vector_derivative_at_within_ivl_diff [of "s" "a" "b" "x" "\" "vector_derivative \ (at x)"] gamma_differentiable x_within_bounds a_leq_b s_subset finite_s by (auto simp add: vector_derivative_works) lemma gamma_deriv_at_within_gen: assumes a_leq_b: "a < b" and x_within_bounds: "x \ s" and s_subset: "s \ {a..b}" and gamma_differentiable: "\x \ s. \ differentiable at x" shows "vector_derivative \ (at x within ({a..b})) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b s_subset by (auto simp add: vector_derivative_works) lemma derivative_component_fun_component_at_within_gen: assumes gamma_differentiable: "\x \ s. \ differentiable at x" and s_subset: "s \ {0..1}" shows "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ s. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ s" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within_gen[of "0" "1"] x_within_bounds gamma_differentiable s_subset by (auto simp add: vector_derivative_works) then have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_deriv_at_within_gen[of "0" "1", where \ = "(\x. \ x \ i)"] x_within_bounds gamma_i_component_smooth s_subset by (auto simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma derivative_component_fun_component_at_within: assumes gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" shows "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ {0..1}" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within[of "0" "1"] x_within_bounds gamma_differentiable by (auto simp add: vector_derivative_works) have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using Derivative.vector_derivative_at_within_ivl[of "(\x. (\ x) \ i)" "vector_derivative (\x. (\ x) \ i) (at x)" "x" "0" "1"] has_vector_derivative_at_within[of "(\x. \ x \ i)" "vector_derivative (\x. \ x \ i) (at x)" "x" "{0..1}"] gamma_i_component_smooth x_within_bounds by (simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma straight_path_diffrentiable_x: fixes b :: "real" and y1 ::"real" assumes gamma_def: "\ = (\x. (b, y2 + y1 * x))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma straight_path_diffrentiable_y: fixes b :: "real" and y1 y2 ::"real" assumes gamma_def: "\ = (\x. (y2 + y1 * x , b))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma piecewise_C1_differentiable_on_imp_continuous_on: assumes "f piecewise_C1_differentiable_on s" shows "continuous_on s f" using assms by (auto simp add: piecewise_C1_differentiable_on_def) lemma boring_lemma1: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (f x, 0)) has_vector_derivative ((D,0::real))) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D *\<^sub>R (1,0))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(1,0)"] by auto have "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D,0)) (at x)" proof - have "(D, 0::'a) = D *\<^sub>R (1, 0)" by simp then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma boring_lemma2: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (0, f x)) has_vector_derivative (0, D)) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative (D *\<^sub>R (0,1))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(0,1)"] by auto then have "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative ((0,D))) (at x)" using scaleR_Pair Real_Vector_Spaces.real_scaleR_def proof - have "(0::'b, D) = D *\<^sub>R (0, 1)" by auto then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma pair_prod_smooth_pw_smooth: assumes "(f::real\real) C1_differentiable_on s" "(g::real\real) piecewise_C1_differentiable_on s" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on s" proof - have f_cont: "continuous_on s f" using assms(1) C1_differentiable_imp_continuous_on by fastforce have g_cont: "continuous_on s g" using assms(2) by (auto simp add: piecewise_C1_differentiable_on_def) obtain t where t:"(finite t \ g C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "(\x. (f x, g x))"] apply (simp add: real_pair_basis) using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t f_cont g_cont by fastforce qed lemma scale_shift_smooth: shows "(\x. a + b * x) C1_differentiable_on s" proof - show "(\x. a + b * x) C1_differentiable_on s" using C1_differentiable_on_mult C1_differentiable_on_add C1_differentiable_on_const C1_differentiable_on_ident by auto qed lemma open_diff: assumes "finite (t::'a::t1_space set)" "open (s::'a set)" shows "open (s -t)" using assms proof(induction "t") show "open s \ open (s - {})" by auto next fix x::"'a::t1_space" fix F::"'a::t1_space set" assume step: "finite F " " x \ F" "open s" then have i: "(s - insert x F) = (s - F) - {x}" by auto assume ind_hyp: " (open s \ open (s - F))" show "open (s - insert x F)" apply (simp only: i) using open_delete[of "s -F"] ind_hyp[OF step(3)] by auto qed lemma has_derivative_transform_within: assumes "0 < d" and "x \ s" and "\x'\s. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s. f x = g x" and "x \ {a..b} -s" and "(f has_derivative f') (at x within {a..b} - s)" shows "(g has_derivative f') (at x within {a..b} - s)" using has_derivative_transform_within[of "b" "x" "{a..b} - s"] assms by auto lemma has_vector_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s . f x = g x" and "x \ {a..b} - s" and "(f has_vector_derivative f') (at x within {a..b} - s)" shows "(g has_vector_derivative f') (at x within {a..b} - s)" using assms has_derivative_transform_within_ivl apply (auto simp add: has_vector_derivative_def) by blast lemma has_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x)" shows "(g has_derivative f') (at x)" using has_derivative_transform_within [of d x UNIV f g f'] assms by simp lemma has_vector_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_vector_derivative f') (at x)" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_at) lemma C1_diff_components_2: assumes "b \ Basis" assumes "f C1_differentiable_on s" shows "(\x. f x \ b) C1_differentiable_on s" proof - obtain D where D:"(\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x))" "continuous_on s D" using assms(2) by (fastforce simp add: C1_differentiable_on_def has_vector_derivative_def) show ?thesis proof (simp add: C1_differentiable_on_def has_vector_derivative_def, intro exI conjI) show "continuous_on s (\x. D x \ b)" using D continuous_on_componentwise assms(1) by fastforce show "(\x\s. ((\x. f x \ b) has_derivative (\y. y * (\x. D x \ b) x)) (at x))" using has_derivative_inner_left D(1) by fastforce qed qed lemma eq_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "f C1_differentiable_on s" shows "g C1_differentiable_on s" proof (simp add: C1_differentiable_on_def) obtain D where D: "(\x\s. (f has_vector_derivative D x) (at x)) \ continuous_on s D" using assms by (auto simp add: C1_differentiable_on_def) then have f: "(\x\s. (g has_vector_derivative D x) (at x))" using assms(1-2) by (metis dist_commute has_vector_derivative_transform_at) have "(\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" using D f by auto then show "\D. (\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" by metis qed lemma eq_pw_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "\x\s. f x = g x" "f piecewise_C1_differentiable_on s" shows "g piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def) have g_cont: "continuous_on s g" using assms piecewise_C1_differentiable_const by (simp add: piecewise_C1_differentiable_on_def) obtain t where t: "finite t \ f C1_differentiable_on s - t" using assms by (auto simp add: piecewise_C1_differentiable_on_def) then have "g C1_differentiable_on s - t" using assms eq_smooth by blast then show "continuous_on s g \ (\t. finite t \ g C1_differentiable_on s - t)" using t g_cont by metis qed lemma scale_piecewise_C1_differentiable_on: assumes "f piecewise_C1_differentiable_on s" shows "(\x. (c::real) * (f x)) piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def, intro conjI) show "continuous_on s (\x. c * f x)" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) show "\t. finite t \ (\x. c * f x) C1_differentiable_on s - t" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) qed lemma eq_smooth_gen: assumes "f C1_differentiable_on s" "\x. f x = g x" shows "g C1_differentiable_on s" using assms unfolding C1_differentiable_on_def by (metis (no_types, lifting) has_vector_derivative_weaken UNIV_I top_greatest) lemma subpath_compose: shows "(subpath a b \) = \ o (\x. (b - a) * x + a)" by (auto simp add: subpath_def) lemma subpath_smooth: assumes "\ C1_differentiable_on {0..1}" "0 \ a" "a < b" "b \ 1" shows "(subpath a b \) C1_differentiable_on {0..1}" proof- have "\ C1_differentiable_on {a..b}" apply (rule C1_differentiable_on_subset) using assms by auto then have "\ C1_differentiable_on (\x. (b - a) * x + a) ` {0..1}" using \a < b\ closed_segment_eq_real_ivl closed_segment_real_eq by auto moreover have "finite ({0..1} \ (\x. (b - a) * x + a) -` {x})" for x proof - have "((\x. (b - a) * x + a) -` {x}) = {(x -a) /(b-a)}" using assms by (auto simp add: divide_simps) then show ?thesis by auto qed ultimately show ?thesis by (force simp add: subpath_compose intro: C1_differentiable_compose derivative_intros) qed lemma has_vector_derivative_divide[derivative_intros]: fixes a :: "'a::real_normed_field" shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" unfolding divide_inverse by (fact has_vector_derivative_mult_left) end diff --git a/thys/Green/Integrals.thy b/thys/Green/Integrals.thy --- a/thys/Green/Integrals.thy +++ b/thys/Green/Integrals.thy @@ -1,946 +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 + using has_real_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 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 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 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 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]) + apply (subst has_real_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) + by (simp only: has_real_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) + by (simp only: has_real_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" 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" 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/Green/Paths.thy b/thys/Green/Paths.thy --- a/thys/Green/Paths.thy +++ b/thys/Green/Paths.thy @@ -1,3584 +1,3584 @@ theory Paths imports Derivs General_Utils Integrals begin (*This has everything related to paths purely*) lemma reverse_subpaths_join: shows " subpath 1 (1 / 2) p +++ subpath (1 / 2) 0 p = reversepath p" using reversepath_subpath join_subpaths_middle pathfinish_subpath pathstart_subpath reversepath_joinpaths by (metis (no_types, lifting)) (*Below F cannot be from 'a \ 'b because the dot product won't work. We have that g returns 'a and then F takes the output of g, so F should start from 'a Then we have to compute the dot product of the vector b with both the derivative of g, and F. Since the derivative of g returns the same type as g, accordingly F should return the same type as g, i.e. 'a. *) definition line_integral:: "('a::euclidean_space \ 'a::euclidean_space) \ (('a) set) \ (real \ 'a) \ real" where "line_integral F basis g \ integral {0 .. 1} (\x. \b\basis. (F(g x) \ b) * (vector_derivative g (at x within {0..1}) \ b))" definition line_integral_exists where "line_integral_exists F basis \ \ (\x. \b\basis. F(\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) integrable_on {0..1}" lemma line_integral_on_pair_straight_path: fixes F::"('a::euclidean_space) \ 'a" and g :: "real \ real" and \ assumes gamma_const: "\x. \(x)\ i = a" and gamma_smooth: "\x \ {0 .. 1}. \ differentiable at x" shows "(line_integral F {i} \) = 0" "(line_integral_exists F {i} \)" proof (simp add: line_integral_def) have *: "F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i) = 0" if "0 \ x \ x \ 1" for x proof - have "((\x. \(x)\ i) has_vector_derivative 0) (at x)" using vector_derivative_const_at[of "a" "x"] and gamma_const by auto then have "(vector_derivative \ (at x) \ i) = 0" using derivative_component_fun_component[ of "\" "x" "i"] and gamma_smooth and that by (simp add: vector_derivative_at) then have "(vector_derivative \ (at x within {0 .. 1}) \ i) = 0" using has_vector_derivative_at_within vector_derivative_at_within_ivl that by (metis atLeastAtMost_iff gamma_smooth vector_derivative_works zero_less_one) then show ?thesis by auto qed then have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) has_integral 0) {0..1}" using has_integral_is_0[of "{0 .. 1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))"] by auto then have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) integrable_on {0..1})" by auto then show "line_integral_exists F {i} \" by (auto simp add:line_integral_exists_def) show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = 0" using * has_integral_is_0[of "{0 .. 1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))"] by auto qed lemma line_integral_on_pair_path_strong: fixes F::"('a::euclidean_space) \ ('a)" and g::"real \ 'a" and \::"(real \ 'a)" and i::'a assumes i_norm_1: "norm i = 1" and g_orthogonal_to_i: "\x. g(x) \ i = 0" and gamma_is_in_terms_of_i: "\ = (\x. f(x) *\<^sub>R i + g(f(x)))" and gamma_smooth: "\ piecewise_C1_differentiable_on {0 .. 1}" and g_continuous_on_f: "continuous_on (f ` {0..1}) g" and path_start_le_path_end: "(pathstart \) \ i \ (pathfinish \) \ i" and field_i_comp_cont: "continuous_on (path_image \) (\x. F x \ i)" shows "line_integral F {i} \ = integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))" "line_integral_exists F {i} \" proof (simp add: line_integral_def) obtain s where gamma_differentiable: "finite s" "(\x \ {0 .. 1} - s. \ differentiable at x)" using gamma_smooth by (auto simp add: C1_differentiable_on_eq piecewise_C1_differentiable_on_def) then have gamma_i_component_smooth: "\x \ {0 .. 1} - s. (\x. \ x \ i) differentiable at x" by auto have field_cont_on_path: "continuous_on ((\x. \ x \ i) ` {0..1}) (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" proof - have 0: "(\x. \ x \ i) = f" proof fix x show "\ x \ i = f x" using g_orthogonal_to_i i_norm_1 by (simp only: gamma_is_in_terms_of_i real_inner_class.inner_add_left g_orthogonal_to_i inner_scaleR_left inner_same_Basis norm_eq_1) qed show ?thesis unfolding 0 apply (rule continuous_on_compose2 [of _ "(\x. F(x) \ i)" "f ` { 0..1}" "(\x. x *\<^sub>R i + g x)"] field_i_comp_cont g_continuous_on_f field_i_comp_cont continuous_intros)+ by (auto simp add: gamma_is_in_terms_of_i path_image_def) qed have path_start_le_path_end': "\ 0 \ i \ \ 1 \ i" using path_start_le_path_end by (auto simp add: pathstart_def pathfinish_def) have gamm_cont: "continuous_on {0..1} (\a. \ a \ i)" apply(rule continuous_on_inner) using gamma_smooth apply (simp add: piecewise_C1_differentiable_on_def) using continuous_on_const by auto then obtain c d where cd: "c \ d" "(\a. \ a \ i) ` {0..1} = {c..d}" by (meson continuous_image_closed_interval zero_le_one) then have subset_cd: "(\a. \ a \ i) ` {0..1} \ {c..d}" by auto have field_cont_on_path_cd: "continuous_on {c..d} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using field_cont_on_path cd by auto have path_vector_deriv_line_integrals: "\x\{0..1} - s. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x)) (at x)" using gamma_i_component_smooth and derivative_component_fun_component and vector_derivative_works by blast then have "\x\{0..1} - s. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x within ({0..1}))) (at x within ({0..1}))" using has_vector_derivative_at_within vector_derivative_at_within_ivl by fastforce then have has_int:"((\x. vector_derivative (\x. \ x \ i) (at x within {0..1}) *\<^sub>R (F ((\ x \ i) *\<^sub>R i + g (\ x \ i)) \ i)) has_integral integral {\ 0 \ i..\ 1 \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using has_integral_substitution_strong[OF gamma_differentiable(1) rel_simps(44) path_start_le_path_end' subset_cd field_cont_on_path_cd gamm_cont, of "(\x. vector_derivative (\x. \(x) \ i) (at x within ({0..1})))"] gamma_is_in_terms_of_i - by (auto simp only: has_field_derivative_iff_has_vector_derivative) + by (auto simp only: has_real_derivative_iff_has_vector_derivative) then have has_int':"((\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within ({0..1})))) has_integral integral {((pathstart \) \ i)..((pathfinish \) \ i)} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using gamma_is_in_terms_of_i i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto simp add: algebra_simps) have substitute: "integral ({((pathstart \) \ i)..((pathfinish \) \ i)}) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i)) = integral ({0..1}) (\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within ({0..1}))))" using gamma_is_in_terms_of_i integral_unique[OF has_int] i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto simp add: algebra_simps) have comp_in_eq_comp_out: "\x \ {0..1} - s. (vector_derivative (\x. \(x) \ i) (at x within {0..1})) = (vector_derivative \ (at x within {0..1})) \ i" proof fix x:: real assume ass:"x \ {0..1} -s" then have x_bounds:"x \ {0..1}" by auto have "\ differentiable at x" using ass gamma_differentiable by auto then have dotprod_in_is_out: "vector_derivative (\x. \(x) \ i) (at x) = (vector_derivative \ (at x)) \ i" using derivative_component_fun_component by force then have 0: "(vector_derivative \ (at x)) \ i = (vector_derivative \ (at x within {0..1})) \ i" proof - have "(\ has_vector_derivative vector_derivative \ (at x)) (at x)" using \\ differentiable at x\ vector_derivative_works by blast moreover have "0 \ x \ x \ 1" using x_bounds by presburger ultimately show ?thesis by (simp add: vector_derivative_at_within_ivl) qed have 1: "vector_derivative (\x. \(x) \ i) (at x) = vector_derivative (\x. \(x) \ i) (at x within {0..1})" using path_vector_deriv_line_integrals and vector_derivative_at_within_ivl and x_bounds by (metis ass atLeastAtMost_iff zero_less_one) show "vector_derivative (\x. \ x \ i) (at x within {0..1}) = vector_derivative \ (at x within {0..1}) \ i" using 0 and 1 and dotprod_in_is_out by auto qed show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using substitute and comp_in_eq_comp_out and negligible_finite Henstock_Kurzweil_Integration.integral_spike [of "s" "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by (metis (no_types, lifting) gamma_differentiable(1)) have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) has_integral integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using has_int' and comp_in_eq_comp_out and negligible_finite Henstock_Kurzweil_Integration.has_integral_spike [of "s" "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))" "integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)"] by (metis (no_types, lifting) gamma_differentiable(1)) then have "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) integrable_on {0..1}" using integrable_on_def by auto then show "line_integral_exists F {i} \" by (auto simp add: line_integral_exists_def) qed lemma line_integral_on_pair_path: fixes F::"('a::euclidean_space) \ ('a)" and g::"real \ 'a" and \::"(real \ 'a)" and i::'a assumes i_norm_1: "norm i = 1" and g_orthogonal_to_i: "\x. g(x) \ i = 0" and gamma_is_in_terms_of_i: "\ = (\x. f(x) *\<^sub>R i + g(f(x)))" and gamma_smooth: "\ C1_differentiable_on {0 .. 1}" and g_continuous_on_f: "continuous_on (f ` {0..1}) g" and path_start_le_path_end: "(pathstart \) \ i \ (pathfinish \) \ i" and field_i_comp_cont: "continuous_on (path_image \) (\x. F x \ i)" shows "(line_integral F {i} \) = integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))" proof (simp add: line_integral_def) have gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" using gamma_smooth C1_differentiable_on_eq by blast then have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" by auto have vec_deriv_i_comp_cont: "continuous_on {0..1} (\x. vector_derivative (\x. \ x \ i) (at x within {0..1}))" proof - have "continuous_on {0..1} (\x. vector_derivative (\x. \ x) (at x within {0..1}))" using gamma_smooth C1_differentiable_on_eq by (smt C1_differentiable_on_def atLeastAtMost_iff continuous_on_eq vector_derivative_at_within_ivl) then have deriv_comp_cont: "continuous_on {0..1} (\x. vector_derivative (\x. \ x) (at x within {0..1}) \ i)" by (simp add: continuous_intros) show ?thesis using derivative_component_fun_component_at_within[OF gamma_differentiable, of "i"] continuous_on_eq[OF deriv_comp_cont, of "(\x. vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by fastforce qed have field_cont_on_path: "continuous_on ((\x. \ x \ i) ` {0..1}) (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" proof - have 0: "(\x. \ x \ i) = f" proof fix x show "\ x \ i = f x" using g_orthogonal_to_i i_norm_1 by (simp only: gamma_is_in_terms_of_i real_inner_class.inner_add_left g_orthogonal_to_i inner_scaleR_left inner_same_Basis norm_eq_1) qed show ?thesis unfolding 0 apply (rule continuous_on_compose2 [of _ "(\x. F(x) \ i)" "f ` { 0..1}" "(\x. x *\<^sub>R i + g x)"] field_i_comp_cont g_continuous_on_f field_i_comp_cont continuous_intros)+ by (auto simp add: gamma_is_in_terms_of_i path_image_def) qed have path_vector_deriv_line_integrals: "\x\{0..1}. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x)) (at x)" using gamma_i_component_smooth and derivative_component_fun_component and vector_derivative_works by blast then have "\x\{0..1}. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x within {0..1})) (at x within {0..1})" using has_vector_derivative_at_within vector_derivative_at_within_ivl by fastforce then have substitute: "integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i)) = integral (cbox 0 1) (\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within {0..1})))" using gauge_integral_by_substitution [of "0" "1" "(\x. (\ x) \ i)" "(\x. vector_derivative (\x. \(x) \ i) (at x within {0..1}))" "(\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))"] and path_start_le_path_end and vec_deriv_i_comp_cont and field_cont_on_path and gamma_is_in_terms_of_i i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto) (*integration by substitution*) have comp_in_eq_comp_out: "\x \ {0..1}. (vector_derivative (\x. \(x) \ i) (at x within {0..1})) = (vector_derivative \ (at x within {0..1})) \ i" proof fix x:: real assume x_bounds: "x \ {0..1}" then have "\ differentiable at x" using gamma_differentiable by auto then have dotprod_in_is_out: "vector_derivative (\x. \(x) \ i) (at x) = (vector_derivative \ (at x)) \ i" using derivative_component_fun_component by force then have 0: "(vector_derivative \ (at x)) \ i = (vector_derivative \ (at x within {0..1})) \ i" using has_vector_derivative_at_within and x_bounds and vector_derivative_at_within_ivl by (smt atLeastAtMost_iff gamma_differentiable inner_commute vector_derivative_works) have 1: "vector_derivative (\x. \(x) \ i) (at x) = vector_derivative (\x. \(x) \ i) (at x within {0..1})" using path_vector_deriv_line_integrals and vector_derivative_at_within_ivl and x_bounds by fastforce show "vector_derivative (\x. \ x \ i) (at x within {0..1}) = vector_derivative \ (at x within {0..1}) \ i" using 0 and 1 and dotprod_in_is_out by auto qed show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using substitute and comp_in_eq_comp_out and Henstock_Kurzweil_Integration.integral_cong [of "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by auto qed lemma content_box_cases: "content (box a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_box_eq inner_diff) lemma content_box_cbox: shows "content (box a b) = content (cbox a b)" by (simp add: content_box_cases content_cbox_cases) lemma content_eq_0: "content (box a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_box_cases not_le intro: less_imp_le antisym eq_refl) lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_lt_nz: "0 < content (box a b) \ content (box a b) \ 0" using Paths.content_eq_0 zero_less_measure_iff by blast lemma content_subset: "cbox a b \ box c d \ content (cbox a b) \ content (box c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq) lemma sum_content_null: assumes "content (box a b) = 0" and "p tagged_division_of (box a b)" shows "sum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x k where xk: "y = (x, k)" using surj_pair[of y] by blast note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] from this(2) obtain c d where k: "k = cbox c d" by blast have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using content_subset[OF assm(1)[unfolded k]] content_pos_le[of "cbox c d"] unfolding assms(1) k by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed lemma has_integral_null [intro]: "content(box a b) = 0 \ (f has_integral 0) (box a b)" by (simp add: content_box_cbox content_eq_0_interior) lemma line_integral_distrib: assumes "line_integral_exists f basis g1" "line_integral_exists f basis g2" "valid_path g1" "valid_path g2" shows "line_integral f basis (g1 +++ g2) = line_integral f basis g1 + line_integral f basis g2" "line_integral_exists f basis (g1 +++ g2)" proof - obtain s1 s2 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) obtain i1 i2 where 1: "((\x. \b\basis. f (g1 x) \ b * (vector_derivative g1 (at x within {0..1}) \ b)) has_integral i1) {0..1}" and 2: "((\x. \b\basis. f (g2 x) \ b * (vector_derivative g2 (at x within {0..1}) \ b)) has_integral i2) {0..1}" using assms by (auto simp: line_integral_exists_def) have i1: "((\x. 2 * (\b\basis. f (g1 (2 * x)) \ b * (vector_derivative g1 (at (2 * x) within {0..1}) \ b))) has_integral i1) {0..1/2}" and i2: "((\x. 2 * (\b\basis. f (g2 (2 * x - 1)) \ b * (vector_derivative g2 (at ((2 * x) - 1) within {0..1}) \ b))) has_integral i2) {1/2..1}" using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z*2) within {0..1})" for z proof - have i:"\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" for z proof- have g1_at_z:"\0 \ z; z*2 < 1; z*2 \ s1\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z*2))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g1(2*x))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have z_ge: "z\ 1" by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" using Derivative.vector_derivative_at_within_ivl[OF g1_at_z[OF ass] ass(1) z_ge] by auto qed assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have "(g1 has_vector_derivative ((vector_derivative g1 (at (z*2))))) (at (z*2))" using s1 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g1 (at (z*2) within {0..1})) = (vector_derivative g1 (at (z*2)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2) within {0..1})" using i[OF ass] ii by auto qed have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1) within {0..1})" for z proof - have i:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" for z proof- have g2_at_z:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g2 (2*x - 1))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have z_le: "z\ 1" by auto have z_ge: "0 \ z" using ass by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" using Derivative.vector_derivative_at_within_ivl[OF g2_at_z[OF ass] z_ge z_le] by auto qed assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have "(g2 has_vector_derivative ((vector_derivative g2 (at (z*2 - 1))))) (at (z*2 - 1))" using s2 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g2 (at (z*2 - 1) within {0..1})) = (vector_derivative g2 (at (z*2 - 1)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1) within {0..1})" using i[OF ass] ii by auto qed have lem1: "((\x. \b\basis. f ((g1+++g2) x) \ b * (vector_derivative (g1+++g2) (at x within {0..1}) \ b)) has_integral i1) {0..1/2}" apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) using s1 apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) by (simp add: sum_distrib_left) moreover have lem2: "((\x. \b\basis. f ((g1+++g2) x) \ b * (vector_derivative (g1+++g2) (at x within {0..1}) \ b)) has_integral i2) {1/2..1}" apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) using s2 apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) by (simp add: sum_distrib_left) ultimately show "line_integral f basis (g1 +++ g2) = line_integral f basis g1 + line_integral f basis g2" apply (simp add: line_integral_def) apply (rule integral_unique [OF has_integral_combine [where c = "1/2"]]) using 1 2 integral_unique apply auto done show "line_integral_exists f basis (g1 +++ g2)" proof (simp add: line_integral_exists_def integrable_on_def) have "(1::real) \ 1 * 2 \ (0::real) \ 1 / 2" by simp then show "\r. ((\r. \a\basis. f ((g1 +++ g2) r) \ a * (vector_derivative (g1 +++ g2) (at r within {0..1}) \ a)) has_integral r) {0..1}" using has_integral_combine [where c = "1/2"] 1 2 divide_le_eq_numeral1(1) lem1 lem2 by blast qed qed lemma line_integral_exists_joinD1: assumes "line_integral_exists f basis (g1 +++ g2)" "valid_path g1" shows "line_integral_exists f basis g1" proof - obtain s1 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. \b\basis. f ((g1 +++ g2) (x/2)) \ b * (vector_derivative (g1 +++ g2) (at (x/2) within {0..1}) \ b)) integrable_on {0..1}" using assms apply (auto simp: line_integral_exists_def) apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) done then have *:"(\x. \b\basis. ((f ((g1 +++ g2) (x/2)) \ b) / 2)* (vector_derivative (g1 +++ g2) (at (x/2) within {0..1}) \ b)) integrable_on {0..1}" by (auto simp: Groups_Big.sum_distrib_left dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z*2) within {0..1})" for z proof - have i:"\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" for z proof- have g1_at_z:"\0 \ z; z*2 < 1; z*2 \ s1\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z*2))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g1(2*x))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have z_ge: "z\ 1" by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" using Derivative.vector_derivative_at_within_ivl[OF g1_at_z[OF ass] ass(1) z_ge] by auto qed assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have "(g1 has_vector_derivative ((vector_derivative g1 (at (z*2))))) (at (z*2))" using s1 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g1 (at (z*2) within {0..1})) = (vector_derivative g1 (at (z*2)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2) within {0..1})" using i[OF ass] ii by auto qed show ?thesis using s1 apply (auto simp: line_integral_exists_def) apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g1) done qed lemma line_integral_exists_joinD2: assumes "line_integral_exists f basis (g1 +++ g2)" "valid_path g2" shows "line_integral_exists f basis g2" proof - obtain s2 where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. \b\basis. f ((g1 +++ g2) (x/2 + 1/2)) \ b * (vector_derivative (g1 +++ g2) (at (x/2 + 1/2) within {0..1}) \ b)) integrable_on {0..1}" using assms apply (auto simp: line_integral_exists_def) apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) apply (simp add: image_affinity_atLeastAtMost_diff) done then have *:"(\x. \b\basis. ((f ((g1 +++ g2) (x/2 + 1/2)) \ b) / 2)* (vector_derivative (g1 +++ g2) (at (x/2 + 1/2) within {0..1}) \ b)) integrable_on {0..1}" by (auto simp: Groups_Big.sum_distrib_left dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1) within {0..1})" for z proof - have i:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" for z proof- have g2_at_z:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g2 (2*x - 1))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have z_le: "z\ 1" by auto have z_ge: "0 \ z" using ass by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" using Derivative.vector_derivative_at_within_ivl[OF g2_at_z[OF ass] z_ge z_le] by auto qed assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have "(g2 has_vector_derivative ((vector_derivative g2 (at (z*2 - 1))))) (at (z*2 - 1))" using s2 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g2 (at (z*2 - 1) within {0..1})) = (vector_derivative g2 (at (z*2 - 1)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1) within {0..1})" using i[OF ass] ii by auto qed show ?thesis using s2 apply (auto simp: line_integral_exists_def) apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g2) done qed lemma has_line_integral_on_reverse_path: assumes g: "valid_path g" and int: "((\x. \b\basis. F (g x) \ b * (vector_derivative g (at x within {0..1}) \ b)) has_integral c){0..1}" shows "((\x. \b\basis. F ((reversepath g) x) \ b * (vector_derivative (reversepath g) (at x within {0..1}) \ b)) has_integral -c){0..1}" proof - { fix s x assume xs: "g C1_differentiable_on ({0..1} - s)" "x \ (-) 1 ` s" "0 \ x" "x \ 1" have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - vector_derivative g (at (1 - x) within {0..1})" proof - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" using xs by (force simp: has_vector_derivative_def C1_differentiable_on_def) have "(g o (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" apply (rule vector_diff_chain_within) apply (intro vector_diff_chain_within derivative_eq_intros | simp)+ apply (rule has_vector_derivative_at_within [OF f']) done then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" by (simp add: o_def) show ?thesis using xs by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) qed } note * = this obtain S where "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using g by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) then show ?thesis using has_integral_affinity01 [OF int, where m= "-1" and c=1] unfolding reversepath_def by (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) (auto simp: * has_integral_neg Groups_Big.sum_negf) qed lemma line_integral_on_reverse_path: assumes "valid_path \" "line_integral_exists F basis \" shows "line_integral F basis \ = - (line_integral F basis (reversepath \))" "line_integral_exists F basis (reversepath \)" proof - obtain i where 0: "((\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) has_integral i){0..1}" using assms unfolding integrable_on_def line_integral_exists_def by auto then have 1: "((\x. \b\basis. F ((reversepath \) x) \ b * (vector_derivative (reversepath \) (at x within {0..1}) \ b)) has_integral -i){0..1}" using has_line_integral_on_reverse_path assms by auto then have rev_line_integral:"line_integral F basis (reversepath \) = -i" using line_integral_def Henstock_Kurzweil_Integration.integral_unique by (metis (no_types)) have line_integral: "line_integral F basis \ = i" using line_integral_def 0 Henstock_Kurzweil_Integration.integral_unique by blast show "line_integral F basis \ = - (line_integral F basis (reversepath \))" using line_integral rev_line_integral by auto show "line_integral_exists F basis (reversepath \)" using 1 line_integral_exists_def by auto qed lemma line_integral_exists_on_degenerate_path: assumes "finite basis" shows "line_integral_exists F basis (\x. c)" proof- have every_component_integrable: "\b\basis. (\x. F ((\x. c) x) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" proof fix b assume b_in_basis: "b \ basis" have cont_field_zero_one: "continuous_on {0..1} (\x. F ((\x. c) x) \ b)" using continuous_on_const by fastforce have cont_path_zero_one: "continuous_on {0..1} (\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)" proof - have "((vector_derivative (\x. c) (at x within {0..1})) \ b) = 0" if "x \ {0..1}" for x proof - have "vector_derivative (\x. c) (at x within {0..1}) = 0" using that gamma_deriv_at_within[of "0" "1"] differentiable_const vector_derivative_const_at by fastforce then show "vector_derivative (\x. c) (at x within {0..1}) \ b = 0" by auto qed then show "continuous_on {0..1} (\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)" using continuous_on_const[of "{0..1}" "0"] continuous_on_eq[of "{0..1}" "\x. 0" "(\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)"] by auto qed show "(\x. F (c) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" using cont_field_zero_one cont_path_zero_one continuous_on_mult integrable_continuous_real by blast qed have integrable_sum': "\t f s. finite t \ \a\t. f a integrable_on s \ (\x. \a\t. f a x) integrable_on s" using integrable_sum by metis have field_integrable_on_basis: "(\x. \b\basis. F (c) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" using integrable_sum'[OF assms(1) every_component_integrable] by auto then show ?thesis using line_integral_exists_def by auto qed lemma degenerate_path_is_valid_path: "valid_path (\x. c)" by (auto simp add: valid_path_def piecewise_C1_differentiable_on_def continuous_on_const) lemma line_integral_degenerate_path: assumes "finite basis" shows "line_integral F basis (\x. c) = 0" proof (simp add: line_integral_def) have "((vector_derivative (\x. c) (at x within {0..1})) \ b) = 0" if "x \ {0..1}" for x b proof - have "vector_derivative (\x. c) (at x within {0..1}) = 0" using that gamma_deriv_at_within[of "0" "1"] differentiable_const vector_derivative_const_at by fastforce then show "vector_derivative (\x. c) (at x within {0..1}) \ b = 0" by auto qed then have 0: "\x. x \ {0..1} \ (\x. \b\basis. F c \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) x = (\x. 0) x" by auto then show "integral {0..1} (\x. \b\basis. F c \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) = 0" using integral_cong[of "{0..1}", OF 0] integral_0 by auto qed definition point_path where "point_path \ \ \c. \ = (\x. c)" lemma line_integral_point_path: assumes "point_path \" assumes "finite basis" shows "line_integral F basis \ = 0" using assms(1) point_path_def line_integral_degenerate_path[OF assms(2)] by force lemma line_integral_exists_point_path: assumes "finite basis" "point_path \" shows "line_integral_exists F basis \" using assms apply(simp add: point_path_def) using line_integral_exists_on_degenerate_path by auto lemma line_integral_exists_subpath: assumes f: "line_integral_exists f basis g" and g: "valid_path g" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(line_integral_exists f basis (subpath u v g))" proof (cases "v=u") case tr: True have zero: "(\b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) = 0" if "x \ {0..1}" for x::real proof - have "(vector_derivative (\x. g u) (at x within {0..1})) = 0" using Deriv.has_vector_derivative_const that Derivative.vector_derivative_at_within_ivl by fastforce then show "(\b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) = 0" by auto qed then have "((\x. \b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) then show ?thesis using f tr by (auto simp add: line_integral_def line_integral_exists_def subpath_def) next case False obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast have *: "((\x. \b\basis. f (g ((v - u) * x + u)) \ b * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ b)) has_integral (1 / (v - u)) * integral {u..v} (\x. \b\basis. f (g (x)) \ b * (vector_derivative g (at x within {0..1}) \ b))) {0..1}" using f uv apply (simp add: line_integral_exists_def subpath_def) apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) apply (simp_all add: has_integral_integral) apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps False) done have vd:"\x. x \ {0..1} \ x \ (\t. (v-u) *\<^sub>R t + u) -` s \ vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within {0..1})" using test2[OF s fs uv] by auto have arg:"\x. (\n\basis. (v - u) * (f (g ((v - u) * x + u)) \ n) * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ n)) = (\b\basis. f (g ((v - u) * x + u)) \ b * (v - u) * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ b))" by (simp add: mult.commute) have"((\x. \b\basis. f (g ((v - u) * x + u)) \ b * (vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) \ b)) has_integral (integral {u..v} (\x. \b\basis. f (g (x)) \ b * (vector_derivative g (at x within {0..1}) \ b)))) {0..1}" apply (cut_tac Henstock_Kurzweil_Integration.has_integral_mult_right [OF *, where c = "v-u"]) using fs assms apply (simp add: False subpath_def line_integral_exists_def) apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) apply (auto simp: inj_on_def False vd finite_vimageI scaleR_conv_of_real Groups_Big.sum_distrib_left mult.assoc[symmetric] arg) done then show "(line_integral_exists f basis (subpath u v g))" by(auto simp add: line_integral_exists_def subpath_def integrable_on_def) qed (*This should have everything that has to do with onecubes*) type_synonym path = "real \ (real * real)" type_synonym one_cube = "(real \ (real * real))" type_synonym one_chain = "(int * path) set" type_synonym two_cube = "(real * real) \ (real * real)" type_synonym two_chain = "two_cube set" definition one_chain_line_integral :: "((real * real) \ (real * real)) => ((real*real) set) \ one_chain \ real" where "one_chain_line_integral F b C \ (\(k,g)\C. k * (line_integral F b g))" definition boundary_chain where "boundary_chain s \ (\(k, \) \ s. k = 1 \ k = -1)" fun coeff_cube_to_path::"(int * one_cube) \ path" where "coeff_cube_to_path (k, \) = (if k = 1 then \ else (reversepath \))" fun rec_join :: "(int*path) list \ path" where "rec_join [] = (\x.0)" | "rec_join [oneC] = coeff_cube_to_path oneC" | "rec_join (oneC # xs) = coeff_cube_to_path oneC +++ (rec_join xs)" fun valid_chain_list where "valid_chain_list [] = True" | "valid_chain_list [oneC] = True" | "valid_chain_list (oneC#l) = (pathfinish (coeff_cube_to_path (oneC)) = pathstart (rec_join l) \ valid_chain_list l)" lemma joined_is_valid: assumes boundary_chain: "boundary_chain (set l)" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and valid_chain_list_ass: "valid_chain_list l" shows "valid_path (rec_join l)" using assms proof(induction l) case Nil then show ?case using C1_differentiable_imp_piecewise[OF C1_differentiable_on_const[of "0" "{0..1}"]] by (auto simp add: valid_path_def) next case (Cons a l) have *: "valid_path (rec_join ((k::int, \) # l))" if "boundary_chain (set (l))" "(\k' \'. (k', \') \ set l \ valid_path \')" "valid_chain_list l" "valid_path (rec_join l) " "(\k' \'. (k', \') \ set ((k, \) # l) \ valid_path \')" "valid_chain_list ((k, \) # l)" "boundary_chain (set ((k,\) # l))" for k \ l proof (cases "l = []") case True with that show "valid_path (rec_join ((k, \) # l))" by auto next case False then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) then show "valid_path (rec_join ((k, \) # l))" proof (simp, intro conjI impI) assume k_eq_1: "k = 1" have 0:"valid_path \" using that by auto have 1:"pathfinish \ = pathstart (rec_join (h#l'))" using that(6) k_eq_1 l_nempty by auto show "valid_path (\ +++ rec_join (h#l'))" using 0 1 valid_path_join that(4) l_nempty by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using that(7) by (auto simp add: boundary_chain_def) have "valid_path \" using that by auto then have 0: "valid_path (reversepath \)" using valid_path_imp_reverse by auto have 1: "pathfinish (reversepath \) = pathstart (rec_join (h#l'))" using that(6) k_eq_neg_1 l_nempty by auto show "valid_path ((reversepath \) +++ rec_join (h#l'))" using 0 1 valid_path_join that(4) l_nempty by blast qed qed have "\ps. valid_chain_list ps \ (\i f p psa. ps = (i, f) # p # psa \ ((i = 1 \ pathfinish f \ pathstart (rec_join (p # psa)) \ i \ 1 \ pathfinish (reversepath f) \ pathstart (rec_join (p # psa))) \ \ valid_chain_list (p # psa)))" by (smt coeff_cube_to_path.elims valid_chain_list.elims(3)) moreover have "boundary_chain (set l)" by (meson Cons.prems(1) boundary_chain_def set_subset_Cons subset_eq) ultimately show ?case using * Cons by (metis (no_types) list.set_intros(2) prod.collapse valid_chain_list.simps(3)) qed lemma pathstart_rec_join_1: "pathstart (rec_join ((1, \) # l)) = pathstart \" proof (cases "l = []") case True then show "pathstart (rec_join ((1, \) # l)) = pathstart \" by simp next case False then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "pathstart (rec_join ((1, \) # l)) = pathstart \" by simp qed lemma pathstart_rec_join_2: "pathstart (rec_join ((-1, \) # l)) = pathstart (reversepath \)" proof cases assume "l = []" then show "pathstart (rec_join ((- 1, \) # l)) = pathstart (reversepath \)" by simp next assume "l \ [] " then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "pathstart (rec_join ((- 1, \) # l)) = pathstart (reversepath \)" by simp qed lemma pathstart_rec_join: "pathstart (rec_join ((1, \) # l)) = pathstart \" "pathstart (rec_join ((-1, \) # l)) = pathstart (reversepath \)" using pathstart_rec_join_1 pathstart_rec_join_2 by auto lemma line_integral_exists_on_rec_join: assumes boundary_chain: "boundary_chain (set l)" and valid_chain_list: "valid_chain_list l" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and line_integral_exists: "\(k, \) \ set l. line_integral_exists F basis \" shows "line_integral_exists F basis (rec_join l)" using assms proof (induction l) case Nil then show ?case proof (simp add: line_integral_exists_def) have "\x. (vector_derivative (\x. 0) (at x)) = 0" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x)" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x within {0..1})" using Derivative.vector_derivative_const_at by auto then have 0: "\x\{0..1}. (vector_derivative (\x. 0) (at x within{0..1})) = 0" by (simp add: gamma_deriv_at_within) have "(\x\{0..1}. (\b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) = 0)" by (simp add: 0) then have " ((\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) then show "(\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) integrable_on {0..1}" by auto qed next case (Cons a l) obtain k \ where aeq: "a = (k,\)" by force show ?case unfolding aeq proof cases assume l_empty: "l = []" then show "line_integral_exists F basis (rec_join ((k, \) # l))" using Cons.prems aeq line_integral_on_reverse_path(2) by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral_exists F basis (rec_join ((k, \) # l))" proof (auto simp add: l_nempty) assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis \" using Cons.prems(4) aeq by auto have 1: "line_integral_exists F basis (rec_join l)" by (metis (mono_tags) Cons boundary_chain_def list.set_intros(2) valid_chain_list.elims(3) valid_chain_list.simps(3)) have 2: "valid_path \" using Cons aeq by auto have 3:"valid_path (rec_join l)" by (metis (no_types) Cons.prems boundary_chain_def joined_is_valid l_nempty set_subset_Cons subsetCE valid_chain_list.simps(3)) show "line_integral_exists F basis (\ +++ rec_join (h#l'))" using line_integral_distrib(2)[OF 0 1 2 3] assms l_nempty by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using Cons aeq by (simp add: boundary_chain_def) have gamma_valid: "valid_path \" using Cons aeq by auto then have 2: "valid_path (reversepath \)" using valid_path_imp_reverse by auto have "line_integral_exists F basis \" using Cons aeq by auto then have 0: "line_integral_exists F basis (reversepath \)" using line_integral_on_reverse_path(2) gamma_valid by auto have 1: "line_integral_exists F basis (rec_join l)" using Cons aeq by (metis (mono_tags) boundary_chain_def insert_iff list.set(2) list.set_intros(2) valid_chain_list.elims(3) valid_chain_list.simps(3)) have 3:"valid_path (rec_join l)" by (metis (no_types) Cons.prems(1) Cons.prems(2) Cons.prems(3) boundary_chain_def joined_is_valid l_nempty set_subset_Cons subsetCE valid_chain_list.simps(3)) show "line_integral_exists F basis ((reversepath \) +++ rec_join (h#l'))" using line_integral_distrib(2)[OF 0 1 2 3] assms l_nempty by auto qed qed qed lemma line_integral_exists_rec_join_cons: assumes "line_integral_exists F basis (rec_join ((1,\) # l))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral_exists F basis (\ +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral_exists F basis (\ +++ rec_join l)" using assms(2) line_integral_distrib(2)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(3)], of "0"] using degenerate_path_is_valid_path by (fastforce simp add: l_empty) next assume "l \ []" then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "line_integral_exists F basis (\ +++ rec_join l)" using assms by auto qed lemma line_integral_exists_rec_join_cons_2: assumes "line_integral_exists F basis (rec_join ((-1,\) # l))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral_exists F basis ((reversepath \) +++ rec_join l)" using assms(2) line_integral_distrib(2)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(3)], of "0"] using degenerate_path_is_valid_path by (auto simp add: l_empty) next assume "l \ []" then obtain h l' where "l = h#l'" by (meson rec_join.elims) with assms show "line_integral_exists F basis ((reversepath \) +++ rec_join l)" using assms by auto qed lemma line_integral_exists_on_rec_join': assumes boundary_chain: "boundary_chain (set l)" and valid_chain_list: "valid_chain_list l" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and line_integral_exists: "line_integral_exists F basis (rec_join l)" and finite_basis: "finite basis" shows "\(k, \) \ set l. line_integral_exists F basis \" using assms proof (induction l) case Nil show ?case by (simp add: line_integral_exists_def) next case ass: (Cons a l) obtain k \ where k_gamma:"a = (k,\)" by fastforce show ?case apply (auto simp add: k_gamma) proof - show "line_integral_exists F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis (\ +++ (rec_join l))" using line_integral_exists_rec_join_cons k_eq_1 k_gamma ass(4) ass(5) ass(6) by auto have 2: "valid_path \" using ass k_gamma by auto show "line_integral_exists F basis \" using line_integral_exists_joinD1[OF 0 2] by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass k_gamma by (simp add: boundary_chain_def) have 0: "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" using line_integral_exists_rec_join_cons_2[OF ] k_eq_neg_1 k_gamma ass(4) ass(5) ass(6) by fastforce have gamma_valid: "valid_path \" using ass k_gamma by auto then have 2: "valid_path (reversepath \)" using valid_path_imp_reverse by auto have "line_integral_exists F basis (reversepath \)" using line_integral_exists_joinD1[OF 0 2] by auto then show "line_integral_exists F basis (\)" using line_integral_on_reverse_path(2)[OF 2] reversepath_reversepath by auto qed next have 0:"boundary_chain (set l)" using ass(2) by (auto simp add: boundary_chain_def) have 1:"valid_chain_list l" using ass(3) apply (auto simp add: k_gamma) by (metis valid_chain_list.elims(3) valid_chain_list.simps(3)) have 2:"(\k \. (k, \) \ set (l) \ valid_path \)" using ass(4) by auto have 3: "valid_path (rec_join l)" using joined_is_valid[OF 0] 1 2 by auto have 4: "line_integral_exists F basis (rec_join l)" proof(cases "k = 1") assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis (\ +++ (rec_join l))" using line_integral_exists_rec_join_cons k_eq_1 k_gamma ass(4) ass(5) ass(6) by auto show "line_integral_exists F basis (rec_join l)" using line_integral_exists_joinD2[OF 0 3] by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass k_gamma by (simp add: boundary_chain_def) have 0: "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" using line_integral_exists_rec_join_cons_2[OF ] k_eq_neg_1 k_gamma ass(4) ass(5) ass(6) by fastforce show "line_integral_exists F basis (rec_join l)" using line_integral_exists_joinD2[OF 0 3] by auto qed show "\a b. (a, b) \ set l \ line_integral_exists F basis b" using 0 1 2 3 4 ass(1)[OF 0 1 2] ass(6) by fastforce qed qed inductive chain_subdiv_path where I: "chain_subdiv_path \ (set l)" if "distinct l" "rec_join l = \" "valid_chain_list l" lemma valid_path_equiv_valid_chain_list: assumes path_eq_chain: "chain_subdiv_path \ one_chain" and "boundary_chain one_chain" "\(k, \) \ one_chain. valid_path \" shows "valid_path \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show "valid_path \" using joined_is_valid assms l_props by blast qed lemma line_integral_rec_join_cons: assumes "line_integral_exists F basis \" "line_integral_exists F basis (rec_join ((l)))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" using assms line_integral_distrib(1)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(4)], of "0"] apply (auto simp add: l_empty) using degenerate_path_is_valid_path line_integral_degenerate_path by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" using assms by (auto simp add: l_nempty) qed lemma line_integral_rec_join_cons_2: assumes "line_integral_exists F basis \" "line_integral_exists F basis (rec_join ((l)))" "(\k' \'. (k', \') \ set ((-1,\) # l) \ valid_path \')" "finite basis" shows "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" proof cases assume l_empty: "l = []" have 0: "line_integral_exists F basis (reversepath \)" using assms line_integral_on_reverse_path(2) by fastforce have 1: "valid_path (reversepath \)" using assms by fastforce show "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" using assms line_integral_distrib(1)[OF 0 line_integral_exists_on_degenerate_path[OF assms(4)], of "0"] apply (auto simp add: l_empty) using degenerate_path_is_valid_path line_integral_degenerate_path by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" using assms by (auto simp add: l_nempty) qed lemma one_chain_line_integral_rec_join: assumes l_props: "set l = one_chain" "distinct l" "valid_chain_list l" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\(k::int, \) \ one_chain. line_integral_exists F basis \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" proof - have 0: "sum_list (map (\((k::int), \). (k::int) * (line_integral F basis \)) l) = one_chain_line_integral F basis one_chain" unfolding one_chain_line_integral_def using l_props Groups_List.comm_monoid_add_class.sum.distinct_set_conv_list[OF l_props(2), of "(\(k, \). (k::int) * (line_integral F basis \))"] by auto have "valid_chain_list l \ boundary_chain (set l) \ (\(k::int, \) \ set l. line_integral_exists F basis \) \ (\(k::int, \) \ set l. valid_path \) \ line_integral F basis (rec_join l) = sum_list (map (\(k::int, \). k * (line_integral F basis \)) l)" proof (induction l) case Nil show ?case unfolding line_integral_def boundary_chain_def apply (auto) proof have "\x. (vector_derivative (\x. 0) (at x)) = 0" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x)" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x within {0..1})" using Derivative.vector_derivative_const_at by auto then have 0: "\x\{0..1}. (vector_derivative (\x. 0) (at x within{0..1})) = 0" by (metis (no_types) box_real(2) vector_derivative_within_cbox zero_less_one) have "(\x\{0..1}. (\b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) = 0)" by (simp add: 0) then show "((\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) qed next case ass: (Cons a l) obtain k::"int" and \::"one_cube" where props: "a = (k,\)" proof let ?k2 = "fst a" let ?\2 = "snd a" show "a = (?k2, ?\2)" by auto qed have "line_integral_exists F basis (rec_join (a # l))" using line_integral_exists_on_rec_join[OF ass(3) ass(2)] ass(5) ass(4) by blast have "boundary_chain (set l)" by (meson ass.prems(2) boundary_chain_def list.set_intros(2)) have val_l: "\f i. (i,f) \ set l \ valid_path f" using ass.prems(4) by fastforce have vcl_l: "valid_chain_list l" by (metis (no_types) ass.prems(1) valid_chain_list.elims(3) valid_chain_list.simps(3)) have line_integral_exists_on_joined: "line_integral_exists F basis (rec_join l)" by (metis \boundary_chain (set l)\ \line_integral_exists F basis (rec_join (a # l))\ emptyE val_l vcl_l joined_is_valid line_integral_exists_joinD2 line_integral_exists_on_rec_join list.set(1) neq_Nil_conv rec_join.simps(3)) have "valid_path (rec_join (a # l))" using joined_is_valid ass(5) ass(3) ass(2) by blast then have joined_is_valid: "valid_path (rec_join l)" using \boundary_chain (set l)\ val_l vcl_l joined_is_valid by blast show ?case proof (clarsimp, cases) assume k_eq_1: "(k::int) = 1" have line_integral_exists_on_gamma: "line_integral_exists F basis \" using ass props by auto have gamma_is_valid: "valid_path \" using ass props by auto have line_int_rw: "line_integral F basis (rec_join ((k, \) # l)) = line_integral F basis (\ +++ rec_join l)" proof - have gam_int: "line_integral_exists F basis \" using ass props by auto have rec_join_int: "line_integral_exists F basis (rec_join l)" using line_integral_exists_on_rec_join using line_integral_exists_on_joined by blast show ?thesis using line_integral_rec_join_cons[OF gam_int rec_join_int] ass k_eq_1 finite_basis props by force qed show "line_integral F basis (rec_join (a # l)) = (case a of (x, \) \ real_of_int x * line_integral F basis \) + (\(x, \)\l. real_of_int x * line_integral F basis \)" apply (simp add: props line_int_rw) using line_integral_distrib[OF line_integral_exists_on_gamma line_integral_exists_on_joined gamma_is_valid joined_is_valid] ass k_eq_1 vcl_l by (auto simp: boundary_chain_def props) next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass props by (auto simp add: boundary_chain_def) have line_integral_exists_on_gamma: "line_integral_exists F basis (reversepath \)" using line_integral_on_reverse_path ass props by auto have gamma_is_valid: "valid_path (reversepath \)" using valid_path_imp_reverse ass props by auto have line_int_rw: "line_integral F basis (rec_join ((k, \) # l)) = line_integral F basis ((reversepath \) +++ rec_join l)" proof - have gam_int: "line_integral_exists F basis \" using ass props by auto have rec_join_int: "line_integral_exists F basis (rec_join l)" using line_integral_exists_on_rec_join using line_integral_exists_on_joined by blast show ?thesis using line_integral_rec_join_cons_2[OF gam_int rec_join_int] using ass k_eq_neg_1 using finite_basis props by blast qed show "line_integral F basis (rec_join (a # l)) = (case a of (x, \) \ real_of_int x * line_integral F basis \) + (\(x, \)\l. real_of_int x * line_integral F basis \)" apply (simp add: props line_int_rw) using line_integral_distrib[OF line_integral_exists_on_gamma line_integral_exists_on_joined gamma_is_valid joined_is_valid] props ass line_integral_on_reverse_path(1)[of "\" "F" "basis"] k_eq_neg_1 using \boundary_chain (set l)\ vcl_l by auto qed qed then have 1:"line_integral F basis (rec_join l) = sum_list (map (\(k::int, \). k * (line_integral F basis \)) l)" using l_props assms by auto then show ?thesis using 0 1 by auto qed lemma line_integral_on_path_eq_line_integral_on_equiv_chain: assumes path_eq_chain: "chain_subdiv_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\(k::int, \) \ one_chain. line_integral_exists F basis \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain = line_integral F basis \" "line_integral_exists F basis \" "valid_path \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show "line_integral_exists F basis \" using line_integral_exists_on_rec_join assms l_props by blast show "valid_path \" using joined_is_valid assms l_props by blast have "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" using one_chain_line_integral_rec_join l_props assms by auto then show "one_chain_line_integral F basis one_chain = line_integral F basis \" using l_props by auto qed lemma line_integral_on_path_eq_line_integral_on_equiv_chain': assumes path_eq_chain: "chain_subdiv_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "line_integral_exists F basis \" and valid_path: "\(k, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain = line_integral F basis \" "\(k, \) \ one_chain. line_integral_exists F basis \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show 0: "\(k, \) \ one_chain. line_integral_exists F basis \" using line_integral_exists_on_rec_join' assms l_props by blast show "one_chain_line_integral F basis one_chain = line_integral F basis \" using line_integral_on_path_eq_line_integral_on_equiv_chain(1)[OF assms(1) assms(2) 0 assms(4) assms(5)] by auto qed definition chain_subdiv_chain where "chain_subdiv_chain one_chain1 subdiv \ \f. (\(f ` one_chain1)) = subdiv \ (\c\one_chain1. chain_subdiv_path (coeff_cube_to_path c) (f c)) \ pairwise (\ p p'. f p \ f p' = {}) one_chain1 \ (\x \ one_chain1. finite (f x))" lemma chain_subdiv_chain_character: shows "chain_subdiv_chain one_chain1 subdiv \ (\f. \(f ` one_chain1) = subdiv \ (\(k, \)\one_chain1. if k = 1 then chain_subdiv_path \ (f (k, \)) else chain_subdiv_path (reversepath \) (f (k, \))) \ (\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {}) \ (\x\one_chain1. finite (f x)))" unfolding chain_subdiv_chain_def by (safe; intro exI conjI iffI; fastforce simp add: pairwise_def) lemma chain_subdiv_chain_imp_finite_subdiv: assumes "finite one_chain1" "chain_subdiv_chain one_chain1 subdiv" shows "finite subdiv" using assms by(auto simp add: chain_subdiv_chain_def) lemma valid_subdiv_imp_valid_one_chain: assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and valid_path: "\(k, \) \ subdiv. valid_path \" shows "\(k, \) \ one_chain1. valid_path \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" using chain1_eq_chain2 chain_subdiv_chain_character by auto have "\ k \. (k,\) \ one_chain1\ valid_path \" proof- fix k \ assume ass: "(k,\) \ one_chain1" show "valid_path \" proof (cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_subdiv_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms apply (simp add: boundary_chain_def) by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast show ?thesis using valid_path_equiv_valid_chain_list[OF i ii iv] by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \) (f(k,\))" using f_props(2) ass using \k \ 1\ by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have "valid_path (reversepath \)" using valid_path_equiv_valid_chain_list[OF i ii iv] by auto then show ?thesis using reversepath_reversepath valid_path_imp_reverse by force qed qed then show valid_path1: "\(k, \) \ one_chain1. valid_path \" by auto qed lemma one_chain_line_integral_eq_line_integral_on_sudivision: assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain2: "\(k, \) \ subdiv. line_integral_exists F basis \" and valid_path: "\(k, \) \ subdiv. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ one_chain1. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 chain_subdiv_chain_character by (auto simp add: pairwise_def chain_subdiv_chain_def) then have 0: "one_chain_line_integral F basis subdiv = one_chain_line_integral F basis (\(f ` one_chain1))" by auto have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) apply (simp add: image_def) using f_props(1) by auto have "\ k \. (k,\) \ one_chain1\ line_integral_exists F basis \" proof- fix k \ assume ass: "(k,\) \ one_chain1" show "line_integral_exists F basis \" proof (cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_subdiv_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have iii:"\(k, \)\f (k, \). line_integral_exists F basis \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain(2)[OF i ii iii iv finite_basis] by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \) (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have iii:"\(k, \)\f (k, \). line_integral_exists F basis \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have x:"line_integral_exists F basis (reversepath \)" using line_integral_on_path_eq_line_integral_on_equiv_chain(2)[OF i ii iii iv finite_basis] by auto have "valid_path (reversepath \)" using line_integral_on_path_eq_line_integral_on_equiv_chain(3)[OF i ii iii iv finite_basis] by auto then show ?thesis using line_integral_on_reverse_path(2) reversepath_reversepath x by fastforce qed qed then show line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" by auto have 1: "one_chain_line_integral F basis (\(f ` one_chain1)) = one_chain_line_integral F basis one_chain1" proof - have 0:"one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" proof - have finite: "\chain \ (f ` one_chain1). finite chain" using f_props(1) finite_chain2 by (meson Sup_upper finite_subset) have disj: " \A\f ` one_chain1. \B\f ` one_chain1. A \ B \ A \ B = {}" by (metis (no_types, opaque_lifting) f_props(3) image_iff) show "one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" using Groups_Big.comm_monoid_add_class.sum.Union_disjoint[OF finite disj] one_chain_line_integral_def by auto qed have 1:"(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = one_chain_line_integral F basis one_chain1" proof - have "(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1. k*(line_integral F basis \))" proof - have i:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \))" proof - have "inj_on f (one_chain1 - {p. f p = {}})" unfolding inj_on_def using f_props(3) by blast then have 0: "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \)))" using Groups_Big.comm_monoid_add_class.sum.reindex by auto have "\ k \. (k, \) \ (one_chain1 - {p. f p = {}}) \ one_chain_line_integral F basis (f(k, \)) = k* (line_integral F basis \)" proof- fix k \ assume ass: "(k, \) \ (one_chain1 - {p. f p = {}})" have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed qed then have "(\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \))) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" by (auto intro!: Finite_Cartesian_Product.sum_cong_aux) then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" using 0 by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ (k, \) \ {(k',\'). k'* (line_integral F basis \') = 0}" proof- fix k::"int" fix \::"one_cube" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast have "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k::int, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed then show "(k, \) \ {(k'::int, \'). k' * line_integral F basis \' = 0}" using zero_line_integral by auto qed then have ii:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" proof - have "\one_chain. one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}})) \ one_chain_line_integral F basis one_chain = 0" proof - fix one_chain assume "one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}}))" then show "one_chain_line_integral F basis one_chain = 0" by (auto simp add: one_chain_line_integral_def) qed then have 0:"(\one_chain \ f ` (one_chain1) - (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" using comm_monoid_add_class.sum.neutral by auto then have "(\one_chain \ f ` (one_chain1). one_chain_line_integral F basis one_chain) - (\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" proof - have finte: "finite (f ` one_chain1)" using finite_chain1 by auto show ?thesis using Groups_Big.sum_diff[OF finte, of " (f ` (one_chain1 - {p. f p = {}}))" "one_chain_line_integral F basis"] 0 by auto qed then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ f (k, \) \ {chain. one_chain_line_integral F basis chain = 0}" proof- fix k::"int" fix \::"one_cube" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have "one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "f (k, \) \ {p'. (one_chain_line_integral F basis p' = 0)}" by auto qed then have iii:"(\(k::int,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \)) = (\(k::int,\)\one_chain1. k*(line_integral F basis \))" proof- have "\ k \. (k,\)\one_chain1 - (one_chain1 - {p. f p = {}}) \ k* (line_integral F basis \) = 0" proof- fix k \ assume ass: "(k,\)\one_chain1 - (one_chain1 - {p. f p = {}})" then have "f(k, \) = {}" by auto then have "one_chain_line_integral F basis (f(k, \)) = 0" by (auto simp add: one_chain_line_integral_def) then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast have "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed then show "k* (line_integral F basis \) = 0" using zero_line_integral by auto qed then have "\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k* (line_integral F basis \) = 0" by auto then have "(\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral [of "one_chain1 - (one_chain1 - {p. f p = {}})" "(\(k::int,\). k* (line_integral F basis \))"] by (simp add: split_beta) then have "(\(k::int,\)\one_chain1. k*(line_integral F basis \)) - (\(k::int,\)\ (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.sum_diff[OF finite_chain1] by (metis (no_types) Diff_subset \(\(k, \)\one_chain1 - (one_chain1 - {p. f p = {}}). k * line_integral F basis \) = 0\ \\f B. B \ one_chain1 \ sum f (one_chain1 - B) = sum f one_chain1 - sum f B\) then show ?thesis by auto qed show ?thesis using i ii iii by auto qed then show ?thesis using one_chain_line_integral_def by auto qed show ?thesis using 0 1 by auto qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using 0 1 by auto qed lemma one_chain_line_integral_eq_line_integral_on_sudivision': assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" and valid_path: "\(k, \) \ subdiv. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ subdiv. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 chain_subdiv_chain_character by (auto simp add: pairwise_def chain_subdiv_chain_def) have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) by blast have "\k \. (k, \) \ subdiv \ line_integral_exists F basis \" proof - fix k \ assume ass: "(k, \) \ subdiv" then obtain k' \' where kp_gammap: "(k',\') \ one_chain1" "(k,\) \ f(k',\')" using f_props(1) by fastforce show "line_integral_exists F basis \" proof (cases "k' = 1") assume k_eq_1: "k' = 1" then have i:"chain_subdiv_path \' (f(k',\'))" using f_props(2) kp_gammap ass by auto have ii:"boundary_chain (f(k',\'))" using f_props(1) ass assms kp_gammap by (meson UN_I boundary_chain_def) have iii:"line_integral_exists F basis \'" using assms kp_gammap by blast have "iv": " \(k, \)\f (k', \'). valid_path \" using f_props(1) ass assms kp_gammap by blast show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain'(2)[OF i ii iii iv finite_basis] kp_gammap by auto next assume "k' \ 1" then have k_eq_neg1: "k' = -1" using boundary_chain1 kp_gammap by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \') (f(k',\'))" using f_props(2) kp_gammap by auto have ii:"boundary_chain (f(k',\'))" using f_props(1) assms kp_gammap by (meson UN_I boundary_chain_def) have iii: " \(k, \)\f (k', \'). valid_path \" using f_props(1) ass assms kp_gammap by blast have iv: "valid_path (reversepath \')" using valid_path_equiv_valid_chain_list[OF i ii iii] by force have "line_integral_exists F basis \'" using assms kp_gammap by blast then have x: "line_integral_exists F basis (reversepath \')" using iv line_integral_on_reverse_path(2) valid_path_reversepath by fastforce show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain'(2)[OF i ii x iii finite_basis] kp_gammap by auto qed qed then show "\(k, \)\subdiv. line_integral_exists F basis \" by auto then show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using one_chain_line_integral_eq_line_integral_on_sudivision(1) assms by auto qed lemma line_integral_sum_gen: assumes finite_basis: "finite basis" and line_integral_exists: "line_integral_exists F basis1 \" "line_integral_exists F basis2 \" and basis_partition: "basis1 \ basis2 = basis" "basis1 \ basis2 = {}" shows "line_integral F basis \ = (line_integral F basis1 \) + (line_integral F basis2 \)" "line_integral_exists F basis \" apply (simp add: line_integral_def) proof - have 0: "integral {0..1} (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) = integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" using Henstock_Kurzweil_Integration.integral_add line_integral_exists by (auto simp add: line_integral_exists_def) have 1: "integral {0..1} (\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = integral {0..1} (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)))" by (metis (mono_tags, lifting) basis_partition finite_Un finite_basis sum.union_disjoint) show "integral {0..1} (\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" using 0 1 by auto have 2: "((\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) has_integral integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) {0..1}" using Henstock_Kurzweil_Integration.has_integral_add line_integral_exists has_integral_integral apply (auto simp add: line_integral_exists_def) by blast have 3: "(\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)))" by (metis (mono_tags, lifting) basis_partition finite_Un finite_basis sum.union_disjoint) show "line_integral_exists F basis \" apply(auto simp add: line_integral_exists_def has_integral_integral) using 2 3 using has_integral_integrable_integral by fastforce qed definition common_boundary_sudivision_exists where "common_boundary_sudivision_exists one_chain1 one_chain2 \ \subdiv. chain_subdiv_chain one_chain1 subdiv \ chain_subdiv_chain one_chain2 subdiv \ (\(k, \) \ subdiv. valid_path \)\ boundary_chain subdiv" lemma common_boundary_sudivision_commutative: "(common_boundary_sudivision_exists one_chain1 one_chain2) = (common_boundary_sudivision_exists one_chain2 one_chain1)" apply (simp add: common_boundary_sudivision_exists_def) by blast lemma common_subdivision_imp_eq_line_integral: assumes "(common_boundary_sudivision_exists one_chain1 one_chain2)" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" "finite one_chain1" "finite one_chain2" "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain2. line_integral_exists F basis \" proof - obtain subdiv where subdiv_props: "chain_subdiv_chain one_chain1 subdiv" "chain_subdiv_chain one_chain2 subdiv" "(\(k, \) \ subdiv. valid_path \)" "(boundary_chain subdiv)" using assms by (auto simp add: common_boundary_sudivision_exists_def) have i: "\(k, \)\subdiv. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision'(2)[OF subdiv_props(1) assms(2) subdiv_props(4) assms(4) subdiv_props(3) assms(5) assms(7)] by auto show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using one_chain_line_integral_eq_line_integral_on_sudivision'(1)[OF subdiv_props(1) assms(2) subdiv_props(4) assms(4) subdiv_props(3) assms(5) assms(7)] one_chain_line_integral_eq_line_integral_on_sudivision(1)[OF subdiv_props(2) assms(3) subdiv_props(4) i subdiv_props(3) assms(6) assms(7)] by auto show "\(k, \)\one_chain2. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision(2)[OF subdiv_props(2) assms(3) subdiv_props(4) i subdiv_props(3) assms(6) assms(7)] by auto qed definition common_sudiv_exists where "common_sudiv_exists one_chain1 one_chain2 \ \subdiv ps1 ps2. chain_subdiv_chain (one_chain1 - ps1) subdiv \ chain_subdiv_chain (one_chain2 - ps2) subdiv \ (\(k, \) \ subdiv. valid_path \) \ (boundary_chain subdiv) \ (\(k, \) \ ps1. point_path \) \ (\(k, \) \ ps2. point_path \)" lemma common_sudiv_exists_comm: shows "common_sudiv_exists C1 C2 = common_sudiv_exists C2 C1" by (auto simp add: common_sudiv_exists_def) lemma line_integral_degenerate_chain: assumes "(\(k, \) \ chain. point_path \)" assumes "finite basis" shows "one_chain_line_integral F basis chain = 0" proof (simp add: one_chain_line_integral_def) have "\(k,g)\chain. line_integral F basis g = 0" using assms line_integral_point_path by blast then have "\(k,g)\chain. real_of_int k * line_integral F basis g = 0" by auto then have "\p. p \ chain \ (case p of (i, f) \ real_of_int i * line_integral F basis f) = 0" by fastforce then show "(\x\chain. case x of (k, g) \ real_of_int k * line_integral F basis g) = 0" by simp qed lemma gen_common_subdiv_imp_common_subdiv: shows "(common_sudiv_exists one_chain1 one_chain2) = (\ps1 ps2. (common_boundary_sudivision_exists (one_chain1 - ps1) (one_chain2 - ps2)) \ (\(k, \)\ps1. point_path \) \ (\(k, \)\ps2. point_path \))" by (auto simp add: common_sudiv_exists_def common_boundary_sudivision_exists_def) lemma common_subdiv_imp_gen_common_subdiv: assumes "(common_boundary_sudivision_exists one_chain1 one_chain2)" shows "(common_sudiv_exists one_chain1 one_chain2)" using assms apply (auto simp add: common_sudiv_exists_def common_boundary_sudivision_exists_def) by (metis Diff_empty all_not_in_conv) lemma one_chain_line_integral_point_paths: assumes "finite one_chain" assumes "finite basis" assumes "(\(k, \)\ps. point_path \)" shows "one_chain_line_integral F basis (one_chain - ps) = one_chain_line_integral F basis (one_chain)" proof- have 0:"(\x\ps. case x of (k, g) \ (real_of_int k * line_integral F basis g) = 0)" using line_integral_point_path assms by force show "one_chain_line_integral F basis (one_chain - ps) = one_chain_line_integral F basis one_chain" unfolding one_chain_line_integral_def using 0 \finite one_chain\ by (force simp add: intro: comm_monoid_add_class.sum.mono_neutral_left) qed lemma boundary_chain_diff: assumes "boundary_chain one_chain" shows "boundary_chain (one_chain - s)" using assms by (auto simp add: boundary_chain_def) lemma gen_common_subdivision_imp_eq_line_integral: assumes "(common_sudiv_exists one_chain1 one_chain2)" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" "finite one_chain1" "finite one_chain2" "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain2. line_integral_exists F basis \" proof - obtain ps1 ps2 where gen_subdiv: "(common_boundary_sudivision_exists (one_chain1 - ps1) (one_chain2 - ps2))" "(\(k, \)\ps1. point_path \)" " (\(k, \)\ps2. point_path \)" using assms(1) gen_common_subdiv_imp_common_subdiv by blast show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using one_chain_line_integral_point_paths gen_common_subdiv_imp_common_subdiv assms(2-7) gen_subdiv common_subdivision_imp_eq_line_integral(1)[OF gen_subdiv(1) boundary_chain_diff[OF assms(2)] boundary_chain_diff[OF assms(3)]] by auto show "\(k, \)\one_chain2. line_integral_exists F basis \" proof- obtain subdiv where subdiv_props: "chain_subdiv_chain (one_chain1-ps1) subdiv" "chain_subdiv_chain (one_chain2-ps2) subdiv" "(\(k, \) \ subdiv. valid_path \)" "(boundary_chain subdiv)" using gen_subdiv(1) by (auto simp add: common_boundary_sudivision_exists_def) have "\(k, \)\subdiv. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision'(2)[OF subdiv_props(1) boundary_chain_diff[OF assms(2)] subdiv_props(4)] assms(4) subdiv_props(3) assms(5) assms(7) by blast then have i: "\(k, \)\one_chain2-ps2. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision(2)[OF subdiv_props(2) boundary_chain_diff[OF assms(3)] subdiv_props(4)] subdiv_props(3) assms(6) assms(7) by blast then show ?thesis using gen_subdiv(3) line_integral_exists_point_path[OF assms(7)] by blast qed qed lemma common_sudiv_exists_refl: assumes "common_sudiv_exists C1 C2" shows "common_sudiv_exists C2 C1" using assms apply(simp add: common_sudiv_exists_def) by auto lemma chain_subdiv_path_singleton: shows "chain_subdiv_path \ {(1,\)}" proof - have "rec_join [(1,\)] = \" by (simp add: joinpaths_def) then have "set [(1,\)] = {(1, \)}" "distinct [(1,\)]" "rec_join [(1,\)] = \" "valid_chain_list [(1,\)]" by auto then show ?thesis by (metis (no_types) chain_subdiv_path.intros) qed lemma chain_subdiv_path_singleton_reverse: shows "chain_subdiv_path (reversepath \) {(-1,\)}" proof - have "rec_join [(-1,\)] = reversepath \" by (simp add: joinpaths_def) then have "set [(-1,\)] = {(- 1, \)}" "distinct [(-1,\)]" "rec_join [(-1,\)] = reversepath \" "valid_chain_list [(-1,\)]" by auto then have "chain_subdiv_path (reversepath \) (set [(- 1, \)])" using chain_subdiv_path.intros by blast then show ?thesis by simp qed lemma chain_subdiv_chain_refl: assumes "boundary_chain C" shows "chain_subdiv_chain C C" using chain_subdiv_path_singleton chain_subdiv_path_singleton_reverse assms unfolding chain_subdiv_chain_def boundary_chain_def pairwise_def using case_prodI2 coeff_cube_to_path.simps by (rule_tac x="\x. {x}" in exI) auto (*path reparam_weaketrization*) definition reparam_weak where "reparam_weak \1 \2 \ \\. (\x\{0..1}. \1 x = (\2 o \) x) \ \ piecewise_C1_differentiable_on {0..1} \ \(0) = 0 \ \(1) = 1 \ \ ` {0..1} = {0..1}" definition reparam where "reparam \1 \2 \ \\. (\x\{0..1}. \1 x = (\2 o \) x) \ \ piecewise_C1_differentiable_on {0..1} \ \(0) = 0 \ \(1) = 1 \ bij_betw \ {0..1} {0..1} \ \ -` {0..1} \ {0..1} \ (\x\{0..1}. finite (\ -` {x}))" lemma reparam_weak_eq_refl: shows "reparam_weak \1 \1" unfolding reparam_weak_def apply (rule_tac x="id" in exI) by (auto simp add: id_def piecewise_C1_differentiable_on_def C1_differentiable_on_def continuous_on_id) lemma line_integral_exists_smooth_one_base: assumes "\ C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) "continuous_on (path_image \) (\x. F x \ b)" shows "line_integral_exists F {b} \" proof- have gamma2_differentiable: "(\x \ {0 .. 1}. \ differentiable at x)" using assms(1) by (auto simp add: valid_path_def C1_differentiable_on_eq) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1}. (\x. (\ x) \ b) differentiable at x)" by auto then have "(\x. (\ x) \ b) differentiable_on {0..1}" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on {0..1} (\x. (\ x) \ b)" using differentiable_imp_continuous_on by auto have gamma2_cont:"continuous_on {0..1} \" using assms(1) C1_differentiable_imp_continuous_on by (auto simp add: valid_path_def) have iii: "continuous_on {0..1} (\x. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" proof- have 0: "continuous_on {0..1} (\x. F (\ x) \ b)" using assms(2) continuous_on_compose[OF gamma2_cont] by (auto simp add: path_image_def) obtain D where D: "(\x\{0..1}. (\ has_vector_derivative D x) (at x)) \ continuous_on {0..1} D" using assms(1) by (auto simp add: C1_differentiable_on_def) then have *:"\x\{0..1}. vector_derivative \ (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on {0..1} (\x. vector_derivative \ (at x within{0..1}))" using continuous_on_eq D by force then have 1: "continuous_on {0..1} (\x. (vector_derivative \ (at x within{0..1})) \ b)" by(auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed then have "(\x. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) integrable_on {0..1}" using integrable_continuous_real by auto then show "line_integral_exists F {b} \" by(auto simp add: line_integral_exists_def) qed lemma contour_integral_primitive_lemma: fixes f :: "complex \ complex" and g :: "real \ complex" assumes "a \ b" and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) has_integral (f(g b) - f(g a))) {a..b}" proof - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {a..b} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "a < x" and b: "x < b" and xk: "x \ k" then have "g differentiable at x within {a..b}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } note * = this show ?thesis apply (rule fundamental_theorem_of_calculus_interior_strong) using k assms cfg * apply (auto simp: at_within_Icc_at) done qed lemma line_integral_primitive_lemma: (*Only for real normed fields, i.e. vectors with products*) fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a::{euclidean_space,real_normed_field}" and g :: "real \ 'a" assumes "\(a::'a). a \ s \ (f has_field_derivative (f' a)) (at a within s)" and "g piecewise_differentiable_on {0::real..1}" "\x. x \ {0..1} \ g x \ s" and "base_vec \ Basis" shows "((\x. ((f'(g x)) * (vector_derivative g (at x within {0..1}))) \ base_vec) has_integral (((f(g 1)) \ base_vec - (f(g 0)) \ base_vec))) {0..1}" proof - obtain k where k: "finite k" "\x\{0..1} - k. g differentiable (at x within {0..1})" and cg: "continuous_on {0..1} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {0..1} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "0 < x" and b: "x < 1" and xk: "x \ k" then have "g differentiable at x within {0..1}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {0..1})) (at x within {0..1})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. of_real u * vector_derivative g (at x within {0..1}))) (at x within {0..1})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {0..1})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {0..1})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {0..1})) (at x within {0..1})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } then have *: "\x. x\{0<..<1} - k \ ((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {0..1})) (at x within {0..1})" by auto have "((\x. ((f'(g x))) * ((vector_derivative g (at x within {0..1})))) has_integral (((f(g 1)) - (f(g 0))))) {0..1}" using fundamental_theorem_of_calculus_interior_strong[OF k(1) zero_le_one _ cfg] using k assms cfg * by (auto simp: at_within_Icc_at) then have "((\x. (((f'(g x))) * ((vector_derivative g (at x within {0..1})))) \ base_vec) has_integral (((f(g 1)) - (f(g 0)))) \ base_vec) {0..1}" using has_integral_componentwise_iff assms(4) by fastforce then show ?thesis using inner_mult_left' by (simp add: inner_mult_left' inner_diff_left) qed lemma reparam_eq_line_integrals: assumes reparam: "reparam \1 \2" and pw_smooth: "\2 piecewise_C1_differentiable_on {0..1}" and (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) cont: "continuous_on (path_image \2) (\x. F x \ b)" and line_integral_ex: "line_integral_exists F {b} \2" (*We need to remove this and work on special cases like conservative fields and field/line combinations that satisfy the improper integrals theorem assumptions*) shows "line_integral F {b} \1 = line_integral F {b} \2" "line_integral_exists F {b} \1" proof- obtain \ where phi: "(\x\{0..1}. \1 x = (\2 o \) x)"" \ piecewise_C1_differentiable_on {0..1}"" \(0) = 0"" \(1) = 1""bij_betw \ {0..1} {0..1}" "\ -` {0..1} \ {0..1}" "\x\{0..1}. finite (\ -` {x})" using reparam by (auto simp add: reparam_def) obtain s where s: "finite s" "\ C1_differentiable_on {0..1} - s" using phi by(auto simp add: reparam_def piecewise_C1_differentiable_on_def) let ?s = "s \ {0..1}" have s_inter: "finite ?s" "\ C1_differentiable_on {0..1} - ?s" using s apply blast by (metis Diff_Compl Diff_Diff_Int Diff_eq inf.commute s(2)) have cont_phi: "continuous_on {0..1} \" using phi by(auto simp add: reparam_def piecewise_C1_differentiable_on_imp_continuous_on) obtain s' D where s'_D: "finite s'" "(\x \ {0 .. 1} - s'. \2 differentiable at x)" "(\x\{0..1} - s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - s') D" using pw_smooth apply (auto simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) by (simp add: vector_derivative_works) let ?s' = "s' \ {0..1}" have gamma2_differentiable: "finite ?s'" "(\x \ {0 .. 1} - ?s'. \2 differentiable at x)" "(\x\{0..1} - ?s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - ?s') D" using s'_D apply blast using s'_D(2) apply auto[1] by (metis Diff_Int2 inf_top.left_neutral s'_D(3)) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1} - ?s'. (\x. (\2 x) \ b) differentiable at x)" using differentiable_inner by force then have "(\x. (\2 x) \ b) differentiable_on {0..1} - ?s'" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on ({0..1} - ?s') (\x. (\2 x) \ b)" using differentiable_imp_continuous_on by auto (**********Additional needed assumptions ************) have s_in01: "?s \ {0..1}" by auto have s'_in01: "?s' \ {0..1}" by auto have phi_backimg_s': "\ -` {0..1} \ {0..1}" using phi by auto have "inj_on \ {0..1}" using phi(5) by (auto simp add: bij_betw_def) have bij_phi: "bij_betw \ {0..1} {0..1}" using phi(5) by auto have finite_bck_img_single: "\x\{0..1}. finite (\ -` {x})" using phi by auto then have finite_bck_img_single_s': " \x\?s'. finite (\ -` {x})" by auto have gamma2_line_integrable: "(\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) integrable_on {0..1}" using line_integral_ex by (simp add: line_integral_exists_def) (****************************************************************) have finite_neg_img: "finite (\ -` ?s')" using finite_bck_img_single by (metis Int_iff finite_Int gamma2_differentiable(1) image_vimage_eq inf_img_fin_dom') have gamma2_cont:"continuous_on ({0..1} - ?s') \2" using gamma2_differentiable by (meson continuous_at_imp_continuous_on differentiable_imp_continuous_within) have iii: "continuous_on ({0..1} - ?s') (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" proof- have 0: "continuous_on ({0..1} - ?s') (\x. F (\2 x) \ b)" using cont continuous_on_compose[OF gamma2_cont] continuous_on_compose2 gamma2_cont unfolding path_image_def by fastforce have D: "(\x\{0..1} - ?s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - ?s') D" using gamma2_differentiable by auto then have *:"\x\{0..1} - ?s'. vector_derivative \2 (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on ({0..1} - ?s') (\x. vector_derivative \2 (at x within{0..1}))" using continuous_on_eq D by metis then have 1: "continuous_on ({0..1} - ?s') (\x. (vector_derivative \2 (at x within{0..1})) \ b)" by (auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed have iv: "\(0) \ \(1)" using phi(3) phi(4) by (simp add: reparam_def) have v: "\`{0..1} \ {0..1}" using phi by (auto simp add: reparam_def bij_betw_def) obtain D where D_props: "(\x\{0..1} - ?s. (\ has_vector_derivative D x) (at x))" using s by (auto simp add: C1_differentiable_on_def) then have "(\x. x \ ({0..1} - ?s) \ (\ has_vector_derivative D x) (at x within {0..1}))" using has_vector_derivative_at_within by blast then have vi: "(\x. x \ ({0..1} - ?s) \ (\ has_real_derivative D x) (at x within {0..1}))" - using has_field_derivative_iff_has_vector_derivative + using has_real_derivative_iff_has_vector_derivative by blast have a:"((\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) ({0..1})" proof- have a: "integral {\ 1..\ 0} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) = 0" using integral_singleton integral_empty iv by (simp add: phi(3) phi(4)) have b: "((\x. D x *\<^sub>R (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) - integral {\ 1..\ 0} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" apply(rule has_integral_substitution_general_'[OF s_inter(1) zero_le_one gamma2_differentiable(1) v gamma2_line_integrable iii cont_phi finite_bck_img_single_s']) proof- have "surj_on {0..1} \" using bij_phi by (metis (full_types) bij_betw_def image_subsetI rangeI) then show "surj_on ?s' \" using bij_phi s'_in01 by blast show "inj_on \ ({0..1} \ (?s \ \ -` ?s'))" proof- have i: "inj_on \ {0..1}" using bij_phi using bij_betw_def by blast have ii: "({0..1} \ (?s \ \ -` ?s')) = {0..1}" using phi_backimg_s' s_in01 s'_in01 by blast show ?thesis using i ii by auto qed show "\x. x \ {0..1} - ?s \ (\ has_real_derivative D x) (at x within {0..1})" using vi by blast qed show ?thesis using a b by auto qed then have b: "integral {0..1} (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" by auto have gamma2_vec_diffable: "\x::real. x \ {0..1} - ((\ -` ?s') \ ?s) \ ((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- fix x::real assume ass: "x \ {0..1} - ((\ -` ?s') \ ?s)" have zer_le_x_le_1:"0\ x \ x \ 1" using ass by simp show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- have **: "\2 differentiable at (\ x)" using gamma2_differentiable(2) ass v by blast have ***: " \ differentiable at x" using s ass by (auto simp add: C1_differentiable_on_eq) then show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" using differentiable_chain_at[OF *** **] by (auto simp add: vector_derivative_works) qed qed then have gamma2_vec_deriv_within: "\x::real. x \ {0..1} - ((\ -` ?s') \ ?s) \ vector_derivative (\2 \ \) (at x) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF gamma2_vec_diffable] by auto have "\x\{0..1} - ((\ -` ?s') \ ?s). D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" proof fix x::real assume ass: "x \ {0..1} -((\ -` ?s') \ ?s)" then have 0: "\ differentiable (at x)" using s by (auto simp add: C1_differentiable_on_def differentiable_def has_vector_derivative_def) obtain D2 where "(\ has_vector_derivative D2) (at x)" using D_props ass by blast have "\ x \ {0..1} - ?s'" using phi(5) ass by (metis Diff_Un Diff_iff Int_iff bij_betw_def image_eqI vimageI) then have 1: "\2 differentiable (at (\ x))" using gamma2_differentiable by auto have 3:" vector_derivative \2 (at (\ x)) = vector_derivative \2 (at (\ x) within {0..1})" proof- have *:"0\ \ x \ \ x \ 1" using phi(5) ass using \\ x \ {0..1} - s' \ {0..1}\ by auto then have **:"(\2 has_vector_derivative (vector_derivative \2 (at (\ x)))) (at (\ x))" using 1 vector_derivative_works by auto show ?thesis using * vector_derivative_at_within_ivl[OF **] by auto qed show "D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = vector_derivative (\2 \ \) (at x within {0..1}) \ b" using vector_derivative_chain_at[OF 0 1] apply (auto simp add: gamma2_vec_deriv_within[OF ass, symmetric] 3[symmetric]) using D_props ass vector_derivative_at by fastforce qed then have c:"\x. x\({0..1} -((\ -` ?s') \ ?s)) \ D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" by auto then have d: "integral ({0..1}) (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b))" proof- have "negligible ((\ -` ?s') \ ?s)" using finite_neg_img s(1) by auto then show ?thesis using c integral_spike by (metis(no_types,lifting)) qed have phi_in_int: "(\x. x \ {0..1} \ \ x \ {0..1})" using phi using v by blast then have e: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))){0..1}" proof- have "negligible ?s" using s_inter(1) by auto have 0: "negligible ((\ -` ?s') \ ?s)" using finite_neg_img s(1) by auto have c':"\ x\ {0..1} - ((\ -` ?s') \ ?s). D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" using c by blast have has_integral_spike_eq': "\s t f g y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using a has_integral_spike_eq'[OF 0 c'] by blast qed then have f: "((\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" proof- assume ass: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" have *:"\x\{0..1} - ( ((\ -` ?s') \ ?s) \ {0,1}). (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) x = (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) x" proof- have "\x\{0<..<1} - (\ -` ?s' \ ?s). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" proof have i: "open ({0<..<1} - ((\ -` ?s') \ ?s))" using open_diff s(1) open_greaterThanLessThan finite_neg_img by (simp add: open_diff) have ii: "\x\{0<..<1::real} - (\ -` ?s' \ ?s). (\2 \ \) x = \1 x" using phi(1) by auto fix x::real assume ass: " x \ {0<..<1::real} - ((\ -` ?s') \ ?s)" then have iii: "(\2 \ \ has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" by (metis (no_types) Diff_iff add.commute add_strict_mono ass atLeastAtMost_iff gamma2_vec_deriv_within gamma2_vec_diffable greaterThanLessThan_iff less_irrefl not_le) (*Most crucial but annoying step*) then have iv:"(\1 has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_derivative_transform_within_open i ii ass apply(auto simp add: has_vector_derivative_def) apply (meson ass has_derivative_transform_within_open i ii) apply (meson ass has_derivative_transform_within_open i ii) by (meson ass has_derivative_transform_within_open i ii) have v: "0 \ x" "x \ 1" using ass by auto have 0: "vector_derivative \1 (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iv v(1) v(2) zero_less_one] by force have 1: "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iii v(1) v(2) zero_less_one] by force then have "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative \1 (at x within {0..1})" using 0 1 by auto then show "vector_derivative (\2 \ \) (at x within {0..1}) \ b = vector_derivative \1 (at x within {0..1}) \ b" by auto qed then have i: "\x\{0..1} - ( ((\ -` ?s') \ ?s)\{0,1}). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" by auto have ii: "\x\{0..1} - (((\ -` ?s') \ ?s)\{0,1}). F (\1 x) \ b = F (\2 (\ x)) \ b" using phi(1) by auto show ?thesis using i ii by metis qed have **: "negligible ((\ -` ?s') \ ?s \ {0, 1})" using s(1) finite_neg_img by auto have has_integral_spike_eq': "\s t g f y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using has_integral_spike_eq'[OF ** *] ass by blast qed then show "line_integral_exists F {b} \1" using phi by(auto simp add: line_integral_exists_def) have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b))" using integral_unique[OF e] integral_unique[OF f] by metis moreover have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" using b d phi by (auto simp add:) ultimately show "line_integral F {b} \1 = line_integral F {b} \2" using phi by(auto simp add: line_integral_def) qed lemma reparam_weak_eq_line_integrals: assumes "reparam_weak \1 \2" "\2 C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) "continuous_on (path_image \2) (\x. F x \ b)" shows "line_integral F {b} \1 = line_integral F {b} \2" "line_integral_exists F {b} \1" proof- obtain \ where phi: "(\x\{0..1}. \1 x = (\2 o \) x)"" \ piecewise_C1_differentiable_on {0..1}"" \(0) = 0"" \(1) = 1"" \ ` {0..1} = {0..1}" using assms(1) by (auto simp add: reparam_weak_def) obtain s where s: "finite s" "\ C1_differentiable_on {0..1} - s" using phi by(auto simp add: reparam_weak_def piecewise_C1_differentiable_on_def) have cont_phi: "continuous_on {0..1} \" using phi by(auto simp add: reparam_weak_def piecewise_C1_differentiable_on_imp_continuous_on) have gamma2_differentiable: "(\x \ {0 .. 1}. \2 differentiable at x)" using assms(2) by (auto simp add: valid_path_def C1_differentiable_on_eq) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1}. (\x. (\2 x) \ b) differentiable at x)" by auto then have "(\x. (\2 x) \ b) differentiable_on {0..1}" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on {0..1} (\x. (\2 x) \ b)" using differentiable_imp_continuous_on by auto have gamma2_cont:"continuous_on {0..1} \2" using assms(2) C1_differentiable_imp_continuous_on by (auto simp add: valid_path_def) have iii: "continuous_on {0..1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" proof- have 0: "continuous_on {0..1} (\x. F (\2 x) \ b)" using assms(3) continuous_on_compose[OF gamma2_cont] by (auto simp add: path_image_def) obtain D where D: "(\x\{0..1}. (\2 has_vector_derivative D x) (at x)) \ continuous_on {0..1} D" using assms(2) by (auto simp add: C1_differentiable_on_def) then have *:"\x\{0..1}. vector_derivative \2 (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on {0..1} (\x. vector_derivative \2 (at x within{0..1}))" using continuous_on_eq D by force then have 1: "continuous_on {0..1} (\x. (vector_derivative \2 (at x within{0..1})) \ b)" by (auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed have iv: "\(0) \ \(1)" using phi(3) phi(4) by (simp add: reparam_weak_def) have v: "\`{0..1} \ {0..1}" using phi(5) by (simp add: reparam_weak_def) obtain D where D_props: "(\x\{0..1} - s. (\ has_vector_derivative D x) (at x))" using s by (auto simp add: C1_differentiable_on_def) then have "(\x. x \ ({0..1} -s) \ (\ has_vector_derivative D x) (at x within {0..1}))" using has_vector_derivative_at_within by blast then have vi: "(\x. x \ ({0..1} - s) \ (\ has_real_derivative D x) (at x within {0..1}))" - using has_field_derivative_iff_has_vector_derivative + using has_real_derivative_iff_has_vector_derivative by blast have a:"((\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" using has_integral_substitution_strong[OF s(1) zero_le_one iv v iii cont_phi vi] by simp then have b: "integral {0..1} (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" by auto have gamma2_vec_diffable: "\x::real. x \ {0..1} -s \ ((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- fix x::real assume ass: "x \ {0..1} -s" have zer_le_x_le_1:"0\ x \ x \ 1" using ass by auto show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- have **: "\2 differentiable at (\ x)" using phi gamma2_differentiable by (auto simp add: zer_le_x_le_1) have ***: " \ differentiable at x" using s ass by (auto simp add: C1_differentiable_on_eq) then show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" using differentiable_chain_at[OF *** **] by (auto simp add: vector_derivative_works) qed qed then have gamma2_vec_deriv_within: "\x::real. x \ {0..1} -s \ vector_derivative (\2 \ \) (at x) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF gamma2_vec_diffable] by auto have "\x\{0..1} - s. D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" proof fix x::real assume ass: "x \ {0..1} -s" then have 0: "\ differentiable (at x)" using s by (auto simp add: C1_differentiable_on_def differentiable_def has_vector_derivative_def) obtain D2 where "(\ has_vector_derivative D2) (at x)" using D_props ass by blast have "\ x \ {0..1}" using phi(5) ass by (auto simp add: reparam_weak_def) then have 1: "\2 differentiable (at (\ x))" using gamma2_differentiable by auto have 3:" vector_derivative \2 (at (\ x)) = vector_derivative \2 (at (\ x) within {0..1})" proof- have *:"0\ \ x \ \ x \ 1" using phi(5) ass by auto then have **:"(\2 has_vector_derivative (vector_derivative \2 (at (\ x)))) (at (\ x))" using 1 vector_derivative_works by auto show ?thesis using * vector_derivative_at_within_ivl[OF **] by auto qed show "D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = vector_derivative (\2 \ \) (at x within {0..1}) \ b" using vector_derivative_chain_at[OF 0 1] apply (auto simp add: gamma2_vec_deriv_within[OF ass, symmetric] 3[symmetric]) using D_props ass vector_derivative_at by fastforce qed then have c:"\x. x\({0..1} -s) \ D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" by auto then have d: "integral ({0..1}) (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b))" proof- have "negligible s" using s(1) by auto then show ?thesis using c integral_spike by (metis(no_types,lifting)) qed have phi_in_int: "(\x. x \ {0..1} \ \ x \ {0..1})" using phi by(auto simp add:) then have e: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))){0..1}" proof- have 0:"negligible s" using s(1) by auto have c':"\ x\ {0..1} -s. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" using c by auto have has_integral_spike_eq': "\s t f g y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using a has_integral_spike_eq'[OF 0 c'] by blast qed then have f: "((\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" proof- assume ass: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" have *:"\x\{0..1} - (s\{0,1}). (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) x = (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) x" proof- have "\x\{0<..<1} - s. (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" proof have i: "open ({0<..<1} - s)" using open_diff s open_greaterThanLessThan by blast have ii: "\x\{0<..<1::real} - s. (\2 \ \) x = \1 x" using phi(1) by auto fix x::real assume ass: " x \ {0<..<1::real} - s" then have iii: "(\2 \ \ has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_vector_derivative_at_within gamma2_vec_deriv_within gamma2_vec_diffable by auto (*Most crucial but annoying step*) then have iv:"(\1 has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_derivative_transform_within_open i ii ass apply(auto simp add: has_vector_derivative_def) by force have v: "0 \ x" "x \ 1" using ass by auto have 0: "vector_derivative \1 (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iv v(1) v(2) zero_less_one] by force have 1: "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iii v(1) v(2) zero_less_one] by force then have "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative \1 (at x within {0..1})" using 0 1 by auto then show "vector_derivative (\2 \ \) (at x within {0..1}) \ b = vector_derivative \1 (at x within {0..1}) \ b" by auto qed then have i: "\x\{0..1} - (s\{0,1}). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" by auto have ii: "\x\{0..1} - (s\{0,1}). F (\1 x) \ b = F (\2 (\ x)) \ b" using phi(1) by auto show ?thesis using i ii by auto qed have **: "negligible (s\{0,1})" using s(1) by auto have has_integral_spike_eq': "\s t g f y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using has_integral_spike_eq'[OF ** *] ass by blast qed then show "line_integral_exists F {b} \1" using phi by(auto simp add: line_integral_exists_def) have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b))" using integral_unique[OF e] integral_unique[OF f] by metis moreover have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" using b d phi by (auto simp add:) ultimately show "line_integral F {b} \1 = line_integral F {b} \2" using phi by(auto simp add: line_integral_def) qed lemma line_integral_sum_basis: assumes "finite (basis::('a::euclidean_space) set)" "\b\basis. line_integral_exists F {b} \" shows "line_integral F basis \ = (\b\basis. line_integral F {b} \)" "line_integral_exists F basis \" using assms proof(induction basis) show "line_integral F {} \ = (\b\{}. line_integral F {b} \)" by(auto simp add: line_integral_def) show "\b\{}. line_integral_exists F {b} \ \ line_integral_exists F {} \" by(simp add: line_integral_exists_def integrable_0) next fix basis::"('a::euclidean_space) set" fix x::"'a::euclidean_space" fix \ assume ind_hyp: "(\b\basis. line_integral_exists F {b} \ \ line_integral_exists F basis \)" "(\b\basis. line_integral_exists F {b} \ \ line_integral F basis \ = (\b\basis. line_integral F {b} \))" assume step: "finite basis" "x \ basis" "\b\insert x basis. line_integral_exists F {b} \" then have 0: "line_integral_exists F {x} \" by auto have 1:"line_integral_exists F basis \" using ind_hyp step by auto then show "line_integral_exists F (insert x basis) \" using step(1) step(2) line_integral_sum_gen(2)[OF _ 0 1] by simp have 3: "finite (insert x basis)" using step(1) by auto have "line_integral F basis \ = (\b\basis. line_integral F {b} \)" using ind_hyp step by auto then show "line_integral F (insert x basis) \ = (\b\insert x basis. line_integral F {b} \)" using step(1) step(2) line_integral_sum_gen(1)[OF 3 0 1] by force qed lemma reparam_weak_eq_line_integrals_basis: assumes "reparam_weak \1 \2" "\2 C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities*) "\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" shows "line_integral F basis \1 = line_integral F basis \2" "line_integral_exists F basis \1" proof- show "line_integral_exists F basis \1" using reparam_weak_eq_line_integrals(2)[OF assms(1) assms(2)] assms(3-4) line_integral_sum_basis(2)[OF assms(4)] by(simp add: subset_iff) show "line_integral F basis \1 = line_integral F basis \2" using reparam_weak_eq_line_integrals[OF assms(1) assms(2)] assms(3-4) line_integral_sum_basis(1)[OF assms(4)] line_integral_exists_smooth_one_base[OF assms(2)] by(simp add: subset_iff) qed lemma reparam_eq_line_integrals_basis: assumes "reparam \1 \2" "\2 piecewise_C1_differentiable_on {0..1}" "\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" "\b\basis. line_integral_exists F {b} \2" (*We need to remove this and work on special cases like conservative fields and field/line combinations that satisfy the improper integrals theorem assumptions*) shows "line_integral F basis \1 = line_integral F basis \2" "line_integral_exists F basis \1" proof- show "line_integral_exists F basis \1" using reparam_eq_line_integrals(2)[OF assms(1) assms(2)] assms(3-5) line_integral_sum_basis(2)[OF assms(4)] by(simp add: subset_iff) show "line_integral F basis \1 = line_integral F basis \2" using reparam_eq_line_integrals[OF assms(1) assms(2)] assms(3-5) line_integral_sum_basis(1)[OF assms(4)] by(simp add: subset_iff) qed lemma line_integral_exists_smooth: assumes "\ C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities*) "\(b::'a::euclidean_space) \basis. continuous_on (path_image \) (\x. F x \ b)" "finite basis" shows "line_integral_exists F basis \" using reparam_weak_eq_line_integrals_basis(2)[OF reparam_weak_eq_refl[where ?\1.0 = \]] assms by fastforce lemma smooth_path_imp_reverse: assumes "g C1_differentiable_on {0..1}" shows "(reversepath g) C1_differentiable_on {0..1}" using assms continuous_on_const apply (auto simp: reversepath_def) apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) apply (auto simp: C1_differentiable_on_eq) apply (simp add: finite_vimageI inj_on_def) done lemma piecewise_smooth_path_imp_reverse: assumes "g piecewise_C1_differentiable_on {0..1}" shows "(reversepath g) piecewise_C1_differentiable_on {0..1}" using assms valid_path_reversepath using valid_path_def by blast definition chain_reparam_weak_chain where "chain_reparam_weak_chain one_chain1 one_chain2 \ \f. bij f \ f ` one_chain1 = one_chain2 \ (\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam_weak \ (snd (f(k,\))) else reparam_weak \ (reversepath (snd (f(k,\)))))" lemma chain_reparam_weak_chain_line_integral: assumes "chain_reparam_weak_chain one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 C1_differentiable_on {0..1}" "\(k2,\2)\one_chain2.\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" and bound1: "boundary_chain one_chain1" and bound2: "boundary_chain one_chain2" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain f where f: "bij f" "(\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam_weak \ (snd (f(k,\))) else reparam_weak \ (reversepath (snd (f(k,\)))))" "f ` one_chain1 = one_chain2" using assms(1) by (auto simp add: chain_reparam_weak_chain_def) have 0:" \x\one_chain1. (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" proof {fix k1 \1 assume ass1: "(k1,\1) \one_chain1" have "real_of_int k1 * line_integral F basis \1 = (case (f (k1,\1)) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" proof(cases) assume ass2: "k1 = 1" let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2\ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam_weak \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam_weak \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) C1_differentiable_on {0..1}" using f(3) assms(2) ass1 smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i line_integral_exists_smooth[OF 1 2 ]] assms by auto qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') next assume "k1 \ 1" then have ass2: "k1 = -1" using bound1 ass1 f(3) unfolding boundary_chain_def by force let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam_weak \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) C1_differentiable_on {0..1}" using f(3) assms(2) ass1 smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i line_integral_exists_smooth[OF 1 2 assms(4)]] by auto qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam_weak \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') qed } then show "\x. x \ one_chain1 \ (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" by (auto simp add: case_prod_beta') qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 by(simp add: one_chain_line_integral_def sum_bij[OF f(1) _ f(3)] split_beta) show "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 by blast qed definition chain_reparam_chain where "chain_reparam_chain one_chain1 one_chain2 \ \f. bij f \ f ` one_chain1 = one_chain2 \ (\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam \ (snd (f(k,\))) else reparam \ (reversepath (snd (f(k,\)))))" definition chain_reparam_weak_path::"((real) \ (real * real)) \ ((int * ((real) \ (real * real))) set) \ bool" where "chain_reparam_weak_path \ one_chain \ \l. set l = one_chain \ distinct l \ reparam \ (rec_join l) \ valid_chain_list l \ l \ []" lemma chain_reparam_chain_line_integral: assumes "chain_reparam_chain one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 piecewise_C1_differentiable_on {0..1}" "\(k2,\2)\one_chain2.\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" and bound1: "boundary_chain one_chain1" and bound2: "boundary_chain one_chain2" and line: "\(k2,\2)\one_chain2. (\b\basis. line_integral_exists F {b} \2)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain f where f: "bij f" "(\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam \ (snd (f(k,\))) else reparam \ (reversepath (snd (f(k,\)))))" "f ` one_chain1 = one_chain2" using assms(1) by (auto simp add: chain_reparam_chain_def) have integ_exist_b: "\(k1,\1)\one_chain1. \b\basis. line_integral_exists F {b} (snd (f (k1, \1)))" using line f by fastforce have valid_cubes: "\(k1,\1)\one_chain1. valid_path (snd (f (k1, \1)))" using assms(2) f(3) valid_path_def by fastforce have integ_rev_exist_b: "\(k1,\1)\one_chain1. \b\basis. line_integral_exists F {b} (reversepath (snd (f (k1, \1))))" using line_integral_on_reverse_path(2) integ_exist_b valid_cubes by blast have 0:" \x\one_chain1. (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" proof {fix k1 \1 assume ass1: "(k1,\1) \one_chain1" have "real_of_int k1 * line_integral F basis \1 = (case (f (k1,\1)) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" proof(cases) assume ass2: "k1 = 1" let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2\ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] integ_exist_b ass1 ass2 ass3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 piecewise_smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) have ii: "line_integral_exists F basis (snd (f (k1, \1)))" using assms(4) line_integral_sum_basis(2) integ_exist_b ass1 by fastforce show ?thesis using i ii line_integral_on_reverse_path(1) valid_path_reversepath by blast qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] integ_rev_exist_b ass1 ass2 ass3 3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') next assume "k1 \ 1" then have ass2: "k1 = -1" using bound1 ass1 f(3) unfolding boundary_chain_def by force let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 piecewise_smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i] integ_rev_exist_b using ass1 assms(4) line_integral_sum_basis(2) by fastforce qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 using ass1 integ_rev_exist_b by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 using ass1 integ_exist_b by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') qed } then show "\x. x \ one_chain1 \ (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" by (auto simp add: case_prod_beta') qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 by (simp add: one_chain_line_integral_def sum_bij[OF f(1) _ f(3)] prod.case_eq_if) show "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 by blast qed lemma path_image_rec_join: fixes \::"real \ (real \ real)" fixes k::int fixes l shows "\k \. (k, \) \ set l \ valid_chain_list l \ path_image \ \ path_image (rec_join l)" proof(induction l) case Nil then show ?case by auto next case ass: (Cons a l) obtain k' \' where a: "a = (k',\')" by force have "path_image \ \ path_image (rec_join ((k',\') # l))" proof(cases) assume "l=[]" then show "path_image \ \ path_image (rec_join ((k',\') # l))" using ass(2-3) a by(auto simp add:) next assume "l\[]" then obtain b l' where b_l: "l = b # l'" by (meson rec_join.elims) obtain k'' \'' where b: "b = (k'',\'')" by force show ?thesis using ass path_image_reversepath b_l path_image_join by(fastforce simp add: a) qed then show ?case using a by auto qed lemma path_image_rec_join_2: fixes l shows "l \ [] \ valid_chain_list l \ path_image (rec_join l) \ (\(k, \) \ set l. path_image \)" proof(induction l) case Nil then show ?case by auto next case ass: (Cons a l) obtain k' \' where a: "a = (k',\')" by force have "path_image (rec_join (a # l)) \ (\(k, y)\set (a # l). path_image y)" proof(cases) assume "l=[]" then show "path_image (rec_join (a # l)) \ (\(k, y)\set (a # l). path_image y)" using step a by(auto simp add:) next assume "l\[]" then obtain b l' where b_l: "l = b # l'" by (meson rec_join.elims) obtain k'' \'' where b: "b = (k'',\'')" by force show ?thesis using ass path_image_reversepath b_l path_image_join apply(auto simp add: a) (*Excuse the ugliness of the next script*) apply blast by fastforce qed then show ?case using a by auto qed lemma continuous_on_closed_UN: assumes "finite S" shows "((\s. s \ S \ closed s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f)" using assms proof(induction S) case empty then show ?case by auto next case (insert x F) then show ?case using continuous_on_closed_Un closed_Union by (simp add: closed_Union continuous_on_closed_Un) qed lemma chain_reparam_weak_path_line_integral: assumes path_eq_chain: "chain_reparam_weak_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\b\basis. \(k::int, \) \ one_chain. line_integral_exists F {b} \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" and cont: "\b\basis. \(k,\2)\one_chain. continuous_on (path_image \2) (\x. F x \ b)" and finite_one_chain: "finite one_chain" shows "line_integral F basis \ = one_chain_line_integral F basis one_chain" "line_integral_exists F basis \" (*"valid_path \" This cannot be proven, see the crappy assumption of derivs/eq_pw_smooth :( *) proof - obtain l where l_props: "set l = one_chain" "distinct l" "reparam \ (rec_join l)" "valid_chain_list l" "l \ []" using chain_reparam_weak_path_def assms by auto have bchain_l: "boundary_chain (set l)" using l_props boundary_chain by (simp add: boundary_chain_def) have cont_forall: "\b\basis. continuous_on (\(k, \)\one_chain. path_image \) (\x. F x \ b)" proof fix b assume ass: "b \ basis" have "continuous_on (\((path_image \ snd) ` one_chain)) (\x. F x \ b)" apply(rule continuous_on_closed_UN[where ?S = "(path_image o snd) ` one_chain" ]) proof show "finite one_chain" using finite_one_chain by auto show "\s. s \ (path_image \ snd) ` one_chain \ closed s" using closed_valid_path_image[OF] valid_path by fastforce show "\s. s \ (path_image \ snd) ` one_chain \ continuous_on s (\x. F x \ b)" using cont ass by force qed then show "continuous_on (\(k, \)\one_chain. path_image \) (\x. F x \ b)" by (auto simp add: Union_eq case_prod_beta) qed show "line_integral_exists F basis \" proof (rule reparam_eq_line_integrals_basis[OF l_props(3) _ _ finite_basis]) let ?\2.0="rec_join l" show "?\2.0 piecewise_C1_differentiable_on {0..1}" apply(simp only: valid_path_def[symmetric]) apply(rule joined_is_valid) using assms l_props by auto show "\b\basis. continuous_on (path_image (rec_join l)) (\x. F x \ b)" using path_image_rec_join_2[OF l_props(5) l_props(4)] l_props(1) using cont_forall continuous_on_subset by blast show "\b\basis. line_integral_exists F {b} (rec_join l)" proof fix b assume ass: "b\basis" show "line_integral_exists F {b} (rec_join l)" proof (rule line_integral_exists_on_rec_join) show "boundary_chain (set l)" using l_props boundary_chain by auto show "valid_chain_list l" using l_props by auto show "\k \. (k, \) \ set l \ valid_path \" using l_props assms by auto show "\(k, \)\set l. line_integral_exists F {b} \" using l_props line_integral_exists ass by blast qed qed qed show "line_integral F basis \ = one_chain_line_integral F basis one_chain" proof- have i: "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" proof (rule one_chain_line_integral_rec_join) show "set l = one_chain" "distinct l" "valid_chain_list l" using l_props by auto show "boundary_chain one_chain" using boundary_chain by auto show "\(k, \)\one_chain. line_integral_exists F basis \" using line_integral_sum_basis(2)[OF finite_basis] line_integral_exists by blast show "\(k, \)\one_chain. valid_path \" using valid_path by auto show "finite basis" using finite_basis by auto qed have ii: "line_integral F basis \ = line_integral F basis (rec_join l)" proof (rule reparam_eq_line_integrals_basis[OF l_props(3) _ _ finite_basis]) let ?\2.0="rec_join l" show "?\2.0 piecewise_C1_differentiable_on {0..1}" apply(simp only: valid_path_def[symmetric]) apply(rule joined_is_valid) using assms l_props by auto show "\b\basis. continuous_on (path_image (rec_join l)) (\x. F x \ b)" using path_image_rec_join_2[OF l_props(5) l_props(4)] l_props(1) using cont_forall continuous_on_subset by blast show "\b\basis. line_integral_exists F {b} (rec_join l)" proof fix b assume ass: "b\basis" show "line_integral_exists F {b} (rec_join l)" proof (rule line_integral_exists_on_rec_join) show "boundary_chain (set l)" using l_props boundary_chain by auto show "valid_chain_list l" using l_props by auto show "\k \. (k, \) \ set l \ valid_path \" using l_props assms by auto show "\(k, \)\set l. line_integral_exists F {b} \" using l_props line_integral_exists ass by blast qed qed qed show "line_integral F basis \ = one_chain_line_integral F basis one_chain" using i ii by auto qed qed definition chain_reparam_chain' where "chain_reparam_chain' one_chain1 subdiv \ \f. ((\(f ` one_chain1)) = subdiv) \ (\cube \ one_chain1. chain_reparam_weak_path (rec_join [cube]) (f cube)) \ (\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {}) \ (\x \ one_chain1. finite (f x))" lemma chain_reparam_chain'_imp_finite_subdiv: assumes "finite one_chain1" "chain_reparam_chain' one_chain1 subdiv" shows "finite subdiv" using assms by(auto simp add: chain_reparam_chain'_def) lemma chain_reparam_chain'_line_integral: assumes chain1_eq_chain2: "chain_reparam_chain' one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain2: "\b\basis. \(k::int, \) \ subdiv. line_integral_exists F {b} \" and valid_path: "\(k, \) \ subdiv. valid_path \" and valid_path_2: "\(k, \) \ one_chain1. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" and cont_field: " \b\basis. \(k, \2)\subdiv. continuous_on (path_image \2) (\x. F x \ b)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ one_chain1. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\cube \ one_chain1. chain_reparam_weak_path (rec_join [cube]) (f cube))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 by (auto simp add: chain_reparam_chain'_def) then have 0: "one_chain_line_integral F basis subdiv = one_chain_line_integral F basis (\(f ` one_chain1))" by auto {fix k \ assume ass:"(k, \) \ one_chain1" have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_reparam_weak_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms unfolding boundary_chain_def by blast have iii:"\b\basis. \(k, \)\f (k, \). line_integral_exists F {b} \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have v: "\b\basis. \(k, \2)\f (k, \). continuous_on (path_image \2) (\x. F x \ b)" using f_props(1) ass assms by blast have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using chain_reparam_weak_path_line_integral[OF i ii iii iv finite_basis v] ass f_props(4) by (auto) then show "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_reparam_weak_path (reversepath \) (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms unfolding boundary_chain_def by blast have iii:"\b\basis. \(k, \)\f (k, \). line_integral_exists F {b} \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have v: "\b\basis. \(k, \2)\f (k, \). continuous_on (path_image \2) (\x. F x \ b)" using f_props(1) ass assms by blast have x:"line_integral_exists F basis (reversepath \) \ one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using chain_reparam_weak_path_line_integral[OF i ii iii iv finite_basis v] ass f_props(4) by auto have "valid_path (reversepath \)" using f_props(1) ass assms by auto then have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path reversepath_reversepath x ass by metis then show "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k::int, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed} note cube_line_integ = this have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) by auto show line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" using cube_line_integ by auto have 1: "one_chain_line_integral F basis (\(f ` one_chain1)) = one_chain_line_integral F basis one_chain1" proof - have 0:"one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" proof - have finite: "\chain \ (f ` one_chain1). finite chain" using f_props(1) finite_chain2 by (meson Sup_upper finite_subset) have disj: " \A\f ` one_chain1. \B\f ` one_chain1. A \ B \ A \ B = {}" apply(simp add:image_def) using f_props(3) by metis show "one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" using Groups_Big.comm_monoid_add_class.sum.Union_disjoint[OF finite disj] one_chain_line_integral_def by auto qed have 1:"(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = one_chain_line_integral F basis one_chain1" proof - have "(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1. k*(line_integral F basis \))" proof - have i:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \))" proof - have "inj_on f (one_chain1 - {p. f p = {}})" unfolding inj_on_def using f_props(3) by force then have 0: "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \)))" using Groups_Big.comm_monoid_add_class.sum.reindex by auto have "\ k \. (k, \) \ (one_chain1 - {p. f p = {}}) \ one_chain_line_integral F basis (f(k, \)) = k* (line_integral F basis \)" using cube_line_integ by auto then have "(\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \))) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" by (auto intro!: Finite_Cartesian_Product.sum_cong_aux) then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" using 0 by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ (k, \) \ {(k',\'). k'* (line_integral F basis \') = 0}" proof- fix k::"int" fix \::"real\real*real" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto show "(k, \) \ {(k'::int, \'). k' * line_integral F basis \' = 0}" using zero_line_integral cube_line_integ ass by force qed then have ii:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" proof - have "\one_chain. one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}})) \ one_chain_line_integral F basis one_chain = 0" proof - fix one_chain assume "one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}}))" then have "one_chain = {}" by auto then show "one_chain_line_integral F basis one_chain = 0" by (auto simp add: one_chain_line_integral_def) qed then have 0:"(\one_chain \ f ` (one_chain1) - (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral by auto then have "(\one_chain \ f ` (one_chain1). one_chain_line_integral F basis one_chain) - (\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" proof - have finte: "finite (f ` one_chain1)" using finite_chain1 by auto show ?thesis using Groups_Big.sum_diff[OF finte, of " (f ` (one_chain1 - {p. f p = {}}))" "one_chain_line_integral F basis"] 0 by auto qed then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ f (k, \) \ {chain. one_chain_line_integral F basis chain = 0}" proof- fix k::"int" fix \::"real\real*real" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have "one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "f (k, \) \ {p'. (one_chain_line_integral F basis p' = 0)}" by auto qed then have iii:"(\(k::int,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \)) = (\(k::int,\)\one_chain1. k*(line_integral F basis \))" proof- have "\ k \. (k,\)\one_chain1 - (one_chain1 - {p. f p = {}}) \ k* (line_integral F basis \) = 0" proof- fix k \ assume ass: "(k,\)\one_chain1 - (one_chain1 - {p. f p = {}})" then have "f(k, \) = {}" by auto then have "one_chain_line_integral F basis (f(k, \)) = 0" by (auto simp add: one_chain_line_integral_def) then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "k* (line_integral F basis \) = 0" using zero_line_integral cube_line_integ ass by force qed then have "\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k* (line_integral F basis \) = 0" by auto then have "(\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral [of "one_chain1 - (one_chain1 - {p. f p = {}})" "(\(k::int,\). k* (line_integral F basis \))"] by (simp add: split_beta) then have "(\(k::int,\)\one_chain1. k*(line_integral F basis \)) - (\(k::int,\)\ (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.sum_diff[OF finite_chain1] by (metis (no_types) Diff_subset \(\(k, \)\one_chain1 - (one_chain1 - {p. f p = {}}). k * line_integral F basis \) = 0\ \\f B. B \ one_chain1 \ sum f (one_chain1 - B) = sum f one_chain1 - sum f B\) then show ?thesis by auto qed show ?thesis using i ii iii by auto qed then show ?thesis using one_chain_line_integral_def by auto qed show ?thesis using 0 1 by auto qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using 0 1 by auto qed lemma chain_reparam_chain'_line_integral_smooth_cubes: assumes "chain_reparam_chain' one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 C1_differentiable_on {0..1}" "\b\basis.\(k2,\2)\one_chain2. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" "finite one_chain1" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k,\)\one_chain1. valid_path \" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- {fix b assume "b \ basis" fix k \ assume "(k, \)\one_chain2" have "line_integral_exists F {b} \" apply(rule line_integral_exists_smooth) using \(k, \) \ one_chain2\ assms(2) apply blast using assms using \(k, \) \ one_chain2\ \b \ basis\ apply blast using \b \ basis\ by blast} then have a: "\b\basis. \(k, \)\one_chain2. line_integral_exists F {b} \" by auto have b: "\(k2,\2)\one_chain2. valid_path \2" using assms(2) by (simp add: C1_differentiable_imp_piecewise case_prod_beta valid_path_def) show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" by (rule chain_reparam_chain'_line_integral[OF assms(1) assms(6) assms(7) a b assms(8) assms(5) assms(4) assms(3)]) show "\(k, \)\one_chain1. line_integral_exists F basis \" by (rule chain_reparam_chain'_line_integral[OF assms(1) assms(6) assms(7) a b assms(8) assms(5) assms(4) assms(3)]) qed lemma chain_subdiv_path_pathimg_subset: assumes "chain_subdiv_path \ subdiv" shows "\(k',\')\subdiv. (path_image \') \ path_image \" using assms chain_subdiv_path.simps path_image_rec_join by(auto simp add: chain_subdiv_path.simps) lemma reparam_path_image: assumes "reparam \1 \2" shows "path_image \1 = path_image \2" using assms apply (auto simp add: reparam_def path_image_def image_def bij_betw_def) apply (force dest!: equalityD2) done lemma chain_reparam_weak_path_pathimg_subset: assumes "chain_reparam_weak_path \ subdiv" shows "\(k',\')\subdiv. (path_image \') \ path_image \" using assms apply(auto simp add: chain_reparam_weak_path_def) using path_image_rec_join reparam_path_image by blast lemma chain_subdiv_chain_pathimg_subset': assumes "chain_subdiv_chain one_chain subdiv" assumes "(k,\)\ subdiv" shows " \k' \'. (k',\')\ one_chain \ path_image \ \ path_image \'" using assms unfolding chain_subdiv_chain_def pairwise_def apply auto by (metis chain_subdiv_path.cases coeff_cube_to_path.simps path_image_rec_join path_image_reversepath) lemma chain_subdiv_chain_pathimg_subset: assumes "chain_subdiv_chain one_chain subdiv" shows "\ (path_image ` {\. \k. (k,\)\ subdiv}) \ \ (path_image ` {\. \k. (k,\)\ one_chain})" using assms unfolding chain_subdiv_chain_def pairwise_def apply auto by (metis UN_iff assms chain_subdiv_chain_pathimg_subset' subsetCE case_prodI2) lemma chain_reparam_chain'_pathimg_subset': assumes "chain_reparam_chain' one_chain subdiv" assumes "(k,\)\ subdiv" shows " \k' \'. (k',\')\ one_chain \ path_image \ \ path_image \'" using assms chain_reparam_weak_path_pathimg_subset apply(auto simp add: chain_reparam_chain'_def set_eq_iff) using path_image_reversepath case_prodE case_prodD old.prod.exhaust apply (simp add: list.distinct(1) list.inject rec_join.elims) by (smt case_prodD coeff_cube_to_path.simps rec_join.simps(2) reversepath_simps(2) surj_pair) definition common_reparam_exists:: "(int \ (real \ real \ real)) set \ (int \ (real \ real \ real)) set \ bool" where "common_reparam_exists one_chain1 one_chain2 \ (\subdiv ps1 ps2. chain_reparam_chain' (one_chain1 - ps1) subdiv \ chain_reparam_chain' (one_chain2 - ps2) subdiv \ (\(k, \)\subdiv. \ C1_differentiable_on {0..1}) \ boundary_chain subdiv \ (\(k, \)\ps1. point_path \) \ (\(k, \)\ps2. point_path \))" lemma common_reparam_exists_imp_eq_line_integral: assumes finite_basis: "finite basis" and "finite one_chain1" "finite one_chain2" "boundary_chain (one_chain1::(int \ (real \ real \ real)) set)" "boundary_chain (one_chain2::(int \ (real \ real \ real)) set)" " \(k2, \2)\one_chain2. \b\basis. continuous_on (path_image \2) (\x. F x \ b)" "(common_reparam_exists one_chain1 one_chain2)" "(\(k, \)\one_chain1. valid_path \)" "(\(k, \)\one_chain2. valid_path \)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain subdiv ps1 ps2 where props: "chain_reparam_chain' (one_chain1 - ps1) subdiv" "chain_reparam_chain' (one_chain2 - ps2) subdiv " "(\(k, \)\subdiv. \ C1_differentiable_on {0..1})" "boundary_chain subdiv" "(\(k, \)\ps1. point_path \)" "(\(k, \)\ps2. point_path \)" using assms by (auto simp add: common_reparam_exists_def) have subdiv_valid: "(\(k, \)\subdiv. valid_path \)" apply(simp add: valid_path_def) using props(3) using C1_differentiable_imp_piecewise by blast have onechain_boundary1: "boundary_chain (one_chain1 - ps1)" using assms(4) by (auto simp add: boundary_chain_def) have onechain_boundary2: "boundary_chain (one_chain2 - ps1)" using assms(5) by (auto simp add: boundary_chain_def) {fix k2 \2 b assume ass: "(k2, \2)\subdiv "" b\basis" have "\ k \. (k, \) \ subdiv \ \k' \'. (k', \') \ one_chain2 \ path_image \ \ path_image \'" by (meson chain_reparam_chain'_pathimg_subset' props Diff_subset subsetCE) then have "continuous_on (path_image \2) (\x. F x \ b)" using assms(6) continuous_on_subset[where ?f = " (\x. F x \ b)"] ass apply(auto simp add: subset_iff) by (metis (mono_tags, lifting) case_prodD)} then have cont_field: "\b\basis. \(k2, \2)\subdiv. continuous_on (path_image \2) (\x. F x \ b)" by auto have one_chain1_ps_valid: "(\(k, \)\one_chain1 - ps1. valid_path \)" using assms by auto have one_chain2_ps_valid: "(\(k, \)\one_chain2 - ps1. valid_path \)" using assms by auto have 0: "one_chain_line_integral F basis (one_chain1 - ps1) = one_chain_line_integral F basis subdiv" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(1) props(3) cont_field finite_basis]) using props assms apply blast using props assms using onechain_boundary1 apply blast using props assms apply blast using one_chain1_ps_valid by blast have 1:"\(k, \)\(one_chain1 - ps1). line_integral_exists F basis \" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(1) props(3) cont_field finite_basis]) using props assms apply blast using props assms using onechain_boundary1 apply blast using props assms apply blast using one_chain1_ps_valid by blast have 2: "one_chain_line_integral F basis (one_chain2 - ps2) = one_chain_line_integral F basis subdiv" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(2) props(3) cont_field finite_basis]) using props assms apply blast apply (simp add: assms(5) boundary_chain_diff) apply (simp add: props(4)) by (simp add: assms(9)) have 3:"\(k, \)\(one_chain2 - ps2). line_integral_exists F basis \" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(2) props(3) cont_field finite_basis]) using props assms apply blast apply (simp add: assms(5) boundary_chain_diff) apply (simp add: props(4)) by (simp add: assms(9)) show line_int_ex_chain1: "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 using "1" finite_basis line_integral_exists_point_path props(5) by fastforce then show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 1 2 3 using assms(2-3) finite_basis one_chain_line_integral_point_paths props(5) props(6) by auto qed definition subcube :: "real \ real \ (int \ (real \ real \ real)) \ (int \ (real \ real \ real))" where "subcube a b cube = (fst cube, subpath a b (snd cube))" lemma subcube_valid_path: assumes "valid_path (snd cube)" "a \ {0..1}" "b \ {0..1}" shows "valid_path (snd (subcube a b cube))" using valid_path_subpath[OF assms] by (auto simp add: subcube_def) end diff --git a/thys/Incompleteness/Coding.thy b/thys/Incompleteness/Coding.thy --- a/thys/Incompleteness/Coding.thy +++ b/thys/Incompleteness/Coding.thy @@ -1,858 +1,860 @@ chapter\De Bruijn Syntax, Quotations, Codes, V-Codes\ theory Coding imports SyntaxN begin declare fresh_Nil [iff] section \de Bruijn Indices (locally-nameless version)\ nominal_datatype dbtm = DBZero | DBVar name | DBInd nat | DBEats dbtm dbtm nominal_datatype dbfm = DBMem dbtm dbtm | DBEq dbtm dbtm | DBDisj dbfm dbfm | DBNeg dbfm | DBEx dbfm declare dbtm.supp [simp] declare dbfm.supp [simp] fun lookup :: "name list \ nat \ name \ dbtm" where "lookup [] n x = DBVar x" | "lookup (y # ys) n x = (if x = y then DBInd n else (lookup ys (Suc n) x))" lemma fresh_imp_notin_env: "atom name \ e \ name \ set e" by (metis List.finite_set fresh_finite_set_at_base fresh_set) lemma lookup_notin: "x \ set e \ lookup e n x = DBVar x" by (induct e arbitrary: n) auto lemma lookup_in: "x \ set e \ \k. lookup e n x = DBInd k \ n \ k \ k < n + length e" by (induction e arbitrary: n) force+ lemma lookup_fresh: "x \ lookup e n y \ y \ set e \ x \ atom y" by (induct arbitrary: n rule: lookup.induct) (auto simp: pure_fresh fresh_at_base) lemma lookup_eqvt[eqvt]: "(p \ lookup xs n x) = lookup (p \ xs) (p \ n) (p \ x)" by (induct xs arbitrary: n) (simp_all add: permute_pure) lemma lookup_inject [iff]: "(lookup e n x = lookup e n y) \ x = y" proof (induction e n x arbitrary: y rule: lookup.induct) case (2 y ys n x z) then show ?case by (metis dbtm.distinct(7) dbtm.eq_iff(3) lookup.simps(2) lookup_in lookup_notin not_less_eq_eq) qed auto nominal_function trans_tm :: "name list \ tm \ dbtm" where "trans_tm e Zero = DBZero" | "trans_tm e (Var k) = lookup e 0 k" | "trans_tm e (Eats t u) = DBEats (trans_tm e t) (trans_tm e u)" by (auto simp: eqvt_def trans_tm_graph_aux_def) (metis tm.strong_exhaust) nominal_termination (eqvt) by lexicographic_order lemma fresh_trans_tm_iff [simp]: "i \ trans_tm e t \ i \ t \ i \ atom ` set e" by (induct t rule: tm.induct, auto simp: lookup_fresh fresh_at_base) lemma trans_tm_forget: "atom i \ t \ trans_tm [i] t = trans_tm [] t" by (induct t rule: tm.induct, auto simp: fresh_Pair) nominal_function (invariant "\(xs, _) y. atom ` set xs \* y") trans_fm :: "name list \ fm \ dbfm" where "trans_fm e (Mem t u) = DBMem (trans_tm e t) (trans_tm e u)" | "trans_fm e (Eq t u) = DBEq (trans_tm e t) (trans_tm e u)" | "trans_fm e (Disj A B) = DBDisj (trans_fm e A) (trans_fm e B)" | "trans_fm e (Neg A) = DBNeg (trans_fm e A)" | "atom k \ e \ trans_fm e (Ex k A) = DBEx (trans_fm (k#e) A)" supply [[simproc del: defined_all]] apply(simp add: eqvt_def trans_fm_graph_aux_def) apply(erule trans_fm_graph.induct) using [[simproc del: alpha_lst]] apply(auto simp: fresh_star_def) apply (metis fm.strong_exhaust fresh_star_insert) apply(erule Abs_lst1_fcb2') apply (simp_all add: eqvt_at_def) apply (simp_all add: fresh_star_Pair perm_supp_eq) apply (simp add: fresh_star_def) done nominal_termination (eqvt) by lexicographic_order lemma fresh_trans_fm [simp]: "i \ trans_fm e A \ i \ A \ i \ atom ` set e" by (nominal_induct A avoiding: e rule: fm.strong_induct, auto simp: fresh_at_base) abbreviation DBConj :: "dbfm \ dbfm \ dbfm" where "DBConj t u \ DBNeg (DBDisj (DBNeg t) (DBNeg u))" lemma trans_fm_Conj [simp]: "trans_fm e (Conj A B) = DBConj (trans_fm e A) (trans_fm e B)" by (simp add: Conj_def) lemma trans_tm_inject [iff]: "(trans_tm e t = trans_tm e u) \ t = u" proof (induct t arbitrary: u rule: tm.induct) case Zero show ?case apply (cases u rule: tm.exhaust, auto) apply (metis dbtm.distinct(1) dbtm.distinct(3) lookup_in lookup_notin) done next case (Var i) show ?case apply (cases u rule: tm.exhaust, auto) apply (metis dbtm.distinct(1) dbtm.distinct(3) lookup_in lookup_notin) apply (metis dbtm.distinct(10) dbtm.distinct(11) lookup_in lookup_notin) done next case (Eats tm1 tm2) thus ?case apply (cases u rule: tm.exhaust, auto) apply (metis dbtm.distinct(12) dbtm.distinct(9) lookup_in lookup_notin) done qed lemma trans_fm_inject [iff]: "(trans_fm e A = trans_fm e B) \ A = B" proof (nominal_induct A avoiding: e B rule: fm.strong_induct) case (Mem tm1 tm2) thus ?case by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: fresh_star_def) next case (Eq tm1 tm2) thus ?case by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: fresh_star_def) next case (Disj fm1 fm2) show ?case by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: Disj fresh_star_def) next case (Neg fm) show ?case by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: Neg fresh_star_def) next case (Ex name fm) thus ?case using [[simproc del: alpha_lst]] proof (cases rule: fm.strong_exhaust [where y=B and c="(e, name)"], simp_all add: fresh_star_def) fix name'::name and fm'::fm assume name': "atom name' \ (e, name)" assume "atom name \ fm' \ name = name'" thus "(trans_fm (name # e) fm = trans_fm (name' # e) fm') = ([[atom name]]lst. fm = [[atom name']]lst. fm')" (is "?lhs = ?rhs") proof (rule disjE) assume "name = name'" thus "?lhs = ?rhs" by (metis fresh_Pair fresh_at_base(2) name') next assume name: "atom name \ fm'" have eq1: "(name \ name') \ trans_fm (name' # e) fm' = trans_fm (name' # e) fm'" by (simp add: flip_fresh_fresh name) have eq2: "(name \ name') \ ([[atom name']]lst. fm') = [[atom name']]lst. fm'" by (rule flip_fresh_fresh) (auto simp: Abs_fresh_iff name) show "?lhs = ?rhs" using name' eq1 eq2 Ex(1) Ex(3) [of "name#e" "(name \ name') \ fm'"] by (simp add: flip_fresh_fresh) (metis Abs1_eq(3)) qed qed qed lemma trans_fm_perm: assumes c: "atom c \ (i,j,A,B)" and t: "trans_fm [i] A = trans_fm [j] B" shows "(i \ c) \ A = (j \ c) \ B" proof - have c_fresh1: "atom c \ trans_fm [i] A" using c by (auto simp: supp_Pair) moreover have i_fresh: "atom i \ trans_fm [i] A" by auto moreover have c_fresh2: "atom c \ trans_fm [j] B" using c by (auto simp: supp_Pair) moreover have j_fresh: "atom j \ trans_fm [j] B" by auto ultimately have "((i \ c) \ (trans_fm [i] A)) = ((j \ c) \ trans_fm [j] B)" by (simp only: flip_fresh_fresh t) then have "trans_fm [c] ((i \ c) \ A) = trans_fm [c] ((j \ c) \ B)" by simp then show "(i \ c) \ A = (j \ c) \ B" by simp qed -section\Characterising the Well-Formed de Bruijn Formulas\ +section\Abstraction and Substitution on de Bruijn Formulas\ + +nominal_function abst_dbtm :: "name \ nat \ dbtm \ dbtm" + where + "abst_dbtm name i DBZero = DBZero" + | "abst_dbtm name i (DBVar name') = (if name = name' then DBInd i else DBVar name')" + | "abst_dbtm name i (DBInd j) = DBInd j" + | "abst_dbtm name i (DBEats t1 t2) = DBEats (abst_dbtm name i t1) (abst_dbtm name i t2)" +apply (simp add: eqvt_def abst_dbtm_graph_aux_def, auto) +apply (metis dbtm.exhaust) +done + +nominal_termination (eqvt) + by lexicographic_order + +nominal_function subst_dbtm :: "dbtm \ name \ dbtm \ dbtm" + where + "subst_dbtm u x DBZero = DBZero" + | "subst_dbtm u x (DBVar name) = (if x = name then u else DBVar name)" + | "subst_dbtm u x (DBInd j) = DBInd j" + | "subst_dbtm u x (DBEats t1 t2) = DBEats (subst_dbtm u x t1) (subst_dbtm u x t2)" +by (auto simp: eqvt_def subst_dbtm_graph_aux_def) (metis dbtm.exhaust) + +nominal_termination (eqvt) + by lexicographic_order + +lemma fresh_iff_non_subst_dbtm: "subst_dbtm DBZero i t = t \ atom i \ t" + by (induct t rule: dbtm.induct) (auto simp: pure_fresh fresh_at_base(2)) + +lemma lookup_append: "lookup (e @ [i]) n j = abst_dbtm i (length e + n) (lookup e n j)" + by (induct e arbitrary: n) (auto simp: fresh_Cons) + +lemma trans_tm_abs: "trans_tm (e@[name]) t = abst_dbtm name (length e) (trans_tm e t)" + by (induct t rule: tm.induct) (auto simp: lookup_notin lookup_append) + +subsection\Well-Formed Formulas\ + +nominal_function abst_dbfm :: "name \ nat \ dbfm \ dbfm" + where + "abst_dbfm name i (DBMem t1 t2) = DBMem (abst_dbtm name i t1) (abst_dbtm name i t2)" + | "abst_dbfm name i (DBEq t1 t2) = DBEq (abst_dbtm name i t1) (abst_dbtm name i t2)" + | "abst_dbfm name i (DBDisj A1 A2) = DBDisj (abst_dbfm name i A1) (abst_dbfm name i A2)" + | "abst_dbfm name i (DBNeg A) = DBNeg (abst_dbfm name i A)" + | "abst_dbfm name i (DBEx A) = DBEx (abst_dbfm name (i+1) A)" +apply (simp add: eqvt_def abst_dbfm_graph_aux_def, auto) +apply (metis dbfm.exhaust) +done + +nominal_termination (eqvt) + by lexicographic_order + +nominal_function subst_dbfm :: "dbtm \ name \ dbfm \ dbfm" + where + "subst_dbfm u x (DBMem t1 t2) = DBMem (subst_dbtm u x t1) (subst_dbtm u x t2)" + | "subst_dbfm u x (DBEq t1 t2) = DBEq (subst_dbtm u x t1) (subst_dbtm u x t2)" + | "subst_dbfm u x (DBDisj A1 A2) = DBDisj (subst_dbfm u x A1) (subst_dbfm u x A2)" + | "subst_dbfm u x (DBNeg A) = DBNeg (subst_dbfm u x A)" + | "subst_dbfm u x (DBEx A) = DBEx (subst_dbfm u x A)" +by (auto simp: eqvt_def subst_dbfm_graph_aux_def) (metis dbfm.exhaust) + +nominal_termination (eqvt) + by lexicographic_order + +lemma fresh_iff_non_subst_dbfm: "subst_dbfm DBZero i t = t \ atom i \ t" + by (induct t rule: dbfm.induct) (auto simp: fresh_iff_non_subst_dbtm) + + +section\Well formed terms and formulas (de Bruijn representation)\ subsection\Well-Formed Terms\ inductive wf_dbtm :: "dbtm \ bool" where Zero: "wf_dbtm DBZero" | Var: "wf_dbtm (DBVar name)" | Eats: "wf_dbtm t1 \ wf_dbtm t2 \ wf_dbtm (DBEats t1 t2)" equivariance wf_dbtm inductive_cases Zero_wf_dbtm [elim!]: "wf_dbtm DBZero" inductive_cases Var_wf_dbtm [elim!]: "wf_dbtm (DBVar name)" inductive_cases Ind_wf_dbtm [elim!]: "wf_dbtm (DBInd i)" inductive_cases Eats_wf_dbtm [elim!]: "wf_dbtm (DBEats t1 t2)" declare wf_dbtm.intros [intro] lemma wf_dbtm_imp_is_tm: assumes "wf_dbtm x" shows "\t::tm. x = trans_tm [] t" using assms proof (induct rule: wf_dbtm.induct) case Zero thus ?case by (metis trans_tm.simps(1)) next case (Var i) thus ?case by (metis lookup.simps(1) trans_tm.simps(2)) next case (Eats dt1 dt2) thus ?case by (metis trans_tm.simps(3)) qed lemma wf_dbtm_trans_tm: "wf_dbtm (trans_tm [] t)" by (induct t rule: tm.induct) auto theorem wf_dbtm_iff_is_tm: "wf_dbtm x \ (\t::tm. x = trans_tm [] t)" by (metis wf_dbtm_imp_is_tm wf_dbtm_trans_tm) -nominal_function abst_dbtm :: "name \ nat \ dbtm \ dbtm" - where - "abst_dbtm name i DBZero = DBZero" - | "abst_dbtm name i (DBVar name') = (if name = name' then DBInd i else DBVar name')" - | "abst_dbtm name i (DBInd j) = DBInd j" - | "abst_dbtm name i (DBEats t1 t2) = DBEats (abst_dbtm name i t1) (abst_dbtm name i t2)" -apply (simp add: eqvt_def abst_dbtm_graph_aux_def, auto) -apply (metis dbtm.exhaust) -done - -nominal_termination (eqvt) - by lexicographic_order - -nominal_function subst_dbtm :: "dbtm \ name \ dbtm \ dbtm" - where - "subst_dbtm u i DBZero = DBZero" - | "subst_dbtm u i (DBVar name) = (if i = name then u else DBVar name)" - | "subst_dbtm u i (DBInd j) = DBInd j" - | "subst_dbtm u i (DBEats t1 t2) = DBEats (subst_dbtm u i t1) (subst_dbtm u i t2)" -by (auto simp: eqvt_def subst_dbtm_graph_aux_def) (metis dbtm.exhaust) - -nominal_termination (eqvt) - by lexicographic_order - -lemma fresh_iff_non_subst_dbtm: "subst_dbtm DBZero i t = t \ atom i \ t" - by (induct t rule: dbtm.induct) (auto simp: pure_fresh fresh_at_base(2)) - -lemma lookup_append: "lookup (e @ [i]) n j = abst_dbtm i (length e + n) (lookup e n j)" - by (induct e arbitrary: n) (auto simp: fresh_Cons) - -lemma trans_tm_abs: "trans_tm (e@[name]) t = abst_dbtm name (length e) (trans_tm e t)" - by (induct t rule: tm.induct) (auto simp: lookup_notin lookup_append) - subsection\Well-Formed Formulas\ -nominal_function abst_dbfm :: "name \ nat \ dbfm \ dbfm" - where - "abst_dbfm name i (DBMem t1 t2) = DBMem (abst_dbtm name i t1) (abst_dbtm name i t2)" - | "abst_dbfm name i (DBEq t1 t2) = DBEq (abst_dbtm name i t1) (abst_dbtm name i t2)" - | "abst_dbfm name i (DBDisj A1 A2) = DBDisj (abst_dbfm name i A1) (abst_dbfm name i A2)" - | "abst_dbfm name i (DBNeg A) = DBNeg (abst_dbfm name i A)" - | "abst_dbfm name i (DBEx A) = DBEx (abst_dbfm name (i+1) A)" -apply (simp add: eqvt_def abst_dbfm_graph_aux_def, auto) -apply (metis dbfm.exhaust) -done - -nominal_termination (eqvt) - by lexicographic_order - -nominal_function subst_dbfm :: "dbtm \ name \ dbfm \ dbfm" - where - "subst_dbfm u i (DBMem t1 t2) = DBMem (subst_dbtm u i t1) (subst_dbtm u i t2)" - | "subst_dbfm u i (DBEq t1 t2) = DBEq (subst_dbtm u i t1) (subst_dbtm u i t2)" - | "subst_dbfm u i (DBDisj A1 A2) = DBDisj (subst_dbfm u i A1) (subst_dbfm u i A2)" - | "subst_dbfm u i (DBNeg A) = DBNeg (subst_dbfm u i A)" - | "subst_dbfm u i (DBEx A) = DBEx (subst_dbfm u i A)" -by (auto simp: eqvt_def subst_dbfm_graph_aux_def) (metis dbfm.exhaust) - -nominal_termination (eqvt) - by lexicographic_order - -lemma fresh_iff_non_subst_dbfm: "subst_dbfm DBZero i t = t \ atom i \ t" - by (induct t rule: dbfm.induct) (auto simp: fresh_iff_non_subst_dbtm) - - -section\Well formed terms and formulas (de Bruijn representation)\ - inductive wf_dbfm :: "dbfm \ bool" where Mem: "wf_dbtm t1 \ wf_dbtm t2 \ wf_dbfm (DBMem t1 t2)" | Eq: "wf_dbtm t1 \ wf_dbtm t2 \ wf_dbfm (DBEq t1 t2)" | Disj: "wf_dbfm A1 \ wf_dbfm A2 \ wf_dbfm (DBDisj A1 A2)" | Neg: "wf_dbfm A \ wf_dbfm (DBNeg A)" | Ex: "wf_dbfm A \ wf_dbfm (DBEx (abst_dbfm name 0 A))" equivariance wf_dbfm lemma atom_fresh_abst_dbtm [simp]: "atom i \ abst_dbtm i n t" by (induct t rule: dbtm.induct) (auto simp: pure_fresh) lemma atom_fresh_abst_dbfm [simp]: "atom i \ abst_dbfm i n A" by (nominal_induct A arbitrary: n rule: dbfm.strong_induct) auto text\Setting up strong induction: "avoiding" for name. Necessary to allow some proofs to go through\ nominal_inductive wf_dbfm avoids Ex: name by (auto simp: fresh_star_def) inductive_cases Mem_wf_dbfm [elim!]: "wf_dbfm (DBMem t1 t2)" inductive_cases Eq_wf_dbfm [elim!]: "wf_dbfm (DBEq t1 t2)" inductive_cases Disj_wf_dbfm [elim!]: "wf_dbfm (DBDisj A1 A2)" inductive_cases Neg_wf_dbfm [elim!]: "wf_dbfm (DBNeg A)" inductive_cases Ex_wf_dbfm [elim!]: "wf_dbfm (DBEx z)" declare wf_dbfm.intros [intro] lemma trans_fm_abs: "trans_fm (e@[name]) A = abst_dbfm name (length e) (trans_fm e A)" apply (nominal_induct A avoiding: name e rule: fm.strong_induct) apply (auto simp: trans_tm_abs fresh_Cons fresh_append) by (metis append_Cons length_Cons) lemma abst_trans_fm: "abst_dbfm name 0 (trans_fm [] A) = trans_fm [name] A" by (metis append_Nil list.size(3) trans_fm_abs) lemma abst_trans_fm2: "i \ j \ abst_dbfm i (Suc 0) (trans_fm [j] A) = trans_fm [j,i] A" using trans_fm_abs [where e="[j]" and name=i] by auto lemma wf_dbfm_imp_is_fm: assumes "wf_dbfm x" shows "\A::fm. x = trans_fm [] A" using assms proof (induct rule: wf_dbfm.induct) case (Mem t1 t2) thus ?case by (metis trans_fm.simps(1) wf_dbtm_imp_is_tm) next case (Eq t1 t2) thus ?case by (metis trans_fm.simps(2) wf_dbtm_imp_is_tm) next case (Disj fm1 fm2) thus ?case by (metis trans_fm.simps(3)) next case (Neg fm) thus ?case by (metis trans_fm.simps(4)) next case (Ex fm name) thus ?case apply auto apply (rule_tac x="Ex name A" in exI) apply (auto simp: abst_trans_fm) done qed lemma wf_dbfm_trans_fm: "wf_dbfm (trans_fm [] A)" apply (nominal_induct A rule: fm.strong_induct) apply (auto simp: wf_dbtm_trans_tm abst_trans_fm) apply (metis abst_trans_fm wf_dbfm.Ex) done lemma wf_dbfm_iff_is_fm: "wf_dbfm x \ (\A::fm. x = trans_fm [] A)" by (metis wf_dbfm_imp_is_fm wf_dbfm_trans_fm) lemma dbtm_abst_ignore [simp]: "abst_dbtm name i (abst_dbtm name j t) = abst_dbtm name j t" by (induct t rule: dbtm.induct) auto lemma abst_dbtm_fresh_ignore [simp]: "atom name \ u \ abst_dbtm name j u = u" by (induct u rule: dbtm.induct) auto lemma dbtm_subst_ignore [simp]: "subst_dbtm u name (abst_dbtm name j t) = abst_dbtm name j t" by (induct t rule: dbtm.induct) auto lemma dbtm_abst_swap_subst: "name \ name' \ atom name' \ u \ subst_dbtm u name (abst_dbtm name' j t) = abst_dbtm name' j (subst_dbtm u name t)" by (induct t rule: dbtm.induct) auto lemma dbfm_abst_swap_subst: "name \ name' \ atom name' \ u \ subst_dbfm u name (abst_dbfm name' j A) = abst_dbfm name' j (subst_dbfm u name A)" by (induct A arbitrary: j rule: dbfm.induct) (auto simp: dbtm_abst_swap_subst) lemma subst_trans_commute [simp]: "atom i \ e \ subst_dbtm (trans_tm e u) i (trans_tm e t) = trans_tm e (subst i u t)" apply (induct t rule: tm.induct) apply (auto simp: lookup_notin fresh_imp_notin_env) apply (metis abst_dbtm_fresh_ignore dbtm_subst_ignore lookup_fresh lookup_notin subst_dbtm.simps(2)) done lemma subst_fm_trans_commute [simp]: "subst_dbfm (trans_tm [] u) name (trans_fm [] A) = trans_fm [] (A (name::= u))" apply (nominal_induct A avoiding: name u rule: fm.strong_induct) apply (auto simp: lookup_notin abst_trans_fm [symmetric]) apply (metis dbfm_abst_swap_subst fresh_at_base(2) fresh_trans_tm_iff) done lemma subst_fm_trans_commute_eq: "du = trans_tm [] u \ subst_dbfm du i (trans_fm [] A) = trans_fm [] (A(i::=u))" by (metis subst_fm_trans_commute) section\Quotations\ fun htuple :: "nat \ hf" where "htuple 0 = \0,0\" | "htuple (Suc k) = \0, htuple k\" fun HTuple :: "nat \ tm" where "HTuple 0 = HPair Zero Zero" | "HTuple (Suc k) = HPair Zero (HTuple k)" lemma eval_tm_HTuple [simp]: "\HTuple n\e = htuple n" by (induct n) auto lemma fresh_HTuple [simp]: "x \ HTuple n" by (induct n) auto lemma HTuple_eqvt[eqvt]: "(p \ HTuple n) = HTuple (p \ n)" by (induct n, auto simp: HPair_eqvt permute_pure) lemma htuple_nonzero [simp]: "htuple k \ 0" by (induct k) auto lemma htuple_inject [iff]: "htuple i = htuple j \ i=j" proof (induct i arbitrary: j) case 0 show ?case by (cases j) auto next case (Suc i) show ?case by (cases j) (auto simp: Suc) qed subsection \Quotations of de Bruijn terms\ definition nat_of_name :: "name \ nat" where "nat_of_name x = nat_of (atom x)" lemma nat_of_name_inject [simp]: "nat_of_name n1 = nat_of_name n2 \ n1 = n2" by (metis nat_of_name_def atom_components_eq_iff atom_eq_iff sort_of_atom_eq) definition name_of_nat :: "nat \ name" where "name_of_nat n \ Abs_name (Atom (Sort ''SyntaxN.name'' []) n)" lemma nat_of_name_Abs_eq [simp]: "nat_of_name (Abs_name (Atom (Sort ''SyntaxN.name'' []) n)) = n" by (auto simp: nat_of_name_def atom_name_def Abs_name_inverse) lemma nat_of_name_name_eq [simp]: "nat_of_name (name_of_nat n) = n" by (simp add: name_of_nat_def) lemma name_of_nat_nat_of_name [simp]: "name_of_nat (nat_of_name i) = i" by (metis nat_of_name_inject nat_of_name_name_eq) lemma HPair_neq_ORD_OF [simp]: "HPair x y \ ORD_OF i" by (metis Not_Ord_hpair Ord_ord_of eval_tm_HPair eval_tm_ORD_OF) text\Infinite support, so we cannot use nominal primrec.\ function quot_dbtm :: "dbtm \ tm" where "quot_dbtm DBZero = Zero" | "quot_dbtm (DBVar name) = ORD_OF (Suc (nat_of_name name))" | "quot_dbtm (DBInd k) = HPair (HTuple 6) (ORD_OF k)" | "quot_dbtm (DBEats t u) = HPair (HTuple 1) (HPair (quot_dbtm t) (quot_dbtm u))" by (rule dbtm.exhaust) auto termination by lexicographic_order lemma quot_dbtm_inject_lemma [simp]: "\quot_dbtm t\e = \quot_dbtm u\e \ t=u" proof (induct t arbitrary: u rule: dbtm.induct) case DBZero show ?case by (induct u rule: dbtm.induct) auto next case (DBVar name) show ?case by (induct u rule: dbtm.induct) (auto simp: hpair_neq_Ord') next case (DBInd k) show ?case by (induct u rule: dbtm.induct) (auto simp: hpair_neq_Ord hpair_neq_Ord') next case (DBEats t1 t2) thus ?case by (induct u rule: dbtm.induct) (simp_all add: hpair_neq_Ord) qed lemma quot_dbtm_inject [iff]: "quot_dbtm t = quot_dbtm u \ t=u" by (metis quot_dbtm_inject_lemma) subsection \Quotations of de Bruijn formulas\ text\Infinite support, so we cannot use nominal primrec.\ function quot_dbfm :: "dbfm \ tm" where "quot_dbfm (DBMem t u) = HPair (HTuple 0) (HPair (quot_dbtm t) (quot_dbtm u))" | "quot_dbfm (DBEq t u) = HPair (HTuple 2) (HPair (quot_dbtm t) (quot_dbtm u))" | "quot_dbfm (DBDisj A B) = HPair (HTuple 3) (HPair (quot_dbfm A) (quot_dbfm B))" | "quot_dbfm (DBNeg A) = HPair (HTuple 4) (quot_dbfm A)" | "quot_dbfm (DBEx A) = HPair (HTuple 5) (quot_dbfm A)" by (rule_tac y=x in dbfm.exhaust, auto) termination by lexicographic_order lemma htuple_minus_1: "n > 0 \ htuple n = \0, htuple (n - 1)\" by (metis Suc_diff_1 htuple.simps(2)) lemma HTuple_minus_1: "n > 0 \ HTuple n = HPair Zero (HTuple (n - 1))" by (metis Suc_diff_1 HTuple.simps(2)) lemmas HTS = HTuple_minus_1 HTuple.simps \ \for freeness reasoning on codes\ lemma quot_dbfm_inject_lemma [simp]: "\quot_dbfm A\e = \quot_dbfm B\e \ A=B" proof (induct A arbitrary: B rule: dbfm.induct) case (DBMem t u) show ?case by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1) next case (DBEq t u) show ?case by (induct B rule: dbfm.induct) (auto simp: htuple_minus_1) next case (DBDisj A B') thus ?case by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1) next case (DBNeg A) thus ?case by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1) next case (DBEx A) thus ?case by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1) qed class quot = fixes quot :: "'a \ tm" ("\_\") instantiation tm :: quot begin definition quot_tm :: "tm \ tm" where "quot_tm t = quot_dbtm (trans_tm [] t)" instance .. end lemma quot_dbtm_fresh [simp]: "s \ (quot_dbtm t)" by (induct t rule: dbtm.induct) auto lemma quot_tm_fresh [simp]: fixes t::tm shows "s \ \t\" by (simp add: quot_tm_def) lemma quot_Zero [simp]: "\Zero\ = Zero" by (simp add: quot_tm_def) lemma quot_Var: "\Var x\ = SUCC (ORD_OF (nat_of_name x))" by (simp add: quot_tm_def) lemma quot_Eats: "\Eats x y\ = HPair (HTuple 1) (HPair \x\ \y\)" by (simp add: quot_tm_def) text\irrelevance of the environment for quotations, because they are ground terms\ lemma eval_quot_dbtm_ignore: "\quot_dbtm t\e = \quot_dbtm t\e'" by (induct t rule: dbtm.induct) auto lemma eval_quot_dbfm_ignore: "\quot_dbfm A\e = \quot_dbfm A\e'" by (induct A rule: dbfm.induct) (auto intro: eval_quot_dbtm_ignore) instantiation fm :: quot begin definition quot_fm :: "fm \ tm" where "quot_fm A = quot_dbfm (trans_fm [] A)" instance .. end lemma quot_dbfm_fresh [simp]: "s \ (quot_dbfm A)" by (induct A rule: dbfm.induct) auto lemma quot_fm_fresh [simp]: fixes A::fm shows "s \ \A\" by (simp add: quot_fm_def) lemma quot_fm_permute [simp]: fixes A:: fm shows "p \ \A\ = \A\" by (metis fresh_star_def perm_supp_eq quot_fm_fresh) lemma quot_Mem: "\x IN y\ = HPair (HTuple 0) (HPair (\x\) (\y\))" by (simp add: quot_fm_def quot_tm_def) lemma quot_Eq: "\x EQ y\ = HPair (HTuple 2) (HPair (\x\) (\y\))" by (simp add: quot_fm_def quot_tm_def) lemma quot_Disj: "\A OR B\ = HPair (HTuple 3) (HPair (\A\) (\B\))" by (simp add: quot_fm_def) lemma quot_Neg: "\Neg A\ = HPair (HTuple 4) (\A\)" by (simp add: quot_fm_def) lemma quot_Ex: "\Ex i A\ = HPair (HTuple 5) (quot_dbfm (trans_fm [i] A))" by (simp add: quot_fm_def) lemma eval_quot_fm_ignore: fixes A:: fm shows "\\A\\e = \\A\\e'" by (metis eval_quot_dbfm_ignore quot_fm_def) lemmas quot_simps = quot_Var quot_Eats quot_Eq quot_Mem quot_Disj quot_Neg quot_Ex section\Definitions Involving Coding\ definition q_Var :: "name \ hf" where "q_Var i \ succ (ord_of (nat_of_name i))" definition q_Ind :: "hf \ hf" where "q_Ind k \ \htuple 6, k\" abbreviation Q_Eats :: "tm \ tm \ tm" where "Q_Eats t u \ HPair (HTuple (Suc 0)) (HPair t u)" definition q_Eats :: "hf \ hf \ hf" where "q_Eats x y \ \htuple 1, x, y\" abbreviation Q_Succ :: "tm \ tm" where "Q_Succ t \ Q_Eats t t" definition q_Succ :: "hf \ hf" where "q_Succ x \ q_Eats x x" lemma quot_Succ: "\SUCC x\ = Q_Succ \x\" by (auto simp: SUCC_def quot_Eats) abbreviation Q_HPair :: "tm \ tm \ tm" where "Q_HPair t u \ Q_Eats (Q_Eats Zero (Q_Eats (Q_Eats Zero u) t)) (Q_Eats (Q_Eats Zero t) t)" definition q_HPair :: "hf \ hf \ hf" where "q_HPair x y \ q_Eats (q_Eats 0 (q_Eats (q_Eats 0 y) x)) (q_Eats (q_Eats 0 x) x)" abbreviation Q_Mem :: "tm \ tm \ tm" where "Q_Mem t u \ HPair (HTuple 0) (HPair t u)" definition q_Mem :: "hf \ hf \ hf" where "q_Mem x y \ \htuple 0, x, y\" abbreviation Q_Eq :: "tm \ tm \ tm" where "Q_Eq t u \ HPair (HTuple 2) (HPair t u)" definition q_Eq :: "hf \ hf \ hf" where "q_Eq x y \ \htuple 2, x, y\" abbreviation Q_Disj :: "tm \ tm \ tm" where "Q_Disj t u \ HPair (HTuple 3) (HPair t u)" definition q_Disj :: "hf \ hf \ hf" where "q_Disj x y \ \htuple 3, x, y\" abbreviation Q_Neg :: "tm \ tm" where "Q_Neg t \ HPair (HTuple 4) t" definition q_Neg :: "hf \ hf" where "q_Neg x \ \htuple 4, x\" abbreviation Q_Conj :: "tm \ tm \ tm" where "Q_Conj t u \ Q_Neg (Q_Disj (Q_Neg t) (Q_Neg u))" definition q_Conj :: "hf \ hf \ hf" where "q_Conj t u \ q_Neg (q_Disj (q_Neg t) (q_Neg u))" abbreviation Q_Imp :: "tm \ tm \ tm" where "Q_Imp t u \ Q_Disj (Q_Neg t) u" definition q_Imp :: "hf \ hf \ hf" where "q_Imp t u \ q_Disj (q_Neg t) u" abbreviation Q_Ex :: "tm \ tm" where "Q_Ex t \ HPair (HTuple 5) t" definition q_Ex :: "hf \ hf" where "q_Ex x \ \htuple 5, x\" abbreviation Q_All :: "tm \ tm" where "Q_All t \ Q_Neg (Q_Ex (Q_Neg t))" definition q_All :: "hf \ hf" where "q_All x \ q_Neg (q_Ex (q_Neg x))" lemmas q_defs = q_Var_def q_Ind_def q_Eats_def q_HPair_def q_Eq_def q_Mem_def q_Disj_def q_Neg_def q_Conj_def q_Imp_def q_Ex_def q_All_def lemma q_Eats_iff [iff]: "q_Eats x y = q_Eats x' y' \ x=x' \ y=y'" by (metis hpair_iff q_Eats_def) lemma quot_subst_eq: "\A(i::=t)\ = quot_dbfm (subst_dbfm (trans_tm [] t) i (trans_fm [] A))" by (metis quot_fm_def subst_fm_trans_commute) lemma Q_Succ_cong: "H \ x EQ x' \ H \ Q_Succ x EQ Q_Succ x'" by (metis HPair_cong Refl) section\Quotations are Injective\ subsection\Terms\ lemma eval_tm_inject [simp]: fixes t::tm shows "\\t\\ e = \\u\\ e \ t=u" proof (induct t arbitrary: u rule: tm.induct) case Zero thus ?case by (cases u rule: tm.exhaust) (auto simp: quot_Var quot_Eats) next case (Var i) thus ?case apply (cases u rule: tm.exhaust, auto) apply (auto simp: quot_Var quot_Eats) done next case (Eats t1 t2) thus ?case apply (cases u rule: tm.exhaust, auto) apply (auto simp: quot_Eats quot_Var) done qed subsection\Formulas\ lemma eval_fm_inject [simp]: fixes A::fm shows "\\A\\ e = \\B\\ e \ A=B" proof (nominal_induct B arbitrary: A rule: fm.strong_induct) case (Mem tm1 tm2) thus ?case by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1) next case (Eq tm1 tm2) thus ?case by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1) next case (Neg \) thus ?case by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1) next case (Disj fm1 fm2) thus ?case by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1) next case (Ex i \) thus ?case apply (induct A arbitrary: i rule: fm.induct) apply (auto simp: trans_fm_perm quot_simps htuple_minus_1 Abs1_eq_iff_all) by (metis (no_types) Abs1_eq_iff_all(3) dbfm.eq_iff(5) fm.eq_iff(5) fresh_Nil trans_fm.simps(5)) qed subsection\The set \\\ of Definition 1.1, constant terms used for coding\ inductive coding_tm :: "tm \ bool" where Ord: "\i. x = ORD_OF i \ coding_tm x" | HPair: "coding_tm x \ coding_tm y \ coding_tm (HPair x y)" declare coding_tm.intros [intro] lemma coding_tm_Zero [intro]: "coding_tm Zero" by (metis ORD_OF.simps(1) Ord) lemma coding_tm_HTuple [intro]: "coding_tm (HTuple k)" by (induct k, auto) inductive_simps coding_tm_HPair [simp]: "coding_tm (HPair x y)" lemma quot_dbtm_coding [simp]: "coding_tm (quot_dbtm t)" apply (induct t rule: dbtm.induct, auto) apply (metis ORD_OF.simps(2) Ord) done lemma quot_dbfm_coding [simp]: "coding_tm (quot_dbfm fm)" by (induct fm rule: dbfm.induct, auto) lemma quot_fm_coding: fixes A::fm shows "coding_tm \A\" by (metis quot_dbfm_coding quot_fm_def) inductive coding_hf :: "hf \ bool" where Ord: "\i. x = ord_of i \ coding_hf x" | HPair: "coding_hf x \ coding_hf y \ coding_hf (\x,y\)" declare coding_hf.intros [intro] lemma coding_hf_0 [intro]: "coding_hf 0" by (metis coding_hf.Ord ord_of.simps(1)) inductive_simps coding_hf_hpair [simp]: "coding_hf (\x,y\)" lemma coding_tm_hf [simp]: "coding_tm t \ coding_hf \t\e" by (induct t rule: coding_tm.induct) auto section \V-Coding for terms and formulas, for the Second Theorem\ text\Infinite support, so we cannot use nominal primrec.\ function vquot_dbtm :: "name set \ dbtm \ tm" where "vquot_dbtm V DBZero = Zero" | "vquot_dbtm V (DBVar name) = (if name \ V then Var name else ORD_OF (Suc (nat_of_name name)))" | "vquot_dbtm V (DBInd k) = HPair (HTuple 6) (ORD_OF k)" | "vquot_dbtm V (DBEats t u) = HPair (HTuple 1) (HPair (vquot_dbtm V t) (vquot_dbtm V u))" by (auto, rule_tac y=b in dbtm.exhaust, auto) termination by lexicographic_order lemma fresh_vquot_dbtm [simp]: "i \ vquot_dbtm V tm \ i \ tm \ i \ atom ` V" by (induct tm rule: dbtm.induct) (auto simp: fresh_at_base pure_fresh) text\Infinite support, so we cannot use nominal primrec.\ function vquot_dbfm :: "name set \ dbfm \ tm" where "vquot_dbfm V (DBMem t u) = HPair (HTuple 0) (HPair (vquot_dbtm V t) (vquot_dbtm V u))" | "vquot_dbfm V (DBEq t u) = HPair (HTuple 2) (HPair (vquot_dbtm V t) (vquot_dbtm V u))" | "vquot_dbfm V (DBDisj A B) = HPair (HTuple 3) (HPair (vquot_dbfm V A) (vquot_dbfm V B))" | "vquot_dbfm V (DBNeg A) = HPair (HTuple 4) (vquot_dbfm V A)" | "vquot_dbfm V (DBEx A) = HPair (HTuple 5) (vquot_dbfm V A)" by (auto, rule_tac y=b in dbfm.exhaust, auto) termination by lexicographic_order lemma fresh_vquot_dbfm [simp]: "i \ vquot_dbfm V fm \ i \ fm \ i \ atom ` V" by (induct fm rule: dbfm.induct) (auto simp: HPair_def HTuple_minus_1) class vquot = fixes vquot :: "'a \ name set \ tm" ("\_\_" [0,1000]1000) instantiation tm :: vquot begin definition vquot_tm :: "tm \ name set \ tm" where "vquot_tm t V = vquot_dbtm V (trans_tm [] t)" instance .. end lemma vquot_dbtm_empty [simp]: "vquot_dbtm {} t = quot_dbtm t" by (induct t rule: dbtm.induct) auto lemma vquot_tm_empty [simp]: fixes t::tm shows "\t\{} = \t\" by (simp add: vquot_tm_def quot_tm_def) lemma vquot_dbtm_eq: "atom ` V \ supp t = atom ` W \ supp t \ vquot_dbtm V t = vquot_dbtm W t" by (induct t rule: dbtm.induct) (auto simp: image_iff, blast+) instantiation fm :: vquot begin definition vquot_fm :: "fm \ name set \ tm" where "vquot_fm A V = vquot_dbfm V (trans_fm [] A)" instance .. end lemma vquot_fm_fresh [simp]: fixes A::fm shows "i \ \A\V \ i \ A \ i \ atom ` V" by (simp add: vquot_fm_def) lemma vquot_dbfm_empty [simp]: "vquot_dbfm {} A = quot_dbfm A" by (induct A rule: dbfm.induct) auto lemma vquot_fm_empty [simp]: fixes A::fm shows "\A\{} = \A\" by (simp add: vquot_fm_def quot_fm_def) lemma vquot_dbfm_eq: "atom ` V \ supp A = atom ` W \ supp A \ vquot_dbfm V A = vquot_dbfm W A" by (induct A rule: dbfm.induct) (auto simp: intro!: vquot_dbtm_eq, blast+) lemma vquot_fm_insert: fixes A::fm shows "atom i \ supp A \ \A\(insert i V) = \A\V" by (auto simp: vquot_fm_def supp_conv_fresh intro: vquot_dbfm_eq) declare HTuple.simps [simp del] end diff --git a/thys/Irrationals_From_THEBOOK/Irrationals_From_THEBOOK.thy b/thys/Irrationals_From_THEBOOK/Irrationals_From_THEBOOK.thy --- a/thys/Irrationals_From_THEBOOK/Irrationals_From_THEBOOK.thy +++ b/thys/Irrationals_From_THEBOOK/Irrationals_From_THEBOOK.thy @@ -1,290 +1,290 @@ section \Some irrational numbers\ text \From Aigner and Ziegler, \emph{Proofs from THE BOOK} (Springer, 2018), Chapter 8, pp. 50--51.\ theory Irrationals_From_THEBOOK imports "Stirling_Formula.Stirling_Formula" begin subsection \Basic definitions and their consequences\ definition hf where "hf \ \n. \x::real. (x^n * (1-x)^n) / fact n" definition cf where "cf \ \n i. if i < n then 0 else (n choose (i-n)) * (-1)^(i-n)" text \Mere knowledge that the coefficients are integers is not enough later on.\ lemma hf_int_poly: fixes x::real shows "hf n = (\x. (1 / fact n) * (\i=0..2*n. real_of_int (cf n i) * x^i))" proof fix x have inj: "inj_on ((+)n) {..n}" by (auto simp: inj_on_def) have [simp]: "((+)n) ` {..n} = {n..2*n}" using nat_le_iff_add by fastforce have "(x^n * (-x + 1)^n) = x ^ n * (\k\n. real (n choose k) * (- x) ^ k)" unfolding binomial_ring by simp also have "\ = x ^ n * (\k\n. real_of_int ((n choose k) * (-1)^k) * x ^ k)" by (simp add: mult.assoc flip: power_minus) also have "\ = (\k\n. real_of_int ((n choose k) * (-1)^k) * x ^ (n+k))" by (simp add: sum_distrib_left mult_ac power_add) also have "\ = (\i=n..2*n. real_of_int (cf n i) * x^i)" by (simp add: sum.reindex [OF inj, simplified] cf_def) finally have "hf n x = (1 / fact n) * (\i = n..2 * n. real_of_int (cf n i) * x^i)" by (simp add: hf_def) moreover have "(\i = 0..i = 0..2 * n. real_of_int (cf n i) * x^i)" using sum.union_disjoint [of "{0..i. real_of_int (cf n i) * x^i"] by (simp add: ivl_disj_int_two(7) ivl_disj_un_two(7) mult_2) qed text \Lemma (ii) in the text has strict inequalities, but that's more work and is less useful.\ lemma assumes "0 \ x" "x \ 1" shows hf_nonneg: "0 \ hf n x" and hf_le_inverse_fact: "hf n x \ 1/fact n" using assms by (auto simp: hf_def divide_simps mult_le_one power_le_one) lemma hf_differt [iff]: "hf n differentiable at x" unfolding hf_int_poly differentiable_def by (intro derivative_eq_intros exI | simp)+ lemma deriv_sum_int: "deriv (\x. \i=0..n. real_of_int (c i) * x^i) x = (if n=0 then 0 else (\i=0..n - Suc 0. real_of_int ((int i + 1) * c (Suc i)) * x^i))" (is "deriv ?f x = (if n=0 then 0 else ?g)") proof - have "(?f has_real_derivative ?g) (at x)" if "n > 0" proof - have "(\i = 0..n. i * x ^ (i - Suc 0) * (c i)) = (\i = Suc 0..n. (real (i - Suc 0) + 1) * real_of_int (c i) * x ^ (i - Suc 0))" using that by (auto simp add: sum.atLeast_Suc_atMost intro!: sum.cong) also have "\ = sum ((\i. (real i + 1) * real_of_int (c (Suc i)) * x^i) \ (\n. n - Suc 0)) {Suc 0..Suc (n - Suc 0)}" using that by simp also have "\ = ?g" by (simp flip: sum.atLeast_atMost_pred_shift [where m=0]) finally have \

: "(\a = 0..n. a * x ^ (a - Suc 0) * (c a)) = ?g" . show ?thesis by (rule derivative_eq_intros \
| simp)+ qed then show ?thesis by (force intro: DERIV_imp_deriv) qed text \We calculate the coefficients of the $k$th derivative precisely.\ lemma hf_deriv_int_poly: "(deriv^^k) (hf n) = (\x. (1/fact n) * (\i=0..2*n-k. of_int (int(\{i<..i+k}) * cf n (i+k)) * x^i))" proof (induction k) case 0 show ?case by (simp add: hf_int_poly) next case (Suc k) define F where "F \ \x. (\i = 0..2*n - k. real_of_int (int(\{i<..i+k}) * cf n (i+k)) * x^i)" have Fd: "F field_differentiable at x" for x unfolding field_differentiable_def F_def by (rule derivative_eq_intros exI | force)+ have [simp]: "prod int {i<..Suc (i + k)} = (1 + int i) * prod int {Suc i<..Suc (i + k)}" for i by (metis Suc_le_mono atLeastSucAtMost_greaterThanAtMost le_add1 of_nat_Suc prod.head) have "deriv (\x. F x / fact n) x = (\i = 0..2 * n - Suc k. of_int (int(\{i<..i+ Suc k}) * cf n (Suc (i+k))) * x^i) / fact n" for x unfolding deriv_cdivide_right [OF Fd] by (fastforce simp add: F_def deriv_sum_int cf_def simp flip: of_int_mult intro: sum.cong) then show ?case by (simp add: Suc F_def) qed lemma hf_deriv_0: "(deriv^^k) (hf n) 0 \ \" proof (cases "n \ k") case True then obtain j where "(fact k::real) = real_of_int j * fact n" by (metis fact_dvd dvd_def mult.commute of_int_fact of_int_mult) moreover have "prod real {0<..k} = fact k" by (simp add: fact_prod atLeastSucAtMost_greaterThanAtMost) ultimately show ?thesis by (simp add: hf_deriv_int_poly dvd_def) next case False then show ?thesis by (simp add: hf_deriv_int_poly cf_def) qed lemma deriv_hf_minus: "deriv (hf n) = (\x. - deriv (hf n) (1-x))" proof fix x have "hf n = hf n \ (\x. (1-x))" by (simp add: fun_eq_iff hf_def mult.commute) then have "deriv (hf n) x = deriv (hf n \ (\x. (1-x))) x" by fastforce also have "\ = deriv (hf n) (1-x) * deriv ((-) 1) x" by (intro real_derivative_chain) auto finally show "deriv (hf n) x = - deriv (hf n) (1-x)" by simp qed lemma deriv_n_hf_diffr [iff]: "(deriv^^k) (hf n) field_differentiable at x" unfolding field_differentiable_def hf_deriv_int_poly by (rule derivative_eq_intros exI | force)+ lemma deriv_n_hf_minus: "(deriv^^k) (hf n) = (\x. (-1)^k * (deriv^^k) (hf n) (1-x))" proof (induction k) case 0 then show ?case by (simp add: fun_eq_iff hf_def) next case (Suc k) have o: "(\x. (deriv ^^ k) (hf n) (1-x)) = (deriv ^^ k) (hf n) \ (-) 1" by auto show ?case proof fix x have [simp]: "((deriv^^k) (hf n) \ (-) 1) field_differentiable at x" by (force intro: field_differentiable_compose) have "(deriv ^^ Suc k) (hf n) x = deriv (\x. (-1) ^ k * (deriv ^^ k) (hf n) (1-x)) x" by simp (metis Suc) also have "\ = (-1) ^ k * deriv (\x. (deriv ^^ k) (hf n) (1-x)) x" using o by fastforce also have "\ = (-1) ^ Suc k * (deriv ^^ Suc k) (hf n) (1-x)" by (subst o, subst deriv_chain, auto) finally show "(deriv ^^ Suc k) (hf n) x = (-1) ^ Suc k * (deriv ^^ Suc k) (hf n) (1-x)" . qed qed subsection \Towards the main result\ lemma hf_deriv_1: "(deriv^^k) (hf n) 1 \ \" by (smt (verit, best) Ints_1 Ints_minus Ints_mult Ints_power deriv_n_hf_minus hf_deriv_0) lemma hf_deriv_eq_0: "k > 2*n \ (deriv^^k) (hf n) = (\x. 0)" by (force simp add: cf_def hf_deriv_int_poly) text \The case for positive integers\ lemma exp_nat_irrational: assumes "s > 0" shows "exp (real_of_int s) \ \" proof assume "exp (real_of_int s) \ \" then obtain a b where ab: "a > 0" "b > 0" "coprime a b" and exp_s: "exp s = of_int a / of_int b" using Rats_cases' div_0 exp_not_eq_zero of_int_0 by (smt (verit, best) exp_gt_zero of_int_0_less_iff zero_less_divide_iff) define n where "n \ nat (max (a^2) (3 * s^3))" then have ns3: "s^3 \ real n / 3" by linarith have "n > 0" using \a > 0\ n_def by (smt (verit, best) zero_less_nat_eq zero_less_power) then have "s ^ (2*n+1) \ s ^ (3*n)" using \a > 0\ assms by (intro power_increasing) auto also have "\ = real_of_int(s^3) ^ n" by (simp add: power_mult) also have "\ \ (n / 3) ^ n" using assms ns3 by (simp add: power_mono) also have "\ \ (n / exp 1) ^ n" using exp_le \n > 0\ by (auto simp add: divide_simps) finally have s_le: "s ^ (2*n+1) \ (n / exp 1) ^ n" by presburger have a_less: "a < sqrt (2*pi*n)" proof - have "2*pi > 1" by (smt (z3) pi_gt_zero sin_gt_zero_02 sin_le_zero) have "a = sqrt (a^2)" by (simp add: ab(1) order_less_imp_le) also have "\ \ sqrt n" unfolding n_def by (smt (verit, ccfv_SIG) int_nat_eq of_nat_less_of_int_iff real_sqrt_le_mono) also have "\ < sqrt (2*pi*n)" by (simp add: \0 < n\ \1 < 2 * pi\) finally show ?thesis . qed have "sqrt (2*pi*n) * (n / exp 1) ^ n > a * s ^ (2*n+1)" using mult_strict_right_mono [OF a_less] mult_left_mono [OF s_le] by (smt (verit, best) s_le ab(1) assms of_int_1 of_int_le_iff of_int_mult zero_less_power) then have n: "fact n > a * s ^ (2*n+1)" using fact_bounds(1) by (smt (verit, best) \0 < n\ of_int_fact of_int_less_iff) define F where "F \ \x. \i\2*n. (-1)^i * s^(2*n-i) * (deriv^^i) (hf n) x" have Fder: "(F has_real_derivative -s * F x + s^(2*n+1) * hf n x) (at x)" for x proof - have *: "sum f {..n+n} = sum f {.. real" by (smt (verit, best) lessThan_Suc_atMost sum.lessThan_Suc that) have [simp]: "(deriv ((deriv ^^ (n+n)) (hf n)) x) = 0" using hf_deriv_eq_0 [where k= "Suc(n+n)"] by simp have \
: "(\k\n+n. (-1) ^ k * ((deriv ^^ Suc k) (hf n) x * of_int s ^ (n+n - k))) + s * (\j=0..n+n. (-1) ^ j * ((deriv ^^ j) (hf n) x * of_int s ^ (n+n - j))) = s * (hf n x * of_int s ^ (n+n))" using \n>0\ apply (subst sum_Suc_reindex) apply (simp add: algebra_simps atLeast0AtMost) apply (force simp add: * mult.left_commute [of "of_int s"] minus_nat.diff_Suc sum_distrib_left simp flip: sum.distrib intro: comm_monoid_add_class.sum.neutral split: nat.split_asm) done show ?thesis unfolding F_def apply (rule derivative_eq_intros field_differentiable_derivI | simp)+ using \
by (simp add: algebra_simps atLeast0AtMost eval_nat_numeral) qed have F01_Ints: "F 0 \ \" "F 1 \ \" by (simp_all add: F_def hf_deriv_0 hf_deriv_1 Ints_sum) define sF where "sF \ \x. exp (of_int s * x) * F x" define sF' where "sF' \ \x. of_int s ^ Suc(2*n) * (exp (of_int s * x) * hf n x)" have sF_der: "(sF has_real_derivative sF' x) (at x)" for x unfolding sF_def sF'_def by (rule refl Fder derivative_eq_intros | force simp: algebra_simps)+ let ?N = "b * integral {0..1} sF'" have sF'_integral: "(sF' has_integral sF 1 - sF 0) {0..1}" - by (smt (verit) fundamental_theorem_of_calculus has_field_derivative_iff_has_vector_derivative + by (smt (verit) fundamental_theorem_of_calculus has_real_derivative_iff_has_vector_derivative has_vector_derivative_at_within sF_der) then have "?N = a * F 1 - b * F 0" using \b > 0\ by (simp add: integral_unique exp_s sF_def algebra_simps) also have "\ \ \" using hf_deriv_1 by (simp add: F01_Ints) finally have N_Ints: "?N \ \" . have "sF' (1/2) > 0" and ge0: "\x. x \ {0..1} \ 0 \ sF' x" using assms by (auto simp add: sF'_def hf_def) moreover have "continuous_on {0..1} sF'" unfolding sF'_def hf_def by (intro continuous_intros) auto ultimately have False if "(sF' has_integral 0) {0..1}" using has_integral_0_cbox_imp_0 [of 0 1 sF' "1/2"] that by auto then have "integral {0..1} sF' > 0" by (metis ge0 has_integral_nonneg integral_unique order_le_less sF'_integral) then have "0 < ?N" by (simp add: \b > 0\) have "integral {0..1} sF' = of_int s ^ Suc(2*n) * integral {0..1} (\x. exp (s*x) * hf n x)" unfolding sF'_def by force also have "\ \ of_int s ^ Suc(2*n) * (exp s * (1 / fact n))" proof (rule mult_left_mono) have "integral {0..1} (\x. exp (s*x) * hf n x) \ integral {0..1} (\x::real. exp s * (1/fact n))" proof (intro mult_mono integral_le) show "(\x. exp (s*x) * hf n x) integrable_on {0..1}" using \0 < ?N\ not_integrable_integral sF'_def by fastforce qed (use assms hf_nonneg hf_le_inverse_fact in auto) also have "\ = exp s * (1 / fact n)" by simp finally show "integral {0..1} (\x. exp (s*x) * hf n x) \ exp s * (1 / fact n)" . qed (use assms in auto) finally have "?N \ b * of_int s ^ Suc(2*n) * exp s * (1 / fact n)" using \b > 0\ by (simp add: sF'_def mult_ac divide_simps) also have "\ < 1" using n apply (simp add: field_simps exp_s) by (metis of_int_fact of_int_less_iff of_int_mult of_int_power) finally show False using \0 < ?N\ Ints_cases N_Ints by force qed theorem exp_irrational: fixes q::real assumes "q \ \" "q \ 0" shows "exp q \ \" proof assume q: "exp q \ \" obtain s t where "s \ 0" "t > 0" "q = of_int s / of_int t" by (metis Rats_cases' assms div_0 of_int_0) then have "(exp q) ^ (nat t) = exp s" by (smt (verit, best) exp_divide_power_eq of_nat_nat zero_less_nat_eq) moreover have "exp q ^ (nat t) \ \" by (simp add: q) ultimately show False by (smt (verit, del_insts) Rats_inverse \s \ 0\ exp_minus exp_nat_irrational of_int_of_nat) qed corollary ln_irrational: fixes q::real assumes "q \ \" "q > 0" "q \ 1" shows "ln q \ \" using assms exp_irrational [of "ln q"] exp_ln_iff [of q] by force end diff --git a/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy b/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy --- a/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy @@ -1,491 +1,491 @@ section \Upper and Lower Solutions\ theory Upper_Lower_Solution imports Flow begin text \Following Walter~\cite{walter} in section 9\ lemma IVT_min: fixes f :: "real \ 'b :: {linorder_topology,real_normed_vector,ordered_real_vector}" \ \generalize?\ assumes y: "f a \ y" "y \ f b" "a \ b" assumes *: "continuous_on {a .. b} f" notes [continuous_intros] = *[THEN continuous_on_subset] obtains x where "a \ x" "x \ b" "f x = y" "\x'. a \ x' \ x' < x \ f x' < y" proof - let ?s = "((\x. f x - y) -` {0..}) \ {a..b}" have "?s \ {}" using assms by auto have "closed ?s" by (rule closed_vimage_Int) (auto intro!: continuous_intros) moreover have "bounded ?s" by (rule bounded_Int) (simp add: bounded_closed_interval) ultimately have "compact ?s" using compact_eq_bounded_closed by blast from compact_attains_inf[OF this \?s \ {}\] obtain x where x: "a \ x" "x \ b" "f x \ y" and min: "\z. a \ z \ z \ b \ f z \ y \ x \ z" by auto have "f x \ y" proof (rule ccontr) assume n: "\ f x \ y" then have "\z\a. z \ x \ (\x. f x - y) z = 0" using x by (intro IVT') (auto intro!: continuous_intros simp: assms) then obtain z where z: "a \ z" "z \ x" "f z = y" by auto then have "a \ z" "z \ b" "f z \ y" using x by auto from min [OF this] z n show False by auto qed then have "a \ x" "x \ b" "f x = y" using x by (auto ) moreover have "f x' < y" if "a \ x'" "x' < x" for x' apply (rule ccontr) using min[of x'] that x by (auto simp: not_less) ultimately show ?thesis .. qed lemma filtermap_at_left_shift: "filtermap (\x. x - d) (at_left a) = at_left (a - d::real)" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) context fixes v v' w w'::"real \ real" and t0 t1 e::real assumes v': "(v has_vderiv_on v') {t0 <.. t1}" and w': "(w has_vderiv_on w') {t0 <.. t1}" assumes pos_ivl: "t0 < t1" assumes e_pos: "e > 0" and e_in: "t0 + e \ t1" assumes less: "\t. t0 < t \ t < t0 + e \ v t < w t" begin lemma first_intersection_crossing_derivatives: assumes na: "t0 < tg" "tg \ t1" "v tg \ w tg" notes [continuous_intros] = vderiv_on_continuous_on[OF v', THEN continuous_on_subset] vderiv_on_continuous_on[OF w', THEN continuous_on_subset] obtains x0 where "t0 < x0" "x0 \ tg" "v' x0 \ w' x0" "v x0 = w x0" "\t. t0 < t \ t < x0 \ v t < w t" proof - have "(v - w) (min tg (t0 + e / 2)) \ 0" "0 \ (v - w) tg" "min tg (t0 + e / 2) \ tg" "continuous_on {min tg (t0 + e / 2)..tg} (v - w)" using less[of "t0 + e / 2"] less[of tg]na \e > 0\ by (auto simp: min_def intro!: continuous_intros) from IVT_min[OF this] obtain x0 where x0: "min tg (t0 + e / 2) \ x0" "x0 \ tg" "v x0 = w x0" "\x'. min tg (t0 + e / 2) \ x' \ x' < x0 \ v x' < w x'" by auto then have x0_in: "t0 < x0" "x0 \ t1" using \e > 0\ na(1,2) by (auto) note \t0 < x0\ \x0 \ tg\ moreover { from v' x0_in have "(v has_derivative (\x. x * v' x0)) (at x0 within {t0<..y. (v y - (v x0 + (y - x0) * v' x0)) / norm (y - x0)) \ 0) (at x0 within {t0<..x. x * w' x0)) (at x0 within {t0<..y. (w y - (w x0 + (y - x0) * w' x0)) / norm (y - x0)) \ 0) (at x0 within {t0<..\<^sub>F x in at x0 within {t0<..\<^sub>F x in at x0 within {t0<.. x < x0" using less na(1) na(3) x0(3) x0_in(1) by (force simp: min_def eventually_at_filter intro!: order_tendstoD[OF tendsto_ident_at])+ then have "\\<^sub>F x in at x0 within {t0<..t0 < t1\ sum_sqs_eq by (auto simp: divide_simps algebra_simps min_def intro!: eventuallyI split: if_split_asm) from this tendsto_diff[OF v w] have 1: "((\x. (v x - w x) / norm (x - x0) + (v' x0 - w' x0)) \ 0) (at x0 within {t0<..\<^sub>F x in at x0 within {t0<.. (v' x0 - w' x0)" by eventually_elim (auto simp: divide_simps intro!: less_imp_le x0(4)) moreover have "at x0 within {t0<.. bot" by (simp add: \t0 < x0\ at_within_eq_bot_iff less_imp_le) ultimately have "0 \ v' x0 - w' x0" by (rule tendsto_upperbound) then have "v' x0 \ w' x0" by simp } moreover note \v x0 = w x0\ moreover have "t0 < t \ t < x0 \ v t < w t" for t by (cases "min tg (t0 + e / 2) \ t") (auto intro: x0 less) ultimately show ?thesis .. qed lemma defect_less: assumes b: "\t. t0 < t \ t \ t1 \ v' t - f t (v t) < w' t - f t (w t)" notes [continuous_intros] = vderiv_on_continuous_on[OF v', THEN continuous_on_subset] vderiv_on_continuous_on[OF w', THEN continuous_on_subset] shows "\t \ {t0 <.. t1}. v t < w t" proof (rule ccontr) assume " \ (\t\{t0 <.. t1}. v t < w t)" then obtain tu where "t0 < tu" "tu \ t1" "v tu \ w tu" by auto from first_intersection_crossing_derivatives[OF this] obtain x0 where "t0 < x0" "x0 \ tu" "w' x0 \ v' x0" "v x0 = w x0" "\t. t0 < t \ t < x0 \ v t < w t" by metis with b[of x0] \tu \ t1\ show False by simp qed end lemma has_derivatives_less_lemma: fixes v v' ::"real \ real" assumes v': "(v has_vderiv_on v') T" assumes y': "(y has_vderiv_on y') T" assumes lu: "\t. t \ T \ t > t0 \ v' t - f t (v t) < y' t - f t (y t)" assumes lower: "v t0 \ y t0" assumes eq_imp: "v t0 = y t0 \ v' t0 < y' t0" assumes t: "t0 < t" "t0 \ T" "t \ T" "is_interval T" shows "v t < y t" proof - have subset: "{t0 .. t} \ T" by (rule atMostAtLeast_subset_convex) (auto simp: assms is_interval_convex) obtain d where "0 < d" "t0 < s \ s \ t \ s < t0 + d \ v s < y s" for s proof cases assume "v t0 = y t0" from this[THEN eq_imp] have *: "0 < y' t0 - v' t0" by simp have "((\t. y t - v t) has_vderiv_on (\t0. y' t0 - v' t0)) {t0 .. t}" by (auto intro!: derivative_intros y' v' has_vderiv_on_subset[OF _ subset]) with \t0 < t\ have d: "((\t. y t - v t) has_real_derivative y' t0 - v' t0) (at t0 within {t0 .. t})" - by (auto simp: has_vderiv_on_def has_field_derivative_iff_has_vector_derivative) + by (auto simp: has_vderiv_on_def has_real_derivative_iff_has_vector_derivative) from has_real_derivative_pos_inc_right[OF d *] \v t0 = y t0\ obtain d where "d > 0" and vy: "h > 0 \ t0 + h \ t \ h < d \ v (t0 + h) < y (t0 + h)" for h by auto have vy: "t0 < s \ s \ t \ s < t0 + d \ v s < y s" for s using vy[of "s - t0"] by simp with \d > 0\ show ?thesis .. next assume "v t0 \ y t0" then have "v t0 < y t0" using lower by simp moreover have "continuous_on {t0 .. t} v" "continuous_on {t0 .. t} y" by (auto intro!: vderiv_on_continuous_on assms has_vderiv_on_subset[OF _ subset]) then have "(v \ v t0) (at t0 within {t0 .. t})" "(y \ y t0) (at t0 within {t0 .. t})" by (auto simp: continuous_on) ultimately have "\\<^sub>F x in at t0 within {t0 .. t}. 0 < y x - v x" by (intro order_tendstoD) (auto intro!: tendsto_eq_intros) then obtain d where "d > 0" "\x. t0 < x \ x \ t \ x < t0 + d \ v x < y x" by atomize_elim (auto simp: eventually_at algebra_simps dist_real_def) then show ?thesis .. qed with \d > 0\ \t0 < t\ obtain e where "e > 0" "t0 + e \ t" "t0 < s \ s < t0 + e \ v s < y s" for s by atomize_elim (auto simp: min_def divide_simps intro!: exI[where x="min (d/2) ((t - t0) / 2)"] split: if_split_asm) from defect_less[ OF has_vderiv_on_subset[OF v'] has_vderiv_on_subset[OF y'] \t0 < t\ this lu] show "v t < y t" using \t0 < t\ subset by (auto simp: subset_iff assms) qed lemma strict_lower_solution: fixes v v' ::"real \ real" assumes sol: "(y solves_ode f) T X" assumes v': "(v has_vderiv_on v') T" assumes lower: "\t. t \ T \ t > t0 \ v' t < f t (v t)" assumes iv: "v t0 \ y t0" "v t0 = y t0 \ v' t0 < f t0 (y t0)" assumes t: "t0 < t" "t0 \ T" "t \ T" "is_interval T" shows "v t < y t" proof - note v' moreover note solves_odeD(1)[OF sol] moreover have 3: "v' t - f t (v t) < f t (y t) - f t (y t)" if "t \ T" "t > t0" for t using lower(1)[OF that] by arith moreover note iv moreover note t ultimately show "v t < y t" by (rule has_derivatives_less_lemma) qed lemma strict_upper_solution: fixes w w'::"real \ real" assumes sol: "(y solves_ode f) T X" assumes w': "(w has_vderiv_on w') T" and upper: "\t. t \ T \ t > t0 \ f t (w t) < w' t" and iv: "y t0 \ w t0" "y t0 = w t0 \ f t0 (y t0) < w' t0" assumes t: "t0 < t" "t0 \ T" "t \ T" "is_interval T" shows "y t < w t" proof - note solves_odeD(1)[OF sol] moreover note w' moreover have "f t (y t) - f t (y t) < w' t - f t (w t)" if "t \ T" "t > t0" for t using upper(1)[OF that] by arith moreover note iv moreover note t ultimately show "y t < w t" by (rule has_derivatives_less_lemma) qed lemma uniform_limit_at_within_subset: assumes "uniform_limit S x l (at t within T)" assumes "U \ T" shows "uniform_limit S x l (at t within U)" by (metis assms(1) assms(2) eventually_within_Un filterlim_iff subset_Un_eq) lemma uniform_limit_le: fixes f::"'c \ 'a \ 'b::{metric_space, linorder_topology}" assumes I: "I \ bot" assumes u: "uniform_limit X f g I" assumes u': "uniform_limit X f' g' I" assumes "\\<^sub>F i in I. \x \ X. f i x \ f' i x" assumes "x \ X" shows "g x \ g' x" proof - have "\\<^sub>F i in I. f i x \ f' i x" using assms by (simp add: eventually_mono) with I tendsto_uniform_limitI[OF u' \x \ X\] tendsto_uniform_limitI[OF u \x \ X\] show ?thesis by (rule tendsto_le) qed lemma uniform_limit_le_const: fixes f::"'c \ 'a \ 'b::{metric_space, linorder_topology}" assumes I: "I \ bot" assumes u: "uniform_limit X f g I" assumes "\\<^sub>F i in I. \x \ X. f i x \ h x" assumes "x \ X" shows "g x \ h x" proof - have "\\<^sub>F i in I. f i x \ h x" using assms by (simp add: eventually_mono) then show ?thesis by (metis tendsto_upperbound I tendsto_uniform_limitI[OF u \x \ X\]) qed lemma uniform_limit_ge_const: fixes f::"'c \ 'a \ 'b::{metric_space, linorder_topology}" assumes I: "I \ bot" assumes u: "uniform_limit X f g I" assumes "\\<^sub>F i in I. \x \ X. h x \ f i x" assumes "x \ X" shows "h x \ g x" proof - have "\\<^sub>F i in I. h x \ f i x" using assms by (simp add: eventually_mono) then show ?thesis by (metis tendsto_lowerbound I tendsto_uniform_limitI[OF u \x \ X\]) qed locale ll_on_open_real = ll_on_open T f X for T f and X::"real set" begin lemma lower_solution: fixes v v' ::"real \ real" assumes sol: "(y solves_ode f) S X" assumes v': "(v has_vderiv_on v') S" assumes lower: "\t. t \ S \ t > t0 \ v' t < f t (v t)" assumes iv: "v t0 \ y t0" assumes t: "t0 \ t" "t0 \ S" "t \ S" "is_interval S" "S \ T" shows "v t \ y t" proof cases assume "v t0 = y t0" have "{t0 -- t} \ S" using t by (simp add: closed_segment_subset is_interval_convex) with sol have "(y solves_ode f) {t0 -- t} X" using order_refl by (rule solves_ode_on_subset) moreover note refl moreover have "{t0 -- t} \ T" using \{t0 -- t} \ S\ \S \ T\ by (rule order_trans) ultimately have t_ex: "t \ existence_ivl t0 (y t0)" by (rule existence_ivl_maximal_segment) have t0_ex: "t0 \ existence_ivl t0 (y t0)" using in_existence_between_zeroI t_ex by blast have "t0 \ T" using assms(9) t(2) by blast from uniform_limit_flow[OF t0_ex t_ex] \t0 \ t\ have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at (y t0))" by simp then have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at_right (y t0))" by (rule uniform_limit_at_within_subset) simp moreover { have "\\<^sub>F i in at (y t0). t \ existence_ivl t0 i" by (rule eventually_mem_existence_ivl) fact then have "\\<^sub>F i in at_right (y t0). t \ existence_ivl t0 i" unfolding eventually_at_filter by eventually_elim simp moreover have "\\<^sub>F i in at_right (y t0). i \ X" proof - have f1: "\r ra rb. r \ existence_ivl ra rb \ rb \ X" by (metis existence_ivl_reverse flow_in_domain flows_reverse) obtain rr :: "(real \ bool) \ (real \ bool) \ real" where "\p f pa fa. (\ eventually p f \ eventually pa f \ p (rr p pa)) \ (\ eventually p fa \ \ pa (rr p pa) \ eventually pa fa)" by (metis (no_types) eventually_mono) then show ?thesis using f1 calculation by meson qed moreover have "\\<^sub>F i in at_right (y t0). y t0 < i" by (simp add: eventually_at_filter) ultimately have "\\<^sub>F i in at_right (y t0). \x\{t0..t}. v x \ flow t0 i x" proof eventually_elim case (elim y') show ?case proof safe fix s assume s: "s \ {t0..t}" show "v s \ flow t0 y' s" proof cases assume "s = t0" with elim iv show ?thesis by (simp add: \t0 \ T\ \y' \ X\) next assume "s \ t0" with s have "t0 < s" by simp have "{t0 -- s} \ S" using \{t0--t} \ S\ closed_segment_eq_real_ivl s by auto from s elim have "{t0..s} \ existence_ivl t0 y'" using ivl_subset_existence_ivl by blast with flow_solves_ode have sol: "(flow t0 y' solves_ode f) {t0 .. s} X" by (rule solves_ode_on_subset) (auto intro!: \y' \ X\ \t0 \ T\) have "{t0 .. s} \ S" using \{t0 -- s} \ S\ by (simp add: closed_segment_eq_real_ivl split: if_splits) with v' have v': "(v has_vderiv_on v') {t0 .. s}" by (rule has_vderiv_on_subset) from \y t0 < y'\ \v t0 = y t0\ have less_init: "v t0 < flow t0 y' t0" by (simp add: flow_initial_time_if \t0 \ T\ \y' \ X\) from strict_lower_solution[OF sol v' lower less_imp_le[OF less_init] _ \t0 < s\] \{t0 .. s} \ S\ less_init \t0 < s\ have "v s < flow t0 y' s" by (simp add: subset_iff is_interval_cc) then show ?thesis by simp qed qed qed } moreover have "t \ {t0 .. t}" using \t0 \ t\ by simp ultimately have "v t \ flow t0 (y t0) t" by (rule uniform_limit_ge_const[OF trivial_limit_at_right_real]) also have "flow t0 (y t0) t = y t" using sol t by (intro maximal_existence_flow) auto finally show ?thesis . next assume "v t0 \ y t0" then have less: "v t0 < y t0" using iv by simp show ?thesis apply (cases "t0 = t") subgoal using iv by blast subgoal using strict_lower_solution[OF sol v' lower iv] less t by force done qed lemma upper_solution: fixes v v' ::"real \ real" assumes sol: "(y solves_ode f) S X" assumes v': "(v has_vderiv_on v') S" assumes upper: "\t. t \ S \ t > t0 \ f t (v t) < v' t" assumes iv: "y t0 \ v t0" assumes t: "t0 \ t" "t0 \ S" "t \ S" "is_interval S" "S \ T" shows "y t \ v t" proof cases assume "v t0 = y t0" have "{t0 -- t} \ S" using t by (simp add: closed_segment_subset is_interval_convex) with sol have "(y solves_ode f) {t0 -- t} X" using order_refl by (rule solves_ode_on_subset) moreover note refl moreover have "{t0 -- t} \ T" using \{t0 -- t} \ S\ \S \ T\ by (rule order_trans) ultimately have t_ex: "t \ existence_ivl t0 (y t0)" by (rule existence_ivl_maximal_segment) have t0_ex: "t0 \ existence_ivl t0 (y t0)" using in_existence_between_zeroI t_ex by blast have "t0 \ T" using assms(9) t(2) by blast from uniform_limit_flow[OF t0_ex t_ex] \t0 \ t\ have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at (y t0))" by simp then have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at_left (y t0))" by (rule uniform_limit_at_within_subset) simp moreover { have "\\<^sub>F i in at (y t0). t \ existence_ivl t0 i" by (rule eventually_mem_existence_ivl) fact then have "\\<^sub>F i in at_left (y t0). t \ existence_ivl t0 i" unfolding eventually_at_filter by eventually_elim simp moreover have "\\<^sub>F i in at_left (y t0). i \ X" proof - have f1: "\r ra rb. r \ existence_ivl ra rb \ rb \ X" by (metis existence_ivl_reverse flow_in_domain flows_reverse) obtain rr :: "(real \ bool) \ (real \ bool) \ real" where "\p f pa fa. (\ eventually p f \ eventually pa f \ p (rr p pa)) \ (\ eventually p fa \ \ pa (rr p pa) \ eventually pa fa)" by (metis (no_types) eventually_mono) then show ?thesis using f1 calculation by meson qed moreover have "\\<^sub>F i in at_left (y t0). i < y t0" by (simp add: eventually_at_filter) ultimately have "\\<^sub>F i in at_left (y t0). \x\{t0..t}. flow t0 i x \ v x" proof eventually_elim case (elim y') show ?case proof safe fix s assume s: "s \ {t0..t}" show "flow t0 y' s \ v s" proof cases assume "s = t0" with elim iv show ?thesis by (simp add: \t0 \ T\ \y' \ X\) next assume "s \ t0" with s have "t0 < s" by simp have "{t0 -- s} \ S" using \{t0--t} \ S\ closed_segment_eq_real_ivl s by auto from s elim have "{t0..s} \ existence_ivl t0 y'" using ivl_subset_existence_ivl by blast with flow_solves_ode have sol: "(flow t0 y' solves_ode f) {t0 .. s} X" by (rule solves_ode_on_subset) (auto intro!: \y' \ X\ \t0 \ T\) have "{t0 .. s} \ S" using \{t0 -- s} \ S\ by (simp add: closed_segment_eq_real_ivl split: if_splits) with v' have v': "(v has_vderiv_on v') {t0 .. s}" by (rule has_vderiv_on_subset) from \y' < y t0\ \v t0 = y t0\ have less_init: "flow t0 y' t0 < v t0" by (simp add: flow_initial_time_if \t0 \ T\ \y' \ X\) from strict_upper_solution[OF sol v' upper less_imp_le[OF less_init] _ \t0 < s\] \{t0 .. s} \ S\ less_init \t0 < s\ have "flow t0 y' s < v s" by (simp add: subset_iff is_interval_cc) then show ?thesis by simp qed qed qed } moreover have "t \ {t0 .. t}" using \t0 \ t\ by simp ultimately have "flow t0 (y t0) t \ v t" by (rule uniform_limit_le_const[OF trivial_limit_at_left_real]) also have "flow t0 (y t0) t = y t" using sol t by (intro maximal_existence_flow) auto finally show ?thesis . next assume "v t0 \ y t0" then have less: "y t0 < v t0" using iv by simp show ?thesis apply (cases "t0 = t") subgoal using iv by blast subgoal using strict_upper_solution[OF sol v' upper iv] less t by force done qed end end diff --git a/thys/Ordinary_Differential_Equations/Library/Gronwall.thy b/thys/Ordinary_Differential_Equations/Library/Gronwall.thy --- a/thys/Ordinary_Differential_Equations/Library/Gronwall.thy +++ b/thys/Ordinary_Differential_Equations/Library/Gronwall.thy @@ -1,268 +1,268 @@ theory Gronwall imports Vector_Derivative_On begin subsection \Gronwall\ lemma derivative_quotient_bound: assumes g_deriv_on: "(g has_vderiv_on g') {a .. b}" assumes frac_le: "\t. t \ {a .. b} \ g' t / g t \ K" assumes g'_cont: "continuous_on {a .. b} g'" assumes g_pos: "\t. t \ {a .. b} \ g t > 0" assumes t_in: "t \ {a .. b}" shows "g t \ g a * exp (K * (t - a))" proof - have g_deriv: "\t. t \ {a .. b} \ (g has_real_derivative g' t) (at t within {a .. b})" using g_deriv_on - by (auto simp: has_vderiv_on_def has_field_derivative_iff_has_vector_derivative[symmetric]) + by (auto simp: has_vderiv_on_def has_real_derivative_iff_has_vector_derivative[symmetric]) from assms have g_nonzero: "\t. t \ {a .. b} \ g t \ 0" by fastforce have frac_integrable: "\t. t \ {a .. b} \ (\t. g' t / g t) integrable_on {a..t}" by (force simp: g_nonzero intro: assms has_field_derivative_subset[OF g_deriv] continuous_on_subset[OF g'_cont] continuous_intros integrable_continuous_real continuous_on_subset[OF vderiv_on_continuous_on[OF g_deriv_on]]) have "\t. t \ {a..b} \ ((\t. g' t / g t) has_integral ln (g t) - ln (g a)) {a .. t}" by (rule fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros assms has_field_derivative_subset[OF g_deriv] - simp: has_field_derivative_iff_has_vector_derivative[symmetric]) + simp: has_real_derivative_iff_has_vector_derivative[symmetric]) hence *: "\t. t \ {a .. b} \ ln (g t) - ln (g a) = integral {a .. t} (\t. g' t / g t)" using integrable_integral[OF frac_integrable] by (rule has_integral_unique[where f = "\t. g' t / g t"]) from * t_in have "ln (g t) - ln (g a) = integral {a .. t} (\t. g' t / g t)" . also have "\ \ integral {a .. t} (\_. K)" using \t \ {a .. b}\ by (intro integral_le) (auto intro!: frac_integrable frac_le integral_le) also have "\ = K * (t - a)" using \t \ {a .. b}\ by simp finally have "ln (g t) \ K * (t - a) + ln (g a)" (is "?lhs \ ?rhs") by simp hence "exp ?lhs \ exp ?rhs" by simp thus ?thesis using \t \ {a .. b}\ g_pos by (simp add: ac_simps exp_add del: exp_le_cancel_iff) qed lemma derivative_quotient_bound_left: assumes g_deriv_on: "(g has_vderiv_on g') {a .. b}" assumes frac_ge: "\t. t \ {a .. b} \ K \ g' t / g t" assumes g'_cont: "continuous_on {a .. b} g'" assumes g_pos: "\t. t \ {a .. b} \ g t > 0" assumes t_in: "t \ {a..b}" shows "g t \ g b * exp (K * (t - b))" proof - have g_deriv: "\t. t \ {a .. b} \ (g has_real_derivative g' t) (at t within {a .. b})" using g_deriv_on - by (auto simp: has_vderiv_on_def has_field_derivative_iff_has_vector_derivative[symmetric]) + by (auto simp: has_vderiv_on_def has_real_derivative_iff_has_vector_derivative[symmetric]) from assms have g_nonzero: "\t. t \ {a..b} \ g t \ 0" by fastforce have frac_integrable: "\t. t \ {a .. b} \ (\t. g' t / g t) integrable_on {t..b}" by (force simp: g_nonzero intro: assms has_field_derivative_subset[OF g_deriv] continuous_on_subset[OF g'_cont] continuous_intros integrable_continuous_real continuous_on_subset[OF vderiv_on_continuous_on[OF g_deriv_on]]) have "\t. t \ {a..b} \ ((\t. g' t / g t) has_integral ln (g b) - ln (g t)) {t..b}" by (rule fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros assms has_field_derivative_subset[OF g_deriv] - simp: has_field_derivative_iff_has_vector_derivative[symmetric]) + simp: has_real_derivative_iff_has_vector_derivative[symmetric]) hence *: "\t. t \ {a..b} \ ln (g b) - ln (g t) = integral {t..b} (\t. g' t / g t)" using integrable_integral[OF frac_integrable] by (rule has_integral_unique[where f = "\t. g' t / g t"]) have "K * (b - t) = integral {t..b} (\_. K)" using \t \ {a..b}\ by simp also have "... \ integral {t..b} (\t. g' t / g t)" using \t \ {a..b}\ by (intro integral_le) (auto intro!: frac_integrable frac_ge integral_le) also have "... = ln (g b) - ln (g t)" using * t_in by simp finally have "K * (b - t) + ln (g t) \ ln (g b)" (is "?lhs \ ?rhs") by simp hence "exp ?lhs \ exp ?rhs" by simp hence "g t * exp (K * (b - t)) \ g b" using \t \ {a..b}\ g_pos by (simp add: ac_simps exp_add del: exp_le_cancel_iff) hence "g t / exp (K * (t - b)) \ g b" by (simp add: algebra_simps exp_diff) thus ?thesis by (simp add: field_simps) qed lemma gronwall_general: fixes g K C a b and t::real defines "G \ \t. C + K * integral {a..t} (\s. g s)" assumes g_le_G: "\t. t \ {a..b} \ g t \ G t" assumes g_cont: "continuous_on {a..b} g" assumes g_nonneg: "\t. t \ {a..b} \ 0 \ g t" assumes pos: "0 < C" "K > 0" assumes "t \ {a..b}" shows "g t \ C * exp (K * (t - a))" proof - have G_pos: "\t. t \ {a..b} \ 0 < G t" by (auto simp: G_def intro!: add_pos_nonneg mult_nonneg_nonneg Henstock_Kurzweil_Integration.integral_nonneg integrable_continuous_real assms intro: less_imp_le continuous_on_subset) have "g t \ G t" using assms by auto also { have "(G has_vderiv_on (\t. K * g t)) {a..b}" by (auto intro!: derivative_eq_intros integral_has_vector_derivative g_cont simp add: G_def has_vderiv_on_def) moreover { fix t assume "t \ {a..b}" hence "K * g t / G t \ K * G t / G t" using pos g_le_G G_pos by (intro divide_right_mono mult_left_mono) (auto intro!: less_imp_le) also have "\ = K" using G_pos[of t] \t \ {a .. b}\ by simp finally have "K * g t / G t \ K" . } ultimately have "G t \ G a * exp (K * (t - a))" apply (rule derivative_quotient_bound) using \t \ {a..b}\ by (auto intro!: continuous_intros g_cont G_pos simp: field_simps pos) } also have "G a = C" by (simp add: G_def) finally show ?thesis by simp qed lemma gronwall_general_left: fixes g K C a b and t::real defines "G \ \t. C + K * integral {t..b} (\s. g s)" assumes g_le_G: "\t. t \ {a..b} \ g t \ G t" assumes g_cont: "continuous_on {a..b} g" assumes g_nonneg: "\t. t \ {a..b} \ 0 \ g t" assumes pos: "0 < C" "K > 0" assumes "t \ {a..b}" shows "g t \ C * exp (-K * (t - b))" proof - have G_pos: "\t. t \ {a..b} \ 0 < G t" by (auto simp: G_def intro!: add_pos_nonneg mult_nonneg_nonneg Henstock_Kurzweil_Integration.integral_nonneg integrable_continuous_real assms intro: less_imp_le continuous_on_subset) have "g t \ G t" using assms by auto also { have "(G has_vderiv_on (\t. -K * g t)) {a..b}" by (auto intro!: derivative_eq_intros g_cont integral_has_vector_derivative' simp add: G_def has_vderiv_on_def) moreover { fix t assume "t \ {a..b}" hence "K * g t / G t \ K * G t / G t" using pos g_le_G G_pos by (intro divide_right_mono mult_left_mono) (auto intro!: less_imp_le) also have "\ = K" using G_pos[of t] \t \ {a .. b}\ by simp finally have "K * g t / G t \ K" . hence "-K \ -K * g t / G t" by simp } ultimately have "G t \ G b * exp (-K * (t - b))" apply (rule derivative_quotient_bound_left) using \t \ {a..b}\ by (auto intro!: continuous_intros g_cont G_pos simp: field_simps pos) } also have "G b = C" by (simp add: G_def) finally show ?thesis by simp qed lemma gronwall_general_segment: fixes a b::real assumes "\t. t \ closed_segment a b \ g t \ C + K * integral (closed_segment a t) g" and "continuous_on (closed_segment a b) g" and "\t. t \ closed_segment a b \ 0 \ g t" and "0 < C" and "0 < K" and "t \ closed_segment a b" shows "g t \ C * exp (K * abs (t - a))" proof cases assume "a \ b" then have *: "abs (t - a) = t -a" using assms by (auto simp: closed_segment_eq_real_ivl) show ?thesis unfolding * using assms by (intro gronwall_general[where b=b]) (auto intro!: simp: closed_segment_eq_real_ivl \a \ b\) next assume "\a \ b" then have *: "K * abs (t - a) = - K * (t - a)" using assms by (auto simp: closed_segment_eq_real_ivl algebra_simps) { fix s :: real assume a1: "b \ s" assume a2: "s \ a" assume a3: "\t. b \ t \ t \ a \ g t \ C + K * integral (if a \ t then {a..t} else {t..a}) g" have "s = a \ s < a" using a2 by (meson less_eq_real_def) then have "g s \ C + K * integral {s..a} g" using a3 a1 by fastforce } then show ?thesis unfolding * using assms \\a \ b\ by (intro gronwall_general_left) (auto intro!: simp: closed_segment_eq_real_ivl) qed lemma gronwall_more_general_segment: fixes a b c::real assumes "\t. t \ closed_segment a b \ g t \ C + K * integral (closed_segment c t) g" and cont: "continuous_on (closed_segment a b) g" and "\t. t \ closed_segment a b \ 0 \ g t" and "0 < C" and "0 < K" and t: "t \ closed_segment a b" and c: "c \ closed_segment a b" shows "g t \ C * exp (K * abs (t - c))" proof - from t c have "t \ closed_segment c a \ t \ closed_segment c b" by (auto simp: closed_segment_eq_real_ivl split_ifs) then show ?thesis proof assume "t \ closed_segment c a" moreover have subs: "closed_segment c a \ closed_segment a b" using t c by (auto simp: closed_segment_eq_real_ivl split_ifs) ultimately show ?thesis by (intro gronwall_general_segment[where b=a]) (auto intro!: assms intro: continuous_on_subset) next assume "t \ closed_segment c b" moreover have subs: "closed_segment c b \ closed_segment a b" using t c by (auto simp: closed_segment_eq_real_ivl) ultimately show ?thesis by (intro gronwall_general_segment[where b=b]) (auto intro!: assms intro: continuous_on_subset) qed qed lemma gronwall: fixes g K C and t::real defines "G \ \t. C + K * integral {0..t} (\s. g s)" assumes g_le_G: "\t. 0 \ t \ t \ a \ g t \ G t" assumes g_cont: "continuous_on {0..a} g" assumes g_nonneg: "\t. 0 \ t \ t \ a \ 0 \ g t" assumes pos: "0 < C" "0 < K" assumes "0 \ t" "t \ a" shows "g t \ C * exp (K * t)" apply(rule gronwall_general[where a=0, simplified, OF assms(2-6)[unfolded G_def]]) using assms(7,8) by simp_all lemma gronwall_left: fixes g K C and t::real defines "G \ \t. C + K * integral {t..0} (\s. g s)" assumes g_le_G: "\t. a \ t \ t \ 0 \ g t \ G t" assumes g_cont: "continuous_on {a..0} g" assumes g_nonneg: "\t. a \ t \ t \ 0 \ 0 \ g t" assumes pos: "0 < C" "0 < K" assumes "a \ t" "t \ 0" shows "g t \ C * exp (-K * t)" apply(simp, rule gronwall_general_left[where b=0, simplified, OF assms(2-6)[unfolded G_def]]) using assms(7,8) by simp_all end \ No newline at end of file diff --git a/thys/Ordinary_Differential_Equations/Library/Vector_Derivative_On.thy b/thys/Ordinary_Differential_Equations/Library/Vector_Derivative_On.thy --- a/thys/Ordinary_Differential_Equations/Library/Vector_Derivative_On.thy +++ b/thys/Ordinary_Differential_Equations/Library/Vector_Derivative_On.thy @@ -1,269 +1,269 @@ theory Vector_Derivative_On imports "HOL-Analysis.Analysis" begin subsection \Vector derivative on a set\ \ \TODO: also for the other derivatives?!\ \ \TODO: move to repository and rewrite assumptions of common lemmas?\ definition has_vderiv_on :: "(real \ 'a::real_normed_vector) \ (real \ 'a) \ real set \ bool" (infix "(has'_vderiv'_on)" 50) where "(f has_vderiv_on f') S \ (\x \ S. (f has_vector_derivative f' x) (at x within S))" lemma has_vderiv_on_empty[intro, simp]: "(f has_vderiv_on f') {}" by (auto simp: has_vderiv_on_def) lemma has_vderiv_on_subset: assumes "(f has_vderiv_on f') S" assumes "T \ S" shows "(f has_vderiv_on f') T" by (meson assms(1) assms(2) contra_subsetD has_vderiv_on_def has_vector_derivative_within_subset) lemma has_vderiv_on_compose: assumes "(f has_vderiv_on f') (g ` T)" assumes "(g has_vderiv_on g') T" shows "(f o g has_vderiv_on (\x. g' x *\<^sub>R f' (g x))) T" using assms unfolding has_vderiv_on_def by (auto intro!: vector_diff_chain_within) lemma has_vderiv_on_open: assumes "open T" shows "(f has_vderiv_on f') T \ (\t \ T. (f has_vector_derivative f' t) (at t))" by (auto simp: has_vderiv_on_def at_within_open[OF _ \open T\]) lemma has_vderiv_on_eq_rhs:\ \TODO: integrate intro \derivative_eq_intros\\ "(f has_vderiv_on g') T \ (\x. x \ T \ g' x = f' x) \ (f has_vderiv_on f') T" by (auto simp: has_vderiv_on_def) lemma [THEN has_vderiv_on_eq_rhs, derivative_intros]: shows has_vderiv_on_id: "((\x. x) has_vderiv_on (\x. 1)) T" and has_vderiv_on_const: "((\x. c) has_vderiv_on (\x. 0)) T" by (auto simp: has_vderiv_on_def intro!: derivative_eq_intros) lemma [THEN has_vderiv_on_eq_rhs, derivative_intros]: fixes f::"real \ 'a::real_normed_vector" assumes "(f has_vderiv_on f') T" shows has_vderiv_on_uminus: "((\x. - f x) has_vderiv_on (\x. - f' x)) T" using assms by (auto simp: has_vderiv_on_def intro!: derivative_eq_intros) lemma [THEN has_vderiv_on_eq_rhs, derivative_intros]: fixes f g::"real \ 'a::real_normed_vector" assumes "(f has_vderiv_on f') T" assumes "(g has_vderiv_on g') T" shows has_vderiv_on_add: "((\x. f x + g x) has_vderiv_on (\x. f' x + g' x)) T" and has_vderiv_on_diff: "((\x. f x - g x) has_vderiv_on (\x. f' x - g' x)) T" using assms by (auto simp: has_vderiv_on_def intro!: derivative_eq_intros) lemma [THEN has_vderiv_on_eq_rhs, derivative_intros]: fixes f::"real \ real" and g::"real \ 'a::real_normed_vector" assumes "(f has_vderiv_on f') T" assumes "(g has_vderiv_on g') T" shows has_vderiv_on_scaleR: "((\x. f x *\<^sub>R g x) has_vderiv_on (\x. f x *\<^sub>R g' x + f' x *\<^sub>R g x)) T" using assms - by (auto simp: has_vderiv_on_def has_field_derivative_iff_has_vector_derivative + by (auto simp: has_vderiv_on_def has_real_derivative_iff_has_vector_derivative intro!: derivative_eq_intros) lemma [THEN has_vderiv_on_eq_rhs, derivative_intros]: fixes f g::"real \ 'a::real_normed_algebra" assumes "(f has_vderiv_on f') T" assumes "(g has_vderiv_on g') T" shows has_vderiv_on_mult: "((\x. f x * g x) has_vderiv_on (\x. f x * g' x + f' x * g x)) T" using assms by (auto simp: has_vderiv_on_def intro!: derivative_eq_intros) lemma has_vderiv_on_ln[THEN has_vderiv_on_eq_rhs, derivative_intros]: fixes g::"real \ real" assumes "\x. x \ s \ 0 < g x" assumes "(g has_vderiv_on g') s" shows "((\x. ln (g x)) has_vderiv_on (\x. g' x / g x)) s" using assms unfolding has_vderiv_on_def - by (auto simp: has_vderiv_on_def has_field_derivative_iff_has_vector_derivative[symmetric] + by (auto simp: has_vderiv_on_def has_real_derivative_iff_has_vector_derivative[symmetric] intro!: derivative_eq_intros) lemma fundamental_theorem_of_calculus': fixes f :: "real \ 'a::banach" shows "a \ b \ (f has_vderiv_on f') {a .. b} \ (f' has_integral (f b - f a)) {a .. b}" by (auto intro!: fundamental_theorem_of_calculus simp: has_vderiv_on_def) lemma has_vderiv_on_If: assumes "U = S \ T" assumes "(f has_vderiv_on f') (S \ (closure T \ closure S))" assumes "(g has_vderiv_on g') (T \ (closure T \ closure S))" assumes "\x. x \ closure T \ x \ closure S \ f x = g x" assumes "\x. x \ closure T \ x \ closure S \ f' x = g' x" shows "((\t. if t \ S then f t else g t) has_vderiv_on (\t. if t \ S then f' t else g' t)) U" using assms by (auto simp: has_vderiv_on_def ac_simps intro!: has_vector_derivative_If_within_closures split del: if_split) lemma mvt_very_simple_closed_segmentE: fixes f::"real\real" assumes "(f has_vderiv_on f') (closed_segment a b)" obtains y where "y \ closed_segment a b" "f b - f a = (b - a) * f' y" proof cases assume "a \ b" with mvt_very_simple[of a b f "\x i. i *\<^sub>R f' x"] assms obtain y where "y \ closed_segment a b" "f b - f a = (b - a) * f' y" by (auto simp: has_vector_derivative_def closed_segment_eq_real_ivl has_vderiv_on_def) thus ?thesis .. next assume "\ a \ b" with mvt_very_simple[of b a f "\x i. i *\<^sub>R f' x"] assms obtain y where "y \ closed_segment a b" "f b - f a = (b - a) * f' y" by (force simp: has_vector_derivative_def has_vderiv_on_def closed_segment_eq_real_ivl algebra_simps) thus ?thesis .. qed lemma mvt_simple_closed_segmentE: fixes f::"real\real" assumes "(f has_vderiv_on f') (closed_segment a b)" assumes "a \ b" obtains y where "y \ open_segment a b" "f b - f a = (b - a) * f' y" proof cases assume "a \ b" with assms have "a < b" by simp with mvt_simple[of a b f "\x i. i *\<^sub>R f' x"] assms obtain y where "y \ open_segment a b" "f b - f a = (b - a) * f' y" by (auto simp: has_vector_derivative_def closed_segment_eq_real_ivl has_vderiv_on_def open_segment_eq_real_ivl) thus ?thesis .. next assume "\ a \ b" then have "b < a" by simp with mvt_simple[of b a f "\x i. i *\<^sub>R f' x"] assms obtain y where "y \ open_segment a b" "f b - f a = (b - a) * f' y" by (force simp: has_vector_derivative_def has_vderiv_on_def closed_segment_eq_real_ivl algebra_simps open_segment_eq_real_ivl) thus ?thesis .. qed lemma differentiable_bound_general_open_segment: fixes a :: "real" and b :: "real" and f :: "real \ 'a::real_normed_vector" and f' :: "real \ 'a" assumes "continuous_on (closed_segment a b) f" assumes "continuous_on (closed_segment a b) g" and "(f has_vderiv_on f') (open_segment a b)" and "(g has_vderiv_on g') (open_segment a b)" and "\x. x \ open_segment a b \ norm (f' x) \ g' x" shows "norm (f b - f a) \ abs (g b - g a)" proof - { assume "a = b" hence ?thesis by simp } moreover { assume "a < b" with assms have "continuous_on {a .. b} f" and "continuous_on {a .. b} g" and "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" and "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" and "\x. x\{a<.. norm (f' x) \ g' x" by (auto simp: open_segment_eq_real_ivl closed_segment_eq_real_ivl has_vderiv_on_def at_within_open[where S="{a<..a < b\ this] have ?thesis by auto } moreover { assume "b < a" with assms have "continuous_on {b .. a} f" and "continuous_on {b .. a} g" and "\x. x\{b<.. (f has_vector_derivative f' x) (at x)" and "\x. x\{b<.. (g has_vector_derivative g' x) (at x)" and "\x. x\{b<.. norm (f' x) \ g' x" by (auto simp: open_segment_eq_real_ivl closed_segment_eq_real_ivl has_vderiv_on_def at_within_open[where S="{b<..b < a\ this] have "norm (f a - f b) \ g a - g b" by simp also have "\ \ abs (g b - g a)" by simp finally have ?thesis by (simp add: norm_minus_commute) } ultimately show ?thesis by arith qed lemma has_vderiv_on_union: assumes "(f has_vderiv_on g) (s \ closure s \ closure t)" assumes "(f has_vderiv_on g) (t \ closure s \ closure t)" shows "(f has_vderiv_on g) (s \ t)" unfolding has_vderiv_on_def proof fix x assume "x \ s \ t" with has_vector_derivative_If_within_closures[of x s t "s \ t" f g f g] assms show "(f has_vector_derivative g x) (at x within s \ t)" by (auto simp: has_vderiv_on_def) qed lemma has_vderiv_on_union_closed: assumes "(f has_vderiv_on g) s" assumes "(f has_vderiv_on g) t" assumes "closed s" "closed t" shows "(f has_vderiv_on g) (s \ t)" using has_vderiv_on_If[OF refl, of f g s t f g] assms by (auto simp: has_vderiv_on_subset) lemma vderiv_on_continuous_on: "(f has_vderiv_on f') S \ continuous_on S f" by (auto intro!: continuous_on_vector_derivative simp: has_vderiv_on_def) lemma has_vderiv_on_cong[cong]: assumes "\x. x \ S \ f x = g x" assumes "\x. x \ S \ f' x = g' x" assumes "S = T" shows "(f has_vderiv_on f') S = (g has_vderiv_on g') T" using assms by (metis has_vector_derivative_transform has_vderiv_on_def) lemma has_vderiv_eq: assumes "(f has_vderiv_on f') S" assumes "\x. x \ S \ f x = g x" assumes "\x. x \ S \ f' x = g' x" assumes "S = T" shows "(g has_vderiv_on g') T" using assms by simp lemma has_vderiv_on_compose': assumes "(f has_vderiv_on f') (g ` T)" assumes "(g has_vderiv_on g') T" shows "((\x. f (g x)) has_vderiv_on (\x. g' x *\<^sub>R f' (g x))) T" using has_vderiv_on_compose[OF assms] by simp lemma has_vderiv_on_compose2: assumes "(f has_vderiv_on f') S" assumes "(g has_vderiv_on g') T" assumes "\t. t \ T \ g t \ S" shows "((\x. f (g x)) has_vderiv_on (\x. g' x *\<^sub>R f' (g x))) T" using has_vderiv_on_compose[OF has_vderiv_on_subset[OF assms(1)] assms(2)] assms(3) by force lemma has_vderiv_on_singleton: "(y has_vderiv_on y') {t0}" by (auto simp: has_vderiv_on_def has_vector_derivative_def has_derivative_within_singleton_iff bounded_linear_scaleR_left) lemma has_vderiv_on_zero_constant: assumes "convex s" assumes "(f has_vderiv_on (\h. 0)) s" obtains c where "\x. x \ s \ f x = c" using has_vector_derivative_zero_constant[of s f] assms by (auto simp: has_vderiv_on_def) lemma bounded_vderiv_on_imp_lipschitz: assumes "(f has_vderiv_on f') X" assumes convex: "convex X" assumes "\x. x \ X \ norm (f' x) \ C" "0 \ C" shows "C-lipschitz_on X f" using assms by (auto simp: has_vderiv_on_def has_vector_derivative_def onorm_scaleR_left onorm_id intro!: bounded_derivative_imp_lipschitz[where f' = "\x d. d *\<^sub>R f' x"]) end diff --git a/thys/Prime_Distribution_Elementary/Moebius_Mu_Sum.thy b/thys/Prime_Distribution_Elementary/Moebius_Mu_Sum.thy --- a/thys/Prime_Distribution_Elementary/Moebius_Mu_Sum.thy +++ b/thys/Prime_Distribution_Elementary/Moebius_Mu_Sum.thy @@ -1,511 +1,511 @@ (* File: Moebius_Mu_Sum.thy Author: Manuel Eberl, TU München Properties of the summatory Möbius \ function, including its relationship to the Prime Number Theorem. *) section \The summatory Möbius \\\ function\ theory Moebius_Mu_Sum imports More_Dirichlet_Misc Dirichlet_Series.Partial_Summation Prime_Number_Theorem.Prime_Counting_Functions Dirichlet_Series.Arithmetic_Summatory_Asymptotics Shapiro_Tauberian Partial_Zeta_Bounds Prime_Number_Theorem.Prime_Number_Theorem_Library Prime_Distribution_Elementary_Library begin text \ In this section, we shall examine the summatory Möbius \\\ function $M(x) := \sum_{n\leq x} \mu(n)$. The main result is that $M(x) \in o(x)$ is equivalent to the Prime Number Theorem. \ context includes prime_counting_notation fixes M H :: "real \ real" defines "M \ sum_upto moebius_mu" defines "H \ sum_upto (\n. moebius_mu n * ln n)" begin lemma sum_upto_moebius_mu_integral: "x > 1 \ ((\t. M t / t) has_integral M x * ln x - H x) {1..x}" and sum_upto_moebius_mu_integrable: "a \ 1 \ (\t. M t / t) integrable_on {a..b}" proof - { fix a b :: real assume ab: "a \ 1" "a < b" have "((\t. M t * (1 / t)) has_integral M b * ln b - M a * ln a - (\n\real -` {a<..b}. moebius_mu n * ln (real n))) {a..b}" unfolding M_def using ab by (intro partial_summation_strong [where X = "{}"]) (auto intro!: derivative_eq_intros continuous_intros - simp flip: has_field_derivative_iff_has_vector_derivative) + simp flip: has_real_derivative_iff_has_vector_derivative) } note * = this { fix x :: real assume x: "x > 1" have "(\n\real -` {1<..x}. moebius_mu n * ln (real n)) = H x" unfolding H_def sum_upto_def by (intro sum.mono_neutral_cong_left) (use x in auto) thus "((\t. M t / t) has_integral M x * ln x - H x) {1..x}" using *[of 1 x] x by simp } { fix a b :: real assume ab: "a \ 1" show "(\t. M t / t) integrable_on {a..b}" using *[of a b] ab by (cases a b rule: linorder_cases) (auto intro: integrable_negligible) } qed lemma sum_moebius_mu_bound: assumes "x \ 0" shows "\M x\ \ x" proof - have "\M x\ \ sum_upto (\n. \moebius_mu n\) x" unfolding M_def sum_upto_def by (rule sum_abs) also have "\ \ sum_upto (\n. 1) x" unfolding sum_upto_def by (intro sum_mono) (auto simp: moebius_mu_def) also have "\ \ x" using assms by (simp add: sum_upto_altdef) finally show ?thesis . qed lemma sum_moebius_mu_aux1: "(\x. M x / x - H x / (x * ln x)) \ O(\x. 1 / ln x)" proof - define R where "R = (\x. integral {1..x} (\t. M t / t))" have "eventually (\x. M x / x - H x / (x * ln x) = R x / (x * ln x)) at_top" using eventually_gt_at_top[of 1] proof eventually_elim case (elim x) thus ?case using sum_upto_moebius_mu_integral[of x] by (simp add: R_def has_integral_iff field_simps) qed hence "(\x. M x / x - H x / (x * ln x)) \ \(\x. R x / (x * ln x))" by (intro bigthetaI_cong) also have "(\x. R x / (x * ln x)) \ O(\x. x / (x * ln x))" proof (intro landau_o.big.divide_right) have "M \ O(\x. x)" using sum_moebius_mu_bound by (intro bigoI[where c = 1] eventually_mono[OF eventually_ge_at_top[of 0]]) auto hence "(\t. M t / t) \ O(\t. 1)" by (simp add: landau_divide_simps) thus "R \ O(\x. x)" unfolding R_def by (intro integral_bigo[where g' = "\_. 1"]) (auto simp: filterlim_ident has_integral_iff intro!: sum_upto_moebius_mu_integrable) qed (intro eventually_mono[OF eventually_gt_at_top[of 1]], auto) also have "(\x::real. x / (x * ln x)) \ \(\x. 1 / ln x)" by real_asymp finally show ?thesis . qed (* 4.13 *) lemma sum_moebius_mu_aux2: "((\x. M x / x - H x / (x * ln x)) \ 0) at_top" proof - have "(\x. M x / x - H x / (x * ln x)) \ O(\x. 1 / ln x)" by (rule sum_moebius_mu_aux1) also have "(\x. 1 / ln x) \ o(\_. 1 :: real)" by real_asymp finally show ?thesis by (auto dest!: smalloD_tendsto) qed lemma sum_moebius_mu_ln_eq: "H = (\x. - dirichlet_prod' moebius_mu \ x)" proof fix x :: real have "fds mangoldt = (fds_deriv (fds moebius_mu) * fds_zeta :: real fds)" using fds_mangoldt' by (simp add: mult_ac) hence eq: "fds_deriv (fds moebius_mu) = fds moebius_mu * (fds mangoldt :: real fds)" by (subst (asm) fds_moebius_inversion [symmetric]) have "-H x = sum_upto (\n. -ln n * moebius_mu n) x" by (simp add: H_def sum_upto_def sum_negf mult_ac) also have "\ = sum_upto (\n. dirichlet_prod moebius_mu mangoldt n) x" using eq by (intro sum_upto_cong) (auto simp: fds_eq_iff fds_nth_deriv fds_nth_mult) also have "\ = dirichlet_prod' moebius_mu \ x" by (subst sum_upto_dirichlet_prod) (simp add: primes_psi_def dirichlet_prod'_def) finally show "H x = -dirichlet_prod' moebius_mu \ x" by simp qed (* 4.14 *) theorem PNT_implies_sum_moebius_mu_sublinear: assumes "\ \[at_top] (\x. x)" shows "M \ o(\x. x)" proof - have "((\x. H x / (x * ln x)) \ 0) at_top" proof (rule tendstoI) fix \' :: real assume \': "\' > 0" define \ where "\ = \' / 2" from \' have \: "\ > 0" by (simp add: \_def) from assms have "((\x. \ x / x) \ 1) at_top" by (elim asymp_equivD_strong) (auto intro!: eventually_mono[OF eventually_gt_at_top[of 0]]) from tendstoD[OF this \] have "eventually (\x. \\ x / x - 1\ < \) at_top" by (simp add: dist_norm) hence "eventually (\x. \\ x - x\ < \ * x) at_top" using eventually_gt_at_top[of 0] by eventually_elim (auto simp: abs_if field_simps) then obtain A' where A': "\x. x \ A' \ \\ x - x\ < \ * x" by (auto simp: eventually_at_top_linorder) define A where "A = max 2 A'" from A' have A: "A \ 2" "\x. x \ A \ \\ x - x\ < \ * x" by (auto simp: A_def) have H_bound: "\H x\ / (x * ln x) \ (1 + \ + \ A) / ln x + \" if "x \ A" for x proof - from \x \ A\ have "x \ 2" using A(1) by linarith note x = \x \ A\ \x \ 2\ define y where "y = nat \floor (x / A)\" have "real y = real_of_int \x / A\" using A x by (simp add: y_def) also have "real_of_int \x / A\ \ x / A" by linarith also have "\ \ x" using x A(1) by (simp add: field_simps) finally have "y \ x" . have "y \ 1" using x A(1) by (auto simp: y_def le_nat_iff le_floor_iff) note y = \y \ 1\ \y \ x\ define S1 where [simp]: "S1 = sum_upto (\m. moebius_mu m * \ (x / m)) y" define S2 where [simp]: "S2 = (\m | m > y \ real m \ x. moebius_mu m * \ (x / m))" have fin: "finite {m. y < m \ real m \ x}" by (rule finite_subset[of _ "{..nat \x\}"]) (auto simp: le_nat_iff le_floor_iff) have "H x = -dirichlet_prod' moebius_mu \ x" by (simp add: sum_moebius_mu_ln_eq) also have "dirichlet_prod' moebius_mu \ x = (\m | m > 0 \ real m \ x. moebius_mu m * \ (x / m))" unfolding dirichlet_prod'_def sum_upto_def .. also have "{m. m > 0 \ real m \ x} = {0<..y} \ {m. y < m \ real m \ x}" using x y A(1) by auto also have "(\m\\. moebius_mu m * \ (x / m)) = S1 + S2" unfolding dirichlet_prod'_def sum_upto_def S1_def S2_def using fin by (subst sum.union_disjoint) (auto intro: sum.cong) finally have abs_H_eq: "\H x\ = \S1 + S2\" by simp define S1_1 where [simp]: "S1_1 = sum_upto (\m. moebius_mu m / m) y" define S1_2 where [simp]: "S1_2 = sum_upto (\m. moebius_mu m * (\ (x / m) - x / m)) y" have "\S1\ = \x * S1_1 + S1_2\" by (simp add: sum_upto_def sum_distrib_left sum_distrib_right mult_ac sum_subtractf ring_distribs) also have "\ \ x * \S1_1\ + \S1_2\" by (rule order.trans[OF abs_triangle_ineq]) (use x in \simp add: abs_mult\) also have "\ \ x * 1 + \ * x * (ln x + 1)" proof (intro add_mono mult_left_mono) show "\S1_1\ \ 1" using abs_sum_upto_moebius_mu_over_n_le[of y] by simp next have "\S1_2\ \ sum_upto (\m. \moebius_mu m * (\ (x / m) - x / m)\) y" unfolding S1_2_def sum_upto_def by (rule sum_abs) also have "\ \ sum_upto (\m. 1 * (\ * (x / m))) y" unfolding abs_mult sum_upto_def proof (intro sum_mono mult_mono less_imp_le[OF A(2)]) fix m assume m: "m \ {i. 0 < i \ real i \ real y}" hence "real m \ real y" by simp also from x A(1) have "\ = of_int \x / A\" by (simp add: y_def) also have "\ \ x / A" by linarith finally show "A \ x / real m" using A(1) m by (simp add: field_simps) qed (auto simp: moebius_mu_def field_simps) also have "\ = \ * x * (\i\{0<..y}. inverse (real i))" by (simp add: sum_upto_altdef sum_distrib_left divide_simps) also have "(\i\{0<..y}. inverse (real i)) = harm (nat \y\)" unfolding harm_def by (intro sum.cong) auto also have "\ \ ln (nat \y\) + 1" by (rule harm_le) (use y in auto) also have "ln (nat \y\) \ ln x" using y by simp finally show "\S1_2\ \ \ * x * (ln x + 1)" using \ x by simp qed (use x in auto) finally have S1_bound: "\S1\ \ x + \ * x * ln x + \ * x" by (simp add: algebra_simps) have "\S2\ \ (\m | y < m \ real m \ x. \moebius_mu m * \ (x / m)\)" unfolding S2_def by (rule sum_abs) also have "\ \ (\m | y < m \ real m \ x. 1 * \ A)" unfolding abs_mult using y proof (intro sum_mono mult_mono) fix m assume m: "m \ {m. y < m \ real m \ x}" hence "y < m" by simp moreover have "y = of_int \x / A\" using x A(1) by (simp add: y_def) ultimately have "\x / A\ < m" by simp hence "x / A \ real m" by linarith hence "\ (x / real m) \ \ A" using m A(1) by (intro \_mono) (auto simp: field_simps) thus "\\ (x / real m)\ \ \ A" by (simp add: \_nonneg) qed (auto simp: moebius_mu_def \_nonneg field_simps intro!: \_mono) also have "\ \ sum_upto (\_. 1 * \ A) x" unfolding sum_upto_def by (intro sum_mono2) auto also have "\ = real (nat \x\) * \ A" by (simp add: sum_upto_altdef) also have "\ \ x * \ A" using x by (intro mult_right_mono) auto finally have S2_bound: "\S2\ \ x * \ A" . have "\H x\ \ \S1\ + \S2\" using abs_H_eq by linarith also have "\ \ x + \ * x * ln x + \ * x + x * \ A" by (intro add_mono S1_bound S2_bound) finally have "\H x\ \ (1 + \ + \ A) * x + \ * x * ln x" by (simp add: algebra_simps) thus "\H x\ / (x * ln x) \ (1 + \ + \ A) / ln x + \" using x by (simp add: field_simps) qed have "eventually (\x. \H x\ / (x * ln x) \ (1 + \ + \ A) / ln x + \) at_top" using eventually_ge_at_top[of A] by eventually_elim (use H_bound in auto) moreover have "eventually (\x. (1 + \ + \ A) / ln x + \ < \') at_top" unfolding \_def using \' by real_asymp moreover have "eventually (\x. \H x\ / (x * ln x) = \H x / (x * ln x)\) at_top" using eventually_gt_at_top[of 1] by eventually_elim (simp add: abs_mult) ultimately have "eventually (\x. \H x / (x * ln x)\ < \') at_top" by eventually_elim simp thus "eventually (\x. dist (H x / (x * ln x)) 0 < \') at_top" by (simp add: dist_norm) qed hence "(\x. H x / (x * ln x)) \ o(\_. 1)" by (intro smalloI_tendsto) auto hence "(\x. H x / (x * ln x) + (M x / x - H x / (x * ln x))) \ o(\_. 1)" proof (rule sum_in_smallo) have "(\x. M x / x - H x / (x * ln x)) \ O(\x. 1 / ln x)" by (rule sum_moebius_mu_aux1) also have "(\x::real. 1 / ln x) \ o(\_. 1)" by real_asymp finally show "(\x. M x / x - H x / (x * ln x)) \ o(\_. 1)" . qed thus ?thesis by (simp add: landau_divide_simps) qed (* 4.15 *) theorem sum_moebius_mu_sublinear_imp_PNT: assumes "M \ o(\x. x)" shows "\ \[at_top] (\x. x)" proof - define \ :: "nat \ real" where [simp]: "\ = (\n. real (divisor_count n))" define C where [simp]: "C = (euler_mascheroni :: real)" define f :: "nat \ real" where "f = (\n. \ n - ln n - 2 * C)" define F where [simp]: "F = sum_upto f" write moebius_mu ("\") \ \The proof is based on the fact that $\psi(x)-x$ can be approximated fairly well by the Dirichlet product $\sum_{n\leq x} \sum_{d\mid n} \mu(d) f(n/d)$:\ have eq: "\ x - x = -sum_upto (dirichlet_prod \ f) x - frac x - 2 * C" if x: "x \ 1" for x proof - have "\x\ - \ x - 2 * C = sum_upto (\_. 1) x - sum_upto mangoldt x - sum_upto (\n. if n = 1 then 2 * C else 0) x" using x by (simp add: sum_upto_altdef \_def le_nat_iff le_floor_iff) also have "\ = sum_upto (\n. 1 - mangoldt n - (if n = 1 then 2 * C else 0)) x" by (simp add: sum_upto_def sum_subtractf) also have "\ = sum_upto (dirichlet_prod \ f) x" by (intro sum_upto_cong refl moebius_inversion) (auto simp: divisor_count_def sum_subtractf mangoldt_sum f_def) finally show "\ x - x = -sum_upto (dirichlet_prod \ f) x - frac x - 2 * C" by (simp add: algebra_simps frac_def) qed \ \We now obtain a bound of the form \\F x\ \ B * sqrt x\.\ have "F \ O(sqrt)" proof - have "F \ \(\x. (sum_upto \ x - (x * ln x + (2 * C - 1) * x)) - (sum_upto ln x - x * ln x + x) + 2 * C * frac x)" (is "_ \ \(?rhs)") by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) (auto simp: sum_upto_altdef sum_subtractf f_def frac_def algebra_simps sum.distrib) also have "?rhs \ O(sqrt)" proof (rule sum_in_bigo, rule sum_in_bigo) show "(\x. sum_upto \ x - (x * ln x + (2 * C - 1) * x)) \ O(sqrt)" unfolding C_def \_def by (rule summatory_divisor_count_asymptotics) show "(\x. sum_upto (\x. ln (real x)) x - x * ln x + x) \ O(sqrt)" by (rule landau_o.big.trans[OF sum_upto_ln_stirling_weak_bigo]) real_asymp qed (use euler_mascheroni_pos in real_asymp) finally show ?thesis . qed hence "(\n. F (real n)) \ O(sqrt)" by (rule landau_o.big.compose) real_asymp from natfun_bigoE[OF this, of 1] obtain B :: real where B: "B > 0" "\n. n \ 1 \ \F (real n)\ \ B * sqrt (real n)" by auto have B': "\F x\ \ B * sqrt x" if "x \ 1" for x proof - have "\F x\ \ B * sqrt (nat \x\)" using B(2)[of "nat \x\"] that by (simp add: sum_upto_altdef le_nat_iff le_floor_iff) also have "\ \ B * sqrt x" using B(1) that by (intro mult_left_mono) auto finally show ?thesis . qed \ \Next, we obtain a good bound for $\sum_{n\leq x} \frac{1}{\sqrt{n}}$.\ from zeta_partial_sum_le_pos''[of "1 / 2"] obtain A where A: "A > 0" "\x. x \ 1 \ \sum_upto (\n. 1 / sqrt n) x\ \ A * sqrt x" by (auto simp: max_def powr_half_sqrt powr_minus field_simps) \ \Finally, we show that $\sum_{n\leq x} \sum_{d\mid n} \mu(d) f(n/d) \in o(x)$.\ have "sum_upto (dirichlet_prod \ f) \ o(\x. x)" proof (rule landau_o.smallI) fix \ :: real assume \: "\ > 0" have *: "eventually (\x. \sum_upto (dirichlet_prod \ f) x\ \ \ * x) at_top" if b: "b \ 1" "A * B / sqrt b \ \ / 3" "B / sqrt b \ \ / 3" for b proof - define K :: real where "K = sum_upto (\n. \f n\ / n) b" have "C \ (1 / 2)" using euler_mascheroni_gt_19_over_33 by auto hence K: "K > 0" unfolding K_def f_def sum_upto_def by (intro sum_pos2[where i = 1]) (use \b \ 1\ in auto) have "eventually (\x. \M x / x\ < \ / 3 / K) at_top" using smalloD_tendsto[OF assms] \ K by (auto simp: tendsto_iff dist_norm) then obtain c' where c': "\x. x \ c' \ \M x / x\ < \ / 3 / K" by (auto simp: eventually_at_top_linorder) define c where "c = max 1 c'" have c: "\M x\ < \ / 3 / K * x" if "x \ c" for x using c'[of x] that by (simp add: c_def field_simps) show "eventually (\x. \sum_upto (dirichlet_prod \ f) x\ \ \ * x) at_top" using eventually_ge_at_top[of "b * c"] eventually_ge_at_top[of 1] eventually_ge_at_top[of b] proof eventually_elim case (elim x) define a where "a = x / b" from elim \b \ 1\ have ab: "a \ 1" "b \ 1" "a * b = x" by (simp_all add: a_def field_simps) from ab have "a * 1 \ a * b" by (intro mult_mono) auto hence "a \ x" by (simp add: ab(3)) from ab have "a * 1 \ a * b" and "1 * b \ a * b" by (intro mult_mono; simp)+ hence "a \ x" "b \ x" by (simp_all add: ab(3)) have "a = x / b" "b = x / a" using ab by (simp_all add: field_simps) have "sum_upto (dirichlet_prod \ f) x = sum_upto (\n. \ n * F (x / n)) a + sum_upto (\n. M (x / n) * f n) b - M a * F b" unfolding M_def F_def by (rule hyperbola_method) (use ab in auto) also have "\\\ \ \ / 3 * x + \ / 3 * x + \ / 3 * x" proof (rule order.trans[OF abs_triangle_ineq4] order.trans[OF abs_triangle_ineq] add_mono)+ have "\sum_upto (\n. \ n * F (x / real n)) a\ \ sum_upto (\n. \\ n * F (x / real n)\) a" unfolding sum_upto_def by (rule sum_abs) also have "\ \ sum_upto (\n. 1 * (B * sqrt (x / real n))) a" unfolding sum_upto_def abs_mult using \a \ x\ by (intro sum_mono mult_mono B') (auto simp: moebius_mu_def) also have "\ = B * sqrt x * sum_upto (\n. 1 / sqrt n) a" by (simp add: sum_upto_def sum_distrib_left real_sqrt_divide) also have "\ \ B * sqrt x * \sum_upto (\n. 1 / sqrt n) a\" using B(1) \x \ 1\ by (intro mult_left_mono) auto also have "\ \ B * sqrt x * (A * sqrt a)" using \a \ 1\ B(1) \x \ 1\ by (intro mult_left_mono A) auto also have "\ = A * B / sqrt b * x" using ab \x \ 1\\x \ 1\ by (subst \a = x / b\) (simp_all add: field_simps real_sqrt_divide) also have "\ \ \ / 3 * x" using \x \ 1\ by (intro mult_right_mono b) auto finally show "\sum_upto (\n. \ n * F (x / n)) a\ \ \ / 3 * x" . next have "\sum_upto (\n. M (x / n) * f n) b\ \ sum_upto (\n. \M (x / n) * f n\) b" unfolding sum_upto_def by (rule sum_abs) also have "\ \ sum_upto (\n. \ / 3 / K * (x / n) * \f n\) b" unfolding sum_upto_def abs_mult proof (intro sum_mono mult_right_mono) fix n assume n: "n \ {n. n > 0 \ real n \ b}" have "c \ 0" by (simp add: c_def) with n have "c * n \ c * b" by (intro mult_left_mono) auto also have "\ \ x" using \b * c \ x\ by (simp add: algebra_simps) finally show "\M (x / real n)\ \ \ / 3 / K * (x / real n)" by (intro less_imp_le[OF c]) (use n in \auto simp: field_simps\) qed auto also have "\ = \ / 3 * x / K * sum_upto (\n. \f n\ / n) b" by (simp add: sum_upto_def sum_distrib_left) also have "\ = \ / 3 * x" unfolding K_def [symmetric] using K by simp finally show "\sum_upto (\n. M (x / real n) * f n) b\ \ \ / 3 * x" . next have "\M a * F b\ \ a * (B * sqrt b)" unfolding abs_mult using ab by (intro mult_mono sum_moebius_mu_bound B') auto also have "\ = B / sqrt b * x" using ab(1,2) by (simp add: real_sqrt_mult \b = x / a\ real_sqrt_divide field_simps) also have "\ \ \ / 3 * x" using \x \ 1\ by (intro mult_right_mono b) auto finally show "\M a * F b\ \ \ / 3 * x" . qed also have "\ = \ * x" by simp finally show ?case . qed qed have "eventually (\b::real. b \ 1 \ A * B / sqrt b \ \ / 3 \ B / sqrt b \ \ / 3) at_top" using \ by (intro eventually_conj; real_asymp) then obtain b where "b \ 1" "A * B / sqrt b \ \ / 3" "B / sqrt b \ \ / 3" by (auto simp: eventually_at_top_linorder) from *[OF this] have "eventually (\x. \sum_upto (dirichlet_prod \ f) x\ \ \ * x) at_top" . thus "eventually (\x. norm (sum_upto (dirichlet_prod \ f) x) \ \ * norm x) at_top" using eventually_ge_at_top[of 0] by eventually_elim simp qed have "(\x. \ x - x) \ \(\x. -(sum_upto (dirichlet_prod \ f) x + (frac x + 2 * C)))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]], subst eq) auto hence "(\x. \ x - x) \ \(\x. sum_upto (dirichlet_prod \ f) x + (frac x + 2 * C))" by (simp only: landau_theta.uminus) also have "(\x. sum_upto (dirichlet_prod \ f) x + (frac x + 2 * C)) \ o(\x. x)" using \sum_upto (dirichlet_prod \ f) \ o(\x. x)\ by (rule sum_in_smallo) real_asymp+ finally show ?thesis by (rule smallo_imp_asymp_equiv) qed text \ We now turn to a related fact: For the weighted sum $A(x) := \sum_{n\leq x} \mu(n)/n$, the asymptotic relation $A(x)\in o(1)$ is also equivalent to the Prime Number Theorem. Like Apostol, we only show one direction, namely that $A(x)\in o(1)$ implies the PNT. \ context fixes A defines "A \ sum_upto (\n. moebius_mu n / n)" begin lemma sum_upto_moebius_mu_integral': "x > 1 \ (A has_integral x * A x - M x) {1..x}" and sum_upto_moebius_mu_integrable': "a \ 1 \ A integrable_on {a..b}" proof - { fix a b :: real assume ab: "a \ 1" "a < b" have "((\t. A t * 1) has_integral A b * b - A a * a - (\n\real -` {a<..b}. moebius_mu n / n * n)) {a..b}" unfolding M_def A_def using ab by (intro partial_summation_strong [where X = "{}"]) (auto intro!: derivative_eq_intros continuous_intros - simp flip: has_field_derivative_iff_has_vector_derivative) + simp flip: has_real_derivative_iff_has_vector_derivative) } note * = this { fix x :: real assume x: "x > 1" have [simp]: "A 1 = 1" by (simp add: A_def) have "(\n\real -` {1<..x}. moebius_mu n / n * n) = (\n\insert 1 (real -` {1<..x}). moebius_mu n / n * n) - 1" using finite_vimage_real_of_nat_greaterThanAtMost[of 1 x] by (subst sum.insert) auto also have "insert 1 (real -` {1<..x}) = {n. n > 0 \ real n \ x}" using x by auto also have "(\n | 0 < n \ real n \ x. moebius_mu n / real n * real n) = M x" unfolding M_def sum_upto_def by (intro sum.cong) auto finally show "(A has_integral x * A x - M x) {1..x}" using *[of 1 x] x by (simp add: mult_ac) } { fix a b :: real assume ab: "a \ 1" show "A integrable_on {a..b}" using *[of a b] ab by (cases a b rule: linorder_cases) (auto intro: integrable_negligible) } qed (* 4.16 *) theorem sum_moebius_mu_div_n_smallo_imp_PNT: assumes smallo: "A \ o(\_. 1)" shows "M \ o(\x. x)" and "\ \[at_top] (\x. x)" proof - have "eventually (\x. M x = x * A x - integral {1..x} A) at_top" using eventually_gt_at_top[of 1] by eventually_elim (use sum_upto_moebius_mu_integral' in \simp add: has_integral_iff\) hence "M \ \(\x. x * A x - integral {1..x} A)" by (rule bigthetaI_cong) also have "(\x. x * A x - integral {1..x} A) \ o(\x. x)" proof (intro sum_in_smallo) from smallo show "(\x. x * A x) \ o(\x. x)" by (simp add: landau_divide_simps) show "(\x. integral {1..x} A) \ o(\x. x)" by (intro integral_smallo[OF smallo] sum_upto_moebius_mu_integrable') (auto intro!: derivative_eq_intros filterlim_ident) qed finally show "M \ o(\x. x)" . thus "\ \[at_top] (\x. x)" by (rule sum_moebius_mu_sublinear_imp_PNT) qed end end end \ No newline at end of file diff --git a/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy b/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy --- a/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy +++ b/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy @@ -1,447 +1,447 @@ (* File: Partial_Zeta_Bounds.thy Author: Manuel Eberl, TU München Asymptotic bounds on sums over k^(-s) for k=1..n, for fixed s and variable n. *) section \Bounds on partial sums of the $\zeta$ function\ theory Partial_Zeta_Bounds imports Euler_MacLaurin.Euler_MacLaurin_Landau Zeta_Function.Zeta_Function Prime_Number_Theorem.Prime_Number_Theorem_Library Prime_Distribution_Elementary_Library begin (* TODO: This does not necessarily require the full complex zeta function. *) text \ We employ Euler--MacLaurin's summation formula to obtain asymptotic estimates for the partial sums of the Riemann $\zeta(s)$ function for fixed real $a$, i.\,e.\ the function \[f(n) = \sum_{k=1}^n k^{-s}\ .\] We distinguish various cases. The case $s = 1$ is simply the Harmonic numbers and is treated apart from the others. \ lemma harm_asymp_equiv: "sum_upto (\n. 1 / n) \[at_top] ln" proof - have "sum_upto (\n. n powr -1) \[at_top] (\x. ln (\x\ + 1))" proof (rule asymp_equiv_sandwich) have "eventually (\x. sum_upto (\n. n powr -1) x \ {ln (\x\ + 1)..ln \x\ + 1}) at_top" using eventually_ge_at_top[of 1] proof eventually_elim case (elim x) have "sum_upto (\n. real n powr -1) x = harm (nat \x\)" unfolding sum_upto_altdef harm_def by (intro sum.cong) (auto simp: field_simps powr_minus) also have "\ \ {ln (\x\ + 1)..ln \x\ + 1}" using elim harm_le[of "nat \x\"] ln_le_harm[of "nat \x\"] by (auto simp: le_nat_iff le_floor_iff) finally show ?case by simp qed thus "eventually (\x. sum_upto (\n. n powr -1) x \ ln (\x\ + 1)) at_top" "eventually (\x. sum_upto (\n. n powr -1) x \ ln \x\ + 1) at_top" by (eventually_elim; simp)+ qed real_asymp+ also have "\ \[at_top] ln" by real_asymp finally show ?thesis by (simp add: powr_minus field_simps) qed lemma fixes s :: real assumes s: "s > 0" "s \ 1" shows zeta_partial_sum_bigo_pos: "(\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - Re (zeta s)) \ O(\x. real x powr -s)" and zeta_partial_sum_bigo_pos': "(\n. \k=1..n. real k powr -s) =o (\n. real n powr (1 - s) / (1 - s) + Re (zeta s)) +o O(\x. real x powr -s)" proof - define F where "F = (\x. x powr (1 - s) / (1 - s))" define f where "f = (\x. x powr -s)" define f' where "f' = (\x. -s * x powr (-s-1))" define z where "z = Re (zeta s)" interpret euler_maclaurin_nat' F f "(!) [f, f']" 1 0 z "{}" proof have "(\b. (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s) - real b powr -s / 2) \ Re (zeta s) - 0" proof (intro tendsto_diff) let ?g = "\b. (\i\<^sub>F b in at_top. Re (?g b) = (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s)" using eventually_ge_at_top[of 1] proof eventually_elim case (elim b) have "(\k=1..b. real k powr -s) = (\kn. n - 1"]) auto also have "\ - real b powr (1 - s) / (1 - s) = Re (?g b)" by (auto simp: powr_Reals_eq add_ac) finally show ?case .. qed moreover have "(\b. Re (?g b)) \ Re (zeta s)" using hurwitz_zeta_critical_strip[of "of_real s" 1] s by (intro tendsto_intros) (simp add: zeta_def) ultimately show "(\b. (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s)) \ Re (zeta s)" by (blast intro: Lim_transform_eventually) qed (use s in real_asymp) thus "(\b. (\k = 1..b. f (real k)) - F (real b) - (\i<2 * 0 + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ([f, f'] ! i) (real b))) \ z" by (simp add: f_def F_def z_def) qed (use \s \ 1\ in \auto intro!: derivative_eq_intros continuous_intros - simp flip: has_field_derivative_iff_has_vector_derivative + simp flip: has_real_derivative_iff_has_vector_derivative simp: F_def f_def f'_def nth_Cons split: nat.splits\) { fix n :: nat assume n: "n \ 1" have "(\k=1..n. real k powr -s) = n powr (1 - s) / (1 - s) + z + 1/2 * n powr -s - EM_remainder 1 f' (int n)" using euler_maclaurin_strong_nat'[of n] n by (simp add: F_def f_def) } note * = this have "(\n. (\k=1..n. real k powr -s) - n powr (1 - s) / (1 - s) - z) \ \(\n. 1/2 * n powr -s - EM_remainder 1 f' (int n))" using * by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\n. 1/2 * n powr -s - EM_remainder 1 f' (int n)) \ O(\n. n powr -s)" proof (intro sum_in_bigo) have "(\x. norm (EM_remainder 1 f' (int x))) \ O(\x. real x powr - s)" proof (rule EM_remainder_strong_bigo_nat[where a = 1 and Y = "{}"]) fix x :: real assume "x \ 1" show "norm (f' x) \ s * x powr (-s-1)" using s by (simp add: f'_def) next from s show "((\x. x powr -s) \ 0) at_top" by real_asymp qed (auto simp: f'_def intro!: continuous_intros derivative_eq_intros) thus "(\x. EM_remainder 1 f' (int x)) \ O(\x. real x powr -s)" by simp qed real_asymp+ finally show "(\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - z) \ O(\x. real x powr -s)" . thus"(\n. \k=1..n. real k powr -s) =o (\n. real n powr (1 - s) / (1 - s) + z) +o O(\x. real x powr -s)" by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def algebra_simps) qed lemma zeta_tail_bigo: fixes s :: real assumes s: "s > 1" shows "(\n. Re (hurwitz_zeta (real n + 1) s)) \ O(\x. real x powr (1 - s))" proof - have [simp]: "complex_of_real (Re (zeta s)) = zeta s" using zeta_real[of s] by (auto elim!: Reals_cases) from s have s': "s > 0" "s \ 1" by auto have "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s) + real n powr (1 - s) / (1 - s)) \ O(\x. real x powr (1 - s))" proof (rule sum_in_bigo) have "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s)) = (\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - Re (zeta s))" (is "?lhs = ?rhs") proof fix n :: nat have "hurwitz_zeta (1 + real n) s = zeta s - (\kauto simp: zeta_def powr_Reals_eq\) also have "(\kk=1..n. real k powr -s)" by (rule sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show "?lhs n = ?rhs n" by (simp add: add_ac) qed also have "\ \ O(\x. real x powr (-s))" by (rule zeta_partial_sum_bigo_pos) (use s in auto) also have "(\x. real x powr (-s)) \ O(\x. real x powr (1-s))" by real_asymp finally show "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s)) \ \" . qed (use s in real_asymp) thus ?thesis by simp qed lemma zeta_tail_bigo': fixes s :: real assumes s: "s > 1" shows "(\n. Re (hurwitz_zeta (real n) s)) \ O(\x. real x powr (1 - s))" proof - have "(\n. Re (hurwitz_zeta (real n) s)) \ \(\n. Re (hurwitz_zeta (real (n - 1) + 1) s))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) (auto simp: of_nat_diff) also have "(\n. Re (hurwitz_zeta (real (n - 1) + 1) s)) \ O(\x. real (x - 1) powr (1 - s))" by (rule landau_o.big.compose[OF zeta_tail_bigo[OF assms]]) real_asymp also have "(\x. real (x - 1) powr (1 - s)) \ O(\x. real x powr (1 - s))" by real_asymp finally show ?thesis . qed lemma fixes s :: real assumes s: "s > 0" shows zeta_partial_sum_bigo_neg: "(\n. (\i=1..n. real i powr s) - n powr (1 + s) / (1 + s)) \ O(\n. n powr s)" and zeta_partial_sum_bigo_neg': "(\n. (\i=1..n. real i powr s)) =o (\n. n powr (1 + s) / (1 + s)) +o O(\n. n powr s)" proof - define F where "F = (\x. x powr (1 + s) / (1 + s))" define f where "f = (\x. x powr s)" define f' where "f' = (\x. s * x powr (s - 1))" have "(\i=1..n. f (real i)) - F n = 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n)" if n: "n \ 1" for n proof - have "(\i\{1<..n}. f (real i)) - integral {real 1..real n} f = (\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (([f, f'] ! k) (real n) - ([f, f'] ! k) (real 1))) + EM_remainder' 1 ([f, f'] ! 1) (real 1) (real n)" by (rule euler_maclaurin_strong_raw_nat[where Y = "{}"]) (use \s > 0\ \n \ 1\ in \auto intro!: derivative_eq_intros continuous_intros - simp flip: has_field_derivative_iff_has_vector_derivative + simp flip: has_real_derivative_iff_has_vector_derivative simp: F_def f_def f'_def nth_Cons split: nat.splits\) also have "(\i\{1<..n}. f (real i)) = (\i\insert 1 {1<..n}. f (real i)) - f 1" using n by (subst sum.insert) auto also from n have "insert 1 {1<..n} = {1..n}" by auto finally have "(\i=1..n. f (real i)) - F n = f 1 + (integral {1..real n} f - F n) + (f (real n) - f 1) / 2 + EM_remainder' 1 f' 1 (real n)" by simp hence "(\i=1..n. f (real i)) - F n = 1 / 2 + (integral {1..real n} f - F n) + f (real n) / 2 + EM_remainder' 1 f' 1 (real n)" using s by (simp add: f_def diff_divide_distrib) also have "(f has_integral (F (real n) - F 1)) {1..real n}" using assms n by (intro fundamental_theorem_of_calculus) - (auto simp flip: has_field_derivative_iff_has_vector_derivative simp: F_def f_def + (auto simp flip: has_real_derivative_iff_has_vector_derivative simp: F_def f_def intro!: derivative_eq_intros continuous_intros) hence "integral {1..real n} f - F n = - F 1" by (simp add: has_integral_iff) also have "1 / 2 + (-F 1) + f (real n) / 2 = 1 / 2 - F 1 + f n / 2" by simp finally show ?thesis . qed hence "(\n. (\i=1..n. f (real i)) - F n) \ \(\n. 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) also have "(\n. 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n)) \ O(\n. real n powr s)" unfolding F_def f_def proof (intro sum_in_bigo) have "(\x. integral {1..real x} (\t. pbernpoly 1 t *\<^sub>R f' t)) \ O(\n. 1 / s * real n powr s)" proof (intro landau_o.big.compose[OF integral_bigo]) have "(\x. pbernpoly 1 x * f' x) \ O(\x. 1 * x powr (s - 1))" by (intro landau_o.big.mult pbernpoly_bigo) (auto simp: f'_def) thus "(\x. pbernpoly 1 x *\<^sub>R f' x) \ O(\x. x powr (s - 1))" by simp from s show "filterlim (\a. 1 / s * a powr s) at_top at_top" by real_asymp next fix a' x :: real assume "a' \ 1" "a' \ x" thus "(\a. pbernpoly 1 a *\<^sub>R f' a) integrable_on {a'..x}" by (intro integrable_EM_remainder') (auto intro!: continuous_intros simp: f'_def) qed (use s in \auto intro!: filterlim_real_sequentially continuous_intros derivative_eq_intros\) thus "(\x. EM_remainder' 1 f' 1 (real x)) \ O(\n. real n powr s)" using \s > 0\ by (simp add: EM_remainder'_def) qed (use \s > 0\ in real_asymp)+ finally show "(\n. (\i=1..n. real i powr s) - n powr (1 + s) / (1 + s)) \ O(\n. n powr s)" by (simp add: f_def F_def) thus "(\n. (\i=1..n. real i powr s)) =o (\n. n powr (1 + s) / (1 + s)) +o O(\n. n powr s)" by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def algebra_simps) qed lemma zeta_partial_sum_le_pos: assumes "s > 0" "s \ 1" defines "z \ Re (zeta (complex_of_real s))" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x - (x powr (1-s) / (1-s) + z)\ \ c * x powr -s" proof (rule sum_upto_asymptotics_lift_nat_real) show "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" using zeta_partial_sum_bigo_pos[OF assms(1,2)] unfolding z_def by (simp add: algebra_simps) from assms have "s < 1 \ s > 1" by linarith thus "(\n. real n powr (1 - s) / (1 - s) + z - (real (Suc n) powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" by standard (use \s > 0\ in \real_asymp+\) show "(\n. real n powr - s) \ O(\n. real (Suc n) powr - s)" by real_asymp show "mono_on (\a. a powr - s) {1..} \ mono_on (\x. - (x powr - s)) {1..}" using assms by (intro disjI2) (auto intro!: mono_onI powr_mono2') from assms have "s < 1 \ s > 1" by linarith hence "mono_on (\a. a powr (1 - s) / (1 - s) + z) {1..}" proof assume "s < 1" thus ?thesis using \s > 0\ by (intro mono_onI powr_mono2 divide_right_mono add_right_mono) auto next assume "s > 1" thus ?thesis by (intro mono_onI le_imp_neg_le add_right_mono divide_right_mono_neg powr_mono2') auto qed thus "mono_on (\a. a powr (1 - s) / (1 - s) + z) {1..} \ mono_on (\x. - (x powr (1 - s) / (1 - s) + z)) {1..}" by blast qed auto lemma zeta_partial_sum_le_pos': assumes "s > 0" "s \ 1" defines "z \ Re (zeta (complex_of_real s))" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x - x powr (1-s) / (1-s)\ \ c" proof - have "\c>0. \x\1. \sum_upto (\n. n powr -s) x - x powr (1-s) / (1-s)\ \ c * 1" proof (rule sum_upto_asymptotics_lift_nat_real) have "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" using zeta_partial_sum_bigo_pos[OF assms(1,2)] unfolding z_def by (simp add: algebra_simps) also have "(\n. real n powr -s) \ O(\n. 1)" using assms by real_asymp finally have "(\n. (\k = 1..n. real k powr - s) - real n powr (1 - s) / (1 - s) - z) \ O(\n. 1)" by (simp add: algebra_simps) hence "(\n. (\k = 1..n. real k powr - s) - real n powr (1 - s) / (1 - s) - z + z) \ O(\n. 1)" by (rule sum_in_bigo) auto thus "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s))) \ O(\n. 1)" by simp from assms have "s < 1 \ s > 1" by linarith thus "(\n. real n powr (1 - s) / (1 - s) - (real (Suc n) powr (1 - s) / (1 - s))) \ O(\n. 1)" by standard (use \s > 0\ in \real_asymp+\) show "mono_on (\a. 1) {1..} \ mono_on (\x::real. -1 :: real) {1..}" using assms by (intro disjI2) (auto intro!: mono_onI powr_mono2') from assms have "s < 1 \ s > 1" by linarith hence "mono_on (\a. a powr (1 - s) / (1 - s)) {1..}" proof assume "s < 1" thus ?thesis using \s > 0\ by (intro mono_onI powr_mono2 divide_right_mono add_right_mono) auto next assume "s > 1" thus ?thesis by (intro mono_onI le_imp_neg_le add_right_mono divide_right_mono_neg powr_mono2') auto qed thus "mono_on (\a. a powr (1 - s) / (1 - s)) {1..} \ mono_on (\x. - (x powr (1 - s) / (1 - s))) {1..}" by blast qed auto thus ?thesis by simp qed lemma zeta_partial_sum_le_pos'': assumes "s > 0" "s \ 1" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x\ \ c * x powr max 0 (1 - s)" proof - from zeta_partial_sum_le_pos'[OF assms] obtain c where c: "c > 0" "\x. x \ 1 \ \sum_upto (\x. real x powr - s) x - x powr (1 - s) / (1 - s)\ \ c" by auto { fix x :: real assume x: "x \ 1" have "\sum_upto (\x. real x powr - s) x\ \ \x powr (1 - s) / (1 - s)\ + c" using c(1) c(2)[OF x] x by linarith also have "\x powr (1 - s) / (1 - s)\ = x powr (1 - s) / \1 - s\" using assms by simp also have "\ \ x powr max 0 (1 - s) / \1 - s\" using x by (intro divide_right_mono powr_mono) auto also have "c = c * x powr 0" using x by simp also have "c * x powr 0 \ c * x powr max 0 (1 - s)" using c(1) x by (intro mult_left_mono powr_mono) auto also have "x powr max 0 (1 - s) / \1 - s\ + c * x powr max 0 (1 - s) = (1 / \1 - s\ + c) * x powr max 0 (1 - s)" by (simp add: algebra_simps) finally have "\sum_upto (\x. real x powr - s) x\ \ (1 / \1 - s\ + c) * x powr max 0 (1 - s)" by simp } moreover have "(1 / \1 - s\ + c) > 0" using c assms by (intro add_pos_pos divide_pos_pos) auto ultimately show ?thesis by blast qed lemma zeta_partial_sum_le_pos_bigo: assumes "s > 0" "s \ 1" shows "(\x. sum_upto (\n. n powr -s) x) \ O(\x. x powr max 0 (1 - s))" proof - from zeta_partial_sum_le_pos''[OF assms] obtain c where "\x\1. \sum_upto (\n. n powr -s) x\ \ c * x powr max 0 (1 - s)" by auto thus ?thesis by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto qed lemma zeta_partial_sum_01_asymp_equiv: assumes "s \ {0<..<1}" shows "sum_upto (\n. n powr -s) \[at_top] (\x. x powr (1 - s) / (1 - s))" proof - from zeta_partial_sum_le_pos'[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr -s) x - x powr (1 - s) / (1 - s)\ \ c" by auto hence "(\x. sum_upto (\x. real x powr -s) x - x powr (1 - s) / (1 - s)) \ O(\_. 1)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\_. 1) \ o(\x. x powr (1 - s) / (1 - s))" using assms by real_asymp finally show ?thesis by (rule smallo_imp_asymp_equiv) qed lemma zeta_partial_sum_gt_1_asymp_equiv: fixes s :: real assumes "s > 1" defines "\ \ Re (zeta s)" shows "sum_upto (\n. n powr -s) \[at_top] (\x. \)" proof - have [simp]: "\ \ 0" using assms zeta_Re_gt_1_nonzero[of s] zeta_real[of s] by (auto elim!: Reals_cases) from zeta_partial_sum_le_pos[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr -s) x - (x powr (1 - s) / (1 - s) + \)\ \ c * x powr -s" by (auto simp: \_def) hence "(\x. sum_upto (\x. real x powr -s) x - \ - x powr (1 - s) / (1 - s)) \ O(\x. x powr -s)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\x. x powr -s) \ o(\_. 1)" using \s > 1\ by real_asymp finally have "(\x. sum_upto (\x. real x powr -s) x - \ - x powr (1 - s) / (1 - s) + x powr (1 - s) / (1 - s)) \ o(\_. 1)" by (rule sum_in_smallo) (use \s > 1\ in real_asymp) thus ?thesis by (simp add: smallo_imp_asymp_equiv) qed lemma zeta_partial_sum_pos_bigtheta: assumes "s > 0" "s \ 1" shows "sum_upto (\n. n powr -s) \ \(\x. x powr max 0 (1 - s))" proof (cases "s > 1") case False thus ?thesis using asymp_equiv_imp_bigtheta[OF zeta_partial_sum_01_asymp_equiv[of s]] assms by (simp add: max_def) next case True have [simp]: "Re (zeta s) \ 0" using True zeta_Re_gt_1_nonzero[of s] zeta_real[of s] by (auto elim!: Reals_cases) show ?thesis using True asymp_equiv_imp_bigtheta[OF zeta_partial_sum_gt_1_asymp_equiv[of s]] by (simp add: max_def) qed lemma zeta_partial_sum_le_neg: assumes "s > 0" shows "\c>0. \x\1. \sum_upto (\n. n powr s) x - x powr (1 + s) / (1 + s)\ \ c * x powr s" proof (rule sum_upto_asymptotics_lift_nat_real) show "(\n. (\k = 1..n. real k powr s) - (real n powr (1 + s) / (1 + s))) \ O(\n. real n powr s)" using zeta_partial_sum_bigo_neg[OF assms(1)] by (simp add: algebra_simps) show "(\n. real n powr (1 + s) / (1 + s) - (real (Suc n) powr (1 + s) / (1 + s))) \ O(\n. real n powr s)" using assms by real_asymp show "(\n. real n powr s) \ O(\n. real (Suc n) powr s)" by real_asymp show "mono_on (\a. a powr s) {1..} \ mono_on (\x. - (x powr s)) {1..}" using assms by (intro disjI1) (auto intro!: mono_onI powr_mono2) show "mono_on (\a. a powr (1 + s) / (1 + s)) {1..} \ mono_on (\x. - (x powr (1 + s) / (1 + s))) {1..}" using assms by (intro disjI1 divide_right_mono powr_mono2 mono_onI) auto qed auto lemma zeta_partial_sum_neg_asymp_equiv: assumes "s > 0" shows "sum_upto (\n. n powr s) \[at_top] (\x. x powr (1 + s) / (1 + s))" proof - from zeta_partial_sum_le_neg[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr s) x - x powr (1 + s) / (1 + s)\ \ c * x powr s" by auto hence "(\x. sum_upto (\x. real x powr s) x - x powr (1 + s) / (1 + s)) \ O(\x. x powr s)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\x. x powr s) \ o(\x. x powr (1 + s) / (1 + s))" using assms by real_asymp finally show ?thesis by (rule smallo_imp_asymp_equiv) qed end \ No newline at end of file diff --git a/thys/Prime_Number_Theorem/Mertens_Theorems.thy b/thys/Prime_Number_Theorem/Mertens_Theorems.thy --- a/thys/Prime_Number_Theorem/Mertens_Theorems.thy +++ b/thys/Prime_Number_Theorem/Mertens_Theorems.thy @@ -1,755 +1,755 @@ (* File: Mertens_Theorems.thy Author: Manuel Eberl (TU München) *) section \Mertens' Theorems\ theory Mertens_Theorems imports Prime_Counting_Functions "Stirling_Formula.Stirling_Formula" begin (*<*) unbundle prime_counting_notation (*>*) text \ In this section, we will prove Mertens' First and Second Theorem. These are weaker results than the Prime Number Theorem, and we will derive them without using it. However, like Mertens himself, we will not only prove them \<^emph>\asymptotically\, but \<^emph>\absolutely\. This means that we will show that the remainder terms are not only ``Big-O'' of some bound, but we will give concrete (and reasonably tight) upper and lower bounds for them that hold on the entire domain. This makes the proofs a bit more tedious. \ subsection \Absolute Bounds for Mertens' First Theorem\ text \ We have already shown the asymptotic form of Mertens' first theorem, i.\,e.\ $\mathfrak{M}(n) = \ln n + O(1)$. We now want to obtain some absolute bounds on the $O(1)$ remainder term using a more careful derivation than before. The precise bounds we will show are $\mathfrak {M}(n) - \ln n \in (-1-\frac{9}{\pi^2}; \ln 4] \approx (-1.9119; 1.3863]$ for \n \ \\. First, we need a simple lemma on the finiteness of exponents to consider in a sum of all prime powers up to a certain point: \ lemma exponents_le_finite: assumes "p > (1 :: nat)" "k > 0" shows "finite {i. real (p ^ (k * i + l)) \ x}" proof (rule finite_subset) show "{i. real (p ^ (k * i + l)) \ x} \ {..nat \x\}" proof safe fix i assume i: "real (p ^ (k * i + l)) \ x" have "i < 2 ^ i" by (rule less_exp) also from assms have "i \ k * i + l" by (cases k) auto hence "2 ^ i \ (2 ^ (k * i + l) :: nat)" using assms by (intro power_increasing) auto also have "\ \ p ^ (k * i + l)" using assms by (intro power_mono) auto also have "real \ \ x" using i by simp finally show "i \ nat \x\" by linarith qed qed auto text \ Next, we need the following bound on $\zeta'(2)$: \ lemma deriv_zeta_2_bound: "Re (deriv zeta 2) > -1" proof - have "((\x::real. ln (x + 3) * (x + 3) powr -2) has_integral (ln 3 + 1) / 3) (interior {0..})" using ln_powr_has_integral_at_top[of 1 0 3 "-2"] by (simp add: interior_real_atLeast powr_minus) hence "((\x::real. ln (x + 3) * (x + 3) powr -2) has_integral (ln 3 + 1) / 3) {0..}" by (subst (asm) has_integral_interior) auto also have "?this \ ((\x::real. ln (x + 3) / (x + 3) ^ 2) has_integral (ln 3 + 1) / 3) {0..}" by (intro has_integral_cong) (auto simp: powr_minus field_simps) finally have int: \ . have "exp (1 / 2 :: real) ^ 2 \ 2 ^ 2" using exp_le by (subst exp_double [symmetric]) simp_all hence exp_half: "exp (1 / 2 :: real) \ 2" by (rule power2_le_imp_le) auto have mono: "ln x / x ^ 2 \ ln y / y ^ 2" if "y \ exp (1/2)" "x \ y" for x y :: real proof (rule DERIV_nonpos_imp_nonincreasing[of _ _ "\x. ln x / x ^ 2"]) fix t assume t: "t \ y" "t \ x" have "y > 0" by (rule less_le_trans[OF _ that(1)]) auto with t that have "ln t \ ln (exp (1 / 2))" by (subst ln_le_cancel_iff) auto hence "ln t \ 1 / 2" by (simp only: ln_exp) from t \y > 0\ have "((\x. ln x / x ^ 2) has_field_derivative ((1 - 2 * ln t) / t ^ 3)) (at t)" by (auto intro!: derivative_eq_intros simp: eval_nat_numeral field_simps) moreover have "(1 - 2 * ln t) / t ^ 3 \ 0" using t that \y > 0\ \ln t \ 1 / 2\ by (intro divide_nonpos_pos) auto ultimately show "\f'. ((\x. ln x / x ^ 2) has_field_derivative f') (at t) \ f' \ 0" by blast qed fact+ have "fds_converges (fds_deriv fds_zeta) (2 :: complex)" by (intro fds_converges_deriv) auto hence "(\n. of_real (-ln (real (Suc n)) / (of_nat (Suc n)) ^ 2)) sums deriv zeta 2" by (auto simp: fds_converges_altdef add_ac eval_fds_deriv_zeta fds_nth_deriv scaleR_conv_of_real simp del: of_nat_Suc) note * = sums_split_initial_segment[OF sums_minus[OF sums_Re[OF this]], of 3] have "(\n. ln (real (n+4)) / real (n+4) ^ 2) sums (-Re (deriv zeta 2) - (ln 2 / 4 + ln 3 / 9))" using * by (simp add: eval_nat_numeral) hence "-Re (deriv zeta 2) - (ln 2 / 4 + ln 3 / 9) = (\n. ln (real (Suc n) + 3) / (real (Suc n) + 3) ^ 2)" by (simp_all add: sums_iff algebra_simps) also have "\ \ (ln 3 + 1) / 3" using int exp_half by (intro decreasing_sum_le_integral divide_nonneg_pos mono) (auto simp: powr_minus field_simps) finally have "-Re (deriv zeta 2) \ (16 * ln 3 + 9 * ln 2 + 12) / 36" by simp also have "ln 3 \ (11 / 10 :: real)" using ln_approx_bounds[of 3 2] by (simp add: power_numeral_reduce numeral_2_eq_2) hence "(16 * ln 3 + 9 * ln 2 + 12) / 36 \ (16 * (11 / 10) + 9 * 25 / 36 + 12) / (36 :: real)" using ln2_le_25_over_36 by (intro add_mono mult_left_mono divide_right_mono) auto also have "\ < 1" by simp finally show ?thesis by simp qed text \ Using the logarithmic derivative of Euler's product formula for $\zeta(s)$ at $s = 2$ and the bound on $\zeta'(2)$ we have just derived, we can obtain the bound \[\sum_{p^i \leq x, i \geq 2} \frac{\ln p}{p^i} < \frac{9}{\pi^2}\ .\] \ lemma mertens_remainder_aux_bound: fixes x :: real defines "R \ (\(p,i) | prime p \ i > 1 \ real (p ^ i) \ x. ln (real p) / p ^ i)" shows "R < 9 / pi\<^sup>2" proof - define S' where "S' = {(p, i). prime p \ i > 1 \ real (p ^ i) \ x}" define S'' where "S'' = {(p, i). prime p \ i > 1 \ real (p ^ Suc i) \ x}" have finite_row: "finite {i. i > 1 \ real (p ^ (i + k)) \ x}" if p: "prime p" for p k proof (rule finite_subset) show "{i. i > 1 \ real (p ^ (i + k)) \ x} \ {..nat \x\}" proof safe fix i assume i: "i > 1" "real (p ^ (i + k)) \ x" have "i < 2 ^ (i + k)" by (induction i) auto also from p have "\ \ p ^ (i + k)" by (intro power_mono) (auto dest: prime_gt_1_nat) also have "real \ \ x" using i by simp finally show "i \ nat \x\" by linarith qed qed auto have "S'' \ S'" unfolding S''_def S'_def proof safe fix p i assume pi: "prime p" "real (p ^ Suc i) \ x" "i > 1" have "real (p ^ i) \ real (p ^ Suc i)" using pi unfolding of_nat_le_iff by (intro power_increasing) (auto dest: prime_gt_1_nat) also have "\ \ x" by fact finally show "real (p ^ i) \ x" . qed have S'_alt: "S' = (SIGMA p:{p. prime p \ real p \ x}. {i. i > 1 \ real (p ^ i) \ x})" unfolding S'_def proof safe fix p i assume "prime p" "real (p ^ i) \ x" "i > 1" hence "p ^ 1 \ p ^ i" by (intro power_increasing) (auto dest: prime_gt_1_nat) also have "real \ \ x" by fact finally show "real p \ x" by simp qed have finite: "finite {p. prime p \ real p \ x}" by (rule finite_subset[OF _ finite_Nats_le_real[of x]]) (auto dest: prime_gt_0_nat) have "finite S'" unfolding S'_alt using finite_row[of _ 0] by (intro finite_SigmaI finite) auto have "R \ 3 / 2 * (\(p, i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i))" proof - have "R = (\y\{0, 1}. \z | z \ S' \ snd z mod 2 = y. ln (real (fst z)) / real (fst z ^ snd z))" using \finite S'\ by (subst sum.group) (auto simp: case_prod_unfold R_def S'_def) also have "\ = (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i)) + (\(p,i) | (p, i) \ S' \ odd i. ln (real p) / real (p ^ i))" unfolding even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one by (simp add: case_prod_unfold) also have "(\(p,i) | (p, i) \ S' \ odd i. ln (real p) / real (p ^ i)) = (\(p,i) | (p, i) \ S'' \ even i. ln (real p) / real (p ^ Suc i))" by (intro sum.reindex_bij_witness[of _ "\(p,i). (p, Suc i)" "\(p,i). (p, i - 1)"]) (auto simp: case_prod_unfold S'_def S''_def elim: oddE simp del: power_Suc) also have "\ \ (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ Suc i))" using \S'' \ S'\ unfolding case_prod_unfold by (intro sum_mono2 divide_nonneg_pos ln_ge_zero finite_subset[OF _ \finite S'\]) (auto simp: S'_def S''_def case_prod_unfold dest: prime_gt_0_nat simp del: power_Suc) also have "\ \ (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (2 * p ^ i))" unfolding case_prod_unfold by (intro sum_mono divide_left_mono) (auto simp: S'_def dest!: prime_gt_1_nat) also have "\ = (1 / 2) * (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i))" by (subst sum_distrib_left) (auto simp: case_prod_unfold) also have "(\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i)) + \ = 3 / 2 * (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i))" by simp finally show ?thesis by simp qed also have "(\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i)) = (\p | prime p \ real p \ x. ln (real p) * (\i | i > 0 \ even i \ real (p ^ i) \ x. (1 / real p) ^ i))" unfolding sum_distrib_left proof (subst sum.Sigma[OF _ ballI]) fix p assume p: "p \ {p. prime p \ real p \ x}" thus "finite {i. 0 < i \ even i \ real (p ^ i) \ x}" by (intro finite_subset[OF _ exponents_le_finite[of p 1 0 x]]) (auto dest: prime_gt_1_nat) qed (auto intro!: sum.cong finite_subset[OF _ finite_Nats_le_real[of x]] dest: prime_gt_0_nat simp: S'_alt power_divide) also have "\ \ (\p | prime p \ real p \ x. ln (real p) / (real p ^ 2 - 1))" proof (rule sum_mono) fix p assume p: "p \ {p. prime p \ real p \ x}" have "p > 1" using p by (auto dest: prime_gt_1_nat) have "(\i | i > 0 \ even i \ real (p ^ i) \ x. (1 / real p) ^ i) = (\i | real (p ^ (2 * i + 2)) \ x. (1 / real p) ^ (2 * i)) / real p ^ 2" (is "_ = ?S / _") unfolding sum_divide_distrib by (rule sum.reindex_bij_witness[of _ "\i. 2 * Suc i" "\i. (i - 2) div 2"]) (insert \p > 1\, auto simp: numeral_3_eq_3 power2_eq_square power_diff algebra_simps elim!: evenE) also have "?S = (\i | real (p ^ (2 * i + 2)) \ x. (1 / real p ^ 2) ^ i)" by (subst power_mult) (simp_all add: algebra_simps power_divide) also have "\ \ (\i. (1 / real p ^ 2) ^ i)" using exponents_le_finite[of p 2 2 x] \p > 1\ by (intro sum_le_suminf) (auto simp: summable_geometric_iff) also have "\ = real p ^ 2 / (real p ^ 2 - 1)" using \p > 1\ by (subst suminf_geometric) (auto simp: field_simps) also have "\ / real p ^ 2 = 1 / (real p ^ 2 - 1)" using \p > 1 \ by (simp add: divide_simps) finally have "(\i | 0 < i \ even i \ real (p ^ i) \ x. (1 / real p) ^ i) \ 1 / (real p ^ 2 - 1)" (is "?lhs \ ?rhs") using \p > 1\ by (simp add: divide_right_mono) thus "ln (real p) * ?lhs \ ln (real p) / (real p ^ 2 - 1)" using \p > 1\ by (simp add: divide_simps) qed also have "\ = (\\<^sub>a p | prime p \ real p \ x. ln (real p) / (real p ^ 2 - 1))" using finite by (intro infsetsum_finite [symmetric]) auto also have "\ \ (\\<^sub>a p | prime p. ln (real p) / (real p ^ 2 - 1))" using eval_fds_logderiv_zeta_real[of 2] finite by (intro infsetsum_mono_neutral_left divide_nonneg_pos) (auto simp: dest: prime_gt_1_nat) also have "\ = -Re (deriv zeta (of_real 2) / zeta (of_real 2))" by (subst eval_fds_logderiv_zeta_real) auto also have "\ = (-Re (deriv zeta 2)) * (6 / pi\<^sup>2)" by (simp add: zeta_even_numeral) also have "\ < 1 * (6 / pi\<^sup>2)" using deriv_zeta_2_bound by (intro mult_strict_right_mono) auto also have "3 / 2 * \ = 9 / pi\<^sup>2" by simp finally show ?thesis by simp qed text \ We now consider the equation \[\ln (n!) = \sum_{k\leq n} \Lambda(k) \left\lfloor\frac{n}{k}\right\rfloor\] and estimate both sides in different ways. The left-hand-side can be estimated using Stirling's formula, and we can simplify the right-hand side to \[\sum_{k\leq n} \Lambda(k) \left\lfloor\frac{n}{k}\right\rfloor = \sum_{p^i \leq x, i \geq 1} \ln p \left\lfloor\frac{n}{p^i}\right\rfloor\] and then split the sum into those $p^i$ with $i = 1$ and those with $i \geq 2$. Applying the bound we have just shown and some more routine estimates, we obtain the following reasonably strong version of Mertens' First Theorem on the naturals: $\mathfrak M(n) - \ln(n) \in (-1-\frac{9}{\pi^2}; \ln 4]$ \ theorem mertens_bound_strong: fixes n :: nat assumes n: "n > 0" shows "\ n - ln n \ {-1 - 9 / pi\<^sup>2<..ln 4}" proof (cases "n \ 3") case False with n consider "n = 1" | "n = 2" by force thus ?thesis proof cases assume [simp]: "n = 1" have "-1 + (-9 / pi\<^sup>2) < 0" by (intro add_neg_neg divide_neg_pos) auto thus ?thesis by simp next assume [simp]: "n = 2" have eq: "\ n - ln n = -ln 2 / 2" by (simp add: eval_\) have "-1 - 9 / pi ^ 2 + ln 2 / 2 \ -1 - 9 / 4 ^ 2 + 25 / 36 / 2" using pi_less_4 ln2_le_25_over_36 by (intro diff_mono add_mono divide_left_mono divide_right_mono power_mono) auto also have "\ < 0" by simp finally have "-ln 2 / 2 > -1 - 9 / pi\<^sup>2" by simp moreover { have "-ln 2 / 2 \ (0::real)" by (intro divide_nonpos_pos) auto also have "\ \ ln 4" by simp finally have "-ln 2 / 2 \ ln (4 :: real)" by simp } ultimately show ?thesis unfolding eq by simp qed next case True hence n: "n \ 3" by simp have finite: "finite {(p, i). prime p \ i \ 1 \ p ^ i \ n}" proof (rule finite_subset) show "{(p, i). prime p \ i \ 1 \ p ^ i \ n} \ {..nat \root 1 (real n)\} \ {..nat \log 2 (real n)\}" using primepows_le_subset[of "real n" 1] n unfolding of_nat_le_iff by auto qed auto define r where "r = prime_sum_upto (\p. ln (real p) * frac (real n / real p)) n" define R where "R = (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) * real (n div (p ^ i)))" define R' where "R' = (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) / p ^ i)" have [simp]: "ln (4 :: real) = 2 * ln 2" using ln_realpow[of 2 2] by simp from pi_less_4 have "ln pi \ ln 4" by (subst ln_le_cancel_iff) auto also have "\ = 2 * ln 2" by simp also have "\ \ 2 * (25 / 36)" by (intro mult_left_mono ln2_le_25_over_36) auto finally have ln_pi: "ln pi \ 25 / 18" by simp have "ln 3 \ ln (4::nat)" by (subst ln_le_cancel_iff) auto also have "\ = 2 * ln 2" by simp also have "\ \ 2 * (25 / 36)" by (intro mult_left_mono ln2_le_25_over_36) auto finally have ln_3: "ln (3::real) \ 25 / 18" by simp have "R / n = (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) * (real (n div (p ^ i)) / n))" by (simp add: R_def sum_divide_distrib field_simps case_prod_unfold) also have "\ \ (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) * (1 / p ^ i))" unfolding R'_def case_prod_unfold using n by (intro sum_mono mult_left_mono) (auto simp: field_simps real_of_nat_div dest: prime_gt_0_nat) also have "\ = R'" by (simp add: R'_def) also have "R' < 9 / pi\<^sup>2" unfolding R'_def using mertens_remainder_aux_bound[of n] by simp finally have "R / n < 9 / pi\<^sup>2" . moreover have "R \ 0" unfolding R_def by (intro sum_nonneg mult_nonneg_nonneg) (auto dest: prime_gt_0_nat) ultimately have R_bounds: "R / n \ {0..<9 / pi\<^sup>2}" by simp have "ln (fact n :: real) \ ln (2 * pi * n) / 2 + n * ln n - n + 1 / (12 * n)" using ln_fact_bounds(2)[of n] n by simp also have "\ / n - ln n = -1 + (ln 2 + ln pi) / (2 * n) + (ln n / n) / 2 + 1 / (12 * real n ^ 2)" using n by (simp add: power2_eq_square field_simps ln_mult) also have "\ \ -1 + (ln 2 + ln pi) / (2 * 3) + (ln 3 / 3) / 2 + 1 / (12 * 3\<^sup>2)" using exp_le n pi_gt3 by (intro add_mono divide_right_mono divide_left_mono mult_mono mult_pos_pos ln_x_over_x_mono power_mono) auto also have "\ \ -1 + (25 / 36 + 25 / 18) / (2 * 3) + (25 / 18 / 3) / 2 + 1 / (12 * 3\<^sup>2)" using ln_pi ln2_le_25_over_36 ln_3 by (intro add_mono divide_left_mono divide_right_mono) auto also have "\ \ 0" by simp finally have "ln n - ln (fact n) / n \ 0" using n by (simp add: divide_right_mono) have "-ln (fact n) \ -ln (2 * pi * n) / 2 - n * ln n + n" using ln_fact_bounds(1)[of n] n by simp also have "ln n + \ / n = -ln (2 * pi) / (2 * n) - (ln n / n) / 2 + 1" using n by (simp add: field_simps ln_mult) also have "\ \ 0 - 0 + 1" using pi_gt3 n by (intro add_mono diff_mono) auto finally have upper: "ln n - ln (fact n) / n \ 1" using n by (simp add: divide_right_mono) with \ln n - ln (fact n) / n \ 0\ have fact_bounds: "ln n - ln (fact n) / n \ {0..1}" by simp have "r \ prime_sum_upto (\p. ln p * 1) n" using less_imp_le[OF frac_lt_1] unfolding r_def \_def prime_sum_upto_def by (intro sum_mono mult_left_mono) (auto simp: dest: prime_gt_0_nat) also have "\ = \ n" by (simp add: \_def) also have "\ < ln 4 * n" using n by (intro \_upper_bound) auto finally have "r / n < ln 4" using n by (simp add: field_simps) moreover have "r \ 0" unfolding r_def prime_sum_upto_def by (intro sum_nonneg mult_nonneg_nonneg) (auto dest: prime_gt_0_nat) ultimately have r_bounds: "r / n \ {0..k. mangoldt k * real (n div k)) (real n)" by (simp add: ln_fact_conv_sum_upto_mangoldt) also have "\ = (\(p,i) | prime p \ i > 0 \ real (p ^ i) \ real n. ln (real p) * real (n div (p ^ i)))" by (intro sum_upto_primepows) (auto simp: mangoldt_non_primepow) also have "{(p, i). prime p \ i > 0 \ real (p ^ i) \ real n} = {(p, i). prime p \ i = 1 \ p \ n} \ {(p, i). prime p \ i > 1 \ (p ^ i) \ n}" unfolding of_nat_le_iff by (auto simp: not_less le_Suc_eq) also have "(\(p,i)\\. ln (real p) * real (n div (p ^ i))) = (\(p,i) | prime p \ i = 1 \ p \ n. ln (real p) * real (n div (p ^ i))) + R" (is "_ = ?S + _") by (subst sum.union_disjoint) (auto intro!: finite_subset[OF _ finite] simp: R_def) also have "?S = prime_sum_upto (\p. ln (real p) * real (n div p)) n" unfolding prime_sum_upto_def by (intro sum.reindex_bij_witness[of _ "\p. (p, 1)" fst]) auto also have "\ = prime_sum_upto (\p. ln (real p) * real n / real p) n - r" unfolding r_def prime_sum_upto_def sum_subtractf[symmetric] using n by (intro sum.cong) (auto simp: frac_def real_of_nat_div algebra_simps) also have "prime_sum_upto (\p. ln (real p) * real n / real p) n = n * \ n" by (simp add: primes_M_def sum_distrib_left sum_distrib_right prime_sum_upto_def field_simps) finally have "\ n - ln n = ln (fact n) / n - ln n + r / n - R / n" using n by (simp add: field_simps) hence "ln n - \ n = ln n - ln (fact n) / n - r / n + R / n" by simp with fact_bounds r_bounds R_bounds show "\ n - ln n \ {-1 - 9 / pi\<^sup>2<..ln 4}" by simp qed text \ As a simple corollary, we obtain a similar bound on the reals. \ lemma mertens_bound_real_strong: fixes x :: real assumes x: "x \ 1" shows "\ x - ln x \ {-1 - 9 / pi ^ 2 - ln (1 + frac x / real (nat \x\)) <.. ln 4}" proof - have "\ x - ln x \ \ (real (nat \x\)) - ln (real (nat \x\))" using assms by simp also have "\ \ ln 4" using mertens_bound_strong[of "nat \x\"] assms by simp finally have "\ x - ln x \ ln 4" . from assms have pos: "real_of_int \x\ \ 0" by linarith have "frac x / real (nat \x\) \ 0" using assms by (intro divide_nonneg_pos) auto moreover have "frac x / real (nat \x\) \ 1 / 1" using assms frac_lt_1[of x] by (intro frac_le) auto ultimately have *: "frac x / real (nat \x\) \ {0..1}" by auto have "ln x - ln (real (nat \x\)) = ln (x / real (nat \x\))" using assms by (subst ln_div) auto also have "x / real (nat \x\) = 1 + frac x / real (nat \x\)" using assms pos by (simp add: frac_def field_simps) finally have "\ x - ln x > -1-9/pi^2-ln (1 + frac x / real (nat \x\))" using mertens_bound_strong[of "nat \x\"] x by simp with \\ x - ln x \ ln 4\ show ?thesis by simp qed text \ We weaken this estimate a bit to obtain nicer bounds: \ lemma mertens_bound_real': fixes x :: real assumes x: "x \ 1" shows "\ x - ln x \ {-2<..25/18}" proof - have "\ x - ln x \ ln 4" using mertens_bound_real_strong[of x] x by simp also have "\ \ 25 / 18" using ln_realpow[of 2 2] ln2_le_25_over_36 by simp finally have "\ x - ln x \ 25 / 18" . have ln2: "ln (2 :: real) \ {2/3..25/36}" using ln_approx_bounds[of 2 1] by (simp add: eval_nat_numeral) have ln3: "ln (3::real) \ {1..10/9}" using ln_approx_bounds[of 3 1] by (simp add: eval_nat_numeral) have ln5: "ln (5::real) \ {4/3..76/45}" using ln_approx_bounds[of 5 1] by (simp add: eval_nat_numeral) have ln7: "ln (7::real) \ {3/2..15/7}" using ln_approx_bounds[of 7 1] by (simp add: eval_nat_numeral) have ln11: "ln (11::real) \ {5/3..290/99}" using ln_approx_bounds[of 11 1] by (simp add: eval_nat_numeral) \ \Choosing the lower bound -2 is somewhat arbitrary here; it is a trade-off between getting a reasonably tight bound and having to make lots of case distinctions. To get -2 as a lower bound, we have to show the cases up to \x = 11\ by case distinction,\ have "\ x - ln x > -2" proof (cases "x \ 11") case False hence "x \ {1..<2} \ x \ {2..<3} \ x \ {3..<5} \ x \ {5..<7} \ x \ {7..<11}" using x by force thus ?thesis proof (elim disjE) assume x: "x \ {1..<2}" hence "ln x - \ x \ ln 2 - 0" by (intro diff_mono) auto also have "\ < 2" using ln2_le_25_over_36 by simp finally show ?thesis by simp next assume x: "x \ {2..<3}" hence [simp]: "\x\ = 2" by (intro floor_unique) auto from x have "ln x - \ x \ ln 3 - ln 2 / 2" by (intro diff_mono) (auto simp: eval_\) also have "\ = ln (9 / 2) / 2" using ln_realpow[of 3 2] by (simp add: ln_div) also have "\ < 2" using ln_approx_bounds[of "9 / 2" 1] by (simp add: eval_nat_numeral) finally show ?thesis by simp next assume x: "x \ {3..<5}" hence "\ 3 = \ x" unfolding primes_M_def by (intro prime_sum_upto_eqI'[where a' = 3 and b' = 4]) (auto simp: nat_le_iff le_numeral_iff nat_eq_iff floor_eq_iff) also have "\ 3 = ln 2 / 2 + ln 3 / 3" by (simp add: eval_\ eval_nat_numeral mark_out_code) finally have [simp]: "\ x = ln 2 / 2 + ln 3 / 3" .. from x have "ln x - \ x \ ln 5 - (ln 2 / 2 + ln 3 / 3)" by (intro diff_mono) auto also have "\ < 2" using ln2 ln3 ln5 by simp finally show ?thesis by simp next assume x: "x \ {5..<7}" hence "\ 5 = \ x" unfolding primes_M_def by (intro prime_sum_upto_eqI'[where a' = 5 and b' = 6]) (auto simp: nat_le_iff le_numeral_iff nat_eq_iff floor_eq_iff) also have "\ 5 = ln 2 / 2 + ln 3 / 3 + ln 5 / 5" by (simp add: eval_\ eval_nat_numeral mark_out_code) finally have [simp]: "\ x = ln 2 / 2 + ln 3 / 3 + ln 5 / 5" .. from x have "ln x - \ x \ ln 7 - (ln 2 / 2 + ln 3 / 3 + ln 5 / 5)" by (intro diff_mono) auto also have "\ < 2" using ln2 ln3 ln5 ln7 by simp finally show ?thesis by simp next assume x: "x \ {7..<11}" hence "\ 7 = \ x" unfolding primes_M_def by (intro prime_sum_upto_eqI'[where a' = 7 and b' = 10]) (auto simp: nat_le_iff le_numeral_iff nat_eq_iff floor_eq_iff) also have "\ 7 = ln 2 / 2 + ln 3 / 3 + ln 5 / 5 + ln 7 / 7" by (simp add: eval_\ eval_nat_numeral mark_out_code) finally have [simp]: "\ x = ln 2 / 2 + ln 3 / 3 + ln 5 / 5 + ln 7 / 7" .. from x have "ln x - \ x \ ln 11 - (ln 2 / 2 + ln 3 / 3 + ln 5 / 5 + ln 7 / 7)" by (intro diff_mono) auto also have "\ < 2" using ln2 ln3 ln5 ln7 ln11 by simp finally show ?thesis by simp qed next case True have "ln x - \ x \ 1 + 9/pi^2 + ln (1 + frac x / real (nat \x\))" using mertens_bound_real_strong[of x] x by simp also have "1 + frac x / real (nat \x\) \ 1 + 1 / 11" using True frac_lt_1[of x] by (intro add_mono frac_le) auto hence "ln (1 + frac x / real (nat \x\)) \ ln (1 + 1 / 11)" using x by (subst ln_le_cancel_iff) (auto intro!: add_pos_nonneg) also have "\ = ln (12 / 11)" by simp also have "\ \ 1585 / 18216" using ln_approx_bounds[of "12 / 11" 1] by (simp add: eval_nat_numeral) also have "9 / pi ^ 2 \ 9 / 3.141592653588 ^ 2" using pi_approx by (intro divide_left_mono power_mono mult_pos_pos) auto also have "1 + \ + 1585 / 18216 < 2" by (simp add: power2_eq_square) finally show ?thesis by simp qed with \\ x - ln x \ 25 / 18\ show ?thesis by simp qed corollary mertens_first_theorem: fixes x :: real assumes x: "x \ 1" shows "\\ x - ln x\ < 2" using mertens_bound_real'[of x] x by (simp add: abs_if) subsection \Mertens' Second Theorem\ text \ Mertens' Second Theorem concerns the asymptotics of the Prime Harmonic Series, namely \[\sum\limits_{p \leq x} \frac{1}{p} = \ln\ln x + M + O\left(\frac{1}{\ln x}\right)\] where $M \approx 0.261497$ is the Meissel--Mertens constant. We define the constant in the following way: \ definition meissel_mertens where "meissel_mertens = 1 - ln (ln 2) + integral {2..} (\t. (\ t - ln t) / (t * ln t ^ 2))" text \ We will require the value of the integral $\int_a^\infty \frac{t}{\ln^2 t} \textrm{d}t = \frac{1}{\ln a}$ as an upper bound on the remainder term: \ lemma integral_one_over_x_ln_x_squared: assumes a: "(a::real) > 1" shows "set_integrable lborel {a<..} (\t. 1 / (t * ln t ^ 2))" (is ?th1) and "set_lebesgue_integral lborel {a<..} (\t. 1 / (t * ln t ^ 2)) = 1 / ln a" (is ?th2) and "((\t. 1 / (t * (ln t)\<^sup>2)) has_integral 1 / ln a) {a<..}" (is ?th3) proof - have cont: "isCont (\t. 1 / (t * (ln t)\<^sup>2)) x" if "x > a" for x using that a by (auto intro!: continuous_intros) have deriv: "((\x. -1 / ln x) has_real_derivative 1 / (x * (ln x)\<^sup>2)) (at x)" if "x > a" for x using that a by (auto intro!: derivative_eq_intros simp: power2_eq_square field_simps) have lim1: "(((\x. - 1 / ln x) \ real_of_ereal) \ -(1 / ln a)) (at_right (ereal a))" unfolding ereal_tendsto_simps using a by (real_asymp simp: field_simps) have lim2: "(((\x. - 1 / ln x) \ real_of_ereal) \ 0) (at_left \)" unfolding ereal_tendsto_simps using a by (real_asymp simp: field_simps) have "set_integrable lborel (einterval a \) (\t. 1 / (t * ln t ^ 2))" by (rule interval_integral_FTC_nonneg[OF _ deriv cont _ lim1 lim2]) (use a in auto) thus ?th1 by simp have "interval_lebesgue_integral lborel (ereal a) \ (\t. 1 / (t * ln t ^ 2)) = 0 - (-(1 / ln a))" by (rule interval_integral_FTC_nonneg[OF _ deriv cont _ lim1 lim2]) (use a in auto) thus ?th2 by (simp add: interval_integral_to_infinity_eq) have "((\t. 1 / (t * ln t ^ 2)) has_integral set_lebesgue_integral lebesgue {a<..} (\t. 1 / (t * ln t ^ 2))) {a<..}" using \?th1\ by (intro has_integral_set_lebesgue) (auto simp: set_integrable_def integrable_completion) also have "set_lebesgue_integral lebesgue {a<..} (\t. 1 / (t * ln t ^ 2)) = 1 / ln a" using \?th2\ unfolding set_lebesgue_integral_def by (subst integral_completion) auto finally show ?th3 . qed text \ We show that the integral in our definition of the Meissel--Mertens constant is well-defined and give an upper bound for its tails: \ lemma assumes "a > (1 :: real)" defines "r \ (\t. (\ t - ln t) / (t * ln t ^ 2))" shows integrable_meissel_mertens: "set_integrable lborel {a<..} r" and meissel_mertens_integral_le: "norm (integral {a<..} r) \ 2 / ln a" proof - have *: "((\t. 2 * (1 / (t * ln t ^ 2))) has_integral 2 * (1 / ln a)) {a<..}" using assms by (intro has_integral_mult_right integral_one_over_x_ln_x_squared) auto show "set_integrable lborel {a<..} r" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "integrable lborel (\t::real. indicator {a<..} t * (2 * (1 / (t * ln t ^ 2))))" using integrable_mult_right[of 2, OF integral_one_over_x_ln_x_squared(1)[of a, unfolded set_integrable_def]] assms by (simp add: algebra_simps) thus "integrable lborel (\t::real. indicator {a<..} t *\<^sub>R (2 / (t * ln t ^ 2)))" by simp fix x :: real show "norm (indicat_real {a<..} x *\<^sub>R r x) \ norm (indicat_real {a<..} x *\<^sub>R (2 / (x * ln x ^ 2)))" proof (cases "x > a") case True thus ?thesis unfolding norm_scaleR norm_mult r_def norm_divide using mertens_first_theorem[of x] assms by (intro mult_mono frac_le divide_nonneg_pos) (auto simp: indicator_def) qed (auto simp: indicator_def) qed (auto simp: r_def) hence "r integrable_on {a<..}" by (simp add: set_borel_integral_eq_integral(1)) hence "norm (integral {a<..} r) \ integral {a<..} (\x. 2 * (1 / (x * ln x ^ 2)))" proof (rule integral_norm_bound_integral) show "(\x. 2 * (1 / (x * (ln x)\<^sup>2))) integrable_on {a<..}" using * by (simp add: has_integral_iff) fix x assume "x \ {a<..}" hence "norm (r x) \ 2 / (x * (ln x)\<^sup>2)" unfolding r_def norm_divide using mertens_first_theorem[of x] assms by (intro mult_mono frac_le divide_nonneg_pos) (auto simp: indicator_def) thus "norm (r x) \ 2* (1 / (x * ln x ^ 2))" by simp qed also have "\ = 2 / ln a" using * by (simp add: has_integral_iff) finally show "norm (integral {a<..} r) \ 2 / ln a" . qed lemma integrable_on_meissel_mertens: assumes "A \ {1..}" "Inf A > 1" "A \ sets borel" shows "(\t. (\ t - ln t) / (t * ln t ^ 2)) integrable_on A" proof - from assms obtain x where x: "1 < x" "x < Inf A" using dense by blast from assms have "bdd_below A" by (intro bdd_belowI[of _ 1]) auto hence "A \ {Inf A..}" by (auto simp: cInf_lower) also have "\ \ {x<..}" using x by auto finally have *: "A \ {x<..}" . have "set_integrable lborel A (\t. (\ t - ln t) / (t * ln t ^ 2))" by (rule set_integrable_subset[OF integrable_meissel_mertens[of x]]) (use x * assms in auto) thus ?thesis by (simp add: set_borel_integral_eq_integral(1)) qed lemma meissel_mertens_bounds: "\meissel_mertens - 1 + ln (ln 2)\ \ 2 / ln 2" proof - have *: "{2..} - {2<..} = {2::real}" by auto also have "negligible \" by simp finally have "integral {2..} (\t. (\ t - ln t) / (t * (ln t)\<^sup>2)) = integral {2<..} (\t. (\ t - ln t) / (t * (ln t)\<^sup>2))" by (intro sym[OF integral_subset_negligible]) auto also have "norm \ \ 2 / ln 2" by (rule meissel_mertens_integral_le) auto finally show "\meissel_mertens - 1 + ln (ln 2)\ \ 2 / ln 2" by (simp add: meissel_mertens_def) qed text \ Finally, obtaining Mertens' second theorem from the first one is nothing but a routine summation by parts, followed by a use of the above bound: \ theorem mertens_second_theorem: defines "f \ prime_sum_upto (\p. 1 / p)" shows "\x. x \ 2 \ \f x - ln (ln x) - meissel_mertens\ \ 4 / ln x" and "(\x. f x - ln (ln x) - meissel_mertens) \ O(\x. 1 / ln x)" proof - define r where "r = (\t. (\ t - ln t) / (t * ln t ^ 2))" { fix x :: real assume x: "x > 2" have "((\t. \ t * (-1 / (t * ln t ^ 2))) has_integral \ x * (1 / ln x) - \ 2 * (1 / ln 2) - (\n\real -` {2<..x}. ind prime n * (ln n / real n) * (1 / ln n))) {2..x}" unfolding primes_M_def prime_sum_upto_altdef1 using x by (intro partial_summation_strong[of "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp: power2_eq_square - simp flip: has_field_derivative_iff_has_vector_derivative) + simp flip: has_real_derivative_iff_has_vector_derivative) also have "\ x * (1 / ln x) - \ 2 * (1 / ln 2) - (\n\real -` {2<..x}. ind prime n * (ln n / n) * (1 / ln n)) = \ x / ln x - (\n\insert 2 (real -` {2<..x}). ind prime n * (ln n / n) * (1 / ln n))" (is "_ = _ - ?S") by (subst sum.insert) (auto simp: primes_M_def finite_vimage_real_of_nat_greaterThanAtMost eval_prime_sum_upto) also have "?S = f x" unfolding f_def prime_sum_upto_altdef1 sum_upto_def using x by (intro sum.mono_neutral_cong_left) (auto simp: not_less numeral_2_eq_2 le_Suc_eq) finally have "((\t. -\ t / (t * ln t ^ 2)) has_integral (\ x / ln x - f x)) {2..x}" by simp from has_integral_neg[OF this] have "((\t. \ t / (t * ln t ^ 2)) has_integral (f x - \ x / ln x)) {2..x}" by simp hence "((\t. \ t / (t * ln t ^ 2) - 1 / (t * ln t)) has_integral (f x - \ x / ln x - (ln (ln x) - ln (ln 2)))) {2..x}" using x by (intro has_integral_diff fundamental_theorem_of_calculus) - (auto simp flip: has_field_derivative_iff_has_vector_derivative + (auto simp flip: has_real_derivative_iff_has_vector_derivative intro!: derivative_eq_intros) also have "?this \ (r has_integral (f x - \ x / ln x - (ln (ln x) - ln (ln 2)))) {2..x}" by (intro has_integral_cong) (auto simp: r_def field_simps power2_eq_square) finally have \ . } note integral = this define R\<^sub>\ where "R\<^sub>\ = (\x. \ x - ln x)" have \: "\ x = ln x + R\<^sub>\ x" for x by (simp add: R\<^sub>\_def) define I where "I = (\x. integral {x..} r)" define C where "C = (1 - ln (ln 2) + I 2)" have C_altdef: "C = meissel_mertens" by (simp add: I_def r_def C_def meissel_mertens_def) show bound: " \f x - ln (ln x) - meissel_mertens\ \ 4 / ln x" if x: "x \ 2" for x proof (cases "x = 2") case True hence "\f x - ln (ln x) - meissel_mertens\ = \1 / 2 - ln (ln 2) - meissel_mertens\" by (simp add: f_def eval_prime_sum_upto ) also have "\ \ 2 / ln 2 + 1 / 2" using meissel_mertens_bounds by linarith also have "\ \ 2 / ln 2 + 2 / ln 2" using ln2_le_25_over_36 by (intro add_mono divide_left_mono) auto finally show ?thesis using True by simp next case False hence x: "x > 2" using x by simp have "integral {2..x} r + I x = integral ({2..x} \ {x..}) r" unfolding I_def r_def using x by (intro integral_Un [symmetric] integrable_on_meissel_mertens) (auto simp: max_def r_def) also have "{2..x} \ {x..} = {2..}" using x by auto finally have *: "integral {2..x} r = I 2 - I x" unfolding I_def by simp have eq: "f x - ln (ln x) - C = R\<^sub>\ x / ln x - I x" using integral[OF x] x by (auto simp: C_def field_simps \ has_integral_iff *) also have "\\\ \ \R\<^sub>\ x / ln x\ + norm (I x)" unfolding real_norm_def by (rule abs_triangle_ineq4) also have "\R\<^sub>\ x / ln x\ \ 2 / \ln x\" unfolding R\<^sub>\_def abs_divide using mertens_first_theorem[of x] x by (intro divide_right_mono) auto also have "{x..} - {x<..} = {x}" and "{x<..} \ {x..}" by auto hence "I x = integral {x<..} r" unfolding I_def by (intro integral_subset_negligible [symmetric]) simp_all also have "norm \ \ 2 / ln x" using meissel_mertens_integral_le[of x] x by (simp add: r_def) finally show "\f x - ln (ln x) - meissel_mertens\ \ 4 / ln x" using x by (simp add: C_altdef) qed have "(\x. f x - ln (ln x) - C) \ O(\x. 1 / ln x)" proof (intro landau_o.bigI[of 4] eventually_mono[OF eventually_ge_at_top[of 2]]) fix x :: real assume x: "x \ 2" with bound[OF x] show "norm (f x - ln (ln x) - C) \ 4 * norm (1 / ln x)" by (simp add: C_altdef) qed (auto intro!: add_pos_nonneg) thus "(\x. f x - ln (ln x) - meissel_mertens) \ O(\x. 1 / ln x)" by (simp add: C_altdef) qed corollary prime_harmonic_asymp_equiv: "prime_sum_upto (\p. 1 / p) \[at_top] (\x. ln (ln x))" proof - define f where "f = prime_sum_upto (\p. 1 / p)" have "(\x. f x - ln (ln x) - meissel_mertens + meissel_mertens) \ o(\x. ln (ln x))" unfolding f_def by (rule sum_in_smallo[OF landau_o.big_small_trans[OF mertens_second_theorem(2)]]) real_asymp+ hence "(\x. f x - ln (ln x)) \ o(\x. ln (ln x))" by simp thus ?thesis unfolding f_def by (rule smallo_imp_asymp_equiv) qed text \ As a corollary, we get the divergence of the prime harmonic series. \ corollary prime_harmonic_diverges: "filterlim (prime_sum_upto (\p. 1 / p)) at_top at_top" using asymp_equiv_symI[OF prime_harmonic_asymp_equiv] by (rule asymp_equiv_at_top_transfer) real_asymp (*<*) unbundle no_prime_counting_notation (*>*) end \ No newline at end of file diff --git a/thys/Prime_Number_Theorem/Newman_Ingham_Tauberian.thy b/thys/Prime_Number_Theorem/Newman_Ingham_Tauberian.thy --- a/thys/Prime_Number_Theorem/Newman_Ingham_Tauberian.thy +++ b/thys/Prime_Number_Theorem/Newman_Ingham_Tauberian.thy @@ -1,680 +1,680 @@ (* File: Newman_Ingham_Tauberian.thy Author: Manuel Eberl (TU München) A proof of Newman's "analytic theorem", originally stated by Ingham *) section \Ingham's Tauberian Theorem\ theory Newman_Ingham_Tauberian imports "HOL-Real_Asymp.Real_Asymp" Prime_Number_Theorem_Library begin text \ In his proof of the Prime Number Theorem, Newman~\cite{newman1998analytic} uses a Tauberian theorem that was first proven by Ingham. Newman gives a nice and straightforward proof of this theorem based on contour integration. This section will be concerned with proving this theorem. This Tauberian theorem is probably the part of the Newman's proof of the Prime Number Theorem where most of the ``heavy lifting'' is done. Its purpose is to extend the summability of a Dirichlet series with bounded coefficients from the region $\mathfrak{R}(s)>1$ to $\mathfrak{R}(s)\geq 1$. In order to show it, we first require a number of auxiliary bounding lemmas. \ lemma newman_ingham_aux1: fixes R :: real and z :: complex assumes R: "R > 0" and z : "norm z = R" shows "norm (1 / z + z / R\<^sup>2) = 2 * \Re z\ / R\<^sup>2" proof - from z and R have [simp]: "z \ 0" by auto have "1 / z + z / R\<^sup>2 = (R\<^sup>2 + z\<^sup>2) * (1 / R\<^sup>2 / z)" using R by (simp add: field_simps power2_eq_square) also have "norm \ = norm (R\<^sup>2 + z\<^sup>2) / R ^ 3" by (simp add: numeral_3_eq_3 z norm_divide norm_mult power2_eq_square) also have "R\<^sup>2 + z\<^sup>2 = z * (z + cnj z)" using complex_norm_square[of z] by (simp add: z power2_eq_square algebra_simps) also have "norm \ = 2 * \Re z\ * R" by (subst complex_add_cnj) (simp_all add: z norm_mult) also have "\ / R ^ 3 = 2 * \Re z\ / R\<^sup>2" using R by (simp add: field_simps numeral_3_eq_3 power2_eq_square) finally show ?thesis . qed lemma newman_ingham_aux2: fixes m :: nat and w z :: complex assumes "1 \ m" "1 \ Re w" "0 < Re z" and f: "\n. 1 \ n \ norm (f n) \ C" shows "norm (\n=1..m. f n / n powr (w - z)) \ C * (m powr Re z) * (1 / m + 1 / Re z)" proof - have [simp]: "C \ 0" by (rule order.trans[OF _ f[of 1]]) auto have "norm (\n=1..m. f n / n powr (w - z)) \ (\n=1..m. C / n powr (1 - Re z))" by (rule sum_norm_le) (insert assms, auto simp: norm_divide norm_powr_real_powr intro!: frac_le assms powr_mono) also have "\ = C * (\n=1..m. n powr (Re z - 1))" by (subst sum_distrib_left) (simp_all add: powr_diff) also have "\ \ C * (m powr Re z * (1 / Re z + 1 / m))" using zeta_partial_sum_le'[of "Re z" m] assms by (intro mult_left_mono) auto finally show ?thesis by (simp add: mult_ac add_ac) qed lemma hurwitz_zeta_real_bound_aux: fixes a x :: real assumes ax: "a > 0" "x > 1" shows "(\i. (a + real (Suc i)) powr (-x)) \ a powr (1 - x) / (x - 1)" proof (rule decreasing_sum_le_integral, goal_cases) have "((\t. (a + t) powr -x) has_integral -(a powr (-x + 1)) / (-x + 1)) (interior {0..})" using powr_has_integral_at_top[of 0 a "-x"] using ax by (simp add: interior_real_atLeast) also have "-(a powr (- x + 1)) / (- x + 1) = a powr (1 - x) / (x - 1)" using ax by (simp add: field_simps) finally show "((\t. (a + t) powr -x) has_integral a powr (1 - x) / (x - 1)) {0..}" by (subst (asm) has_integral_interior) auto qed (insert ax, auto intro!: powr_mono2') text \ Given a function that is analytic on some vertical line segment, we can find a rectangle around that line segment on which the function is also analytic. \ lemma analytic_on_axis_extend: fixes y1 y2 x :: real defines "S \ {z. Re z = x \ Im z \ {y1..y2}}" assumes "y1 \ y2" assumes "f analytic_on S" obtains x1 x2 :: real where "x1 < x" "x2 > x" "f analytic_on cbox (Complex x1 y1) (Complex x2 y2)" proof - define C where "C = {box a b |a b z. f analytic_on box a b \ z \ box a b \ z \ S}" have "S = cbox (Complex x y1) (Complex x y2)" by (auto simp: S_def in_cbox_complex_iff) also have "compact \" by simp finally have 1: "compact S" . have 2: "S \ \C" proof (intro subsetI) fix z assume "z \ S" from \f analytic_on S\ and this obtain a b where "z \ box a b" "f analytic_on box a b" by (blast elim: analytic_onE_box) with \z \ S\ show "z \ \C" unfolding C_def by blast qed have 3: "open X" if "X \ C" for X using that by (auto simp: C_def) from compactE[OF 1 2 3] obtain T where T: "T \ C" "finite T" "S \ \T" by blast define x1 where "x1 = Max (insert (x - 1) ((\X. x + (Inf (Re ` X) - x) / 2) ` T))" define x2 where "x2 = Min (insert (x + 1) ((\X. x + (Sup (Re ` X) - x) / 2) ` T))" have *: "x + (Inf (Re ` X) - x) / 2 < x \ x + (Sup (Re ` X) - x) / 2 > x" if "X \ T" for X proof - from that and T obtain a b s where [simp]: "X = box a b" and s: "s \ box a b" "s \ S" by (force simp: C_def) hence le: "Re a < Re b" "Im a < Im b" by (auto simp: in_box_complex_iff) show ?thesis using le s unfolding \X = box a b\ Re_image_box[OF le] Im_image_box[OF le] by (auto simp: S_def in_box_complex_iff) qed from * T have "x1 < x" unfolding x1_def by (subst Max_less_iff) auto from * T have "x2 > x" unfolding x2_def by (subst Min_gr_iff) auto have "f analytic_on (\T)" using T by (subst analytic_on_Union) (auto simp: C_def) moreover have "z \ \T" if "z \ cbox (Complex x1 y1) (Complex x2 y2)" for z proof - from that have "Complex x (Im z) \ S" by (auto simp: in_cbox_complex_iff S_def) with T obtain X where X: "X \ T" "Complex x (Im z) \ X" by auto with T obtain a b where [simp]: "X = box a b" by (auto simp: C_def) from X have le: "Re a < Re b" "Im a < Im b" by (auto simp: in_box_complex_iff) from that have "Re z \ x2" by (simp add: in_cbox_complex_iff) also have "\ \ x + (Sup (Re ` X) - x) / 2" unfolding x2_def by (rule Min.coboundedI)(use T X in auto) also have "\ = (x + Re b) / 2" using le unfolding \X = box a b\ Re_image_box[OF le] by (simp add: field_simps) also have "\ < (Re b + Re b) / 2" using X by (intro divide_strict_right_mono add_strict_right_mono) (auto simp: in_box_complex_iff) also have "\ = Re b" by simp finally have [simp]: "Re z < Re b" . have "Re a = (Re a + Re a) / 2" by simp also have "\ < (x + Re a) / 2" using X by (intro divide_strict_right_mono add_strict_right_mono) (auto simp: in_box_complex_iff) also have "\ = x + (Inf (Re ` X) - x) / 2" using le unfolding \X = box a b\ Re_image_box[OF le] by (simp add: field_simps) also have "\ \ x1" unfolding x1_def by (rule Max.coboundedI)(use T X in auto) also have "\ \ Re z" using that by (simp add: in_cbox_complex_iff) finally have [simp]: "Re z > Re a" . from X have "z \ X" by (simp add: in_box_complex_iff) with T X show ?thesis by blast qed hence "cbox (Complex x1 y1) (Complex x2 y2) \ \T" by blast ultimately have "f analytic_on cbox (Complex x1 y1) (Complex x2 y2)" by (rule analytic_on_subset) with \x1 < x\ and \x2 > x\ and that[of x1 x2] show ?thesis by blast qed text \ We will now prove the theorem. The precise setting is this: Consider a Dirichlet series $F(s) = \sum a_n n^{-s}$ with bounded coefficients. Clearly, this converges to an analytic function $f(s)$ on $\{s\mid \mathfrak R(s)>1\}$. If $f(s)$ is analytic on the larger set $\{s\mid \mathfrak R(s)\geq 1\}$, $F$ converges to $f(s)$ for all $\mathfrak R(s) \geq 1$. The proof follows Newman's argument very closely, but some of the precise bounds we use are a bit different from his. Also, like Harrison, we choose a combination of a semicircle and a rectangle as our contour, whereas Newman uses a circle with a vertical cut-off. The result of the Residue theorem is the same in both cases, but the bounding of the contributions of the different parts is somewhat different. The reason why we picked Harrison's contour over Newman's is because we could not understand how his bounding of the different contributions fits to his contour, and it seems likely that this is also the reason why Harrison altered the contour in the first place. \ lemma Newman_Ingham_1: fixes F :: "complex fds" and f :: "complex \ complex" assumes coeff_bound: "fds_nth F \ O(\_. 1)" assumes f_analytic: "f analytic_on {s. Re s \ 1}" assumes F_conv_f: "\s. Re s > 1 \ eval_fds F s = f s" assumes w: "Re w \ 1" shows "fds_converges F w" and "eval_fds F w = f w" proof - \ \We get a bound on our coefficients and call it \C\.\ obtain C where C: "C \ 1" "\n. norm (fds_nth F n) \ C" using natfun_bigo_1E[OF coeff_bound, where lb = 1] by blast write contour_integral ("\[_]") \ \We show convergence directly by showing that the difference between the partial sums and the limit vanishes.\ have "(\N. eval_fds (fds_truncate N F) w) \ f w" unfolding tendsto_iff dist_norm norm_minus_commute[of "eval_fds F s" for F s] proof safe fix \ :: real assume \: "\ > 0" \ \We choose an integration radius that is big enough for the error to be sufficiently small.\ define R where "R = max 1 (3 * C / \)" have R: "R \ 3 * C / \" "R \ 1" by (auto simp: R_def) \ \Next, we extend the analyticity of \f (w + z)\ to the left of the complex plane within a thin rectangle that is at least as high as the circle.\ obtain l where l: "l > 0" "(\z. f (w + z)) analytic_on {s. Re s > 0 \ Im s \ {-R-1<.. Re s > -l}" proof - have f_analytic': "(\z. f (w + z)) analytic_on {s. Re s \ 0}" by (rule analytic_on_compose_gen[OF _ f_analytic, unfolded o_def]) (insert w, auto intro: analytic_intros) hence "(\z. f (w + z)) analytic_on {s. Re s = 0 \ Im s \ {-R-1..R+1}}" by (rule analytic_on_subset) auto from analytic_on_axis_extend[OF _ this] obtain x1 x2 where x12: "x1 < 0" "x2 > 0" "(\z. f (w + z)) analytic_on cbox (Complex x1 (-R-1)) (Complex x2 (R+1))" using \R \ 1\ by auto from this(3) have "(\z. f (w + z)) analytic_on {s. Re s \ {x1..0} \ Im s \ {-R-1..R+1}}" by (rule analytic_on_subset) (insert x12, auto simp: in_cbox_complex_iff) with f_analytic' have "(\z. f (w + z)) analytic_on ({s. Re s \ 0} \ {s. Re s \ {x1..0} \ Im s \ {-R-1..R+1}})" by (subst analytic_on_Un) auto hence "(\z. f (w + z)) analytic_on {s. Re s > 0 \ Im s \ {-R-1<.. Re s > x1}" by (rule analytic_on_subset) auto with \x1 < 0\ and that[of "-x1"] show ?thesis by auto qed \ \The function \f (w + z)\ is now analytic on the open box $(-l; R+1) + i(-R+1; R+1)$. We call this region \X\.\ define X where "X = box (Complex (-l) (-R-1)) (Complex (R+1) (R+1))" have [simp, intro]: "open X" "convex X" by (simp_all add: X_def open_box) from R l have [simp]: "0 \ X" by (auto simp: X_def in_box_complex_iff) have analytic: "(\z. f (w + z)) analytic_on X" by (rule analytic_on_subset[OF l(2)]) (auto simp: X_def in_box_complex_iff) note f_analytic' [analytic_intros] = analytic_on_compose_gen[OF _ analytic, unfolded o_def] note f_holo [holomorphic_intros] = holomorphic_on_compose_gen[OF _ analytic_imp_holomorphic[OF analytic], unfolded o_def] note f_cont [continuous_intros] = continuous_on_compose2[OF holomorphic_on_imp_continuous_on[OF analytic_imp_holomorphic[OF analytic]]] \ \We now pick a smaller closed box \X'\ inside the big open box \X\. This is because we need a compact set for the next step. our integration path still lies entirely within \X'\, and since \X'\ is compact, \f (w + z)\ is bounded on it, so we obtain such a bound and call it \M\.\ define \ where "\ = min (1/2) (l/2)" from l have \: "\ > 0" "\ \ 1/2" "\ < l" by (auto simp: \_def) define X' where "X' = cbox (Complex (-\) (-R)) (Complex R R)" have "X' \ X" unfolding X'_def X_def using l(1) R \ by (intro subset_box_imp) (auto simp: Basis_complex_def) have [intro]: "compact X'" by (simp add: X'_def) moreover have "continuous_on X' (\z. f (w + z))" using w \X' \ X\ by (auto intro!: continuous_intros) ultimately obtain M where M: "M \ 0" "\z. z \ X' \ norm (f (w + z)) \ M" using continuous_on_compact_bound by blast \ \Our objective is now to show that the difference between the \N\-th partial sum and the limit is below a certain bound (depending on \N\) which tends to \0\ for \N \ \\. We use the following bound:\ define bound where "bound = (\N::nat. (2*C/R + C/N + 3*M / (pi*R*ln N) + 3*R*M / (\*pi * N powr \)))" have "2 * C / R < \" using M(1) R C(1) \(1) \ by (auto simp: field_simps) \ \Evidently this is below @{term \} for sufficiently large \N\.\ hence "eventually (\N::nat. bound N < \) at_top" using M(1) R C(1) \(1) \ unfolding bound_def by real_asymp \ \It now only remains to show that the difference is indeed less than the claimed bound.\ thus "eventually (\N. norm (f w - eval_fds (fds_truncate N F) w) < \) at_top" using eventually_gt_at_top[of 1] proof eventually_elim case (elim N) note N = this \ \Like Harrison (and unlike Newman), our integration path \\\ consists of a semicircle \A\ of radius \R\ in the right-halfplane and a box of width \\\ and height \2R\ on the left halfplane. The latter consists of three straight lines, which we call \B1\ to \B3\.\ define A where "A = part_circlepath 0 R (-pi/2) (pi/2)" define B2 where "B2 = linepath (Complex (-\) R) (Complex (-\) (-R))" define B1 where "B1 = linepath (R * \) (R * \ - \)" define B3 where "B3 = linepath (-R * \ - \) (-R * \)" define \ where "\ = A +++ B1 +++ B2 +++ B3" \ \We first need to show some basic facts about the geometry of our integration path.\ have [simp, intro]: "path A" "path B1" "path B3" "path B2" "valid_path A" "valid_path B1" "valid_path B3" "valid_path B2" "arc A" "arc B1" "arc B3" "arc B2" "pathstart A = -\ * R" "pathfinish A = \ * R" "pathstart B1 = \ * R" "pathfinish B1 = R * \ - \" "pathstart B3 = -R * \ - \" "pathfinish B3 = -\ * R" "pathstart B2 = R * \ - \" "pathfinish B2 = -R * \ - \" using R \ by (simp_all add: A_def B1_def B3_def exp_eq_polar B2_def Complex_eq arc_part_circlepath) hence [simp, intro]: "valid_path \" by (simp add: \_def A_def B1_def B3_def B2_def exp_eq_polar Complex_eq) hence [simp, intro]: "path \" using valid_path_imp_path by blast have [simp]: "pathfinish \ = pathstart \" by (simp add: \_def exp_eq_polar) have image_B2: "path_image B2 = {s. Re s = -\ \ Im s \ {-R..R}}" using R by (auto simp: closed_segment_same_Re closed_segment_eq_real_ivl B2_def) have image_B1: "path_image B1 = {s. Re s \ {-\..0} \ Im s = R}" and image_B3: "path_image B3 = {s. Re s \ {-\..0} \ Im s = -R}" using \ by (auto simp: B1_def B3_def closed_segment_same_Im closed_segment_eq_real_ivl) have image_A: "path_image A = {s. Re s \ 0 \ norm s = R}" unfolding A_def using R by (subst path_image_semicircle_Re_ge) auto also have "z \ \ \ z \ X' - {0}" for z using complex_Re_le_cmod[of z] abs_Im_le_cmod[of z] \ R by (auto simp: X'_def in_cbox_complex_iff) hence "{s. Re s \ 0 \ norm s = R} \ X' - {0}" by auto finally have "path_image B2 \ X' - {0}" "path_image A \ X' - {0}" "path_image B1 \ X' - {0}" "path_image B3 \ X' - {0}" using \\ > 0\ by (auto simp: X'_def in_cbox_complex_iff image_B2 image_B1 image_B3) note path_images = this \X' \ X\ \ \\\\ is a simple path, which, combined with its simple geometric shape, makes reasoning about its winding numbers trivial.\ from R have "simple_path A" unfolding A_def by (subst simple_path_part_circlepath) auto have "simple_path \" unfolding \_def proof (intro simple_path_join_loop subsetI arc_join, goal_cases) fix z assume z: "z \ path_image A \ path_image (B1 +++ B2 +++ B3)" with image_A have "Re z \ 0" "norm z = R" by auto with z R \ show "z \ {pathstart A, pathstart (B1 +++ B2 +++ B3)}" by (auto simp: path_image_join image_B1 image_B2 image_B3 complex_eq_iff) qed (insert R, auto simp: image_B1 image_B3 path_image_join image_B2 complex_eq_iff) \ \We define the integrands in the same fashion as Newman:\ define g where "g = (\z::complex. f (w + z) * N powr z * (1 / z + z / R\<^sup>2))" define S where "S = eval_fds (fds_truncate N F)" define g_S where "g_S = (\z::complex. S (w + z) * N powr z * (1 / z + z / R\<^sup>2))" define rem where "rem = (\z::complex. f z - S z)" define g_rem where "g_rem = (\z::complex. rem (w + z) * N powr z * (1 / z + z / R\<^sup>2))" have g_holo: "g holomorphic_on X - {0}" unfolding g_def by (auto intro!: holomorphic_intros analytic_imp_holomorphic'[OF analytic]) have rem_altdef: "rem z = eval_fds (fds_remainder N F) z" if "Re z > 1" for z proof - have abscissa: "abs_conv_abscissa F \ 1" using assms by (intro bounded_coeffs_imp_abs_conv_abscissa_le_1) (simp_all add: natfun_bigo_iff_Bseq) from assms and that have "f z = eval_fds F z" by auto also have "F = fds_truncate N F + fds_remainder N F" by (rule fds_truncate_plus_remainder [symmetric]) also from that have "eval_fds \ z = S z + eval_fds (fds_remainder N F) z" unfolding S_def by (subst eval_fds_add) (auto intro!: fds_abs_converges_imp_converges fds_abs_converges[OF le_less_trans[OF abscissa]]) finally show ?thesis by (simp add: rem_def) qed \ \We now come to the first application of the residue theorem along the path \\\:\ have "\[\] g = 2 * pi * \ * winding_number \ 0 * residue g 0" proof (subst Residue_theorem) show "g holomorphic_on X - {0}" by fact show "path_image \ \ X - {0}" using path_images by (auto simp: \_def path_image_join) thus "\z. z \ X \ winding_number \ z = 0" by (auto intro!: simply_connected_imp_winding_number_zero[of X] convex_imp_simply_connected) qed (insert path_images, auto intro: convex_connected) also have "winding_number \ 0 = 1" proof (rule simple_closed_path_winding_number_pos) from R \ have "\g\{A, B1, B2, B3}. Re (winding_number g 0) > 0" unfolding A_def B1_def B2_def B3_def by (auto intro!: winding_number_linepath_pos_lt winding_number_part_circlepath_pos_less) hence "valid_path \ \ 0 \ path_image \ \ Re (winding_number \ 0) > 0" unfolding \_def using path_images(1-4) by (intro winding_number_join_pos_combined') auto thus "Re (winding_number \ 0) > 0" by simp qed (insert path_images \simple_path \\, auto simp: \_def path_image_join) also have "residue g 0 = f w" proof - have "g = (\z::complex. f (w + z) * N powr z * (1 + z\<^sup>2 / R\<^sup>2) / z)" by (auto simp: g_def divide_simps fun_eq_iff power2_eq_square simp del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) moreover from N have "residue \ 0 = f w" by (subst residue_simple'[of X]) (auto intro!: holomorphic_intros analytic_imp_holomorphic[OF analytic]) ultimately show ?thesis by (simp only:) qed finally have "2 * pi * \ * f w = \[\] g" by simp also have "\ = \[A] g + \[B2] g + \[B1] g + \[B3] g" unfolding \_def by (subst contour_integral_join, (insert path_images, auto intro!: contour_integral_join contour_integrable_holomorphic_simple g_holo)[4])+ (simp_all add: add_ac) finally have integral1: "2 * pi * \ * f w = \[A] g + \[B2] g + \[B1] g + \[B3] g" . \ \Next, we apply the residue theorem along a circle of radius \R\ to another integrand that is related to the partial sum:\ have "\[circlepath 0 R] g_S = 2 * pi * \ * residue g_S 0" proof (subst Residue_theorem) show "g_S holomorphic_on UNIV - {0}" by (auto simp: g_S_def S_def intro!: holomorphic_intros) qed (insert R, auto simp: winding_number_circlepath_centre) also have "residue g_S 0 = S w" proof - have "g_S = (\z::complex. S (w + z) * N powr z * (1 + z\<^sup>2 / R\<^sup>2) / z)" by (auto simp: g_S_def divide_simps fun_eq_iff power2_eq_square simp del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) moreover from N have "residue \ 0 = S w" by (subst residue_simple'[of X]) (auto intro!: holomorphic_intros simp: S_def) ultimately show ?thesis by (simp only:) qed finally have "2 * pi * \ * S w = \[circlepath 0 R] g_S" .. \ \We split this integral into integrals along two semicircles in the left and right half-plane, respectively:\ also have "\ = \[part_circlepath 0 R (-pi/2) (3*pi/2)] g_S" proof (rule Cauchy_theorem_homotopic_loops) show "homotopic_loops (-{0}) (circlepath 0 R) (part_circlepath 0 R (- pi / 2) (3 * pi / 2))" unfolding circlepath_def using R by (intro homotopic_loops_part_circlepath[where k = 1]) auto qed (auto simp: g_S_def S_def intro!: holomorphic_intros) also have "\ = \[A +++ -A] g_S" proof (rule Cauchy_theorem_homotopic_paths) have *: "-A = part_circlepath 0 R (pi/2) (3*pi/2)" unfolding A_def by (intro part_circlepath_mirror[where k = 0]) auto from R show "homotopic_paths (-{0}) (part_circlepath 0 R (-pi/2) (3*pi/2)) (A +++ -A)" unfolding * unfolding A_def by (intro homotopic_paths_part_circlepath) (auto dest!: in_path_image_part_circlepath) qed (auto simp: g_S_def S_def A_def exp_eq_polar intro!: holomorphic_intros) also have "\ = \[A] g_S + \[-A] g_S" using R by (intro contour_integral_join contour_integrable_holomorphic_simple[of _ "-{0}"]) (auto simp: A_def g_S_def S_def path_image_mirror dest!: in_path_image_part_circlepath intro!: holomorphic_intros) also have "\[-A] g_S = -\[A] (\x. g_S (-x))" by (simp add: A_def contour_integral_mirror contour_integral_neg) finally have integral2: "2 * pi * \ * S w = \[A] g_S - \[A] (\x. g_S (-x))" by simp \ \Next, we show a small bounding lemma that we will need for the final estimate:\ have circle_bound: "norm (1 / z + z / R\<^sup>2) \ 2 / R" if [simp]: "norm z = R" for z :: complex proof - have "norm (1 / z + z / R\<^sup>2) \ 1 / R + 1 / R" by (intro order.trans[OF norm_triangle_ineq] add_mono) (insert R, simp_all add: norm_divide norm_mult power2_eq_square) thus ?thesis by simp qed \ \The next bound differs somewhat from Newman's, but it works just as well. Its purpose is to bound the contribution of the two short horizontal line segments.\ have B12_bound: "norm (integral {- \..0} (\x. g (x + R' * \))) \ 3 * M / R / ln N" (is "?I \ _") if "\R'\ = R" for R' proof - have "?I \ integral {-\..0} (\x. 3 * M / R * N powr x)" proof (rule integral_norm_bound_integral) fix x assume x: "x \ {-\..0}" define z where "z = x + \ * R'" from R that have [simp]: "z \ 0" "Re z = x" "Im z = R'" by (auto simp: z_def complex_eq_iff) from x R that have "z \ X'" by (auto simp: z_def X'_def in_cbox_complex_iff) from x R that have "norm z \ \ + R" by (intro order.trans[OF cmod_le add_mono]) auto hence "norm (1 / z + z / R\<^sup>2) \ 1 / R + (\ / R + 1) / R" using R that abs_Im_le_cmod[of z] by (intro order.trans[OF norm_triangle_ineq add_mono]) (auto simp: norm_divide norm_mult power2_eq_square field_simps ) also have "\ / R \ 1" using \ R by auto finally have "norm (1 / z + z / R\<^sup>2) \ 3 / R" using R by (simp add: divide_right_mono) hence "norm (g z) \ M * N powr x * (3 / R)" unfolding g_def norm_mult using \M \ 0\ \z \ X'\ by (intro mult_mono mult_nonneg_nonneg M) (auto simp: norm_powr_real_powr) thus "norm (g (x + R' * \)) \ 3 * M / R * N powr x" by (simp add: mult_ac z_def) qed (insert N R l that \, auto intro!: integrable_continuous_real continuous_intros simp: g_def X_def complex_eq_iff in_box_complex_iff) also have "\ = 3 * M / R * integral {-\..0} (\x. N powr x)" by simp also have "((\x. N powr x) has_integral (N powr 0 / ln N - N powr (-\) / ln N)) {-\..0}" using \ N by (intro fundamental_theorem_of_calculus) - (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] powr_def + (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric] powr_def intro!: derivative_eq_intros) hence "integral {-\..0} (\x. N powr x) = 1 / ln (real N) - real N powr - \ / ln (real N)" using N by (simp add: has_integral_iff) also have "\ \ 1 / ln (real N)" using N by simp finally show ?thesis using M R by (simp add: mult_left_mono divide_right_mono) qed \ \We combine the two results from the residue theorem and obtain an integral representation of the difference between the partial sums and the limit:\ have "2 * pi * \ * (f w - S w) = \[A] g - \[A] g_S + \[A] (\x. g_S (-x)) + \[B1] g + \[B3] g + \[B2] g" unfolding ring_distribs integral1 integral2 by (simp add: algebra_simps) also have "\[A] g - \[A] g_S = \[A] (\x. g x - g_S x)" using path_images by (intro contour_integral_diff [symmetric]) (auto intro!: contour_integrable_holomorphic_simple[of _ "X - {0}"] holomorphic_intros simp: g_S_def g_holo S_def) also have "\ = \[A] g_rem" by (simp add: g_rem_def g_def g_S_def algebra_simps rem_def) finally have "2 * pi * \ * (f w - S w) = \[A] g_rem + \[A] (\x. g_S (- x)) + \[B1] g + \[B3] g + \[B2] g" . \ \We now bound each of these integrals individually:\ also have "norm \ \ 2 * C * pi / R + 2 * C * pi * (1 / N + 1 / R) + 3 * M / R / ln N + 3 * M / R / ln N + 6 * R * M * N powr (-\) / \" proof (rule order.trans[OF norm_triangle_ineq] add_mono)+ have "\[B1] g = -\[reversepath B1] g" by (simp add: contour_integral_reversepath) also have "\[reversepath B1] g = integral {-\..0} (\x. g (x + R * \))" unfolding B1_def reversepath_linepath using \ by (subst contour_integral_linepath_same_Im) auto also have "norm (-\) = norm \" by simp also have "\ \ 3 * M / R / ln N" using R by (intro B12_bound) auto finally show "norm (\[B1] g) \ \" . next have "\[B3] g = integral {-\..0} (\x. g (x + (-R) * \))" unfolding B3_def using \ by (subst contour_integral_linepath_same_Im) auto also have "norm \ \ 3 * M / R / ln N" using R by (intro B12_bound) auto finally show "norm (\[B3] g) \ \" . next have "norm (\[B2] g) \ M * N powr (-\) * (3 / \) * norm (Complex (- \) (-R) - Complex (- \) R)" unfolding B2_def proof ((rule contour_integral_bound_linepath; (fold B2_def)?), goal_cases) case (3 z) from 3 \ R have [simp]: "z \ 0" and Re_z: "Re z = -\" and Im_z: "Im z \ {-R..R}" by (auto simp: closed_segment_same_Re closed_segment_eq_real_ivl) from 3 have "z \ X'" using R \ path_images by (auto simp: B2_def) from 3 \ R have "norm z \ sqrt (\^2 + R^2)" unfolding cmod_def using Re_z Im_z by (intro real_sqrt_le_mono add_mono) (auto simp: power2_le_iff_abs_le) from power_mono[OF this, of 2] have norm_sqr: "norm z ^ 2 \ \\<^sup>2 + R\<^sup>2" by simp have "norm (1 / z + z / R\<^sup>2) \ (1 + (norm z)\<^sup>2 / R\<^sup>2) / \" unfolding add_divide_distrib using \ R abs_Re_le_cmod[of z] by (intro order.trans[OF norm_triangle_ineq] add_mono) (auto simp: norm_divide norm_mult field_simps power2_eq_square Re_z) also have "\ \ (1 + (1 + \\<^sup>2 / R\<^sup>2)) / \" using \ R \z \ X'\ norm_sqr unfolding X'_def by (intro divide_right_mono add_left_mono) (auto simp: field_simps in_cbox_complex_iff intro!: power_mono) also have "\\<^sup>2 / R\<^sup>2 \ 1" using \ R by (auto simp: field_simps intro!: power_mono) finally have "norm (1 / z + z / R\<^sup>2) \ 3 / \" using \ by(simp add: divide_right_mono) with \z \ X'\ show "norm (g z) \ M * N powr (-\) * (3 / \)" unfolding g_def norm_mult by (intro mult_mono mult_nonneg_nonneg M) (auto simp: norm_powr_real_powr Re_z) qed (insert path_images M \, auto intro!: contour_integrable_holomorphic_simple[OF g_holo]) thus "norm (\[B2] g) \ 6 * R * M * N powr (-\) / \" using R by (simp add: field_simps cmod_def real_sqrt_mult) next have "norm (\[A] (\x. g_S (- x))) \ (2 * C / (real N * R) + 2 * C / R\<^sup>2) * R * ((pi/2) - (-pi/2))" unfolding A_def proof ((rule contour_integral_bound_part_circlepath_strong[where k = "{R * \, -R*\}"]; (fold A_def)?), goal_cases) case (6 z) hence [simp]: "z \ 0" and "norm z = R" using R by (auto simp: A_def dest!: in_path_image_part_circlepath) from 6 have "Re z \ 0" using \norm z = R\ by (auto simp: cmod_def abs_if complex_eq_iff split: if_splits) with 6 have "Re z > 0" using image_A by auto have "S (w - z) = (\k = 1..N. fds_nth F k / of_nat k powr (w - z))" by (simp add: S_def eval_fds_truncate) also have "norm \ \ C * N powr Re z * (1 / N + 1 / Re z)" using \Re z > 0\ w N by (intro newman_ingham_aux2 C) auto finally have "norm (S (w - z)) \ \" . hence "norm (g_S (-z)) \ (C * N powr (Re z) * (1 / N + 1 / Re z)) * N powr (-Re z) * (2 * Re z / R\<^sup>2)" unfolding g_S_def norm_mult using newman_ingham_aux1[OF _ \norm z = R\] \Re z > 0\ \C \ 1\ R by (intro mult_mono mult_nonneg_nonneg circle_bound) (auto simp: norm_powr_real_powr norm_uminus_minus) also have "\ = 2 * C * (Re z / N + 1) / R\<^sup>2" using R N \Re z > 0\ by (simp add: powr_minus algebra_simps) also have "\ \ 2 * C / (N * R) + 2 * C / R\<^sup>2" unfolding add_divide_distrib ring_distribs using R N abs_Re_le_cmod[of z] \norm z = R\ \Re z > 0\ \C \ 1\ by (intro add_mono) (auto simp: power2_eq_square field_simps mult_mono) finally show ?case . qed (insert R N image_A C, auto intro!: contour_integrable_holomorphic_simple[of _ "-{0}"] holomorphic_intros simp: g_S_def S_def) also have "\ = 2 * C * pi * (1 / N + 1 / R)" using R N by (simp add: power2_eq_square field_simps) finally show "norm (\[A] (\x. g_S (- x))) \ \" . next have "norm (\[A] g_rem) \ (2 * C / R\<^sup>2) * R * ((pi/2) - (-pi/2))" unfolding A_def proof ((rule contour_integral_bound_part_circlepath_strong[where k = "{R * \, -R*\}"]; (fold A_def)?), goal_cases) case (6 z) hence [simp]: "z \ 0" and "norm z = R" using R by (auto simp: A_def dest!: in_path_image_part_circlepath) from 6 have "Re z \ 0" using \norm z = R\ by (auto simp: cmod_def abs_if complex_eq_iff split: if_splits) with 6 have "Re z > 0" using image_A by auto have summable: "summable (\n. C * (1 / (Suc n + N) powr (Re w + Re z)))" using summable_hurwitz_zeta_real[of "Re w + Re z" "Suc N"] \Re z > 0\ w unfolding powr_minus by (intro summable_mult) (auto simp: field_simps) have "rem (w + z) = (\n. fds_nth F (Suc n + N) / (Suc n + N) powr (w + z))" using \Re z > 0\ w by (simp add: rem_altdef eval_fds_remainder) also have "norm \ \ (\n. C / (Suc n + N) powr Re (w + z))" using summable by (intro norm_suminf_le) (auto simp: norm_divide norm_powr_real_powr intro!: divide_right_mono C) also have "\ = (\n. C * (Suc n + N) powr -Re (w + z))" unfolding powr_minus by (simp add: field_simps) also have "\ = C * (\n. (Suc n + N) powr -Re (w + z))" using summable_hurwitz_zeta_real[of "Re w + Re z" "Suc N"] \Re z > 0\ w by (subst suminf_mult) (auto simp: add_ac) also have "(\n. (Suc n + N) powr -Re (w + z)) \ N powr (1 - Re (w + z)) / (Re (w + z) - 1)" using \Re z > 0\ w N hurwitz_zeta_real_bound_aux[of N "Re (w + z)"] by (auto simp: add_ac) also have "\ \ N powr -Re z / Re z" using w N \Re z > 0\ by (intro frac_le powr_mono) auto finally have "norm (rem (w + z)) \ C / (Re z * N powr Re z)" using C by (simp add: mult_left_mono mult_right_mono powr_minus field_simps) hence "norm (g_rem z) \ (C / (Re z * N powr Re z)) * N powr (Re z) * (2 * Re z / R\<^sup>2)" unfolding g_rem_def norm_mult using newman_ingham_aux1[OF _ \norm z = R\] R \Re z > 0\ C by (intro mult_mono mult_nonneg_nonneg circle_bound) (auto simp: norm_powr_real_powr norm_uminus_minus) also have "\ = 2 * C / R\<^sup>2" using R N \Re z > 0\ by (simp add: powr_minus field_simps) finally show ?case . next show "g_rem contour_integrable_on A" using path_images by (auto simp: g_rem_def rem_def S_def intro!: contour_integrable_holomorphic_simple[of _ "X-{0}"] holomorphic_intros) qed (insert R N C, auto) also have "(2 * C / R\<^sup>2) * R * ((pi/2) - (-pi/2)) = 2 * C * pi / R" using R by (simp add: power2_eq_square field_simps) finally show "norm (\[A] g_rem) \ \" . qed also have "\ = 4*C*pi/R + 2*C*pi/N + 6*M/R / ln N + 6*R*M*N powr - \ / \" by (simp add: algebra_simps) also have "\ = 2*pi * (2*C/R + C/N + 3*M / (pi*R*ln N) + 3*R*M / (\*pi * N powr \))" by (simp add: field_simps powr_minus ) also have "norm (2 * pi * \ * (f w - S w)) = 2 * pi * norm (f w - S w)" by (simp add: norm_mult) finally have "norm (f w - S w) \ bound N" by (simp add: bound_def) also have "bound N < \" by fact finally show "norm (f w - S w) < \" . qed qed thus "fds_converges F w" by (auto simp: fds_converges_altdef2 intro: convergentI) thus "eval_fds F w = f w" using \(\N. eval_fds (fds_truncate N F) w) \ f w\ by (intro tendsto_unique[OF _ tendsto_eval_fds_truncate]) auto qed text \ The theorem generalises in a trivial way; we can replace the requirement that the coefficients of $f(s)$ be $O(1)$ by $O(n^{\sigma-1})$ for some $\sigma\in\mathbb{R}$, then $f(s)$ converges for $\mathfrak{R}(s)>\sigma$. If it can be analytically continued to $\mathfrak{R}(s)\geq\sigma$, it is also convergent there. \ theorem Newman_Ingham: fixes F :: "complex fds" and f :: "complex \ complex" assumes coeff_bound: "fds_nth F \ O(\n. n powr of_real (\ - 1))" assumes f_analytic: "f analytic_on {s. Re s \ \}" assumes F_conv_f: "\s. Re s > \ \ eval_fds F s = f s" assumes w: "Re w \ \" shows "fds_converges F w" and "eval_fds F w = f w" proof - define F' where "F' = fds_shift (-of_real (\ - 1)) F" define f' where "f' = f \ (\s. s + of_real (\ - 1))" have "fds_nth F' = (\n. fds_nth F n * of_nat n powr -of_real(\ - 1))" by (auto simp: fun_eq_iff F'_def) also have "\ \ O(\n. of_nat n powr of_real (\ - 1) * of_nat n powr -of_real(\ - 1))" by (intro landau_o.big.mult_right assms) also have "(\n. of_nat n powr of_real (\ - 1) * of_nat n powr -of_real (\ - 1)) \ \(\_. 1)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) (auto simp: powr_minus powr_diff) finally have bigo: "fds_nth F' \ O(\_. 1)" . from f_analytic have analytic: "f' analytic_on {s. Re s \ 1}" unfolding f'_def by (intro analytic_on_compose_gen[OF _ f_analytic]) (auto intro!: analytic_intros) have F'_f: "eval_fds F' s = f' s" if "Re s > 1" for s using assms that by (auto simp: F'_def f'_def algebra_simps) have w': "1 \ Re (w - of_real (\ - 1))" using w by simp have 1: "fds_converges F' (w - of_real (\ - 1))" using bigo analytic F'_f w' by (rule Newman_Ingham_1) thus "fds_converges F w" by (auto simp: F'_def) have 2: "eval_fds F' (w - of_real (\ - 1)) = f' (w - of_real (\ - 1))" using bigo analytic F'_f w' by (rule Newman_Ingham_1) thus "eval_fds F w = f w" using assms by (simp add: F'_def f'_def) qed 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) + simp flip: has_real_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) + simp flip: has_real_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 + (auto simp flip: has_real_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 + (auto simp flip: has_real_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 + (auto simp flip: has_real_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 + (auto simp flip: has_real_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 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 + (auto simp flip: has_real_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.thy b/thys/Prime_Number_Theorem/Prime_Number_Theorem.thy --- a/thys/Prime_Number_Theorem/Prime_Number_Theorem.thy +++ b/thys/Prime_Number_Theorem/Prime_Number_Theorem.thy @@ -1,826 +1,826 @@ (* File: Prime Number Theorem.thy Authors: Manuel Eberl (TU München), Larry Paulson (University of Cambridge) A proof of the Prime Number Theorem and some related properties *) section \The Prime Number Theorem\ theory Prime_Number_Theorem imports Newman_Ingham_Tauberian Prime_Counting_Functions begin (*<*) unbundle prime_counting_notation (*>*) subsection \Constructing Newman's function\ text \ Starting from Mertens' first theorem, i.\,e.\ $\mathfrak M(x) = \ln x + O(1)$, we now want to derive that $\mathfrak M(x) = \ln x + c + o(1)$. This result is considerably stronger and it implies the Prime Number Theorem quite directly. In order to do this, we define the Dirichlet series \[f(s) = \sum_{n=1}^\infty \frac{\mathfrak{M}(n)}{n^s}\ .\] We will prove that this series extends meromorphically to $\mathfrak{R}(s)\geq 1$ and apply Ingham's theorem to it (after we subtracted its pole at $s = 1$). \ definition fds_newman where "fds_newman = fds (\n. complex_of_real (\ n))" lemma fds_nth_newman: "fds_nth fds_newman n = of_real (\ n)" by (simp add: fds_newman_def fds_nth_fds) lemma norm_fds_nth_newman: "norm (fds_nth fds_newman n) = \ n" unfolding fds_nth_newman norm_of_real by (intro abs_of_nonneg sum_nonneg divide_nonneg_pos) (auto dest: prime_ge_1_nat) text \ The Dirichlet series $f(s) + \zeta'(s)$ has the coefficients $\mathfrak{M}(n) - \ln n$, so by Mertens' first theorem, $f(s) + \zeta'(s)$ has bounded coefficients. \ lemma bounded_coeffs_newman_minus_deriv_zeta: defines "f \ fds_newman + fds_deriv fds_zeta" shows "Bseq (\n. fds_nth f n)" proof - have "(\n. \ (real n) - ln (real n)) \ O(\_. 1)" using mertens_bounded by (rule landau_o.big.compose) real_asymp from natfun_bigo_1E[OF this, of 1] obtain c where c: "c \ 1" "\n. \\ (real n) - ln (real n)\ \ c" by auto show ?thesis proof (intro BseqI[of c] allI) fix n :: nat show "norm (fds_nth f n) \ c" proof (cases "n = 0") case False hence "fds_nth f n = of_real (\ n - ln n)" by (simp add: f_def fds_nth_newman fds_nth_deriv fds_nth_zeta scaleR_conv_of_real) also from \n \ 0\ have "norm \ \ c" using c(2)[of n] by (simp add: in_Reals_norm) finally show ?thesis . qed (insert c, auto) qed (insert c, auto) qed text \ A Dirichlet series with bounded coefficients converges for all $s$ with $\mathfrak{R}(s)>1$ and so does $\zeta'(s)$, so we can conclude that $f(s)$ does as well. \ lemma abs_conv_abscissa_newman: "abs_conv_abscissa fds_newman \ 1" and conv_abscissa_newman: "conv_abscissa fds_newman \ 1" proof - define f where "f = fds_newman + fds_deriv fds_zeta" have "abs_conv_abscissa f \ 1" using bounded_coeffs_newman_minus_deriv_zeta unfolding f_def by (rule bounded_coeffs_imp_abs_conv_abscissa_le_1) hence "abs_conv_abscissa (f - fds_deriv fds_zeta) \ 1" by (intro abs_conv_abscissa_diff_leI) (auto simp: abs_conv_abscissa_deriv) also have "f - fds_deriv fds_zeta = fds_newman" by (simp add: f_def) finally show "abs_conv_abscissa fds_newman \ 1" . from conv_le_abs_conv_abscissa and this show "conv_abscissa fds_newman \ 1" by (rule order.trans) qed text \ We now change the order of summation to obtain an alternative form of $f(s)$ in terms of a sum of Hurwitz $\zeta$ functions. \ lemma eval_fds_newman_conv_infsetsum: assumes s: "Re s > 1" shows "eval_fds fds_newman s = (\\<^sub>ap | prime p. (ln (real p) / real p) * hurwitz_zeta p s)" "(\p. ln (real p) / real p * hurwitz_zeta p s) abs_summable_on {p. prime p}" proof - from s have conv: "fds_abs_converges fds_newman s" by (intro fds_abs_converges le_less_trans[OF abs_conv_abscissa_newman]) auto define f where "f = (\n p. ln (real p) / real p / of_nat n powr s)" have eq: "(\\<^sub>an\{p..}. f n p) = ln (real p) / real p * hurwitz_zeta p s" if "prime p" for p proof - have "(\\<^sub>an\{p..}. f n p) = (\\<^sub>ax\{p..}. (ln (real p) / of_nat p) * (1 / of_nat x powr s))" by (simp add: f_def) also have "\ = (ln (real p) / of_nat p) * (\\<^sub>ax\{p..}. 1 / of_nat x powr s)" using abs_summable_hurwitz_zeta[of s 0 p] that s by (intro infsetsum_cmult_right) (auto dest: prime_gt_0_nat) also have "(\\<^sub>ax\{p..}. 1 / of_nat x powr s) = hurwitz_zeta p s" using s that by (subst hurwitz_zeta_nat_conv_infsetsum(2)) (auto dest: prime_gt_0_nat simp: field_simps powr_minus) finally show ?thesis . qed have norm_f: "norm (f n p) = ln p / p / n powr Re s" if "prime p" for n p :: nat by (auto simp: f_def norm_divide norm_mult norm_powr_real_powr) from conv have "(\n. norm (fds_nth fds_newman n / n powr s)) abs_summable_on UNIV" by (intro abs_summable_on_normI) (simp add: fds_abs_converges_altdef') also have "(\n. norm (fds_nth fds_newman n / n powr s)) = (\n. \p | prime p \ p \ n. norm (f n p))" by (auto simp: norm_divide norm_fds_nth_newman sum_divide_distrib primes_M_def prime_sum_upto_def norm_mult norm_f norm_powr_real_powr intro!: sum.cong) finally have summable1: "(\(n,p). f n p) abs_summable_on (SIGMA n:UNIV. {p. prime p \ p \ n})" using conv by (subst abs_summable_on_Sigma_iff) auto also have "?this \ (\(p,n). f n p) abs_summable_on (\(n,p). (p,n)) ` (SIGMA n:UNIV. {p. prime p \ p \ n})" by (subst abs_summable_on_reindex_iff [symmetric]) (auto simp: case_prod_unfold inj_on_def) also have "(\(n,p). (p,n)) ` (SIGMA n:UNIV. {p. prime p \ p \ n}) = (SIGMA p:{p. prime p}. {p..})" by auto finally have summable2: "(\(p,n). f n p) abs_summable_on \" . from abs_summable_on_Sigma_project1'[OF this] have "(\p. \\<^sub>an\{p..}. f n p) abs_summable_on {p. prime p}" by auto also have "?this \ (\p. ln (real p) / real p * hurwitz_zeta p s) abs_summable_on {p. prime p}" by (intro abs_summable_on_cong eq) auto finally show \ . have "eval_fds fds_newman s = (\\<^sub>an. \p | prime p \ p \ n. ln (real p) / real p / of_nat n powr s)" using conv by (simp add: eval_fds_altdef fds_nth_newman sum_divide_distrib primes_M_def prime_sum_upto_def) also have "\ = (\\<^sub>an. \\<^sub>ap | prime p \ p \ n. f n p)" unfolding f_def by (subst infsetsum_finite) auto also have "\ = (\\<^sub>a(n, p) \ (SIGMA n:UNIV. {p. prime p \ p \ n}). f n p)" using summable1 by (subst infsetsum_Sigma) auto also have "\ = (\\<^sub>a(p, n) \ (\(n,p). (p, n)) ` (SIGMA n:UNIV. {p. prime p \ p \ n}). f n p)" by (subst infsetsum_reindex) (auto simp: case_prod_unfold inj_on_def) also have "(\(n,p). (p, n)) ` (SIGMA n:UNIV. {p. prime p \ p \ n}) = (SIGMA p:{p. prime p}. {p..})" by auto also have "(\\<^sub>a(p,n)\\. f n p) = (\\<^sub>ap | prime p. \\<^sub>an\{p..}. f n p)" using summable2 by (subst infsetsum_Sigma) auto also have "(\\<^sub>ap | prime p. \\<^sub>an\{p..}. f n p) = (\\<^sub>ap | prime p. ln (real p) / real p * hurwitz_zeta p s)" by (intro infsetsum_cong eq) auto finally show "eval_fds fds_newman s = (\\<^sub>ap | prime p. (ln (real p) / real p) * hurwitz_zeta p s)" . qed text \ We now define a meromorphic continuation of $f(s)$ on $\mathfrak{R}(s) > \frac{1}{2}$. To construct $f(s)$, we express it as \[f(s) = \frac{1}{z-1}\left(\bar f(s) - \frac{\zeta'(s)}{\zeta(s)}\right)\ ,\] where $\bar f(s)$ (which we shall call \pre_newman\) is a function that is analytic on $\Re(s) > \frac{1}{2}$, which can be shown fairly easily using the Weierstra{\ss} M test. $\zeta'(s)/\zeta(s)$ is meromorphic except for a single pole at $s = 1$ and one $k$-th order pole for any $k$-th order zero of $\zeta$, but for the Prime Number Theorem, we are only concerned with the area $\mathfrak{R}(s) \geq 1$, where $\zeta$ does not have any zeros. Taken together, this means that $f(s)$ is analytic for $\mathfrak{R}(s)\geq 1$ except for a double pole at $s = 1$, which we will take care of later. \ context fixes A :: "nat \ complex \ complex" and B :: "nat \ complex \ complex" defines "A \ (\p s. (s - 1) * pre_zeta (real p) s - of_nat p / (of_nat p powr s * (of_nat p powr s - 1)))" defines "B \ (\p s. of_real (ln (real p)) / of_nat p * A p s)" begin definition pre_newman :: "complex \ complex" where "pre_newman s = (\p. if prime p then B p s else 0)" definition newman where "newman s = 1 / (s - 1) * (pre_newman s - deriv zeta s / zeta s)" text \ The sum used in the definition of \pre_newman\ converges uniformly on any disc within the half-space with $\mathfrak{R}(s) > \frac{1}{2}$ by the Weierstra{\ss} M test. \ lemma uniform_limit_pre_newman: assumes r: "r \ 0" "Re s - r > 1 / 2" shows "uniform_limit (cball s r) (\n s. \p 1 / 2" if "dist s z \ r" for z using abs_Re_le_cmod[of "s - z"] r that by (auto simp: dist_norm abs_if split: if_splits) define x where "x = Re s - r" \ \The lower bound for the real part in the disc\ from r Re have "x > 1 / 2" by (auto simp: x_def) \ \The following sequence \M\ bounds the summand, and it is obviously $O(n^{-1-\epsilon})$ and therefore summable\ define C where "C = (norm s + r + 1) * (norm s + r) / x" define M where "M = (\p::nat. ln p * (C / p powr (x + 1) + 1 / (p powr x * (p powr x - 1))))" show ?thesis unfolding pre_newman_def proof (intro Weierstrass_m_test_ev[OF eventually_mono[OF eventually_gt_at_top[of 1]]] ballI) show "summable M" proof (rule summable_comparison_test_bigo) define \ where "\ = min (2 * x - 1) x / 2" from \x > 1 / 2\ have \: "\ > 0" "1 + \ < 2 * x" "1 + \ < x + 1" by (auto simp: \_def min_def field_simps) show "M \ O(\n. n powr (- 1 - \))" unfolding M_def distrib_left by (intro sum_in_bigo) (use \ in real_asymp)+ from \ show "summable (\n. norm (n powr (- 1 - \)))" by (simp add: summable_real_powr_iff) qed next fix p :: nat and z assume p: "p > 1" and z: "z \ cball s r" from z r Re[of z] have x: "Re z \ x" "x > 1 / 2" and "Re z > 1 / 2" using abs_Re_le_cmod[of "s - z"] by (auto simp: x_def algebra_simps dist_norm) have norm_z: "norm z \ norm s + r" using z norm_triangle_ineq2[of z s] r by (auto simp: dist_norm norm_minus_commute) from \p > 1\ and x and r have "M p \ 0" by (auto simp: C_def M_def intro!: mult_nonneg_nonneg add_nonneg_nonneg divide_nonneg_pos) have bound: "norm ((z - 1) * pre_zeta p z) \ norm (z - 1) * (norm z / (Re z * p powr Re z))" using pre_zeta_bound'[of z p] p \Re z > 1 / 2\ unfolding norm_mult by (intro mult_mono pre_zeta_bound) auto have "norm (B p z) = ln p / p * norm (A p z)" unfolding B_def using \p > 1\ by (simp add: B_def norm_mult norm_divide) also have "\ \ ln p / p * (norm (z - 1) * norm z / Re z / p powr Re z + p / (p powr Re z * (p powr Re z - 1)))" unfolding A_def using \p > 1\ and \Re z > 1 / 2\ and bound by (intro mult_left_mono order.trans[OF norm_triangle_ineq4 add_mono] mult_left_mono) (auto simp: norm_divide norm_mult norm_powr_real_powr intro!: divide_left_mono order.trans[OF _ norm_triangle_ineq2]) also have "\ = ln p * (norm (z - 1) * norm z / Re z / p powr (Re z + 1) + 1 / (p powr Re z * (p powr Re z - 1)))" using \p > 1\ by (simp add: field_simps powr_add powr_minus) also have "norm (z - 1) * norm z / Re z / p powr (Re z + 1) \ C / p powr (x + 1)" unfolding C_def using r \Re z > 1 / 2\ norm_z p x by (intro mult_mono frac_le powr_mono order.trans[OF norm_triangle_ineq4]) auto also have "1 / (p powr Re z * (p powr Re z - 1)) \ 1 / (p powr x * (p powr x - 1))" using \p > 1\ x by (intro divide_left_mono mult_mono powr_mono diff_right_mono mult_pos_pos) (auto simp: ge_one_powr_ge_zero) finally have "norm (B p z) \ M p" using \p > 1\ by (simp add: mult_left_mono M_def) with \M p \ 0\ show "norm (if prime p then B p z else 0) \ M p" by simp qed qed lemma sums_pre_newman: "Re s > 1 / 2 \ (\p. if prime p then B p s else 0) sums pre_newman s" using tendsto_uniform_limitI[OF uniform_limit_pre_newman[of 0 s]] by (auto simp: sums_def) lemma analytic_pre_newman [THEN analytic_on_subset, analytic_intros]: "pre_newman analytic_on {s. Re s > 1 / 2}" proof - have holo: "(\s::complex. if prime p then B p s else 0) holomorphic_on X" if "X \ {s. Re s > 1 / 2}" for X and p :: nat using that by (cases "prime p") (auto intro!: holomorphic_intros simp: B_def A_def dest!: prime_gt_1_nat) have holo': "pre_newman holomorphic_on ball s r" if r: "r \ 0" "Re s - r > 1 / 2" for s r proof - from r have Re: "Re z > 1 / 2" if "dist s z \ r" for z using abs_Re_le_cmod[of "s - z"] r that by (auto simp: dist_norm abs_if split: if_splits) show ?thesis by (rule holomorphic_uniform_limit[OF _ uniform_limit_pre_newman[of r s]]) (insert that Re, auto intro!: always_eventually holomorphic_on_imp_continuous_on holomorphic_intros holo) qed show ?thesis unfolding analytic_on_def proof safe fix s assume "Re s > 1 / 2" thus "\r>0. pre_newman holomorphic_on ball s r" by (intro exI[of _ "(Re s - 1 / 2) / 2"] conjI holo') (auto simp: field_simps) qed qed lemma holomorphic_pre_newman [holomorphic_intros]: "X \ {s. Re s > 1 / 2} \ pre_newman holomorphic_on X" using analytic_pre_newman by (rule analytic_imp_holomorphic) lemma eval_fds_newman: assumes s: "Re s > 1" shows "eval_fds fds_newman s = newman s" proof - have eq: "(ln (real p) / real p) * hurwitz_zeta p s = 1 / (s - 1) * (ln (real p) / (p powr s - 1) + B p s)" if p: "prime p" for p proof - have "(ln (real p) / real p) * hurwitz_zeta p s = ln (real p) / real p * (p powr (1 - s) / (s - 1) + pre_zeta p s)" using s by (auto simp add: hurwitz_zeta_def) also have "\ = 1 / (s - 1) * (ln (real p) / (p powr s - 1) + B p s)" using p s by (simp add: divide_simps powr_diff B_def) (auto simp: A_def field_simps dest: prime_gt_1_nat)? finally show ?thesis . qed have "(\p. (ln (real p) / real p) * hurwitz_zeta p s) abs_summable_on {p. prime p}" using s by (intro eval_fds_newman_conv_infsetsum) hence "(\p. 1 / (s - 1) * (ln (real p) / (p powr s - 1) + B p s)) abs_summable_on {p. prime p}" by (subst (asm) abs_summable_on_cong[OF eq refl]) auto hence summable: "(\p. ln (real p) / (p powr s - 1) + B p s) abs_summable_on {p. prime p}" using s by (subst (asm) abs_summable_on_cmult_right_iff) auto from s have [simp]: "s \ 1" by auto have "eval_fds fds_newman s = (\\<^sub>ap | prime p. (ln (real p) / real p) * hurwitz_zeta p s)" using s by (rule eval_fds_newman_conv_infsetsum) also have "\ = (\\<^sub>ap | prime p. 1 / (s - 1) * (ln (real p) / (p powr s - 1) + B p s))" by (intro infsetsum_cong eq) auto also have "\ = 1 / (s - 1) * (\\<^sub>ap | prime p. ln (real p) / (p powr s - 1) + B p s)" (is "_ = _ * ?S") by (rule infsetsum_cmult_right[OF summable]) also have "?S = (\p. if prime p then ln (real p) / (p powr s - 1) + B p s else 0)" by (subst infsetsum_nat[OF summable]) auto also have "\ = (\p. (if prime p then ln (real p) / (p powr s - 1) else 0) + (if prime p then B p s else 0))" by (intro suminf_cong) auto also have "\ = pre_newman s - deriv zeta s / zeta s" using sums_pre_newman[of s] sums_logderiv_zeta[of s] s by (subst suminf_add [symmetric]) (auto simp: sums_iff) finally show ?thesis by (simp add: newman_def) qed end text \ Next, we shall attempt to get rid of the pole by subtracting suitable multiples of $\zeta(s)$ and $\zeta'(s)$. To this end, we shall first prove the following alternative definition of $\zeta'(s)$: \ lemma deriv_zeta_eq': assumes "0 < Re s" "s \ 1" shows "deriv zeta s = deriv (\z. pre_zeta 1 z * (z - 1)) s / (s - 1) - (pre_zeta 1 s * (s - 1) + 1) / (s - 1)\<^sup>2" (is "_ = ?rhs") proof (rule DERIV_imp_deriv) 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 *: "deriv (\z. pre_zeta 1 z * (z - 1)) s = deriv (pre_zeta 1) s * (s - 1) + pre_zeta 1 s" by (subst deriv_mult) (auto intro!: holomorphic_on_imp_differentiable_at[of _ UNIV] holomorphic_intros) hence "((\s. pre_zeta 1 s + 1 / (s - 1)) has_field_derivative deriv (pre_zeta 1) s - 1 / ((s - 1) * (s - 1))) (at s)" using assms by (auto intro!: derivative_eq_intros) also have "deriv (pre_zeta 1) s - 1 / ((s - 1) * (s - 1)) = ?rhs" using * assms by (simp add: divide_simps power2_eq_square, simp add: field_simps) also have "((\s. pre_zeta 1 s + 1 / (s - 1)) has_field_derivative ?rhs) (at s) \ (zeta has_field_derivative ?rhs) (at s)" using assms by (intro has_field_derivative_cong_ev eventually_mono[OF t1_space_nhds[of _ 1]]) (auto simp: zeta_def hurwitz_zeta_def) finally show \ . qed text \ From this, it follows that $(s - 1) \zeta'(s) - \zeta'(s) / \zeta(s)$ is analytic for $\mathfrak{R}(s) \geq 1$: \ lemma analytic_zeta_derivdiff: obtains a where "(\z. if z = 1 then a else (z - 1) * deriv zeta z - deriv zeta z / zeta z) analytic_on {s. Re s \ 1}" proof have neq: "pre_zeta 1 z * (z - 1) + 1 \ 0" if "Re z \ 1" for z using zeta_Re_ge_1_nonzero[of z] that by (cases "z = 1") (auto simp: zeta_def hurwitz_zeta_def divide_simps) let ?g = "\z. (1 - inverse (pre_zeta 1 z * (z - 1) + 1)) * ((z - 1) * deriv ((\u. pre_zeta 1 u * (u - 1))) z - (pre_zeta 1 z * (z - 1) + 1))" show "(\z. if z = 1 then deriv ?g 1 else (z - 1) * deriv zeta z - deriv zeta z / zeta z) analytic_on {s. Re s \ 1}" (is "?f analytic_on _") proof (rule pole_theorem_analytic_0) show "?g analytic_on {s. 1 \ Re s}" using neq by (auto intro!: analytic_intros) next show "\d>0. \w\ball z d - {1}. ?g w = (w - 1) * ?f w" if z: "z \ {s. 1 \ Re s}" for z proof - have *: "isCont (\z. pre_zeta 1 z * (z - 1) + 1) z" by (auto intro!: continuous_intros) obtain e where "e > 0" and e: "\y. dist z y < e \ pre_zeta (Suc 0) y * (y-1) + 1 \ 0" using continuous_at_avoid [OF * neq[of z]] z by auto show ?thesis proof (intro exI ballI conjI) fix w assume w: "w \ ball z (min e 1) - {1}" then have "Re w > 0" using complex_Re_le_cmod [of "z-w"] z by (simp add: dist_norm) with w show "?g w = (w - 1) * (if w = 1 then deriv ?g 1 else (w - 1) * deriv zeta w - deriv zeta w / zeta w)" by (subst (1 2) deriv_zeta_eq', simp_all add: zeta_def hurwitz_zeta_def divide_simps e power2_eq_square) (simp_all add: algebra_simps)? qed (use \e > 0\ in auto) qed qed auto qed text \ Finally, $f(s) + \zeta'(s) + c\zeta(s)$ is analytic. \ lemma analytic_newman_variant: obtains c a where "(\z. if z = 1 then a else newman z + deriv zeta z + c * zeta z) analytic_on {s. Re s \ 1}" proof - obtain c where (* -euler_mascheroni *) c: "(\z. if z = 1 then c else (z - 1) * deriv zeta z - deriv zeta z / zeta z) analytic_on {s. Re s \ 1}" using analytic_zeta_derivdiff by blast let ?g = "\z. pre_newman z + (if z = 1 then c else (z - 1) * deriv zeta z - deriv zeta z / zeta z) - (c + pre_newman 1) * (pre_zeta 1 z * (z - 1) + 1)" have "(\z. if z = 1 then deriv ?g 1 else newman z + deriv zeta z + (-(c + pre_newman 1)) * zeta z) analytic_on {s. Re s \ 1}" (is "?f analytic_on _") proof (rule pole_theorem_analytic_0) show "?g analytic_on {s. 1 \ Re s}" by (intro c analytic_intros) auto next show "\d>0. \w\ball z d - {1}. ?g w = (w - 1) * ?f w" if "z \ {s. 1 \ Re s}" for z using that by (intro exI[of _ 1], simp_all add: newman_def divide_simps zeta_def hurwitz_zeta_def) (auto simp: field_simps)? qed auto with that show ?thesis by blast qed subsection \The asymptotic expansion of \\\\ text \ Our next goal is to show the key result that $\mathfrak{M}(x) = \ln n + c + o(1)$. As a first step, we invoke Ingham's Tauberian theorem on the function we have just defined and obtain that the sum \[\sum\limits_{n=1}^\infty \frac{\mathfrak{M}(n) - \ln n + c}{n}\] exists. \ lemma mertens_summable: obtains c :: real where "summable (\n. (\ n - ln n + c) / n)" proof - (* c = euler_mascheroni - pre_newman 1 *) from analytic_newman_variant obtain c a where analytic: "(\z. if z = 1 then a else newman z + deriv zeta z + c * zeta z) analytic_on {s. Re s \ 1}" . define f where "f = (\z. if z = 1 then a else newman z + deriv zeta z + c * zeta z)" have analytic: "f analytic_on {s. Re s \ 1}" using analytic by (simp add: f_def) define F where "F = fds_newman + fds_deriv fds_zeta + fds_const c * fds_zeta" note le = conv_abscissa_add_leI conv_abscissa_deriv_le conv_abscissa_newman conv_abscissa_mult_const_left note intros = le le[THEN le_less_trans] le[THEN order.trans] fds_converges have eval_F: "eval_fds F s = f s" if s: "Re s > 1" for s proof - have "eval_fds F s = eval_fds (fds_newman + fds_deriv fds_zeta) s + eval_fds (fds_const c * fds_zeta) s" unfolding F_def using s by (subst eval_fds_add) (auto intro!: intros) also have "\ = f s" using s unfolding f_def by (subst eval_fds_add) (auto intro!: intros simp: eval_fds_newman eval_fds_deriv_zeta eval_fds_mult eval_fds_zeta) finally show ?thesis . qed have conv: "fds_converges F s" if "Re s \ 1" for s proof (rule Newman_Ingham_1) have "(\n. \ (real n) - ln (real n)) \ O(\_. 1)" using mertens_bounded by (rule landau_o.big.compose) real_asymp from natfun_bigo_1E[OF this, of 1] obtain c' where c': "c' \ 1" "\n. \\ (real n) - ln (real n)\ \ c'" by auto have "Bseq (fds_nth F)" proof (intro BseqI allI) fix n :: nat show "norm (fds_nth F n) \ (c' + norm c)" unfolding F_def using c' by (auto simp: fds_nth_zeta fds_nth_deriv fds_nth_newman scaleR_conv_of_real in_Reals_norm intro!: order.trans[OF norm_triangle_ineq] add_mono) qed (insert c', auto intro: add_pos_nonneg) thus "fds_nth F \ O(\_. 1)" by (simp add: natfun_bigo_iff_Bseq) next show "f analytic_on {s. Re s \ 1}" by fact next show "eval_fds F s = f s" if "Re s > 1" for s using that by (rule eval_F) qed (insert that, auto simp: F_def intro!: intros) from conv[of 1] have "summable (\n. fds_nth F n / of_nat n)" unfolding fds_converges_def by auto also have "?this \ summable (\n. (\ n - Ln n + c) / n)" by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) (auto simp: F_def fds_nth_newman fds_nth_deriv fds_nth_zeta scaleR_conv_of_real intro!: sum.cong dest: prime_gt_0_nat) finally have "summable (\n. (\ n - Re (Ln (of_nat n)) + Re c) / n)" by (auto dest: summable_Re) also have "?this \ summable (\n. (\ n - ln n + Re c) / n)" by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) (auto intro!: sum.cong) finally show ?thesis using that[of "Re c"] by blast qed text \ Next, we prove a lemma given by Newman stating that if the sum $\sum a_n / n$ exists and $a_n + \ln n$ is nondecreasing, then $a_n$ must tend to 0. Unfortunately, the proof is rather tedious, but so is the paper version by Newman. \ lemma sum_goestozero_lemma: fixes d::real assumes d: "\\i = M..N. a i / i\ < d" and le: "\n. a n + ln n \ a (Suc n) + ln (Suc n)" and "0 < M" "M < N" shows "a M \ d * N / (real N - real M) + (real N - real M) / M \ -a N \ d * N / (real N - real M) + (real N - real M) / M" proof - have "0 \ d" using assms by linarith+ then have "0 \ d * N / (N - M + 1)" by simp then have le_dN: "\0 \ x \ x \ d * N / (N - M + 1)\ \ x \ d * N / (N - M + 1)" for x::real by linarith have le_a_ln: "a m + ln m \ a n + ln n" if "n \ m" for n m by (rule transitive_stepwise_le) (use le that in auto) have *: "x \ b \ y \ b" if "a \ b" "x \ a" "y \ a" for a b x y::real using that by linarith show ?thesis proof (rule *) show "d * N / (N - M) + ln (N / M) \ d * N / (real N - real M) + (real N - real M) / M" using \0 < M\ \M < N\ ln_le_minus_one [of "N / M"] by (simp add: of_nat_diff) (simp add: divide_simps) next have "a M - ln (N / M) \ (d * N) / (N - M + 1)" proof (rule le_dN) assume 0: "0 \ a M - ln (N / M)" have "(Suc N - M) * (a M - ln (N / M)) / N = (\i = M..N. (a M - ln (N / M)) / N)" by simp also have "\ \ (\i = M..N. a i / i)" proof (rule sum_mono) fix i assume i: "i \ {M..N}" with \0 < M\ have "0 < i" by auto have "(a M - ln (N / M)) / N \ (a M - ln (N / M)) / i" using 0 using i \0 < M\ by (simp add: frac_le_eq divide_simps mult_left_mono) also have "a M + ln (real M) \ a i + ln (real N)" by (rule order.trans[OF le_a_ln[of M i]]) (use i assms in auto) hence "(a M - ln (N / M)) / i \ a i / real i" using assms i by (intro divide_right_mono) (auto simp: ln_div field_simps) finally show "(a M - ln (N / M)) / real N \ a i / real i" . qed finally have "((Suc N) - M) * (a M - ln (N / M)) / N \ \\i = M..N. a i / i\" by simp also have "\ \ d" using d by simp finally have "((Suc N) - M) * (a M - ln (N / M)) / N \ d" . then show ?thesis using \M < N\ by (simp add: of_nat_diff field_simps) qed also have "\ \ d * N / (N - M)" using assms(1,4) by (simp add: field_simps) finally show "a M \ d * N / (N - M) + ln (N / M)" by simp next have "- a N - ln (N / M) \ (d * N) / (N - M + 1)" proof (rule le_dN) assume 0: "0 \ - a N - ln (N / M)" have "(\i = M..N. a i / i) \ (\i = M..N. (a N + ln (N / M)) / N)" proof (rule sum_mono) fix i assume i: "i \ {M..N}" with \0 < M\ have "0 < i" by auto have "a i + ln (real M) \ a N + ln (real N)" by (rule order.trans[OF _ le_a_ln[of i N]]) (use i assms in auto) hence "a i / i \ (a N + ln (N / M)) / i" using assms(3,4) by (intro divide_right_mono) (auto simp: field_simps ln_div) also have "\ \ (a N + ln (N / M)) / N" using i \i > 0\ 0 by (intro divide_left_mono_neg) auto finally show "a i / i \ (a N + ln (N / M)) / N" . qed also have "\ = ((Suc N) - M) * (a N + ln (N / M)) / N" by simp finally have "(\i = M..N. a i / i) \ (real (Suc N) - real M) * (a N + ln (N / M)) / N" using \M < N\ by (simp add: of_nat_diff) then have "-((real (Suc N) - real M) * (a N + ln (N / M)) / N) \ \\i = M..N. a i / i\" by linarith also have "\ \ d" using d by simp finally have "- ((real (Suc N) - real M) * (a N + ln (N / M)) / N) \ d" . then show ?thesis using \M < N\ by (simp add: of_nat_diff field_simps) qed also have "\ \ d * N / real (N - M)" using \0 < M\ \M < N\ \0 \ d\ by (simp add: field_simps) finally show "-a N \ d * N / real (N - M) + ln (N / M)" by simp qed qed proposition sum_goestozero_theorem: assumes summ: "summable (\i. a i / i)" and le: "\n. a n + ln n \ a (Suc n) + ln (Suc n)" shows "a \ 0" proof (clarsimp simp: lim_sequentially) fix r::real assume "r > 0" have *: "\n0. \n\n0. \a n\ < \" if \: "0 < \" "\ < 1" for \ proof - have "0 < (\ / 8)\<^sup>2" using \0 < \\ by simp then obtain N0 where N0: "\m n. m \ N0 \ norm (\k=m..n. (\i. a i / i) k) < (\ / 8)\<^sup>2" by (metis summable_partial_sum_bound summ) obtain N1 where "real N1 > 4 / \" using reals_Archimedean2[of "4 / \"] \ by auto hence "N1 \ 0" and N1: "1 / real N1 < \ / 4" using \ by (auto simp: divide_simps mult_ac intro: Nat.gr0I) have "\a n\ < \" if n: "n \ 2 * N0 + N1 + 7" for n proof - define k where "k = \n * \/4\" have "n * \ / 4 > 1" and "n * \ / 4 \ n / 4" and "n / 4 < n" using less_le_trans[OF N1, of "n / N1 * \ / 4"] \N1 \ 0\ \ n by (auto simp: field_simps) hence k: "k > 0" "4 * k \ n" "nat k < n" "(n * \ / 4) - 1 < k" "k \ (n * \ / 4)" unfolding k_def by linarith+ have "-a n < \" proof - have "N0 \ n - nat k" using n k by linarith then have *: "\\k = n - nat k .. n. a k / k\ < (\ / 8)\<^sup>2" using N0 [of "n - nat k" n] by simp have "-a n \ (\ / 8)\<^sup>2 * n / \n * \ / 4\ + \n * \ / 4\ / (n - k)" using sum_goestozero_lemma [OF * le, THEN conjunct2] k by (simp add: of_nat_diff k_def) also have "\< \" proof - have "\ / 16 * n / k < 2" using k by (auto simp: field_simps) then have "\ * (\ / 16 * n / k) < \ * 2" using \ mult_less_cancel_left_pos by blast then have "(\ / 8)\<^sup>2 * n / k < \ / 2" by (simp add: field_simps power2_eq_square) moreover have "k / (n - k) < \ / 2" proof - have "(\ + 2) * k < 4 * k" using k \ by simp also have "\ \ \ * real n" using k by (auto simp: field_simps) finally show ?thesis using k by (auto simp: field_simps) qed ultimately show ?thesis unfolding k_def by linarith qed finally show ?thesis . qed moreover have "a n < \" proof - have "N0 \ n" using n k by linarith then have *: "\\k = n .. n + nat k. a k / k\ < (\/8)\<^sup>2" using N0 [of n "n + nat k"] by simp have "a n \ (\/8)\<^sup>2 * (n + nat k) / k + k / n" using sum_goestozero_lemma [OF * le, THEN conjunct1] k by (simp add: of_nat_diff) also have "\< \" proof - have "4 \ 28 * real_of_int k" using k by linarith then have "\/16 * n / k < 2" using k by (auto simp: field_simps) have "\ * (real n + k) < 32 * k" proof - have "\ * n / 4 < k + 1" by (simp add: mult.commute k_def) then have "\ * n < 4 * k + 4" by (simp add: divide_simps) also have "\ \ 8 * k" using k by auto finally have 1: "\ * real n < 8 * k" . have 2: "\ * k < k" using k \ by simp show ?thesis using k add_strict_mono [OF 1 2] by (simp add: algebra_simps) qed then have "(\ / 8)\<^sup>2 * real (n + nat k) / k < \ / 2" using \ k by (simp add: divide_simps mult_less_0_iff power2_eq_square) moreover have "k / n < \ / 2" using k \ by (auto simp: k_def field_simps) ultimately show ?thesis by linarith qed finally show ?thesis . qed ultimately show ?thesis by force qed then show ?thesis by blast qed show "\n0. \n\n0. \a n\ < r" using * [of "min r (1/5)"] \0 < r\ by force qed text \ This leads us to the main intermediate result: \ lemma Mertens_convergent: "convergent (\n::nat. \ n - ln n)" proof - obtain c where c: "summable (\n. (\ n - ln n + c) / n)" by (blast intro: mertens_summable) then obtain l where l: "(\n. (\ n - ln n + c) / n) sums l" by (auto simp: summable_def) have *: "(\n. \ n - ln n + c) \ 0" by (rule sum_goestozero_theorem[OF c]) auto hence "(\n. \ n - ln n) \ -c" by (simp add: tendsto_iff dist_norm) thus ?thesis by (rule convergentI) qed corollary \_minus_ln_limit: obtains c where "((\x::real. \ x - ln x) \ c) at_top" proof - from Mertens_convergent obtain c where "(\n. \ n - ln n) \ c" by (auto simp: convergent_def) hence 1: "((\x::real. \ (nat \x\) - ln (nat \x\)) \ c) at_top" by (rule filterlim_compose) real_asymp have 2: "((\x::real. ln (nat \x\) - ln x) \ 0) at_top" by real_asymp have 3: "((\x. \ x - ln x) \ c) at_top" using tendsto_add[OF 1 2] by simp with that show ?thesis by blast qed subsection \The asymptotics of the prime-counting functions\ text \ We will now use the above result to prove the asymptotics of the prime-counting functions $\vartheta(x) \sim x$, $\psi(x) \sim x$, and $\pi(x) \sim x / \ln x$. The last of these is typically called the Prime Number Theorem, but since these functions can be expressed in terms of one another quite easily, knowing the asymptotics of any of them immediately gives the asymptotics of the other ones. In this sense, all of the above are equivalent formulations of the Prime Number Theorem. The one we shall tackle first, due to its strong connection to the $\mathfrak{M}$ function, is $\vartheta(x) \sim x$. We know that $\mathfrak{M}(x)$ has the asymptotic expansion $\mathfrak{M}(x) = \ln x + c + o(1)$. We also know that \[\vartheta(x) = x\mathfrak{M}(x) - \int\nolimits_2^x \mathfrak{M}(t) \,\mathrm{d}t\ .\] Substituting in the above asymptotic equation, we obtain: \begin{align*} \vartheta(x) &= x\ln x + cx + o(x) - \int\nolimits_2^x \ln t + c + o(1) \,\mathrm{d}t\\ &= x\ln x + cx + o(x) - (x\ln x - x + cx + o(x))\\ &= x + o(x) \end{align*} In conclusion, $\vartheta(x) \sim x$. \ theorem \_asymptotics: "\ \[at_top] (\x. x)" proof - from \_minus_ln_limit obtain c where c: "((\x. \ x - ln x) \ c) at_top" by auto define r where "r = (\x. \ x - ln x - c)" have \_expand: "\ = (\x. ln x + c + r x)" by (simp add: r_def) have r: "r \ o(\_. 1)" unfolding r_def using tendsto_add[OF c tendsto_const[of "-c"]] by (intro smalloI_tendsto) auto define r' where "r' = (\x. integral {2..x} r)" have integrable_r: "r integrable_on {x..y}" if "2 \ x" for x y :: real using that unfolding r_def by (intro integrable_diff integrable_primes_M) (auto intro!: integrable_continuous_real continuous_intros) hence integral: "(r has_integral r' x) {2..x}" if "x \ 2" for x by (auto simp: has_integral_iff r'_def) have r': "r' \ o(\x. x)" using integrable_r unfolding r'_def by (intro integral_smallo[OF r]) (auto simp: filterlim_ident) define C where "C = 2 * (c + ln 2 - 1)" have "\ \[at_top] (\x. x + (r x * x + C - r' x))" proof (intro asymp_equiv_refl_ev eventually_mono[OF eventually_gt_at_top]) fix x :: real assume x: "x > 2" have "(\ has_integral ((x * ln x - x + c * x) - (2 * ln 2 - 2 + c * 2) + r' x)) {2..x}" unfolding \_expand using x by (intro has_integral_add[OF fundamental_theorem_of_calculus integral]) - (auto simp flip: has_field_derivative_iff_has_vector_derivative + (auto simp flip: has_real_derivative_iff_has_vector_derivative intro!: derivative_eq_intros continuous_intros) from has_integral_unique[OF \_conv_\_integral this] show "\ x = x + (r x * x + C - r' x)" using x by (simp add: field_simps \_expand C_def) qed also have "(\x. r x * x + C - r' x) \ o(\x. x)" proof (intro sum_in_smallo r) show "(\_. C) \ o(\x. x)" by real_asymp qed (insert landau_o.small_big_mult[OF r, of "\x. x"] r', simp_all) hence "(\x. x + (r x * x + C - r' x)) \[at_top] (\x. x)" by (subst asymp_equiv_add_right) auto finally show ?thesis by auto qed text \ The various other forms of the Prime Number Theorem follow as simple corollaries. \ corollary \_asymptotics: "\ \[at_top] (\x. x)" using \_asymptotics PNT4_imp_PNT5 by simp corollary prime_number_theorem: "\ \[at_top] (\x. x / ln x)" using \_asymptotics PNT4_imp_PNT1 by simp corollary ln_\_asymptotics: "(\x. ln (\ x)) \[at_top] ln" using prime_number_theorem PNT1_imp_PNT1' by simp corollary \_ln_\_asymptotics: "(\x. \ x * ln (\ x)) \[at_top] (\x. x)" using prime_number_theorem PNT1_imp_PNT2 by simp corollary nth_prime_asymptotics: "(\n. real (nth_prime n)) \[at_top] (\n. real n * ln (real n))" using \_ln_\_asymptotics PNT2_imp_PNT3 by simp text \ The following versions use a little less notation. \ corollary prime_number_theorem': "((\x. \ x / (x / ln x)) \ 1) at_top" using prime_number_theorem by (rule asymp_equivD_strong[OF _ eventually_mono[OF eventually_gt_at_top[of 1]]]) auto corollary prime_number_theorem'': "(\x. card {p. prime p \ real p \ x}) \[at_top] (\x. x / ln x)" proof - have "\ = (\x. card {p. prime p \ real p \ x})" by (intro ext) (simp add: \_def prime_sum_upto_def) with prime_number_theorem show ?thesis by simp qed corollary prime_number_theorem''': "(\n. card {p. prime p \ p \ n}) \[at_top] (\n. real n / ln (real n))" proof - have "(\n. card {p. prime p \ real p \ real n}) \[at_top] (\n. real n / ln (real n))" using prime_number_theorem'' by (rule asymp_equiv_compose') (simp add: filterlim_real_sequentially) thus ?thesis by simp qed (*<*) unbundle no_prime_counting_notation (*>*) end \ No newline at end of file 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,2042 +1,2042 @@ 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 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: vector_add_divide_simps) (simp add: field_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 assms obtain a b where "cbox a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" by (rule open_contains_cbox) 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 + (insert x m, auto simp flip: has_real_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 of_bool_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 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 + (auto simp flip: has_real_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 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 + (auto simp flip: has_real_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 diff --git a/thys/Safe_Distance/Safe_Distance.thy b/thys/Safe_Distance/Safe_Distance.thy --- a/thys/Safe_Distance/Safe_Distance.thy +++ b/thys/Safe_Distance/Safe_Distance.thy @@ -1,1472 +1,1472 @@ \<^marker>\creator "Albert Rizaldi" "Fabian Immler\ section \Safe Distance\ theory Safe_Distance imports "HOL-Analysis.Multivariate_Analysis" "HOL-Decision_Procs.Approximation" Sturm_Sequences.Sturm begin text \ This theory is about formalising the safe distance rule. The safe distance rule is obtained from Vienna Convention which basically states the following thing. ``The car at all times must maintain a safe distance towards the vehicle in front of it, such that whenever the vehicle in front and the ego vehicle apply maximum deceleration, there will not be a collision.'' To formalise this safe distance rule we have to define first what is a safe distance. To define this safe distance, we have to model the physics of the movement of the vehicle. The following model is sufficient. \s = s\<^sub>0 + v\<^sub>0 * t + 1 / 2 * a\<^sub>0 * t\<^sup>2\ Assumptions in this model are : \<^item> Both vehicles are assumed to be point mass. The exact location of the ego vehicle is the front-most occupancy of the ego vehicle. Similarly for the other vehicle, its exact location is the rearmost occupancy of the other vehicle. \<^item> Both cars can never drive backward. \ lemmas [simp del] = div_mult_self1 div_mult_self2 div_mult_self3 div_mult_self4 subsection \Quadratic Equations\ lemma discriminant: "a * x\<^sup>2 + b * x + c = (0::real) \ 0 \ b\<^sup>2 - 4 * a * c" by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))") lemma quadratic_eq_factoring: assumes D : "D = b\<^sup>2 - 4 * a * c" assumes nn: "0 \ D" assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)" assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)" assumes a : "a \ 0" shows "a * x\<^sup>2 + b * x + c = a * (x - x\<^sub>1) * (x - x\<^sub>2)" using nn by (simp add: D x1 x2) (simp add: assms power2_eq_square power3_eq_cube field_split_simps) lemma quadratic_eq_zeroes_iff: assumes D : "D = b\<^sup>2 - 4 * a * c" assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)" assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)" assumes a : "a \ 0" shows "a * x\<^sup>2 + b * x + c = 0 \ (D \ 0 \ (x = x\<^sub>1 \ x = x\<^sub>2))" (is "?z \ _") using quadratic_eq_factoring[OF D _ x1 x2 a, of x] discriminant[of a x b c] a by (auto simp: D) subsection \Convexity Condition\ lemma p_convex: fixes a b c x y z :: real assumes p_def: "p = (\x. a * x\<^sup>2 + b * x + c)" assumes less : "x < y" "y < z" and ge : "p x > p y" "p y \ p z" shows "a > 0" using less ge unfolding p_def by (sos "((((A<0 * (A<1 * A<2)) * R<1) + (((A<2 * R<1) * (R<1/4 * [y + ~1*z]^2)) + (((A<=1 * R<1) * (R<1 * [x + ~1*y]^2)) + (((A<=1 * (A<0 * (A<1 * R<1))) * (R<1/4 * [1]^2)) + (((A<=0 * R<1) * (R<1/4 * [~1*y^2 + x*y + ~1*x*z + y*z]^2)) + ((A<=0 * (A<0 * (A<1 * R<1))) * (R<1 * [x + ~1/2*y + ~1/2*z]^2))))))))") definition root_in :: "real \ real \ (real \ real) \ bool" where "root_in m M f = (\x\{m .. M}. f x = 0)" definition quadroot_in :: "real \ real \ real \ real \ real \ bool" where "quadroot_in m M a b c = root_in m M (\x. a * x^2 + b * x + c)" lemma card_iff_exists: "0 < card X \ finite X \ (\x. x \ X)" by (auto simp: card_gt_0_iff) lemma quadroot_in_sturm[code]: "quadroot_in m M a b c \ (a = 0 \ b = 0 \ c = 0 \ m \ M) \ (m \ M \ poly [:c, b, a:] m = 0) \ count_roots_between [:c, b, a:] m M > 0" apply (cases "a = 0 \ b = 0 \ c = 0 \ m \ M") apply (force simp: quadroot_in_def root_in_def) apply (cases "m \ M \ poly [:c, b, a:] m = 0") apply (force simp: quadroot_in_def root_in_def algebra_simps power2_eq_square count_roots_between_correct card_iff_exists) proof - assume H: "\ (a = 0 \ b = 0 \ c = 0 \ m \ M)" "\ (m \ M \ poly [:c, b, a:] m = 0)" hence "poly [:c, b, a:] m \ 0 \ m > M" by auto then have "quadroot_in m M a b c \ 0 < count_roots_between [:c, b, a:] m M" proof (rule disjE) assume pnz: "poly [:c, b, a:] m \ 0" then have nz: "[:c, b, a:] \ 0" by auto show ?thesis unfolding count_roots_between_correct card_iff_exists apply safe apply (rule finite_subset[where B="{x. poly [:c, b, a:] x = 0}"]) apply force apply (rule poly_roots_finite) apply (rule nz) using pnz apply (auto simp add: count_roots_between_correct quadroot_in_def root_in_def card_iff_exists algebra_simps power2_eq_square) apply (case_tac "x = m") apply (force simp: algebra_simps) apply force done qed (auto simp: quadroot_in_def count_roots_between_correct root_in_def card_eq_0_iff) then show "quadroot_in m M a b c = (a = 0 \ b = 0 \ c = 0 \ m \ M \ m \ M \ poly [:c, b, a:] m = 0 \ 0 < count_roots_between [:c, b, a:] m M)" using H by metis qed lemma check_quadroot_linear: fixes a b c :: real assumes "a = 0" shows "\ quadroot_in m M a b c \ ((b = 0 \ c = 0 \ M < m) \ (b = 0 \ c \ 0) \ (b \ 0 \ (let x = - c / b in m > x \ x > M)))" proof - have "quadroot_in m M a b c \ (b = 0 \ quadroot_in m M a b c) \ (b \ 0 \ quadroot_in m M a b c)" by auto also have "(b = 0 \ quadroot_in m M a b c) \ ((b = 0 \ c = 0 \ m \ M) \ (b \ 0 \ c = 0))" by (auto simp: quadroot_in_def Let_def root_in_def assms field_split_simps intro!: bexI[where x="-c / b"]) also have "(b \ 0 \ quadroot_in m M a b c) \ (b = 0 \ (let x = -c / b in m \ x \ x \ M))" apply (auto simp: quadroot_in_def Let_def root_in_def assms field_split_simps intro!: bexI[where x="-c / b"]) by (metis mult.commute mult_le_cancel_left_neg add_eq_0_iff mult_le_cancel_iff2)+ finally show ?thesis by (simp add: Let_def not_less not_le) qed lemma check_quadroot_nonlinear: assumes "a \ 0" shows "quadroot_in m M a b c = (let D = b^2 - 4 * a * c in D \ 0 \ ((let x = (-b + sqrt D)/(2*a) in m \ x \ x \ M) \ (let x = (-b - sqrt D)/(2*a) in m \ x \ x \ M)))" by (auto simp: quadroot_in_def Let_def root_in_def quadratic_eq_zeroes_iff[OF refl refl refl assms]) lemma ncheck_quadroot: shows "\quadroot_in m M a b c \ (a = 0 \\quadroot_in m M a b c) \ (a = 0 \ \quadroot_in m M a b c)" by auto subsection \Movement\ locale movement = fixes a v s0 :: real begin text \ Function to compute the distance using the equation \s(t) = s\<^sub>0 + v\<^sub>0 * t + 1 / 2 * a\<^sub>0 * t\<^sup>2\ Input parameters: \<^item> \s\<^sub>0\: initial distance \<^item> \v\<^sub>0\: initial velocity (positive means forward direction and the converse is true) \<^item> \a\: acceleration (positive for increasing and negative for decreasing) \<^item> \t\: time For the time \t < 0\, we assume the output of the function is \s\<^sub>0\. Otherwise, the output is calculated according to the equation above. \ subsubsection \Continuous Dynamics\ definition p :: "real \ real" where "p t = s0 + v * t + 1/2 * a * t\<^sup>2" lemma p_all_zeroes: assumes D: "D = v\<^sup>2 - 2 * a * s0" shows "p t = 0 \ ((a \ 0 \ 0 \ D \ ((t = (- v + sqrt D) / a) \ t = (- v - sqrt D) / a)) \ (a = 0 \ v = 0 \ s0 = 0) \ (a = 0 \ v \ 0 \ t = (- s0 / v)))" using quadratic_eq_zeroes_iff[OF refl refl refl, of "a / 2" t v s0] by (auto simp: movement.p_def D power2_eq_square field_split_simps) lemma p_zero[simp]: "p 0 = s0" by (simp add: p_def) lemma p_continuous[continuous_intros]: "continuous_on T p" unfolding p_def by (auto intro!: continuous_intros) lemma isCont_p[continuous_intros]: "isCont p x" using p_continuous[of UNIV] by (auto simp: continuous_on_eq_continuous_at) definition p' :: "real \ real" where "p' t = v + a * t" lemma p'_zero: "p' 0 = v" by (simp add: p'_def) lemma p_has_vector_derivative[derivative_intros]: "(p has_vector_derivative p' t) (at t within s)" by (auto simp: p_def[abs_def] p'_def has_vector_derivative_def algebra_simps intro!: derivative_eq_intros) lemma p_has_real_derivative[derivative_intros]: "(p has_real_derivative p' t) (at t within s)" using p_has_vector_derivative - by (simp add: has_field_derivative_iff_has_vector_derivative) + by (simp add: has_real_derivative_iff_has_vector_derivative) definition p'' :: "real \ real" where "p'' t = a" lemma p'_has_vector_derivative[derivative_intros]: "(p' has_vector_derivative p'' t) (at t within s)" by (auto simp: p'_def[abs_def] p''_def has_vector_derivative_def algebra_simps intro!: derivative_eq_intros) lemma p'_has_real_derivative[derivative_intros]: "(p' has_real_derivative p'' t) (at t within s)" using p'_has_vector_derivative - by (simp add: has_field_derivative_iff_has_vector_derivative) + by (simp add: has_real_derivative_iff_has_vector_derivative) definition t_stop :: real where "t_stop = - v / a" lemma p'_stop_zero: "p' t_stop = (if a = 0 then v else 0)" by (auto simp: p'_def t_stop_def) lemma p'_pos_iff: "p' x > 0 \ (if a > 0 then x > -v / a else if a < 0 then x < -v / a else v > 0)" by (auto simp: p'_def field_split_simps) lemma le_t_stop_iff: "a \ 0 \ x \ t_stop \ (if a < 0 then p' x \ 0 else p' x \ 0)" by (auto simp: p'_def field_split_simps t_stop_def) lemma p'_continuous[continuous_intros]: "continuous_on T p'" unfolding p'_def by (auto intro: continuous_intros) lemma isCont_p'[continuous_intros]: "isCont p' x" using p'_continuous[of UNIV] by (auto simp: continuous_on_eq_continuous_at) definition p_max :: real where "p_max = p t_stop" lemmas p_t_stop = p_max_def[symmetric] lemma p_max_eq: "p_max = s0 - v\<^sup>2 / a / 2" by (auto simp: p_max_def p_def t_stop_def field_split_simps power2_eq_square) subsubsection \Hybrid Dynamics\ definition s :: "real \ real" where "s t = ( if t \ 0 then s0 else if t \ t_stop then p t else p_max)" definition q :: "real \ real" where "q t = s0 + v * t" definition q' :: "real \ real" where "q' t = v" lemma init_q: "q 0 = s0" unfolding q_def by auto lemma q_continuous[continuous_intros]: "continuous_on T q" unfolding q_def by (auto intro: continuous_intros) lemma isCont_q[continuous_intros]: "isCont q x" using q_continuous[of UNIV] by (auto simp:continuous_on_eq_continuous_at) lemma q_has_vector_derivative[derivative_intros]: "(q has_vector_derivative q' t) (at t within u)" by (auto simp: q_def[abs_def] q'_def has_vector_derivative_def algebra_simps intro!: derivative_eq_intros) lemma q_has_real_derivative[derivative_intros]: "(q has_real_derivative q' t) (at t within u)" using q_has_vector_derivative - by (simp add:has_field_derivative_iff_has_vector_derivative) + by (simp add:has_real_derivative_iff_has_vector_derivative) lemma s_cond_def: "t \ 0 \ s t = s0" "0 \ t \ t \ t_stop \ s t = p t" by (simp_all add: s_def) end locale braking_movement = movement + assumes decel: "a < 0" assumes nonneg_vel: "v \ 0" begin lemma t_stop_nonneg: "0 \ t_stop" using decel nonneg_vel by (auto simp: t_stop_def divide_simps) lemma t_stop_pos: assumes "v \ 0" shows "0 < t_stop" using decel nonneg_vel assms by (auto simp: t_stop_def divide_simps) lemma t_stop_zero: assumes "t_stop = 0" shows "v = 0" using assms decel by (auto simp: t_stop_def) lemma t_stop_zero_not_moving: "t_stop = 0 \ q t = s0" unfolding q_def using t_stop_zero by auto abbreviation "s_stop \ s t_stop" lemma s_t_stop: "s_stop = p_max" using t_stop_nonneg by (auto simp: s_def t_stop_def p_max_def p_def) lemma s0_le_s_stop: "s0 \ s_stop" proof (rule subst[where t="s_stop" and s="p_max"]) show "p_max = s_stop" by (rule sym[OF s_t_stop]) next show "s0 \ p_max" proof (rule subst[where t="p_max" and s="s0 - v\<^sup>2 / a / 2"]) show " s0 - v\<^sup>2 / a / 2 = p_max" using p_max_eq by auto next have "0 \ - v\<^sup>2 / a / 2" using decel zero_le_square[of v] proof - have f1: "a \ 0" using \a < 0\ by linarith have "(- 1 * v\<^sup>2 \ 0) = (0 \ v\<^sup>2)" by auto then have "0 \ - 1 * v\<^sup>2 / a" using f1 by (meson zero_le_divide_iff zero_le_power2) then show ?thesis by force qed thus "s0 \ s0 - v\<^sup>2 / a / 2" by auto qed qed lemma p_mono: "x \ y \ y \ t_stop \ p x \ p y" using decel proof - assume "x \ y" and "y \ t_stop" and "a < 0" hence "x + y \ - 2 * v / a" unfolding t_stop_def by auto hence "-1 / 2 * a * (x + y) \ v" (is "?lhs0 \ ?rhs0") using \a < 0\ by (auto simp add: field_simps) hence "?lhs0 * (x- y) \ ?rhs0 * (x - y)" using \x \ y\ by sos hence "v * x + 1 / 2 * a * x\<^sup>2 \ v * y + 1 / 2 * a * y\<^sup>2" by (auto simp add: field_simps power_def) thus " p x \ p y" unfolding p_max_def p_def t_stop_def by auto qed lemma p_antimono: "x \ y \ t_stop \ x \ p y \ p x" using decel proof - assume "x \ y" and "t_stop \ x" and "a < 0" hence "- 2 * v / a \ x + y" unfolding t_stop_def by auto hence "v \ - 1/ 2 * a * (x + y)" using \a < 0\ by (auto simp add: field_simps) hence "v * (x - y) \ -1/2 * a * (x\<^sup>2 - y\<^sup>2) " using \x \ y\ by sos hence "v * y + 1/2 * a * y\<^sup>2 \ v * x + 1/2 * a * x\<^sup>2" by (auto simp add: field_simps) thus "p y \ p x" unfolding p_max_def p_def t_stop_def by auto qed lemma p_strict_mono: "x < y \ y \ t_stop \ p x < p y" using decel proof - assume "x < y" and "y \ t_stop" and "a < 0" hence "x + y < - 2 * v / a" unfolding t_stop_def by auto hence "-1 / 2 * a * (x + y) < v" (is "?lhs0 < ?rhs0") using \a < 0\ by (auto simp add: field_simps) hence "?lhs0 * (x- y) > ?rhs0 * (x - y)" using \x < y\ by sos hence "v * x + 1 / 2 * a * x\<^sup>2 < v * y + 1 / 2 * a * y\<^sup>2" by (auto simp add: field_simps power_def) thus " p x < p y" unfolding p_max_def p_def t_stop_def by auto qed lemma p_strict_antimono: "x < y \ t_stop \ x\ p y < p x" using decel proof - assume "x < y" and "t_stop \ x" and "a < 0" hence "- 2 * v / a < x + y" unfolding t_stop_def by auto hence "v < - 1/ 2 * a * (x + y)" using \a < 0\ by (auto simp add: field_simps) hence "v * (x - y) > -1/2 * a * (x\<^sup>2 - y\<^sup>2) " using \x < y\ by sos hence "v * y + 1/2 * a * y\<^sup>2 < v * x + 1/2 * a * x\<^sup>2" by (auto simp add: field_simps) thus "p y < p x" unfolding p_max_def p_def t_stop_def by auto qed lemma p_max: "p x \ p_max" unfolding p_max_def by (cases "x \ t_stop") (auto intro: p_mono p_antimono) lemma continuous_on_s[continuous_intros]: "continuous_on T s" unfolding s_def[abs_def] using t_stop_nonneg by (intro continuous_on_subset[where t=T and s = "{.. 0}\({0 .. t_stop} \ {t_stop ..})"] continuous_on_If) (auto simp: p_continuous p_max_def antisym_conv[where x=0]) lemma isCont_s[continuous_intros]: "isCont s x" using continuous_on_s[of UNIV] by (auto simp: continuous_on_eq_continuous_at) definition s' :: "real \ real" where "s' t = (if t \ t_stop then p' t else 0)" lemma s_has_real_derivative: assumes "t \ 0" "v / a \ 0" "a \ 0" shows "(s has_real_derivative s' t) (at t within {0..})" proof - from assms have *: "t \ t_stop \ t \ {0 .. t_stop}" by simp from assms have "0 \ t_stop" by (auto simp: t_stop_def) have "((\t. if t \ {0 .. t_stop} then p t else p_max) has_real_derivative (if t \ {0..t_stop} then p' t else 0)) (at t within {0..})" unfolding s_def[abs_def] s'_def - has_field_derivative_iff_has_vector_derivative + has_real_derivative_iff_has_vector_derivative apply (rule has_vector_derivative_If_within_closures[where T = "{t_stop ..}"]) using \0 \ t_stop\ \a \ 0\ by (auto simp: assms p'_stop_zero p_t_stop max_def insert_absorb intro!: p_has_vector_derivative) from _ _ this show ?thesis - unfolding has_vector_derivative_def has_field_derivative_iff_has_vector_derivative + unfolding has_vector_derivative_def has_real_derivative_iff_has_vector_derivative s'_def s_def[abs_def] * by (rule has_derivative_transform) (auto simp: assms s_def p_max_def t_stop_def) qed lemma s_has_vector_derivative[derivative_intros]: assumes "t \ 0" "v / a \ 0" "a \ 0" shows "(s has_vector_derivative s' t) (at t within {0..})" using s_has_real_derivative[OF assms] - by (simp add: has_field_derivative_iff_has_vector_derivative) + by (simp add: has_real_derivative_iff_has_vector_derivative) lemma s_has_field_derivative[derivative_intros]: assumes "t \ 0" "v / a \ 0" "a \ 0" shows "(s has_field_derivative s' t) (at t within {0..})" using s_has_vector_derivative[OF assms] - by(simp add: has_field_derivative_iff_has_vector_derivative) + by(simp add: has_real_derivative_iff_has_vector_derivative) lemma s_has_real_derivative_at: assumes "0 < x" "0 \ v" "a < 0" shows "(s has_real_derivative s' x) (at x)" proof - from assms have "(s has_real_derivative s' x) (at x within {0 ..})" by (intro s_has_real_derivative) (auto intro!: divide_nonneg_nonpos) then have "(s has_real_derivative s' x) (at x within {0<..})" by (rule DERIV_subset) auto then show "(s has_real_derivative s' x) (at x)" using assms by (subst (asm) at_within_open) auto qed lemma s_delayed_has_field_derivative[derivative_intros]: assumes "\ < t" "0 \ v" "a < 0" shows "((\x. s (x - \)) has_field_derivative s' (t - \)) (at t within {\<..})" proof - from assms have "((\x. s (x + - \)) has_real_derivative s' (t - \)) (at t)" using DERIV_shift[of "s" "(s' (t - \))" t "-\"] s_has_real_derivative_at by auto thus "((\x. s (x - \)) has_field_derivative s' (t - \)) (at t within {\<..})" using has_field_derivative_at_within by auto qed lemma s_delayed_has_vector_derivative[derivative_intros]: assumes "\ < t" "0 \ v" "a < 0" shows "((\x. s (x - \)) has_vector_derivative s' (t - \)) (at t within {\<..})" using s_delayed_has_field_derivative[OF assms] - by(simp add:has_field_derivative_iff_has_vector_derivative) + by(simp add:has_real_derivative_iff_has_vector_derivative) lemma s'_nonneg: "0 \ v \ a \ 0 \ 0 \ s' x" by (auto simp: s'_def p'_def t_stop_def field_split_simps) lemma s'_pos: "0 \ x \ x < t_stop \ 0 \ v \ a \ 0 \ 0 < s' x" by (intro le_neq_trans s'_nonneg) (auto simp: s'_def p'_def t_stop_def field_split_simps) subsubsection \Monotonicity of Movement\ lemma s_mono: assumes "t \ u" "u \ 0" shows "s t \ s u" using p_mono[of u t] assms p_max[of u] by (auto simp: s_def) lemma s_strict_mono: assumes "u < t" "t \ t_stop" "u \ 0" shows "s u < s t" using p_strict_mono[of u t] assms p_max[of u] by (auto simp: s_def) lemma s_antimono: assumes "x \ y" assumes "t_stop \ x" shows "s y \ s x" proof - from assms have "t_stop \ y" by auto hence "s y \ p_max" unfolding s_def p_max_eq using p_max_def p_max_eq s0_le_s_stop s_t_stop by auto also have "... \ s x" using \t_stop \ x\ s_mono s_t_stop t_stop_nonneg by fastforce ultimately show "s y \ s x" by auto qed lemma init_s : "t \ 0 \ s t = s0" using decel nonneg_vel by (simp add: divide_simps s_def) lemma q_min: "0 \ t \ s0 \ q t" unfolding q_def using nonneg_vel by auto lemma q_mono: "x \ y \ q x \ q y" unfolding q_def using nonneg_vel by (auto simp: mult_left_mono) subsubsection \Maximum at Stopping Time\ lemma s_max: "s x \ s_stop" using p_max[of x] p_max[of 0] unfolding s_t_stop by (auto simp: s_def) lemma s_eq_s_stop: "NO_MATCH t_stop x \ x \ t_stop \ s x = s_stop" using t_stop_nonneg by (auto simp: s_def p_max_def) end subsection \Safe Distance\ locale safe_distance = fixes a\<^sub>e v\<^sub>e s\<^sub>e :: real fixes a\<^sub>o v\<^sub>o s\<^sub>o :: real assumes nonneg_vel_ego : "0 \ v\<^sub>e" assumes nonneg_vel_other : "0 \ v\<^sub>o" assumes decelerate_ego : "a\<^sub>e < 0" assumes decelerate_other : "a\<^sub>o < 0" assumes in_front : "s\<^sub>e < s\<^sub>o" begin lemmas hyps = nonneg_vel_ego nonneg_vel_other decelerate_ego decelerate_other in_front sublocale ego: braking_movement a\<^sub>e v\<^sub>e s\<^sub>e by (unfold_locales; rule hyps) sublocale other: braking_movement a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; rule hyps) sublocale ego_other: movement "a\<^sub>o - a\<^sub>e" "v\<^sub>o - v\<^sub>e" "s\<^sub>o - s\<^sub>e" by unfold_locales subsubsection \Collision\ definition collision :: "real set \ bool" where "collision time_set \ (\t \ time_set. ego.s t = other.s t )" abbreviation no_collision :: "real set \ bool" where "no_collision time_set \ \ collision time_set" lemma no_collision_initially : "no_collision {.. 0}" using decelerate_ego nonneg_vel_ego using decelerate_other nonneg_vel_other in_front by (auto simp: divide_simps collision_def ego.s_def other.s_def) lemma no_collisionI: "(\t. t \ S \ ego.s t \ other.s t) \ no_collision S" unfolding collision_def by blast theorem cond_1: "ego.s_stop < s\<^sub>o \ no_collision {0..}" proof (rule no_collisionI, simp) fix t::real assume "t \ 0" have "ego.s t \ ego.s_stop" by (rule ego.s_max) also assume "\ < s\<^sub>o" also have "\ = other.s 0" by (simp add: other.init_s) also have "\ \ other.s t" using \0 \ t\ hyps by (intro other.s_mono) auto finally show "ego.s t \ other.s t" by simp qed lemma ego_other_strict_ivt: assumes "ego.s t > other.s t" shows "collision {0 ..< t}" proof cases have [simp]: "s\<^sub>e < s\<^sub>o \ ego.s 0 \ other.s 0" by (simp add: movement.s_def) assume "0 \ t" with assms in_front have "\x\0. x \ t \ other.s x - ego.s x = 0" by (intro IVT2, auto simp: continuous_diff other.isCont_s ego.isCont_s) then show ?thesis using assms by (auto simp add: algebra_simps collision_def Bex_def order.order_iff_strict) qed (insert assms hyps, auto simp: collision_def ego.init_s other.init_s intro!: bexI[where x=0]) lemma collision_subset: "collision s \ s \ t \ collision t" by (auto simp: collision_def) lemma ego_other_ivt: assumes "ego.s t \ other.s t" shows "collision {0 .. t}" proof cases assume "ego.s t > other.s t" from ego_other_strict_ivt[OF this] show ?thesis by (rule collision_subset) auto qed (insert hyps assms; cases "t \ 0"; force simp: collision_def ego.init_s other.init_s) theorem cond_2: assumes "ego.s_stop \ other.s_stop" shows "collision {0 ..}" using assms apply (intro collision_subset[where t="{0 ..}" and s = "{0 .. max ego.t_stop other.t_stop}"]) apply (intro ego_other_ivt[where t = "max ego.t_stop other.t_stop"]) apply (auto simp: ego.s_eq_s_stop other.s_eq_s_stop) done abbreviation D2 :: "real" where "D2 \ (v\<^sub>o - v\<^sub>e)^2 - 2 * (a\<^sub>o - a\<^sub>e) * (s\<^sub>o - s\<^sub>e)" abbreviation t\<^sub>D' :: "real" where "t\<^sub>D' \ sqrt (2 * (ego.s_stop - other.s_stop) / a\<^sub>o)" lemma pos_via_half_dist: "dist a b < b / 2 \ b > 0 \ a > 0" by (auto simp: dist_real_def abs_real_def split: if_splits) lemma collision_within_p: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "collision {0..} \ (\t\0. ego.p t = other.p t \ t < ego.t_stop \ t < other.t_stop)" proof (auto simp: collision_def, goal_cases) case (2 t) then show ?case by (intro bexI[where x = t]) (auto simp: ego.s_def other.s_def) next case (1 t) then show ?case using assms hyps ego.t_stop_nonneg other.t_stop_nonneg apply (auto simp: ego.s_def other.s_def ego.s_t_stop other.s_t_stop ego.p_t_stop other.p_t_stop not_le split: if_splits) defer proof goal_cases case 1 from 1 have le: "ego.t_stop \ other.t_stop" by auto from 1 have "ego.t_stop < t" by simp from other.s_strict_mono[OF this] 1 have "other.s ego.t_stop < other.s t" by auto also have "\ = ego.s ego.t_stop" using ego.s_t_stop ego.t_stop_nonneg 1 other.s_def by auto finally have "other.s ego.t_stop < ego.s ego.t_stop" . from ego_other_strict_ivt[OF this] le in_front show ?case by (auto simp add: collision_def) (auto simp: movement.s_def split: if_splits) next case 2 from 2 have "other.p_max = ego.p t" by simp also have "\ \ ego.p ego.t_stop" using 2 by (intro ego.p_mono) auto also have "\ = ego.p_max" by (simp add: ego.p_t_stop) also note \\ < other.p_max\ finally show ?case by arith next case 3 thus "\t\0. ego.p t = other.p t \ t < ego.t_stop \ t < other.t_stop" apply (cases "t = other.t_stop") apply (simp add: other.p_t_stop ) apply (metis (no_types) ego.p_max not_le) apply (cases "t = ego.t_stop") apply (simp add: ego.p_t_stop) defer apply force proof goal_cases case (1) let ?d = "\t. other.p' t - ego.p' t" define d' where "d' = ?d ego.t_stop / 2" have d_cont: "isCont ?d ego.t_stop" unfolding ego.t_stop_def other.p'_def ego.p'_def by simp have "?d ego.t_stop > 0" using 1 by (simp add: ego.p'_stop_zero other.p'_pos_iff) (simp add: ego.t_stop_def other.t_stop_def) then have "d' > 0" by (auto simp: d'_def) from d_cont[unfolded continuous_at_eps_delta, THEN spec, rule_format, OF \d' > 0\] obtain e where e: "e > 0" "\x. dist x ego.t_stop < e \ ?d x > 0" unfolding d'_def using \?d ego.t_stop > 0\ pos_via_half_dist by force define t' where "t' = ego.t_stop - min (ego.t_stop / 2) (e / 2)" have "0 < ego.t_stop" using 1 by auto have "other.p t' - ego.p t' < other.p ego.t_stop - ego.p ego.t_stop" apply (rule DERIV_pos_imp_increasing[of t']) apply (force simp: t'_def e min_def \0 < ego.t_stop\) apply (auto intro!: exI[where x = "?d x" for x] intro!: derivative_intros e) using \e > 0\ apply (auto simp: t'_def dist_real_def algebra_simps) done also have "\ = 0" using 1 by (simp add: ego.p_t_stop) finally have less: "other.p t' < ego.p t'" by simp have "t' > 0" using 1 by (auto simp: t'_def algebra_simps min_def) have "t' < ego.t_stop" by (auto simp: t'_def \e > 0\ \ego.t_stop > 0\) from less_le_trans[OF \t' < ego.t_stop\ \ego.t_stop \ other.t_stop\] have "t' < other.t_stop" . from ego_other_strict_ivt[of t'] less have "collision {0..t' > 0\ \t' < ego.t_stop\ \t' < other.t_stop\ by (auto simp: other.s_def ego.s_def split: if_splits) thus ?case using \t' > 0\ \t' < ego.t_stop\ \t' < other.t_stop\ apply (auto simp: collision_def ego.s_def other.s_def movement.p_def split: if_splits) apply (rule_tac x = t in exI) apply (auto simp: movement.p_def)[] done qed qed qed lemma collision_within_eq: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "collision {0..} \ collision {0 ..< min ego.t_stop other.t_stop}" unfolding collision_within_p[OF assms] unfolding collision_def by (safe; force simp: ego.s_def other.s_def movement.p_def ego.t_stop_def other.t_stop_def split: if_splits) lemma collision_excluded: "(\t. t \ T \ ego.s t \ other.s t) \ collision S \ collision (S - T)" by (auto simp: collision_def) lemma collision_within_less: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "collision {0..} \ collision {0 <..< min ego.t_stop other.t_stop}" proof - note collision_within_eq[OF assms] also have "collision {0 ..< min ego.t_stop other.t_stop} \ collision ({0 ..< min ego.t_stop other.t_stop} - {0})" using hyps assms by (intro collision_excluded) (auto simp: ego.s_def other.s_def) also have "{0 ..< min ego.t_stop other.t_stop} - {0} = {0 <..< min ego.t_stop other.t_stop}" by auto finally show ?thesis unfolding collision_def by (safe; force simp: ego.s_def other.s_def movement.p_def ego.t_stop_def other.t_stop_def split: if_splits) qed theorem cond_3: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "collision {0..} \ (a\<^sub>o > a\<^sub>e \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)" proof - have "v\<^sub>o \ 0" using assms(1) assms(2) movement.s_def movement.t_stop_def by auto with hyps have "v\<^sub>o > 0" by auto note hyps = hyps this define t1 where "t1 = (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e)" define t2 where "t2 = (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e)" define bounded where "bounded \ \t. (0 \ t \ t \ ego.t_stop \ t \ other.t_stop)" have ego_other_conv: "\t. bounded t \ ego.p t = other.p t \ ego_other.p t = 0" by (auto simp: movement.p_def field_split_simps) let ?r = "{0 <..< min ego.t_stop other.t_stop}" have D2: "D2 = (v\<^sub>o - v\<^sub>e)\<^sup>2 - 4 * ((a\<^sub>o - a\<^sub>e) / 2) * (s\<^sub>o - s\<^sub>e)" by simp define D where "D = D2" note D = D_def[symmetric] define x1 where "x1 \ (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (2 * ((a\<^sub>o - a\<^sub>e) / 2))" define x2 where "x2 \ (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (2 * ((a\<^sub>o - a\<^sub>e) / 2))" have x2: "x2 =(- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e)" by (simp add: x2_def field_split_simps) have x1: "x1 =(- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e)" by (simp add: x1_def field_split_simps) from collision_within_less[OF assms] have coll_eq: "collision {0..} = collision ?r" by (auto simp add: bounded_def) also have "\ \ (a\<^sub>o > a\<^sub>e \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)" proof safe assume H: "a\<^sub>e < a\<^sub>o" "v\<^sub>o < v\<^sub>e" "0 \ D2" assume sqrt: "sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o" have nz: "(a\<^sub>o - a\<^sub>e) / 2 \ 0" using \a\<^sub>e < a\<^sub>o\ by simp note sol = quadratic_eq_zeroes_iff[OF D2 x1_def[THEN meta_eq_to_obj_eq] x2_def[THEN meta_eq_to_obj_eq] nz] from sol[of x2] \0 \ D2\ have "other.p x2 = ego.p x2" by (auto simp: ego.p_def other.p_def field_split_simps) moreover have "x2 > 0" proof (rule ccontr) assume "\ 0 < x2" then have "ego_other.p x2 \ ego_other.p 0" using H hyps by (intro DERIV_nonpos_imp_nonincreasing[of x2]) (auto intro!: exI[where x="ego_other.p' x" for x] derivative_eq_intros simp: ego_other.p'_def add_nonpos_nonpos mult_nonneg_nonpos) also have "ego_other.p 0 > 0" using hyps by (simp add: ego_other.p_def) finally (xtrans) show False using \other.p x2 = ego.p x2\ by (simp add: movement.p_def field_split_simps power2_eq_square) qed moreover have "x2 < other.t_stop" using sqrt H hyps by (auto simp: x2 other.t_stop_def field_split_simps power2_eq_square) ultimately show "collision {0<.. \ ego.p ego.t_stop" by (auto intro!: ego.p_antimono) also have "\ = ego.s x2" using 2 by (auto simp: ego.s_def ego.p_t_stop) finally have "other.s x2 \ ego.s x2" . from ego_other_ivt[OF this] show ?thesis unfolding coll_eq[symmetric] by (rule collision_subset) auto qed (auto simp: collision_def ego.s_def other.s_def not_le intro!: bexI[where x=x2]) next let ?max = "max ego.t_stop other.t_stop" let ?min = "min ego.t_stop other.t_stop" assume "collision ?r" then obtain t where t: "ego.p t = other.p t" "0 < t" "t < ?min" by (auto simp: collision_def ego.s_def other.s_def) then have "t < - (v\<^sub>e / a\<^sub>e)" "t < - (v\<^sub>o / a\<^sub>o)" "t < other.t_stop" by (simp_all add: ego.t_stop_def other.t_stop_def) from t have "ego_other.p t = 0" by (auto simp: movement.p_def field_split_simps) from t have "t < ?max" by auto from hyps assms have "0 < ego_other.p 0" by simp from ego_other.p_def[abs_def, THEN meta_eq_to_obj_eq] have eop_eq: "ego_other.p = (\t. 1 / 2 * (a\<^sub>o - a\<^sub>e) * t\<^sup>2 + (v\<^sub>o - v\<^sub>e) * t + (s\<^sub>o - s\<^sub>e))" by (simp add: algebra_simps) show "a\<^sub>o > a\<^sub>e" proof - have "ego.p other.t_stop \ ego.p_max" by (rule ego.p_max) also have "... \ other.p other.t_stop" using hyps assms by (auto simp:other.s_def ego.s_def ego.p_t_stop split:if_splits) finally have "0 \ ego_other.p other.t_stop" by (auto simp add:movement.p_def field_simps) from p_convex[OF eop_eq, of 0 t other.t_stop, simplified \ego_other.p t = 0\, OF \0 < t\ \t < other.t_stop\ \0 < ego_other.p 0\ \0 \ ego_other.p other.t_stop\] show "a\<^sub>o > a\<^sub>e" by (simp add: algebra_simps) qed have rewr: "4 * ((a\<^sub>o - a\<^sub>e) / 2) = 2 * (a\<^sub>o - a\<^sub>e)" by simp from \a\<^sub>o > a\<^sub>e\ \ego_other.p t = 0\ ego_other.p_all_zeroes[OF D2[symmetric], of t] have "0 \ D2" and disj: "(t = (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e) \ t = (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e))" using hyps assms unfolding rewr by simp_all show "0 \ D2" by fact from add_strict_mono[OF \t < - (v\<^sub>e / a\<^sub>e)\ \t < - (v\<^sub>o / a\<^sub>o)\] \0 < t\ \a\<^sub>o > a\<^sub>e\ have "0 < - (v\<^sub>e / a\<^sub>e) + - (v\<^sub>o / a\<^sub>o)" by (simp add: divide_simps) then have "0 > v\<^sub>e * a\<^sub>o + a\<^sub>e * v\<^sub>o" using hyps by (simp add: field_split_simps split: if_splits) show "v\<^sub>o < v\<^sub>e" using \a\<^sub>e < a\<^sub>o\ \movement.p (a\<^sub>o - a\<^sub>e) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e) t = 0\ in_front t(2) apply (auto simp: movement.p_def divide_less_0_iff algebra_simps power2_eq_square) by (smt divide_less_0_iff mult_le_cancel_right mult_mono mult_nonneg_nonneg nonneg_vel_ego) from disj have "x2 < ?min" proof rule assume "t = (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e)" then show ?thesis using \t < ?min\ by (simp add: x2) next assume "t = (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e)" also have "\ \ x2" unfolding x2 apply (rule divide_right_mono) apply (subst (2) diff_conv_add_uminus) apply (rule add_left_mono) using \a\<^sub>o > a\<^sub>e\ \D2 \ 0\ by auto also (xtrans) note \t < ?min\ finally show ?thesis . qed then show "sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o" using hyps \a\<^sub>o > a\<^sub>e\ by (auto simp: x2 field_split_simps other.t_stop_def) qed finally show ?thesis . qed subsubsection \Formalising Safe Distance\ text \First definition for Safe Distance based on \cond_1\.\ definition absolute_safe_distance :: real where "absolute_safe_distance = - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)" lemma absolute_safe_distance: assumes "s\<^sub>o - s\<^sub>e > absolute_safe_distance" shows "no_collision {0..}" proof - from assms hyps absolute_safe_distance_def have "ego.s_stop < s\<^sub>o" by (auto simp add:ego.s_def ego.p_def ego.t_stop_def power_def) thus ?thesis by (rule cond_1) qed text \First Fallback for Safe Distance.\ definition fst_safe_distance :: real where "fst_safe_distance = v\<^sub>o\<^sup>2 / (2 * a\<^sub>o) - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)" definition distance_leq_d2 :: real where "distance_leq_d2 = (a\<^sub>e + a\<^sub>o) / (2 * a\<^sub>o\<^sup>2) * v\<^sub>o\<^sup>2 - v\<^sub>o * v\<^sub>e / a\<^sub>o" lemma snd_leq_fst_exp: "distance_leq_d2 \ fst_safe_distance" proof - have "0 \ (other.t_stop - ego.t_stop)\<^sup>2" by auto hence "- ego.t_stop\<^sup>2 \ other.t_stop\<^sup>2 - 2 * other.t_stop * ego.t_stop" by (simp add:power_def algebra_simps) with hyps(3) have "- ego.t_stop\<^sup>2 * (a\<^sub>e / 2) \ (other.t_stop\<^sup>2 - 2 * other.t_stop * ego.t_stop) * (a\<^sub>e / 2)" by (smt half_gt_zero_iff mult_le_cancel_right) with ego.t_stop_def other.t_stop_def hyps have "- v\<^sub>e\<^sup>2 / (2 * a\<^sub>e) \ a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o" by (simp add:power_def algebra_simps) with fst_safe_distance_def distance_leq_d2_def have 1: "fst_safe_distance \ a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o + v\<^sub>o\<^sup>2 / (2 * a\<^sub>o)" by (auto simp add:algebra_simps) have "a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o + v\<^sub>o\<^sup>2 / (2 * a\<^sub>o) = distance_leq_d2" (is "?LHS = _") proof - have "?LHS = a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o + a\<^sub>o * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2)" by (auto simp add: algebra_simps power_def) also have "... = distance_leq_d2" by (auto simp add: power_def field_split_simps distance_leq_d2_def) finally show ?thesis by auto qed with 1 show ?thesis by auto qed lemma sqrt_D2_leq_stop_time_diff: assumes "a\<^sub>e < a\<^sub>o" assumes "0 \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o " assumes "s\<^sub>o - s\<^sub>e \ distance_leq_d2" shows "sqrt D2 \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o" proof - from assms have "- 2 * (a\<^sub>o - a\<^sub>e) * (s\<^sub>o - s\<^sub>e) \ - 2 * (a\<^sub>o - a\<^sub>e) * distance_leq_d2" (is "?L \ ?R") by simp hence "D2 \ (v\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * (a\<^sub>o - a\<^sub>e) * distance_leq_d2" by (simp add: algebra_simps) also have "... = (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)\<^sup>2" proof - from distance_leq_d2_def have 1: "(v\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * (a\<^sub>o - a\<^sub>e) * distance_leq_d2 = (v\<^sub>o - v\<^sub>e)\<^sup>2 - (a\<^sub>o - a\<^sub>e) * (a\<^sub>e + a\<^sub>o) / a\<^sub>o\<^sup>2 * v\<^sub>o\<^sup>2 + 2 * (a\<^sub>o - a\<^sub>e) * v\<^sub>o * v\<^sub>e / a\<^sub>o" by (auto simp add: field_split_simps) with hyps(4) have "... = (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)\<^sup>2" by (auto simp add: power_def field_split_simps) with 1 show ?thesis by auto qed finally show ?thesis by (smt assms(2) real_le_lsqrt real_sqrt_le_0_iff) qed lemma cond2_imp_pos_vo: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "v\<^sub>o \ 0" proof (rule ccontr) assume "\ v\<^sub>o \ 0" with other.s_def other.t_stop_def have "other.s_stop = s\<^sub>o" by auto with assms(2) have "ego.s_stop < s\<^sub>o" by auto with assms(1) show "False" by auto qed lemma cond2_imp_gt_fst_sd: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "fst_safe_distance < s\<^sub>o - s\<^sub>e" proof (cases "v\<^sub>e \ 0") case True from fst_safe_distance_def assms ego.s_def ego.t_stop_pos[OF \v\<^sub>e \ 0\] ego.p_def ego.t_stop_def other.s_def other.t_stop_pos[OF cond2_imp_pos_vo[OF assms]] other.p_def other.t_stop_def hyps show ?thesis by (simp add: power_def algebra_simps) next case False with fst_safe_distance_def have "fst_safe_distance = v\<^sub>o\<^sup>2 / (2 * a\<^sub>o)" by auto also have "... \ 0" by (simp add: divide_nonneg_neg hyps) also have "... < s\<^sub>o - s\<^sub>e" by (simp add: algebra_simps hyps) finally show ?thesis by auto qed text \Second Fallback for Safe Distance.\ definition snd_safe_distance :: real where "snd_safe_distance = (v\<^sub>o - v\<^sub>e)\<^sup>2 / (2 * (a\<^sub>o - a\<^sub>e))" lemma fst_leq_snd_safe_distance: assumes "a\<^sub>e < a\<^sub>o" shows"fst_safe_distance \ snd_safe_distance" proof - have "0 \ (v\<^sub>o / a\<^sub>o - v\<^sub>e / a\<^sub>e)\<^sup>2" by auto hence 1: "0 \ (v\<^sub>o / a\<^sub>o)\<^sup>2 - 2 * v\<^sub>o * v\<^sub>e / (a\<^sub>o * a\<^sub>e) + (v\<^sub>e / a\<^sub>e)\<^sup>2" by (auto simp add: power_def algebra_simps) from hyps have "0 \ a\<^sub>o * a\<^sub>e" by (simp add: mult_nonpos_nonpos) from mult_right_mono[OF 1 this] hyps have "0 \ v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o - 2 * v\<^sub>o * v\<^sub>e + v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e" by (auto simp add: power_def algebra_simps) with hyps have 2: "(v\<^sub>o\<^sup>2 / (2 * a\<^sub>o) - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)) * (2 * (a\<^sub>o - a\<^sub>e)) \ (v\<^sub>o - v\<^sub>e)\<^sup>2" by (auto simp add: power_def field_split_simps) from assms have "0 \ 2 * (a\<^sub>o - a\<^sub>e)" by auto from divide_right_mono[OF 2 this] assms fst_safe_distance_def snd_safe_distance_def show ?thesis by auto qed lemma snd_safe_distance_iff_nonneg_D2: assumes "a\<^sub>e < a\<^sub>o" shows "s\<^sub>o - s\<^sub>e \ snd_safe_distance \ 0 \ D2" proof - from snd_safe_distance_def assms pos_le_divide_eq[of "2 * (a\<^sub>o - a\<^sub>e)"] have "s\<^sub>o - s\<^sub>e \ snd_safe_distance \ (s\<^sub>o - s\<^sub>e) * (2 * (a\<^sub>o - a\<^sub>e)) \ (v\<^sub>o - v\<^sub>e)\<^sup>2" by auto also have "... \ 0 \ D2" by (auto simp add: algebra_simps) finally show ?thesis by auto qed lemma t_stop_diff_neg_means_leq_D2: assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" "a\<^sub>e < a\<^sub>o" "0 \ D2" shows "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0 \ sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o" proof assume only_if: "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0" from assms have "... \ sqrt D2" by auto with only_if show "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2" by linarith next assume if_part: "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2" from cond2_imp_gt_fst_sd[OF assms(1) assms(2)] snd_leq_fst_exp have "distance_leq_d2 \ s\<^sub>o - s\<^sub>e" by auto from if_part and sqrt_D2_leq_stop_time_diff [OF \a\<^sub>e < a\<^sub>o\ _ \distance_leq_d2 \ s\<^sub>o - s\<^sub>e\] show " v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0" by linarith qed theorem cond_3': assumes "s\<^sub>o \ ego.s_stop" "ego.s_stop < other.s_stop" shows "collision {0..} \ (a\<^sub>o > a\<^sub>e \ v\<^sub>o < v\<^sub>e \ s\<^sub>o - s\<^sub>e \ snd_safe_distance \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0)" proof (cases "a\<^sub>o \ a\<^sub>e \ v\<^sub>o \ v\<^sub>e") case True with cond_3[OF assms] show ?thesis by auto next case False from \\ (a\<^sub>o \ a\<^sub>e \ v\<^sub>e \ v\<^sub>o)\ have "a\<^sub>o > a\<^sub>e" by auto from \\ (a\<^sub>o \ a\<^sub>e \ v\<^sub>e \ v\<^sub>o)\ have "v\<^sub>o < v\<^sub>e" by auto show ?thesis proof - from snd_safe_distance_iff_nonneg_D2 [OF \a\<^sub>o > a\<^sub>e\] have 1: "(a\<^sub>e < a\<^sub>o \ v\<^sub>o < v\<^sub>e \ s\<^sub>o - s\<^sub>e \ snd_safe_distance \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0) \ (a\<^sub>e < a\<^sub>o \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0)" by auto from t_stop_diff_neg_means_leq_D2[OF assms \a\<^sub>e < a\<^sub>o\] have "... = (a\<^sub>e < a\<^sub>o \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)" by auto with 1 cond_3[OF assms] show ?thesis by blast qed qed definition d :: "real \ real" where "d t = ( if t \ 0 then s\<^sub>o - s\<^sub>e else if t \ ego.t_stop \ t \ other.t_stop then ego_other.p t else if ego.t_stop \ t \ t \ other.t_stop then other.p t - ego.s_stop else if other.t_stop \ t \ t \ ego.t_stop then other.s_stop - ego.p t else other.s_stop - ego.s_stop )" lemma d_diff: "d t = other.s t - ego.s t" by (auto simp: d_def ego.s_eq_s_stop other.s_eq_s_stop ego.s_cond_def other.s_cond_def movement.p_def field_split_simps) lemma collision_d: "collision S \ (\t\S. d t = 0)" by (force simp: d_diff collision_def ) lemma collision_restrict: "collision {0..} \ collision {0..max ego.t_stop other.t_stop}" by (auto simp: max.coboundedI1 ego.t_stop_nonneg min_def ego.s_eq_s_stop other.s_eq_s_stop collision_def intro!: bexI[where x = "min t (max (movement.t_stop a\<^sub>e v\<^sub>e) (movement.t_stop a\<^sub>o v\<^sub>o))" for t]) lemma collision_union: "collision (A \ B) \ collision A \ collision B" by (auto simp: collision_def) lemma symbolic_checker: "collision {0..} \ (quadroot_in 0 (min ego.t_stop other.t_stop) (1/2 * (a\<^sub>o - a\<^sub>e)) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e)) \ (quadroot_in ego.t_stop other.t_stop (1/2 * a\<^sub>o) v\<^sub>o (s\<^sub>o - ego.s_stop)) \ (quadroot_in other.t_stop ego.t_stop (1/2 * a\<^sub>e) v\<^sub>e (s\<^sub>e - other.s_stop))" (is "_ \ ?q1 \ ?q2 \ ?q3") proof - have *: "{0..max ego.t_stop other.t_stop} = {0 .. min ego.t_stop other.t_stop} \ {ego.t_stop .. other.t_stop} \ {other.t_stop .. ego.t_stop}" using ego.t_stop_nonneg other.t_stop_nonneg by auto have "collision {0..min (movement.t_stop a\<^sub>e v\<^sub>e) (movement.t_stop a\<^sub>o v\<^sub>o)} = ?q1" by (force simp: collision_def quadroot_in_def root_in_def d_def power2_eq_square field_split_simps movement.p_def movement.s_cond_def) moreover have "collision {ego.t_stop .. other.t_stop} = ?q2" using ego.t_stop_nonneg by (force simp: collision_def quadroot_in_def root_in_def d_def ego.s_eq_s_stop movement.s_cond_def movement.p_def) moreover have "collision {other.t_stop .. ego.t_stop} = ?q3" using other.t_stop_nonneg by (force simp: collision_def quadroot_in_def root_in_def d_def other.s_eq_s_stop movement.s_cond_def movement.p_def) ultimately show ?thesis unfolding collision_restrict * collision_union by auto qed end subsection \Checker Design\ definition rel_dist_to_stop :: "real \ real \ real" where "rel_dist_to_stop v a \ - v\<^sup>2 / (2 * a)" context includes floatarith_notation begin definition rel_dist_to_stop_expr :: "nat \ nat \ floatarith" where "rel_dist_to_stop_expr v a = Mult (Minus (Power (Var v) 2)) (Inverse (Mult (Num 2) (Var a)))" definition rel_dist_to_stop' :: "nat \ float interval option \ float interval option \ float interval option" where "rel_dist_to_stop' p v a = approx p (rel_dist_to_stop_expr 0 1) [v, a]" lemma rel_dist_to_stop': "interpret_floatarith (rel_dist_to_stop_expr 0 1) [v, a] = rel_dist_to_stop v a" by (simp add: rel_dist_to_stop_def rel_dist_to_stop_expr_def inverse_eq_divide) definition first_safe_dist :: "real \ real \ real" where "first_safe_dist v\<^sub>e a\<^sub>e \ rel_dist_to_stop v\<^sub>e a\<^sub>e" definition second_safe_dist :: "real \ real \ real \ real \ real" where "second_safe_dist v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o \ rel_dist_to_stop v\<^sub>e a\<^sub>e - rel_dist_to_stop v\<^sub>o a\<^sub>o" definition second_safe_dist_expr :: "nat \ nat \ nat \ nat \ floatarith" where "second_safe_dist_expr ve ae vo ao = Add (rel_dist_to_stop_expr ve ae) (Minus (rel_dist_to_stop_expr vo ao))" definition second_safe_dist' :: "nat \ float interval option \ float interval option \ float interval option \ float interval option \ float interval option" where "second_safe_dist' p v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o = approx p (second_safe_dist_expr 0 1 2 3) [v\<^sub>e, a\<^sub>e, v\<^sub>o, a\<^sub>o]" lemma second_safe_dist': "interpret_floatarith (second_safe_dist_expr 0 1 2 3) [v, a, v', a'] = second_safe_dist v a v' a'" by (simp add: second_safe_dist_def second_safe_dist_expr_def rel_dist_to_stop_def rel_dist_to_stop_expr_def inverse_eq_divide) definition t_stop :: "real \ real \ real" where "t_stop v a \ - v / a" definition t_stop_expr :: "nat \ nat \ floatarith" where "t_stop_expr v a = Minus (Mult (Var v) (Inverse (Var a)))" end definition s_stop :: "real \ real \ real \ real" where "s_stop s v a \ s + rel_dist_to_stop v a" definition discriminant :: "real \ real \ real \ real \ real \ real \ real" where "discriminant s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ (v\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * (a\<^sub>o - a\<^sub>e) * (s\<^sub>o - s\<^sub>e)" definition suff_cond_safe_dist2 :: "real \ real \ real \ real \ real \ real \ bool" where "suff_cond_safe_dist2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ let D2 = discriminant s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o in \ (a\<^sub>e < a\<^sub>o \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2 )" lemma less_sqrt_iff: "y \ 0 \ x < sqrt y \ (x \ 0 \ x\<^sup>2 < y)" by (smt real_le_lsqrt real_less_rsqrt real_sqrt_ge_zero) lemma suff_cond_safe_dist2_code[code]: "suff_cond_safe_dist2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = (let D2 = discriminant s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o in (a\<^sub>e < a\<^sub>o \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o \ 0 \ (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)\<^sup>2 \ D2)))" using real_sqrt_ge_zero real_less_rsqrt less_sqrt_iff by (auto simp: suff_cond_safe_dist2_def Let_def) text \ There are two expressions for safe distance. The first safe distance \first_safe_dist\ is always valid. Whenever the distance is bigger than \first_safe_dist\, it is guarantee to be collision free. The second one is \second_safe_dist\. If the sufficient condition \suff_cond_safe_dist2\ is satisfied and the distance is bigger than \second_safe_dist\, it is guaranteed to be collision free. \ definition check_precond :: "real \ real \ real \ real \ real \ real \ bool" where "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ s\<^sub>o > s\<^sub>e \ 0 \ v\<^sub>e \ 0 \ v\<^sub>o \ a\<^sub>e < 0 \ a\<^sub>o < 0 " lemma check_precond_safe_distance: "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o" proof assume "safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o" then interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o . show "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" by (auto simp: check_precond_def in_front nonneg_vel_ego other.nonneg_vel ego.decel other.decel) qed (unfold_locales; auto simp: check_precond_def) subsubsection \Prescriptive Checker\ definition checker :: "real \ real \ real \ real \ real \ real \ bool" where "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ let distance = s\<^sub>o - s\<^sub>e; precond = check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o; safe_dist1 = first_safe_dist v\<^sub>e a\<^sub>e; safe_dist2 = second_safe_dist v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o; cond2 = suff_cond_safe_dist2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o in precond \ (safe_dist1 < distance \ (safe_dist2 < distance \ distance \ safe_dist1 \ cond2))" lemma aux_logic: assumes "a \ b" assumes "b \ a \ c" shows "a \ b \ c" using assms by blast theorem soundness_correctness: "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ safe_distance.no_collision a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o {0..}" proof (rule aux_logic, simp add: checker_def Let_def) assume cp: "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" then have in_front': "s\<^sub>o > s\<^sub>e" and nonneg_vel_ego: "0 \ v\<^sub>e" and nonneg_vel_other: "0 \ v\<^sub>o" and decelerate_ego: "a\<^sub>e < 0" and decelerate_other: "a\<^sub>o < 0" by (auto simp: check_precond_def) from in_front' have in_front: "0 < s\<^sub>o - s\<^sub>e" by arith interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; fact) interpret ego: braking_movement a\<^sub>e v\<^sub>e s\<^sub>e by (unfold_locales; fact) interpret other: braking_movement a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; fact) have "ego.p_max < s\<^sub>o \ other.p_max \ ego.p_max \ s\<^sub>o \ ego.p_max \ ego.p_max < other.p_max" by arith then show "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = safe_distance.no_collision a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o {0..}" proof (elim disjE) assume "ego.p_max < s\<^sub>o" then have "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" using \a\<^sub>e < 0\ cp by (simp add: checker_def Let_def first_safe_dist_def rel_dist_to_stop_def ego.p_max_def ego.p_def ego.t_stop_def algebra_simps power2_eq_square) moreover have "no_collision {0..}" using \ego.p_max < s\<^sub>o\ by (intro cond_1) (auto simp: ego.s_t_stop) ultimately show ?thesis by auto next assume "other.p_max \ ego.p_max" then have "\ checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" using \a\<^sub>e < 0\ \a\<^sub>o < 0\ other.nonneg_vel by (auto simp add: checker_def Let_def first_safe_dist_def second_safe_dist_def rel_dist_to_stop_def movement.p_max_def movement.p_def movement.t_stop_def algebra_simps power2_eq_square) (smt divide_nonneg_neg mult_nonneg_nonneg) moreover have "collision {0..}" using \other.p_max \ ego.p_max\ by (intro cond_2) (auto simp: other.s_t_stop ego.s_t_stop) ultimately show ?thesis by auto next assume H: "s\<^sub>o \ ego.p_max \ ego.p_max < other.p_max" then have "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = (\ (a\<^sub>e < a\<^sub>o \ v\<^sub>o < v\<^sub>e \ 0 \ D2 \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2))" using \a\<^sub>e < 0\ \a\<^sub>o < 0\ cp by (simp add: checker_def Let_def first_safe_dist_def rel_dist_to_stop_def ego.p_max_def ego.p_def ego.t_stop_def algebra_simps power2_eq_square second_safe_dist_def suff_cond_safe_dist2_def discriminant_def not_less other.p_max_def other.p_def other.t_stop_def) also have "\ = no_collision {0..}" using H unfolding Not_eq_iff by (intro cond_3[symmetric]) (auto simp: ego.s_t_stop other.s_t_stop) finally show ?thesis by auto qed qed definition checker2 :: "real \ real \ real \ real \ real \ real \ bool" where "checker2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ let distance = s\<^sub>o - s\<^sub>e; precond = check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o; safe_dist1 = first_safe_dist v\<^sub>e a\<^sub>e; safe_dist2 = second_safe_dist v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o; safe_dist3 = - rel_dist_to_stop (v\<^sub>o - v\<^sub>e) (a\<^sub>o - a\<^sub>e) in if \ precond then False else if distance > safe_dist1 then True else if a\<^sub>o > a\<^sub>e \ v\<^sub>o < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0 then distance > safe_dist3 else distance > safe_dist2" theorem checker_eq_checker2: "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ checker2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" proof (cases "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o") case False with checker_def checker2_def show ?thesis by auto next case True with check_precond_def safe_distance_def have "safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o" by (simp add: check_precond_safe_distance) from this interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o by auto interpret ego: braking_movement a\<^sub>e v\<^sub>e s\<^sub>e by (unfold_locales; fact) interpret other: braking_movement a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; fact) from \check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o\ cond_3 cond_3'[symmetric] fst_leq_snd_safe_distance ego.s_t_stop ego.p_max_def ego.p_def ego.t_stop_def hyps other.s_t_stop other.p_max_def other.p_def other.t_stop_def checker2_def checker_def suff_cond_safe_dist2_def fst_safe_distance_def first_safe_dist_def snd_safe_distance_def second_safe_dist_def rel_dist_to_stop_def discriminant_def show ?thesis by (auto simp add:power_def Let_def split: if_splits) qed definition checker3 :: "real \ real \ real \ real \ real \ real \ bool" where "checker3 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ let distance = s\<^sub>o - s\<^sub>e; precond = check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o; s_stop_e = s\<^sub>e + rel_dist_to_stop v\<^sub>e a\<^sub>e; s_stop_o = s\<^sub>o + rel_dist_to_stop v\<^sub>o a\<^sub>o in precond \ (s_stop_e < s\<^sub>o \ (s\<^sub>o \ s_stop_e \ s_stop_e < s_stop_o \ (\(a\<^sub>o > a\<^sub>e \ v\<^sub>o < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0 \ distance * (a\<^sub>o - a\<^sub>e) \ (v\<^sub>o - v\<^sub>e)\<^sup>2 / 2))))" theorem checker2_eq_checker3: "checker2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ checker3 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" apply (auto simp: checker2_def checker3_def Let_def first_safe_dist_def not_less suff_cond_safe_dist2_def second_safe_dist_def rel_dist_to_stop_def check_precond_def) proof goal_cases case 1 then interpret safe_distance by unfold_locales auto from fst_leq_snd_safe_distance 1 show ?case by (auto simp: fst_safe_distance_def snd_safe_distance_def) next case 2 then interpret safe_distance by unfold_locales auto from fst_leq_snd_safe_distance 2 show ?case by (auto simp: fst_safe_distance_def snd_safe_distance_def field_split_simps) next case 3 then interpret safe_distance by unfold_locales auto from fst_leq_snd_safe_distance 3 show ?case by (auto simp: fst_safe_distance_def snd_safe_distance_def field_split_simps) qed subsubsection \Approximate Checker\ lemma checker2_def': "checker2 a b c d e f = ( let distance = d - a; precond = check_precond a b c d e f; safe_dist1 = first_safe_dist b c; safe_dist2 = second_safe_dist b c e f; C = c < f \ e < b \ b * f > c * e; P1 = (e - b)\<^sup>2 < 2 * distance * (f - c); P2 = - b\<^sup>2 / c + e\<^sup>2 / f < 2 * distance in precond \ (safe_dist1 < distance \ safe_dist1 \ distance \ (C \ P1 \ \C \ P2)))" unfolding checker2_def by (auto simp: Let_def field_split_simps check_precond_def second_safe_dist_def rel_dist_to_stop_def) lemma power2_less_sqrt_iff: "(x::real)\<^sup>2 < y \ (y \ 0 \ abs x < sqrt y)" apply (auto simp: real_less_rsqrt abs_real_def less_sqrt_iff) apply (meson le_less le_less_trans not_less power2_less_0)+ done schematic_goal checker_form: "interpret_form ?x ?y \ checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" unfolding checker_eq_checker2 checker2_eq_checker3 checker3_def check_precond_def first_safe_dist_def second_safe_dist_def suff_cond_safe_dist2_def Let_def t_stop_def s_stop_def rel_dist_to_stop_def discriminant_def not_le not_less de_Morgan_conj de_Morgan_disj power2_less_sqrt_iff apply (tactic \(Reification.tac @{context} @{thms interpret_form.simps interpret_floatarith.simps interpret_floatarith_divide interpret_floatarith_diff}) NONE 1\) apply assumption done context includes floatarith_notation begin definition "checker' p s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = approx_form p (Conj (Conj (Less (Var (Suc (Suc 0))) (Var (Suc (Suc (Suc 0))))) (Conj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Var (Suc (Suc (Suc (Suc (Suc 0))))))) (Conj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Var (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Conj (Less (Var 0) (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))) (Less (Var (Suc 0)) (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))))))) (Disj (Less (Add (Var (Suc (Suc 0))) (Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc 0)))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var 0))))) (Var (Suc (Suc (Suc 0))))) (Conj (LessEqual (Var (Suc (Suc (Suc 0)))) (Add (Var (Suc (Suc 0))) (Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc 0)))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var 0)))))) (Conj (Less (Add (Var (Suc (Suc 0))) (Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc 0)))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var 0))))) (Add (Var (Suc (Suc (Suc 0)))) (Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var (Suc 0))))))) (Disj (LessEqual (Var (Suc 0)) (Var 0)) (Disj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc 0)))))) (Var (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Disj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Add (Var (Suc (Suc (Suc (Suc (Suc 0)))))) (Minus (Mult (Mult (Var 0) (Inverse (Var (Suc 0)))) (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))))) (Less (Mult (Power (Add (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) (Minus (Var (Suc (Suc (Suc (Suc (Suc 0)))))))) 2) (Inverse (Var (Suc (Suc (Suc (Suc 0))))))) (Mult (Add (Var (Suc (Suc (Suc 0)))) (Minus (Var (Suc (Suc 0))))) (Add (Var (Suc 0)) (Minus (Var 0)))))))))))) ([a\<^sub>e, a\<^sub>o, s\<^sub>e, s\<^sub>o, Interval' (Float 2 0) (Float 2 0), v\<^sub>e, v\<^sub>o, Interval' (Float 0 1) (Float 0 1)]) (replicate 8 0)" lemma less_Suc_iff_disj: "i < Suc x \ i = x \ i < x" by auto lemma checker'_soundness_correctness: assumes "a \ {real_of_float al .. real_of_float au}" assumes "b \ {real_of_float bl .. real_of_float bu}" assumes "c \ {real_of_float cl .. real_of_float cu}" assumes "d \ {real_of_float dl .. real_of_float du}" assumes "e \ {real_of_float el .. real_of_float eu}" assumes "f \ {real_of_float fl .. real_of_float fu}" assumes chk: "checker' p (Interval' al au) (Interval' bl bu) (Interval' cl cu) (Interval' dl du) (Interval' el eu) (Interval' fl fu)" shows "checker a b c d e f" apply (rule checker_form) apply (rule approx_form_aux) apply (rule chk[unfolded checker'_def]) using assms(1-6) unfolding bounded_by_def proof (auto split: option.splits) fix i x2 assume *: "[Interval' cl cu, Interval' fl fu, Interval' al au, Interval' dl du, Interval' (Float 2 0) (Float 2 0), Interval' bl bu, Interval' el eu, Interval' 0 0] ! i = Some x2" assume " i < Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))" then consider "i = 0" | "i = 1" | "i = 2" | "i = 3" | "i = 4" | "i = 5" | "i = 6" | "i = 7" by linarith thus " [c, f, a, d, 2, b, e, 0] ! i \\<^sub>r x2" apply cases using assms(1-6) * by (auto intro!: in_real_intervalI dest!: Interval'_eq_Some) qed lemma approximate_soundness_correctness: assumes "a \ {real_of_float al .. real_of_float au}" assumes "b \ {real_of_float bl .. real_of_float bu}" assumes "c \ {real_of_float cl .. real_of_float cu}" assumes "d \ {real_of_float dl .. real_of_float du}" assumes "e \ {real_of_float el .. real_of_float eu}" assumes "f \ {real_of_float fl .. real_of_float fu}" assumes chk: "checker' p (Interval' al au) (Interval' bl bu) (Interval' cl cu) (Interval' dl du) (Interval' el eu) (Interval' fl fu)" shows checker'_precond: "check_precond a b c d e f" and checker'_no_collision: "safe_distance.no_collision c b a f e d {0..}" unfolding atomize_conj apply (subst soundness_correctness[symmetric]) using checker'_soundness_correctness[OF assms] by (auto simp: checker_def Let_def) subsubsection \Symbolic Checker\ definition symbolic_checker :: "real \ real \ real \ real \ real \ real \ bool" where "symbolic_checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ let e_stop = - v\<^sub>e / a\<^sub>e; o_stop = - v\<^sub>o / a\<^sub>o in check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ (\quadroot_in 0 (min e_stop o_stop) (1/2 * (a\<^sub>o - a\<^sub>e)) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e) \ \quadroot_in e_stop o_stop (1/2 * a\<^sub>o) v\<^sub>o (s\<^sub>o - movement.p a\<^sub>e v\<^sub>e s\<^sub>e e_stop) \ \quadroot_in o_stop e_stop (1/2 * a\<^sub>e) v\<^sub>e (s\<^sub>e - movement.p a\<^sub>o v\<^sub>o s\<^sub>o o_stop))" theorem symbolic_soundness_correctness: "symbolic_checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ safe_distance.no_collision a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o {0..}" proof - { assume c: "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o" then interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o by (simp add: check_precond_safe_distance) have "symbolic_checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = no_collision {0..}" using c unfolding symbolic_checker symbolic_checker_def ego.s_t_stop other.s_t_stop ego.p_max_def other.p_max_def by (auto simp: Let_def movement.t_stop_def) } then show ?thesis by (auto simp: symbolic_checker_def Let_def) qed end end \ No newline at end of file diff --git a/thys/Safe_Distance/Safe_Distance_Reaction.thy b/thys/Safe_Distance/Safe_Distance_Reaction.thy --- a/thys/Safe_Distance/Safe_Distance_Reaction.thy +++ b/thys/Safe_Distance/Safe_Distance_Reaction.thy @@ -1,1757 +1,1757 @@ \<^marker>\creator "Albert Rizaldi" "Fabian Immler\ section \Safe Distance with Reaction Time\ theory Safe_Distance_Reaction imports Safe_Distance begin subsection \Normal Safe Distance\ locale safe_distance_normal = safe_distance + fixes \ :: real assumes pos_react: "0 < \" begin sublocale ego2: braking_movement a\<^sub>e v\<^sub>e "(ego.q \)" .. lemma ego2_s_init: "ego2.s 0 = ego.q \" unfolding ego2.s_def by auto definition \ :: "real \ real" where "\ t = t - \" definition \' :: "real \ real" where "\' t = 1" lemma \_continuous[continuous_intros]: "continuous_on T \" unfolding \_def by (auto intro: continuous_intros) lemma isCont_\[continuous_intros]: "isCont \ x" using \_continuous[of UNIV] by (auto simp: continuous_on_eq_continuous_at) lemma del_has_vector_derivative[derivative_intros]: "(\ has_vector_derivative \' t) (at t within u)" by (auto simp: \_def[abs_def] \'_def has_vector_derivative_def algebra_simps intro!: derivative_eq_intros) lemma del_has_real_derivative[derivative_intros]: "(\ has_real_derivative \' t) (at t within u)" using del_has_vector_derivative - by (simp add:has_field_derivative_iff_has_vector_derivative) + by (simp add:has_real_derivative_iff_has_vector_derivative) lemma delay_image: "\ ` {\..} = {0..}" proof (rule subset_antisym, unfold image_def, unfold \_def) show "{y. \x\{\..}. y = x - \} \ {0..}" by auto next show "{0..} \ {y. \x\{\..}. y = x - \}" proof (rule subsetI) fix a assume "(a::real) \ {0..}" hence "0 \ a" by simp hence "\x\{\..}. a = x - \" using bexI[where x = "a + \"] by auto thus "a \ {y. \x\{\..}. y = x - \}" by auto qed qed lemma s_delayed_has_real_derivative[derivative_intros]: assumes "\ \ t" shows "((ego2.s \ \) has_field_derivative ego2.s' (t - \) * \' t) (at t within {\..})" proof (rule DERIV_image_chain) from assms have 0: "0 \ t - \" by simp from ego2.t_stop_nonneg have 1: "v\<^sub>e / a\<^sub>e \ 0" unfolding ego2.t_stop_def by simp from ego2.decel have 2: "a\<^sub>e \ 0" by simp show "(ego2.s has_real_derivative ego2.s' (t - \)) (at (\ t) within \ ` {\..})" using ego2.s_has_real_derivative[OF 0 1 2] sym[OF delay_image] unfolding \_def by simp next from del_has_real_derivative show "(\ has_real_derivative \' t) (at t within {\..})" by auto qed lemma s_delayed_has_real_derivative' [derivative_intros]: assumes "\ \ t" shows "((ego2.s \ \) has_field_derivative (ego2.s' \ \) t) (at t within {\..})" proof - from s_delayed_has_real_derivative[OF assms] have "((ego2.s \ \) has_field_derivative ego2.s' (t - \) * \' t) (at t within {\..})" by auto hence "((ego2.s \ \) has_field_derivative ego2.s' (t - \) * 1) (at t within {\..})" using \'_def[of t] by metis hence "((ego2.s \ \) has_field_derivative ego2.s' (t - \)) (at t within {\..})" by (simp add:algebra_simps) thus ?thesis unfolding comp_def \_def by auto qed lemma s_delayed_has_vector_derivative' [derivative_intros]: assumes "\ \ t" shows "((ego2.s \ \) has_vector_derivative (ego2.s' \ \) t) (at t within {\..})" using s_delayed_has_real_derivative'[OF assms] - by (simp add:has_field_derivative_iff_has_vector_derivative) + by (simp add:has_real_derivative_iff_has_vector_derivative) definition u :: "real \ real" where "u t = ( if t \ 0 then s\<^sub>e else if t \ \ then ego.q t else (ego2.s \ \) t)" lemma init_u: "t \ 0 \ u t = s\<^sub>e" unfolding u_def by auto lemma u_delta: "u \ = ego2.s 0" proof - have "u \ = ego.q \" using pos_react unfolding u_def by auto also have "... = ego2.s 0" unfolding ego2.s_def by auto finally show "u \ = ego2.s 0" . qed lemma q_delta: "ego.q \ = ego2.s 0" using u_delta pos_react unfolding u_def by auto definition u' :: "real \ real" where "u' t = (if t \ \ then ego.q' t else ego2.s' (t - \))" lemma u'_delta: "u' \ = ego2.s' 0" proof - have "u' \ = ego.q' \" unfolding u'_def by auto also have "... = v\<^sub>e" unfolding ego2.q'_def by simp also have "... = ego2.p' 0" unfolding ego2.p'_def by simp also have "... = ego2.s' 0" using ego2.t_stop_nonneg unfolding ego2.s'_def by auto finally show "u' \ = ego.s' 0" . qed lemma q'_delta: "ego.q' \ = ego2.s' 0" using u'_delta unfolding u'_def by auto lemma u_has_real_derivative[derivative_intros]: assumes nonneg_t: "t \ 0" shows "(u has_real_derivative u' t) (at t within {0..})" proof - from pos_react have "0 \ \" by simp have temp: "((\t. if t \ {0 .. \} then ego.q t else (ego2.s \ \) t) has_real_derivative (if t \ {0..\} then ego.q' t else (ego2.s' \ \) t)) (at t within {0..})" (is "(?f1 has_real_derivative ?f2) (?net)") unfolding u_def[abs_def] u'_def - has_field_derivative_iff_has_vector_derivative + has_real_derivative_iff_has_vector_derivative apply (rule has_vector_derivative_If_within_closures[where T = "{\..}"]) using \0 \ \\ q_delta q'_delta ego.s_has_vector_derivative[OF assms] ego.decel ego.t_stop_nonneg s_delayed_has_vector_derivative'[of "t"] \_def unfolding comp_def by (auto simp: assms max_def insert_absorb intro!: ego.q_has_vector_derivative) show ?thesis - unfolding has_vector_derivative_def has_field_derivative_iff_has_vector_derivative + unfolding has_vector_derivative_def has_real_derivative_iff_has_vector_derivative u'_def u_def[abs_def] proof (rule has_derivative_transform[where f="(\t. if t \ {0..\} then ego.q t else (ego2.s \ \) t)"]) from nonneg_t show " t \ {0..}" by auto next fix x assume "(x::real) \ {0..}" hence "x \ \ \ x \ {0 .. \}" by simp thus " (if x \ 0 then s\<^sub>e else if x \ \ then ego.q x else (ego2.s \ \) x) = (if x \ {0..\} then ego.q x else (ego2.s \ \) x)" using pos_react unfolding ego.q_def by auto next from temp have "(?f1 has_vector_derivative ?f2) ?net" - using has_field_derivative_iff_has_vector_derivative by auto + using has_real_derivative_iff_has_vector_derivative by auto moreover with assms have "t \ {0 .. \} \ t \ \" by auto ultimately show " ((\t. if t \ {0..\} then ego.q t else (ego2.s \ \) t) has_derivative (\x. x *\<^sub>R (if t \ \ then ego2.q' t else ego2.s' (t - \)))) (at t within {0..})" unfolding comp_def \_def has_vector_derivative_def by auto qed qed definition t_stop :: real where "t_stop = ego2.t_stop + \" lemma t_stop_nonneg: "0 \ t_stop" unfolding t_stop_def using ego2.t_stop_nonneg pos_react by auto lemma t_stop_pos: "0 < t_stop" unfolding t_stop_def using ego2.t_stop_nonneg pos_react by auto lemma t_stop_zero: assumes "t_stop \ x" assumes "x \ \" shows "v\<^sub>e = 0" using assms unfolding t_stop_def using ego2.t_stop_nonneg pos_react ego2.t_stop_zero by auto lemma u'_stop_zero: "u' t_stop = 0" unfolding u'_def t_stop_def ego2.q'_def ego2.s'_def using ego2.t_stop_nonneg ego2.p'_stop_zero decelerate_ego ego2.t_stop_zero by auto definition u_max :: real where "u_max = u (ego2.t_stop + \)" lemma u_max_eq: "u_max = ego.q \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" proof (cases "ego2.t_stop = 0") assume "ego2.t_stop = 0" hence "v\<^sub>e = 0" using ego2.t_stop_zero by simp with \ego2.t_stop = 0\ show "u_max = ego.q \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" unfolding u_max_def u_def using pos_react by auto next assume "ego2.t_stop \ 0" hence "u_max = (ego2.s \ \) (ego2.t_stop + \)" unfolding u_max_def u_def using ego2.t_stop_nonneg pos_react by auto moreover have "... = ego2.s ego2.t_stop" unfolding comp_def \_def by auto moreover have "... = ego2.p_max" unfolding ego2.s_def ego2.p_max_def using \ego2.t_stop \ 0\ ego2.t_stop_nonneg by auto moreover have "... = ego.q \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" using ego2.p_max_eq . ultimately show ?thesis by auto qed lemma u_mono: assumes "x \ y" "y \ t_stop" shows "u x \ u y" proof - have "y \ 0 \ (0 < y \ y \ \) \ \ < y" by auto moreover { assume "y \ 0" with assms have "x \ 0" by auto with \y \ 0\ have "u x \ u y" unfolding u_def by auto } moreover { assume "0 < y \ y \ \" with assms have "x \ \" by auto hence "u x \ u y" proof (cases "x \ 0") assume "x \ 0" with \x \ \\ and \0 < y \ y \ \\ show "u x \ u y" unfolding u_def using ego.q_min by auto next assume "\ x \ 0" with \0 < y \ y \ \\ and assms show "u x \ u y" unfolding u_def using ego.q_mono by auto qed } moreover { assume "\ < y" have "u x \ u y" proof (cases "\ < x") assume "\ < x" with pos_react have "\ x \ 0" by auto moreover from \\ < y\ and pos_react have "\ y \ 0" by auto ultimately show "u x \ u y" unfolding u_def comp_def using assms ego2.s_mono[of "x - \" "y - \"] \\ < y\ \\ < x\ by (auto simp:\_def) next assume "\ \ < x" hence "x \ \" by simp hence "u x \ ego.q \" unfolding u_def using pos_react nonneg_vel_ego by (auto simp add:ego.q_def mult_left_mono) also have "... = ego2.s (\ \)" unfolding ego2.s_def unfolding \_def by auto also have "... \ ego2.s (\ y)" unfolding \_def using \\ < y\ by (auto simp add:ego2.s_mono) also have "... = u y" unfolding u_def using \\ < y\ pos_react by auto ultimately show "u x \ u y" by auto qed } ultimately show "u x \ u y" by auto qed lemma u_antimono: "x \ y \ t_stop \ x \ u y \ u x" proof - assume 1: "x \ y" assume 2: "t_stop \ x" hence "\ \ x" unfolding \_def t_stop_def using pos_react ego2.t_stop_nonneg by auto with 1 have "\ \ y" by auto from 1 and 2 have 3: "t_stop \ y" by auto show "u y \ u x" proof (cases "x \ \ \ y \ \") assume "x \ \ \ y \ \" hence "x \ \" and "y \ \" by auto have "u y \ (ego2.s \ \) y" unfolding u_def using \\ \ y\ \y \ \\ pos_react by auto also have "... \ (ego2.s \ \) x" unfolding comp_def proof (intro ego2.s_antimono) show "\ x \ \ y" unfolding \_def using \x \ y\ by auto next show "ego2.t_stop \ \ x" unfolding \_def using \t_stop \ x\ by (auto simp: t_stop_def) qed also have "... \ u x" unfolding u_def using \\ \ x\\x \ \\ pos_react by auto ultimately show "u y \ u x" by auto next assume "\ (x \ \ \ y \ \)" have "x \ \ \ y \ \" proof (rule impI; erule contrapos_pp[where Q="\ x = \"]) assume "\ y \ \" hence "y = \" by simp with \t_stop \ y\ have "ego2.t_stop = 0" unfolding t_stop_def using ego2.t_stop_nonneg by auto with \t_stop \ x\ have "x = \" unfolding t_stop_def using \x \ y\ \y = \\ by auto thus "\ x \ \" by auto qed with \\ (x \ \ \ y \ \)\ have "(x = \ \ y = \) \ (x = \)" by auto moreover { assume "x = \ \ y = \" hence "x = \" and "y = \" by auto hence "u y \ ego.q \" unfolding u_def using pos_react by auto also have "... \ u x" unfolding u_def using \x = \\ pos_react by auto ultimately have "u y \ u x" by auto } moreover { assume "x = \" hence "ego2.t_stop = 0" using \t_stop \ x\ ego2.t_stop_nonneg by (auto simp:t_stop_def) hence "v\<^sub>e = 0" by (rule ego2.t_stop_zero) hence "u y \ ego.q \" using pos_react \x = \\ \x \ y\ \v\<^sub>e = 0\ unfolding u_def comp_def \_def ego2.s_def ego2.p_def ego2.p_max_def ego2.t_stop_def by auto also have "... \ u x" using \x = \\ pos_react unfolding u_def by auto ultimately have "u y \ u x" by auto } ultimately show ?thesis by auto qed qed lemma u_max: "u x \ u_max" unfolding u_max_def using t_stop_def by (cases "x \ t_stop") (auto intro: u_mono u_antimono) lemma u_eq_u_stop: "NO_MATCH t_stop x \ x \ t_stop \ u x = u_max" proof - assume "t_stop \ x" with t_stop_pos have "0 < x" by auto from \t_stop \ x\ have "\ \ x" unfolding t_stop_def using ego2.t_stop_nonneg by auto show "u x = u_max" proof (cases "x \ \") assume "x \ \" with \t_stop \ x\ have "v\<^sub>e = 0" by (rule t_stop_zero) also have "x = \" using \x \ \\ and \\ \ x\ by auto ultimately have "u x = ego.q \" unfolding u_def using pos_react by auto also have "... = u_max" unfolding u_max_eq using \v\<^sub>e = 0\ by auto ultimately show "u x = u_max" by simp next assume "\ x \ \" hence "\ < x" by auto hence "u x = (ego2.s \ \) x" unfolding u_def using pos_react by auto also have "... = ego2.s ego2.t_stop" proof (unfold comp_def; unfold \_def; intro order.antisym) have "x - \ \ ego2.t_stop" using \t_stop \ x\ unfolding t_stop_def by auto thus "ego2.s (x - \) \ ego2.s ego2.t_stop" by (rule ego2.s_antimono) simp next have "x - \ \ ego2.t_stop" using \t_stop \ x\ unfolding t_stop_def by auto thus "ego2.s ego2.t_stop \ ego2.s (x - \)" using ego2.t_stop_nonneg by (rule ego2.s_mono) qed also have "... = u_max" unfolding u_max_eq ego2.s_t_stop ego2.p_max_eq by auto ultimately show "u x = u_max" by auto qed qed lemma at_least_delta: assumes "x \ \" assumes "t_stop \ x" shows "ego.q x = ego2.s (x - \)" using assms ego2.t_stop_nonneg unfolding t_stop_def ego2.s_def less_eq_real_def by auto lemma continuous_on_u[continuous_intros]: "continuous_on T u" unfolding u_def[abs_def] using t_stop_nonneg pos_react at_least_delta proof (intro continuous_on_subset[where t=T and s = "{..0} \ ({0..\} \ ({\ .. t_stop} \ {t_stop ..}))"] continuous_on_If continuous_intros) fix x assume " \ x \ \" assume "x \ {0..\}" hence "0 \ x" and "x \ \" by auto thus "ego.q x = (ego2.s \ \) x" unfolding comp_def \_def ego2.s_def using \\ x \ \\ by auto next fix x assume "x \ {\..t_stop} \ {t_stop..}" hence "\ \ x" unfolding t_stop_def using pos_react ego.t_stop_nonneg by auto also assume "x \ \" ultimately have "x = \" by auto thus "ego.q x = (ego2.s \ \) x" unfolding comp_def \_def ego2.s_def by auto next fix t::real assume "t \ {.. 0}" hence "t \ 0" by auto also assume "\ t \ 0" ultimately have "t = 0" by auto hence "s\<^sub>e = ego.q t" unfolding ego.q_def by auto with pos_react \t = 0\ show "s\<^sub>e = (if t \ \ then ego.q t else (ego2.s \ \) t)" by auto next fix t::real assume "t \ {0..\} \ ({\..t_stop} \ {t_stop..})" hence "0 \ t" using pos_react ego2.t_stop_nonneg by (auto simp: t_stop_def) also assume "t \ 0" ultimately have "t = 0" by auto hence " s\<^sub>e = (if t \ \ then ego.q t else (ego2.s \ \) t)" using pos_react ego.init_q by auto thus "s\<^sub>e = (if t \ \ then ego.q t else (ego2.s \ \) t)" by auto next show "T \ {..0} \ ({0..\} \ ({\..t_stop} \ {t_stop..}))" by auto qed lemma isCont_u[continuous_intros]: "isCont u x" using continuous_on_u[of UNIV] by (auto simp:continuous_on_eq_continuous_at) definition collision_react :: "real set \ bool" where "collision_react time_set \ (\t\time_set. u t = other.s t )" abbreviation no_collision_react :: "real set \ bool" where "no_collision_react time_set \ \ collision_react time_set" lemma no_collision_reactI: assumes "\t. t \ S \ u t \ other.s t" shows "no_collision_react S" using assms unfolding collision_react_def by blast lemma no_collision_union: assumes "no_collision_react S" assumes "no_collision_react T" shows "no_collision_react (S \ T)" using assms unfolding collision_react_def by auto lemma collision_trim_subset: assumes "collision_react S" assumes "no_collision_react T" assumes "T \ S" shows "collision_react (S - T)" using assms unfolding collision_react_def by auto theorem cond_1r : "u_max < s\<^sub>o \ no_collision_react {0..}" proof (rule no_collision_reactI, simp) fix t :: real assume "0 \ t" have "u t \ u_max" by (rule u_max) also assume "... < s\<^sub>o" also have "... = other.s 0" by (simp add: other.init_s) also have "... \ other.s t" using \0 \ t\ hyps by (intro other.s_mono) auto finally show "u t \ other.s t" by simp qed definition safe_distance_1r :: real where "safe_distance_1r = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" lemma sd_1r_eq: "(s\<^sub>o - s\<^sub>e > safe_distance_1r) = (u_max < s\<^sub>o)" proof - have "(s\<^sub>o - s\<^sub>e > safe_distance_1r) = (s\<^sub>o - s\<^sub>e > v\<^sub>e * \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2)" unfolding safe_distance_1r_def by auto moreover have "... = (s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2 < s\<^sub>o)" by auto ultimately show ?thesis using u_max_eq ego.q_def by auto qed lemma sd_1r_correct: assumes "s\<^sub>o - s\<^sub>e > safe_distance_1r" shows "no_collision_react {0..}" proof - from assms have "u_max < s\<^sub>o" using sd_1r_eq by auto thus ?thesis by (rule cond_1r) qed lemma u_other_strict_ivt: assumes "u t > other.s t" shows "collision_react {0.. t" with assms in_front have "\x\0. x \ t \ other.s x - u x = 0" by (intro IVT2) (auto simp: init_u other.init_s continuous_diff isCont_u other.isCont_s) then show ?thesis using assms by (auto simp add: algebra_simps collision_react_def Bex_def order.order_iff_strict) qed(insert assms hyps, auto simp: collision_react_def init_u other.init_s) lemma collision_react_subset: "collision_react s \ s \ t \ collision_react t" by (auto simp:collision_react_def) lemma u_other_ivt: assumes "u t \ other.s t" shows "collision_react {0 .. t}" proof cases assume "u t > other.s t" from u_other_strict_ivt[OF this] show ?thesis by (rule collision_react_subset) auto qed (insert hyps assms; cases "t \ 0"; force simp: collision_react_def init_u other.init_s) theorem cond_2r: assumes "u_max \ other.s_stop" shows "collision_react {0 ..}" using assms apply(intro collision_react_subset[where t="{0..}" and s ="{0 .. max t_stop other.t_stop}"]) apply(intro u_other_ivt[where t ="max t_stop other.t_stop"]) apply(auto simp: u_eq_u_stop other.s_eq_s_stop) done definition ego_other2 :: "real \ real" where "ego_other2 t = other.s t - u t" lemma continuous_on_ego_other2[continuous_intros]: "continuous_on T ego_other2" unfolding ego_other2_def[abs_def] by (intro continuous_intros) lemma isCont_ego_other2[continuous_intros]: "isCont ego_other2 x" using continuous_on_ego_other2[of UNIV] by (auto simp: continuous_on_eq_continuous_at) definition ego_other2' :: "real \ real" where "ego_other2' t = other.s' t - u' t" lemma ego_other2_has_real_derivative[derivative_intros]: assumes "0 \ t" shows "(ego_other2 has_real_derivative ego_other2' t) (at t within {0..})" using assms other.t_stop_nonneg decelerate_other unfolding other.t_stop_def by (auto simp: ego_other2_def[abs_def] ego_other2'_def algebra_simps intro!: derivative_eq_intros) theorem cond_3r_1: assumes "u \ \ other.s \" shows "collision_react {0 .. \}" proof (unfold collision_react_def) have 1: "\t\0. t \ \ \ ego_other2 t = 0" proof (intro IVT2) show "ego_other2 \ \ 0" unfolding ego_other2_def using assms by auto next show "0 \ ego_other2 0" unfolding ego_other2_def using other.init_s[of 0] init_u[of 0] in_front by auto next show "0 \ \" using pos_react by auto next show "\t. 0 \ t \ t \ \ \ isCont ego_other2 t" using isCont_ego_other2 by auto qed then obtain t where "0 \ t \ t \ \ \ ego_other2 t = 0" by auto hence "t \ {0 .. \}" and "u t = other.s t" unfolding ego_other2_def by auto thus "\t\{0..\}. u t = other.s t" by (intro bexI) qed definition distance0 :: real where "distance0 = v\<^sub>e * \ - v\<^sub>o * \ - a\<^sub>o * \\<^sup>2 / 2" definition distance0_2 :: real where "distance0_2 = v\<^sub>e * \ + 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o" theorem cond_3r_1': assumes "s\<^sub>o - s\<^sub>e \ distance0" assumes "\ \ other.t_stop" shows "collision_react {0 .. \}" proof - from assms have "u \ \ other.s \" unfolding distance0_def other.s_def other.p_def u_def ego.q_def using pos_react by auto thus ?thesis using cond_3r_1 by auto qed theorem distance0_2_eq: assumes "\ > other.t_stop" shows "(u \ < other.s \) = (s\<^sub>o - s\<^sub>e > distance0_2)" proof - from assms have "(u \ < other.s \) = (ego.q \ < other.p_max)" using u_def other.s_def pos_react by auto also have "... = (s\<^sub>e + v\<^sub>e * \ < s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2)" using ego.q_def other.p_max_def other.p_def other.t_stop_def by auto also have "... = (v\<^sub>e * \ - v\<^sub>o * (- v\<^sub>o / a\<^sub>o) - 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2 < s\<^sub>o - s\<^sub>e)" by linarith also have "... = (v\<^sub>e * \ + v\<^sub>o\<^sup>2 / a\<^sub>o - 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)" using other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto also have "... = (v\<^sub>e * \ + 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)" by linarith thus ?thesis using distance0_2_def by (simp add: calculation) qed theorem cond_3r_1'_2: assumes "s\<^sub>o - s\<^sub>e \ distance0_2" assumes "\ > other.t_stop" shows "collision_react {0 .. \}" proof - from assms distance0_2_eq have "u \ \ other.s \" unfolding distance0_def other.s_def other.p_def u_def ego.q_def using pos_react by auto thus ?thesis using cond_3r_1 by auto qed definition safe_distance_3r :: real where "safe_distance_3r = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2" lemma distance0_at_most_sd3r: "distance0 \ safe_distance_3r" unfolding distance0_def safe_distance_3r_def using nonneg_vel_ego decelerate_ego by (auto simp add:field_simps) definition safe_distance_4r :: real where "safe_distance_4r = (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 + v\<^sub>e * \" lemma distance0_at_most_sd4r: assumes "a\<^sub>o > a\<^sub>e" shows "distance0 \ safe_distance_4r" proof - from assms have "a\<^sub>o \ a\<^sub>e" by auto have "0 \ (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / (2 * a\<^sub>o - 2 * a\<^sub>e)" by (rule divide_nonneg_nonneg) (auto simp add:assms \a\<^sub>e \ a\<^sub>o\) thus ?thesis unfolding distance0_def safe_distance_4r_def by auto qed definition safe_distance_2r :: real where "safe_distance_2r = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" lemma vo_start_geq_ve: assumes "\ \ other.t_stop" assumes "other.s' \ \ v\<^sub>e" shows "u \ < other.s \" proof - from assms have "v\<^sub>e \ v\<^sub>o + a\<^sub>o * \" unfolding other.s'_def other.p'_def by auto with mult_right_mono[OF this, of "\"] have "v\<^sub>e * \ \ v\<^sub>o * \ + a\<^sub>o * \\<^sup>2" (is "?l0 \ ?r0") using pos_react by (auto simp add:field_simps power_def) hence "s\<^sub>e + ?l0 \ s\<^sub>e + ?r0" by auto also have "... < s\<^sub>o + ?r0" using in_front by auto also have "... < s\<^sub>o + v\<^sub>o * \ + a\<^sub>o * \\<^sup>2 / 2" using decelerate_other pos_react by auto finally show ?thesis using pos_react assms(1) unfolding u_def ego.q_def other.s_def other.t_stop_def other.p_def by auto qed theorem so_star_stop_leq_se_stop: assumes "\ \ other.t_stop" assumes "other.s' \ < v\<^sub>e" assumes "\ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" shows "0 \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2 + (v\<^sub>o + a\<^sub>o * \)\<^sup>2 / a\<^sub>o / 2" proof - consider "v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ \ 0" | "\ (v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ \ 0)" by auto thus ?thesis proof (cases) case 1 hence "v\<^sub>e - a\<^sub>e / a\<^sub>o * (v\<^sub>o + a\<^sub>o * \) \ 0" unfolding other.s'_def other.p'_def by (auto simp add:assms(1)) hence "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o - a\<^sub>e * \ \ 0" (is "?l0 \ 0") using decelerate_other by (auto simp add:field_simps) hence "?l0 / a\<^sub>e \ 0" using divide_right_mono_neg[OF \?l0 \ 0\] decelerate_ego by auto hence "0 \ v\<^sub>e / a\<^sub>e - v\<^sub>o / a\<^sub>o - \" using decelerate_ego by (auto simp add:field_simps) hence *: "- v\<^sub>e / a\<^sub>e \ - (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o" using decelerate_other by (auto simp add:field_simps) from assms have **: "v\<^sub>o + a\<^sub>o * \ \ v\<^sub>e" unfolding other.s'_def other.p'_def by auto have vo_star_nneg: "v\<^sub>o + a\<^sub>o * \ \ 0" proof - from assms(1) have "- v\<^sub>o \ a\<^sub>o * \" unfolding other.t_stop_def using decelerate_other by (auto simp add:field_simps) thus ?thesis by auto qed from mult_mono[OF * ** _ \0 \ v\<^sub>o + a\<^sub>o * \\] have "- (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o * (v\<^sub>o + a\<^sub>o * \) \ - v\<^sub>e / a\<^sub>e * v\<^sub>e" using nonneg_vel_ego decelerate_ego by (auto simp add:field_simps) hence "- (v\<^sub>o + a\<^sub>o * \)\<^sup>2 / a\<^sub>o \ - v\<^sub>e\<^sup>2 / a\<^sub>e " by (auto simp add: field_simps power_def) thus ?thesis by (auto simp add:field_simps) next case 2 with assms have "a\<^sub>o \ a\<^sub>e" by auto from assms(2) have "(v\<^sub>o + a\<^sub>o * \) \ v\<^sub>e" unfolding other.s'_def using assms unfolding other.p'_def by auto have vo_star_nneg: "v\<^sub>o + a\<^sub>o * \ \ 0" proof - from assms(1) have "- v\<^sub>o \ a\<^sub>o * \" unfolding other.t_stop_def using decelerate_other by (auto simp add:field_simps) thus ?thesis by auto qed with mult_mono[OF \v\<^sub>o + a\<^sub>o * \ \ v\<^sub>e\ \v\<^sub>o + a\<^sub>o * \ \ v\<^sub>e\] have *: "(v\<^sub>o + a\<^sub>o * \)\<^sup>2 \ v\<^sub>e\<^sup>2" using nonneg_vel_ego by (auto simp add:power_def) from \a\<^sub>o \ a\<^sub>e\ have "- 1 /a\<^sub>o \ - 1 / a\<^sub>e" using decelerate_ego decelerate_other by (auto simp add:field_simps) from mult_mono[OF * this] have "(v\<^sub>o + a\<^sub>o * \)\<^sup>2 * (- 1 / a\<^sub>o) \ v\<^sub>e\<^sup>2 * (- 1 / a\<^sub>e)" using nonneg_vel_ego decelerate_other by (auto simp add:field_simps) then show ?thesis by auto qed qed theorem distance0_at_most_distance2r: assumes "\ \ other.t_stop" assumes "other.s' \ < v\<^sub>e" assumes "\ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" shows "distance0 \ safe_distance_2r" proof - from so_star_stop_leq_se_stop[OF assms] have " 0 \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2 + (v\<^sub>o + a\<^sub>o * \)\<^sup>2 / a\<^sub>o / 2 " (is "0 \ ?term") by auto have "safe_distance_2r = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" unfolding safe_distance_2r_def by auto also have "... = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + (v\<^sub>o + a\<^sub>o * \)\<^sup>2 / 2 / a\<^sub>o - v\<^sub>o * \ - a\<^sub>o * \\<^sup>2 / 2" using decelerate_other by (auto simp add:field_simps power_def) also have "... = v\<^sub>e * \ - v\<^sub>o * \ - a\<^sub>o * \\<^sup>2 / 2 + ?term" (is "_ = ?left + ?term") by (auto simp add:field_simps) finally have "safe_distance_2r = distance0 + ?term" unfolding distance0_def by auto with \0 \ ?term\ show "distance0 \ safe_distance_2r" by auto qed theorem dist0_sd2r_1: assumes "\ \ other.t_stop" assumes "\ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r" shows "s\<^sub>o - s\<^sub>e > distance0" proof (cases "other.s' \ \ v\<^sub>e") assume "v\<^sub>e \ other.s' \" from vo_start_geq_ve[OF assms(1) this] have "u \ < other.s \" by auto thus ?thesis unfolding distance0_def u_def using pos_react assms(1) unfolding ego.q_def other.s_def other.p_def by auto next assume "\ v\<^sub>e \ other.s' \" hence "v\<^sub>e > other.s' \" by auto from distance0_at_most_distance2r[OF assms(1) this assms(2)] have "distance0 \ safe_distance_2r" by auto with assms(3) show ?thesis by auto qed theorem sd2r_eq: assumes "\ > other.t_stop" shows "(u_max < other.s \) = (s\<^sub>o - s\<^sub>e > safe_distance_2r)" proof - from assms have "(u_max < other.s \) = (ego2.s (- v\<^sub>e / a\<^sub>e) < other.p_max)" using u_max_def ego2.t_stop_def u_def other.s_def \_def pos_react ego2.p_max_eq ego2.s_t_stop u_max_eq by auto also have "... = (s\<^sub>e + v\<^sub>e * \ + v\<^sub>e * (- v\<^sub>e / a\<^sub>e) + 1 / 2 * a\<^sub>e * (- v\<^sub>e / a\<^sub>e)\<^sup>2 < s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2)" using ego2.s_def ego2.p_def ego.q_def other.p_max_def other.p_def other.t_stop_def ego2.p_max_def ego2.s_t_stop ego2.t_stop_def by auto also have "... = (v\<^sub>e * \ + v\<^sub>e * (- v\<^sub>e / a\<^sub>e) + 1 / 2 * a\<^sub>e * (- v\<^sub>e / a\<^sub>e)\<^sup>2 - v\<^sub>o * (- v\<^sub>o / a\<^sub>o) - 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2 < s\<^sub>o - s\<^sub>e)" by linarith also have "... = (v\<^sub>e * \ - v\<^sub>e\<^sup>2 / a\<^sub>e + 1 / 2 * v\<^sub>e\<^sup>2 / a\<^sub>e + v\<^sub>o\<^sup>2 / a\<^sub>o - 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)" using ego2.p_def ego2.p_max_def ego2.p_max_eq ego2.t_stop_def other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto also have "... = (v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)" by linarith thus ?thesis using distance0_2_def by (simp add: calculation safe_distance_2r_def) qed theorem dist0_sd2r_2: assumes "\ > - v\<^sub>o / a\<^sub>o" assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r" shows "s\<^sub>o - s\<^sub>e > distance0_2" proof - have "- v\<^sub>e\<^sup>2 / 2 / a\<^sub>e \ 0" using zero_le_power2 hyps(3) divide_nonneg_neg by (auto simp add:field_simps) hence "v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o \ v\<^sub>e * \ + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" by simp hence "safe_distance_2r \ distance0_2" using safe_distance_2r_def distance0_2_def by auto thus ?thesis using assms(2) by linarith qed end subsection \Safe Distance Delta\ locale safe_distance_no_collsion_delta = safe_distance_normal + assumes no_collision_delta: "u \ < other.s \" begin sublocale delayed_safe_distance: safe_distance a\<^sub>e v\<^sub>e "ego.q \" a\<^sub>o "other.s' \" "other.s \" proof (unfold_locales) from nonneg_vel_ego show "0 \ v\<^sub>e" by auto next from nonneg_vel_other show "0 \ other.s' \" unfolding other.s'_def other.p'_def other.t_stop_def using decelerate_other by (auto simp add: field_split_simps) next from decelerate_ego show "a\<^sub>e < 0" by auto next from decelerate_other show "a\<^sub>o < 0" by auto next from no_collision_delta show "ego.q \ < other.s \" unfolding u_def using pos_react by auto qed lemma no_collision_react_initially_strict: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" shows "no_collision_react {0 <..< \}" proof (rule no_collision_reactI) fix t::real assume "t \ {0 <..< \}" show "u t \ other.s t" proof (rule ccontr) assume "\ u t \ other.s t" hence "ego_other2 t = 0" unfolding ego_other2_def by auto from \t \ {0 <..< \}\ have "ego_other2 t = other.s t - ego.q t" unfolding ego_other2_def u_def using ego.init_q by auto have "\ \ other.t_stop \ other.t_stop < \" by auto moreover { assume le_t_stop: "\ \ other.t_stop" with \ego_other2 t = other.s t - ego.q t\ have "ego_other2 t = other.p t - ego.q t" unfolding other.s_def using \t \ {0 <..< \}\ by auto with \ego_other2 t = 0\ have "other.p t - ego.q t = 0" by auto hence eq: "(s\<^sub>o- s\<^sub>e) + (v\<^sub>o - v\<^sub>e) * t + (1/2 * a\<^sub>o) * t\<^sup>2 = 0" unfolding other.p_def ego.q_def by (auto simp: algebra_simps) define p where "p \ \x. (1/2 * a\<^sub>o) * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e)" have "0 < 1/2 * a\<^sub>o" proof (intro p_convex[where p=p and b="v\<^sub>o - v\<^sub>e" and c="s\<^sub>o - s\<^sub>e"]) show "0 < t" using \t \ {0 <..< \}\ by auto next show "t < \" using \t \ {0 <..< \}\ by auto next show "p t < p 0" unfolding p_def using eq in_front by (auto simp: algebra_simps) next from eq have "p t = 0" unfolding p_def by auto also have "... < p \" using no_collision_delta pos_react le_t_stop unfolding p_def u_def other.s_def ego.q_def other.p_def by (auto simp:algebra_simps) finally have "p t < p \" by simp thus "p t \ p \" by auto next show "p = (\x. 1 / 2 * a\<^sub>o * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e))" unfolding p_def by (rule refl) qed hence "0 < a\<^sub>o" by auto with decelerate_other have False by simp } moreover { assume gt_t_stop: "\ > other.t_stop" have t_lt_t_stop: "t < other.t_stop" proof (rule ccontr) assume "\ t < other.t_stop" hence "other.t_stop \ t" by simp from \ego_other2 t = 0\ have "ego.q t = other.p_max" unfolding ego_other2_def u_def other.s_def comp_def \_def other.p_max_def using \t \ {0 <..< \}\ \other.t_stop \ t\ gt_t_stop by (auto split:if_splits) have "ego.q t = u t" unfolding u_def using \t \ {0 <..< \}\ by auto also have "... \ u_max" using u_max by auto also have "... < other.p_max" using assms(2) other.s_t_stop by auto finally have "ego.q t < other.p_max" by auto with \ego.q t = other.p_max\ show False by auto qed with \ego_other2 t = other.s t - ego.q t\ have "ego_other2 t = other.p t - ego.q t" unfolding other.s_def using \t \ {0 <..< \}\ by auto with \ego_other2 t = 0\ have "other.p t - ego.q t = 0" by auto hence eq: "(s\<^sub>o- s\<^sub>e) + (v\<^sub>o - v\<^sub>e) * t + (1/2 * a\<^sub>o) * t\<^sup>2 = 0" unfolding other.p_def ego.q_def by (auto simp: algebra_simps) define p where "p \ \x. (1/2 * a\<^sub>o) * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e)" have "0 < 1/2 * a\<^sub>o" proof (intro p_convex[where p=p and b="v\<^sub>o - v\<^sub>e" and c="s\<^sub>o - s\<^sub>e"]) show "0 < t" using \t \ {0 <..< \}\ by auto next show "t < other.t_stop" using t_lt_t_stop by auto next show "p t < p 0" unfolding p_def using eq in_front by (auto simp: algebra_simps) next from eq have zero: "p t = 0" unfolding p_def by auto have eq: "p other.t_stop = ego_other2 other.t_stop" unfolding ego_other2_def other.s_t_stop u_def ego.q_def other.s_def other.p_def p_def using \\ > other.t_stop\ other.t_stop_nonneg other.t_stop_def by (auto simp: diff_divide_distrib left_diff_distrib') have "u other.t_stop \ u_max" using u_max by auto also have "... < other.s_stop" using assms by auto finally have "0 \ other.s_stop - u other.t_stop" by auto hence "0 \ ego_other2 other.t_stop" unfolding ego_other2_def by auto hence "0 \ p other.t_stop" using eq by auto with zero show "p t \ p other.t_stop" by auto next show "p = (\x. 1 / 2 * a\<^sub>o * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e))" unfolding p_def by (rule refl) qed hence False using decelerate_other by auto } ultimately show False by auto qed qed lemma no_collision_react_initially: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" shows "no_collision_react {0 .. \}" proof - have "no_collision_react {0 <..< \}" by (rule no_collision_react_initially_strict[OF assms]) have "u 0 \ other.s 0" using init_u other.init_s in_front by auto hence "no_collision_react {0}" unfolding collision_react_def by auto with \no_collision_react {0 <..< \}\ have "no_collision_react ({0} \ {0 <..< \})" using no_collision_union[of "{0}" "{0 <..< \}"] by auto moreover have "{0} \ {0 <..< \} = {0 ..< \}" using pos_react by auto ultimately have "no_collision_react {0 ..< \}" by auto have "u \ \ other.s \" using no_collision_delta by auto hence "no_collision_react {\}" unfolding collision_react_def by auto with \no_collision_react {0 ..< \}\ have "no_collision_react ({\} \ {0 ..< \})" using no_collision_union[of "{\}" "{0 ..< \}"] by auto moreover have "{\} \ {0 ..< \} = {0 .. \}" using pos_react by auto ultimately show "no_collision_react {0 .. \}" by auto qed lemma collision_after_delta: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" shows "collision_react {0 ..} \ collision_react {\..}" proof assume "collision_react {0 ..}" have "no_collision_react {0 .. \}" by (rule no_collision_react_initially[OF assms]) with \collision_react {0..}\ have "collision_react ({0..} - {0 .. \})" using pos_react by (auto intro: collision_trim_subset) moreover have "{0..} - {0 .. \} = {\ <..}" using pos_react by auto ultimately have "collision_react {\ <..}" by auto thus "collision_react {\ ..}" by (auto intro:collision_react_subset) next assume "collision_react {\..}" moreover have "{\..} \ {0 ..}" using pos_react by auto ultimately show "collision_react {0 ..}" by (rule collision_react_subset) qed lemma collision_react_strict: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" shows "collision_react {\ ..} \ collision_react {\ <..}" proof assume asm: "collision_react {\ ..}" have "no_collision_react {\}" using no_collision_delta unfolding collision_react_def by auto moreover have "{\ <..} \ {\ ..}" by auto ultimately have "collision_react ({\ ..} - {\})" using asm collision_trim_subset by simp moreover have "{\ <..} = {\ ..} - {\}" by auto ultimately show "collision_react {\ <..}" by auto next assume "collision_react {\ <..}" thus "collision_react {\ ..}" using collision_react_subset[where t="{\ ..}" and s="{\ <..}"] by fastforce qed lemma delayed_other_s_stop_eq: "delayed_safe_distance.other.s_stop = other.s_stop" proof (unfold other.s_t_stop; unfold delayed_safe_distance.other.s_t_stop; unfold movement.p_max_eq) have "\ \ other.t_stop \ other.t_stop < \" by auto moreover { assume "\ \ other.t_stop" hence "other.s \ - (other.s' \)\<^sup>2 / a\<^sub>o / 2 = s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o / 2" unfolding other.s_def other.s'_def using pos_react decelerate_other by (auto simp add: other.p_def other.p'_def power2_eq_square field_split_simps) } moreover { assume "other.t_stop < \" hence "other.s \ - (other.s' \)\<^sup>2 / a\<^sub>o / 2 = s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o / 2" unfolding other.s_def other.s'_def other.p_max_eq using pos_react decelerate_other by (auto) } ultimately show "other.s \ - (other.s' \)\<^sup>2 / a\<^sub>o / 2 = s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o / 2" by auto qed lemma delayed_cond3': assumes "other.s \ \ u_max" assumes "u_max < other.s_stop" shows "delayed_safe_distance.collision {0 ..} \ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ other.s \ - ego.q \ \ delayed_safe_distance.snd_safe_distance \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" proof (rule delayed_safe_distance.cond_3') have "other.s \ \ u_max" using \other.s \ \ u_max\ . also have "... = ego2.s_stop" unfolding u_max_eq ego2.s_t_stop ego2.p_max_eq by (rule refl) finally show "other.s \ \ ego2.s_stop" by auto next have "ego2.s_stop = u_max" unfolding ego2.s_t_stop ego2.p_max_eq u_max_eq by (rule refl) also have "... < other.s_stop" using assms by auto also have "... \ delayed_safe_distance.other.s_stop" using delayed_other_s_stop_eq by auto finally show "ego2.s_stop < delayed_safe_distance.other.s_stop" by auto qed lemma delayed_other_t_stop_eq: assumes "\ \ other.t_stop" shows "delayed_safe_distance.other.t_stop + \ = other.t_stop" using assms decelerate_other unfolding delayed_safe_distance.other.t_stop_def other.t_stop_def other.s'_def movement.t_stop_def other.p'_def by (auto simp add: field_split_simps) lemma delayed_other_s_eq: assumes "0 \ t" shows "delayed_safe_distance.other.s t = other.s (t + \)" proof (cases "\ \ other.t_stop") assume 1: "\ \ other.t_stop" have "t + \ \ other.t_stop \ other.t_stop < t + \" by auto moreover { assume "t + \ \ other.t_stop" hence "delayed_safe_distance.other.s t = delayed_safe_distance.other.p t" using delayed_other_t_stop_eq [OF 1] assms unfolding delayed_safe_distance.other.s_def by auto also have "... = other.p (t + \)" unfolding movement.p_def other.s_def other.s'_def other.p'_def using pos_react 1 by (auto simp add: power2_eq_square field_split_simps) also have "... = other.s (t + \)" unfolding other.s_def using assms pos_react \t + \ \ other.t_stop\ by auto finally have "delayed_safe_distance.other.s t = other.s (t + \)" by auto } moreover { assume "other.t_stop < t + \" hence "delayed_safe_distance.other.s t = delayed_safe_distance.other.p_max" using delayed_other_t_stop_eq [OF 1] assms delayed_safe_distance.other.t_stop_nonneg unfolding delayed_safe_distance.other.s_def by auto also have "... = other.p_max" unfolding movement.p_max_eq other.s_def other.s'_def other.p_def other.p'_def using pos_react 1 decelerate_other by (auto simp add: power2_eq_square field_split_simps) also have "... = other.s (t + \)" unfolding other.s_def using assms pos_react \other.t_stop < t + \\ by auto finally have "delayed_safe_distance.other.s t = other.s (t + \)" by auto } ultimately show ?thesis by auto next assume "\ \ \ other.t_stop" hence "other.t_stop < \" by auto hence "other.s' \ = 0" and "other.s \ = other.p_max" unfolding other.s'_def other.s_def using pos_react by auto hence "delayed_safe_distance.other.s t = delayed_safe_distance.other.p_max" unfolding delayed_safe_distance.other.s_def using assms decelerate_other by (auto simp add:movement.p_max_eq movement.p_def movement.t_stop_def) also have "... = other.p_max" unfolding movement.p_max_eq using \other.s' \ = 0\ \other.s \ = other.p_max\ using other.p_max_eq by auto also have "... = other.s (t + \)" unfolding other.s_def using pos_react assms \other.t_stop < \\ by auto finally show "delayed_safe_distance.other.s t = other.s (t + \)" by auto qed lemma translate_collision_range: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" shows "delayed_safe_distance.collision {0 ..} \ collision_react {\ ..}" proof assume "delayed_safe_distance.collision {0 ..}" then obtain t where eq: "ego2.s t = delayed_safe_distance.other.s t" and "0 \ t" unfolding delayed_safe_distance.collision_def by auto have "ego2.s t = (ego2.s \ \) (t + \)" unfolding comp_def \_def by auto also have "... = u (t + \)" unfolding u_def using \0 \ t\ pos_react by (auto simp: \_def ego2.init_s) finally have left:"ego2.s t = u (t + \)" by auto have right: "delayed_safe_distance.other.s t = other.s (t + \)" using delayed_other_s_eq pos_react \0 \ t\ by auto with eq and left have "u (t + \) = other.s (t + \)" by auto moreover have "\ \ t + \" using \0 \ t\ by auto ultimately show "collision_react {\ ..}" unfolding collision_react_def by auto next assume "collision_react {\ ..}" hence "collision_react {\ <..}" using collision_react_strict[OF assms] by simp then obtain t where eq: "u t = other.s t" and "\ < t" unfolding collision_react_def by auto moreover hence "u t = (ego2.s \ \) t" unfolding u_def using pos_react by auto moreover have "other.s t = delayed_safe_distance.other.s (t - \)" using delayed_other_s_eq \\ < t\ by auto ultimately have "ego2.s (t - \) = delayed_safe_distance.other.s (t - \)" unfolding comp_def \_def by auto with \\ < t\ show "delayed_safe_distance.collision {0 ..}" unfolding delayed_safe_distance.collision_def by auto qed theorem cond_3r_2: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" assumes "other.s \ \ u_max" shows "collision_react {0 ..} \ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ other.s \ - ego.q \ \ delayed_safe_distance.snd_safe_distance \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" proof - have "collision_react {0 ..} \ collision_react {\ ..}" by (rule collision_after_delta[OF assms(1) assms(2)]) also have "... \ delayed_safe_distance.collision {0 ..}" by (simp add: translate_collision_range[OF assms(1) assms(2)]) also have "... \ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ other.s \ - ego.q \ \ delayed_safe_distance.snd_safe_distance \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" by (rule delayed_cond3'[OF assms(3) assms(2)]) finally show "collision_react {0 ..} \ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ other.s \ - ego.q \ \ delayed_safe_distance.snd_safe_distance \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" by auto qed lemma sd_2r_correct_for_3r_2: assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r" assumes "other.s \ \ u_max" assumes "\ (a\<^sub>o > a\<^sub>e \ other.s' \ < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \ < 0)" shows "no_collision_react {0..}" proof - from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" unfolding safe_distance_2r_def by auto hence "s\<^sub>o - v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto hence "s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto hence "s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1/2 * a\<^sub>o * (-v\<^sub>o / a\<^sub>o)\<^sup>2 > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" using other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto hence "other.s_stop > u_max" unfolding other.s_def using u_max_eq other.t_stop_def using ego.q_def other.p_def other.p_max_def other.s_def other.s_t_stop by auto thus ?thesis using assms(2) assms(3) collision_after_delta cond_1r delayed_cond3' translate_collision_range by linarith qed lemma sd2_at_most_sd4: assumes "a\<^sub>o > a\<^sub>e" shows "safe_distance_2r \ safe_distance_4r" proof - have "a\<^sub>o \ 0" and "a\<^sub>e \ 0" and "a\<^sub>o - a\<^sub>e \ 0" and "0 < 2 * (a\<^sub>o - a\<^sub>e)" using hyps assms(1) by auto have "0 \ (- v\<^sub>e * a\<^sub>o + v\<^sub>o * a\<^sub>e + a\<^sub>o * a\<^sub>e * \) * (- v\<^sub>e * a\<^sub>o + v\<^sub>o * a\<^sub>e + a\<^sub>o * a\<^sub>e * \)" (is "0 \ (?l1 + ?l2 + ?l3) * ?r") by auto also have "... = v\<^sub>e\<^sup>2 * a\<^sub>o\<^sup>2 + v\<^sub>o\<^sup>2 * a\<^sub>e\<^sup>2 + a\<^sub>o\<^sup>2 * a\<^sub>e\<^sup>2 * \\<^sup>2 - 2 * v\<^sub>e * a\<^sub>o * v\<^sub>o * a\<^sub>e - 2 * a\<^sub>o\<^sup>2 * a\<^sub>e * \ * v\<^sub>e + 2 * a\<^sub>o * a\<^sub>e\<^sup>2 * \ * v\<^sub>o" (is "?lhs = ?rhs") by (auto simp add:algebra_simps power_def) finally have "0 \ ?rhs" by auto hence "(- v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e - v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o) * (a\<^sub>o * a\<^sub>e) \ (a\<^sub>o * a\<^sub>e * \\<^sup>2 - 2 * v\<^sub>e * v\<^sub>o - 2 * a\<^sub>o * \ * v\<^sub>e + 2 * a\<^sub>e * \ * v\<^sub>o) * (a\<^sub>o * a\<^sub>e)" by (auto simp add: algebra_simps power_def) hence "2 * v\<^sub>e * \ * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e + v\<^sub>e\<^sup>2 + v\<^sub>o\<^sup>2 - v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o \ v\<^sub>o\<^sup>2 + a\<^sub>o\<^sup>2 * \\<^sup>2 + v\<^sub>e\<^sup>2 + 2 * v\<^sub>o * \ * a\<^sub>o - 2 * v\<^sub>e * v\<^sub>o - 2 * a\<^sub>o * \ * v\<^sub>e - 2 * v\<^sub>o * \ * a\<^sub>o + 2 * a\<^sub>e * \ * v\<^sub>o - a\<^sub>o\<^sup>2 * \\<^sup>2 + a\<^sub>o * a\<^sub>e * \\<^sup>2 + 2 * v\<^sub>e * \ * (a\<^sub>o - a\<^sub>e)" by (auto simp add: ego2.decel other.decel) hence "2 * v\<^sub>e * \ * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e + v\<^sub>e\<^sup>2 + v\<^sub>o\<^sup>2 - v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o \ (v\<^sub>o + \ * a\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * v\<^sub>o * \ * a\<^sub>o + 2 * a\<^sub>e * \ * v\<^sub>o - a\<^sub>o\<^sup>2 * \\<^sup>2 + a\<^sub>o * a\<^sub>e * \\<^sup>2 + 2 * v\<^sub>e * \ * (a\<^sub>o - a\<^sub>e)" by (auto simp add: algebra_simps power_def) hence "v\<^sub>e * \ * 2 * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e * 2 * a\<^sub>o + v\<^sub>e\<^sup>2 / 2 / a\<^sub>e * 2 * a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o * 2 * a\<^sub>o - v\<^sub>o\<^sup>2 / 2 / a\<^sub>o * 2 * a\<^sub>e \ (v\<^sub>o + \ * a\<^sub>o - v\<^sub>e)\<^sup>2 - v\<^sub>o * \ * 2 * a\<^sub>o - v\<^sub>o * \ * 2 * -a\<^sub>e - a\<^sub>o * \\<^sup>2 / 2 * 2 * a\<^sub>o - a\<^sub>o * \\<^sup>2 / 2 * 2 * -a\<^sub>e + v\<^sub>e * \ * 2 * (a\<^sub>o - a\<^sub>e)" (is "?lhs1 \ ?rhs1") by (simp add: \a\<^sub>o \ 0\ \a\<^sub>e \ 0\ power2_eq_square algebra_simps) hence "v\<^sub>e * \ * 2 * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e * 2 * (a\<^sub>o - a\<^sub>e) + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o * 2 * (a\<^sub>o - a\<^sub>e) \ (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) * 2 * (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ * 2 * (a\<^sub>o - a\<^sub>e) - a\<^sub>o * \\<^sup>2 / 2 * 2 * (a\<^sub>o - a\<^sub>e) + v\<^sub>e * \ * 2 *(a\<^sub>o - a\<^sub>e)" (is "?lhs2 \ ?rhs2") proof - assume "?lhs1 \ ?rhs1" have "?lhs1 = ?lhs2" by (auto simp add:field_simps) moreover have "?rhs1 = ?rhs2" using \a\<^sub>o - a\<^sub>e \ 0\ by (auto simp add:field_simps) ultimately show ?thesis using \?lhs1 \ ?rhs1\ by auto qed hence "(v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o) * 2 * (a\<^sub>o - a\<^sub>e) \ ((v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 + v\<^sub>e * \) * 2 *(a\<^sub>o - a\<^sub>e)" by (simp add: algebra_simps) hence "v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o \ (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 + v\<^sub>e * \" using \a\<^sub>o > a\<^sub>e\ mult_le_cancel_iff1[OF \0 < 2 * (a\<^sub>o - a\<^sub>e)\, of "(v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o)" "(v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 + v\<^sub>e * \"] semiring_normalization_rules(18) by (metis (no_types, lifting)) thus ?thesis using safe_distance_2r_def safe_distance_4r_def by auto qed lemma sd_4r_correct: assumes "s\<^sub>o - s\<^sub>e > safe_distance_4r" assumes "other.s \ \ u_max" assumes "\ \ other.t_stop" assumes "a\<^sub>o > a\<^sub>e" shows "no_collision_react {0..}" proof - from assms have "s\<^sub>o - s\<^sub>e > (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 + v\<^sub>e * \" unfolding safe_distance_4r_def by auto hence "s\<^sub>o + v\<^sub>o * \ + 1/2 * a\<^sub>o * \\<^sup>2 - s\<^sub>e - v\<^sub>e * \ > (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)" by linarith hence "other.s \ - ego.q \ > (other.s' \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)" using assms(3) ego.q_def other.p_def other.s_def other.p'_def other.s'_def pos_react by auto hence "other.s \ - ego.q \ > delayed_safe_distance.snd_safe_distance" by (simp add: delayed_safe_distance.snd_safe_distance_def) hence c: "\ (other.s \ - ego.q \ \ delayed_safe_distance.snd_safe_distance)" by linarith have "u_max < other.s_stop" unfolding u_max_eq other.s_t_stop other.p_max_eq ego.q_def using assms(1) sd2_at_most_sd4[OF assms(4)] unfolding safe_distance_4r_def safe_distance_2r_def by auto consider "s\<^sub>o \ u_max" | "s\<^sub>o > u_max" by linarith thus ?thesis proof (cases) case 1 from cond_3r_2[OF this \u_max < other.s_stop\ assms(2)] show ?thesis using c by auto next case 2 then show ?thesis using cond_1r by auto qed qed text \Irrelevant, this Safe Distance is unreachable in the checker.\ definition safe_distance_5r :: real where "safe_distance_5r = v\<^sub>e\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o + v\<^sub>e * \" lemma sd_5r_correct: assumes "s\<^sub>o - s\<^sub>e > safe_distance_5r" assumes "u_max < other.s_stop" assumes "other.s \ \ u_max" assumes "\ > other.t_stop" shows "no_collision_react {0..}" proof - from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o + v\<^sub>e * \" unfolding safe_distance_5r_def by auto hence "s\<^sub>o + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o - s\<^sub>e - v\<^sub>e * \ > (0 - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)" using assms(2-4) unfolding other.s_def other.s_t_stop apply (auto simp: movement.p_t_stop split: if_splits) using cond_1r cond_2r other.s_t_stop by linarith+ hence "other.s \ - ego.q \ > (other.s' \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)" using assms(2) assms(3) assms(4) other.s_def other.s_t_stop by auto hence "other.s \ - ego.q \ > delayed_safe_distance.snd_safe_distance" by (simp add: delayed_safe_distance.snd_safe_distance_def) hence "\ (other.s \ - ego.q \ \ delayed_safe_distance.snd_safe_distance)" by linarith thus ?thesis using assms(2) assms(3) cond_1r cond_3r_2 by linarith qed lemma translate_no_collision_range: "delayed_safe_distance.no_collision {0 ..} \ no_collision_react {\ ..}" proof assume left: "delayed_safe_distance.no_collision {0 ..}" show "no_collision_react {\ ..}" proof (unfold collision_react_def; simp; rule ballI) fix t::real assume "t \ {\ ..}" hence "\ \ t" by simp with pos_react have "0 \ t - \" by simp with left have ineq: "ego2.s (t - \) \ delayed_safe_distance.other.s (t - \)" unfolding delayed_safe_distance.collision_def by auto have "ego2.s (t - \) = (ego2.s \ \) t" unfolding comp_def \_def by auto also have "... = u t" unfolding u_def using \\ \ t\ pos_react by (auto simp: \_def ego2.init_s) finally have "ego2.s (t - \) = u t" by auto moreover have "delayed_safe_distance.other.s (t - \) = other.s t" using delayed_other_s_eq pos_react \\ \ t\ by auto ultimately show "u t \ other.s t" using ineq by auto qed next assume right: "no_collision_react {\ ..}" show "delayed_safe_distance.no_collision {0 ..}" proof (unfold delayed_safe_distance.collision_def; simp; rule ballI) fix t ::real assume "t \ {0 ..}" hence "0 \ t" by auto hence "\ \ t + \" by auto with right have ineq: "u (t + \) \ other.s (t + \)" unfolding collision_react_def by auto have "u (t + \) = ego2.s t" unfolding u_def comp_def \_def using \0 \ t\ pos_react \\ \ t+ \\ by (auto simp add:ego2.init_s) moreover have "other.s (t + \) = delayed_safe_distance.other.s t" using delayed_other_s_eq[of t] using \0 \ t\ by auto ultimately show "ego2.s t \ delayed_safe_distance.other.s t" using ineq by auto qed qed lemma delayed_cond1: assumes "other.s \ > u_max" shows "delayed_safe_distance.no_collision {0 ..}" proof - have "ego2.s_stop = u_max" unfolding ego2.s_t_stop ego2.p_max_eq u_max_eq by auto also have "... < other.s \" using assms by simp finally have "ego2.s_stop < other.s \" by auto thus "delayed_safe_distance.no_collision {0 ..}" by (simp add: delayed_safe_distance.cond_1) qed theorem cond_3r_3: assumes "s\<^sub>o \ u_max" assumes "u_max < other.s_stop" assumes "other.s \ > u_max" shows "no_collision_react {0 ..}" proof - have eq: "{0 ..} = {0 .. \} \ {\ ..}" using pos_react by auto show ?thesis unfolding eq proof (intro no_collision_union) show "no_collision_react {0 .. \}" by (rule no_collision_react_initially[OF assms(1) assms(2)]) next have "delayed_safe_distance.no_collision {0 ..}" by (rule delayed_cond1[OF assms(3)]) with translate_no_collision_range show "no_collision_react {\ ..}" by auto qed qed lemma sd_2r_correct_for_3r_3: assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r" assumes "other.s \ > u_max" shows "no_collision_react {0..}" proof - from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" unfolding safe_distance_2r_def by auto hence "s\<^sub>o - v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto hence "s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto hence "s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1/2 * a\<^sub>o * (-v\<^sub>o / a\<^sub>o)\<^sup>2 > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" using other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto hence "other.s_stop > u_max" unfolding other.s_def using u_max_eq other.t_stop_def using ego.q_def other.p_def other.p_max_def other.s_def other.s_t_stop by auto thus ?thesis using assms(2) cond_1r cond_3r_3 by linarith qed lemma sd_3r_correct: assumes "s\<^sub>o - s\<^sub>e > safe_distance_3r" assumes "\ \ other.t_stop" shows "no_collision_react {0 ..}" proof - from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e - v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2" unfolding safe_distance_3r_def by auto hence "s\<^sub>o + v\<^sub>o * \ + 1/2 * a\<^sub>o * \\<^sup>2 > s\<^sub>e + v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto hence "other.s \ > u_max" using other.s_def u_max_eq assms(2) ego.q_def other.p_def pos_react by auto thus ?thesis using cond_1r cond_3r_3 delayed_other_s_stop_eq delayed_safe_distance.other.s0_le_s_stop by linarith qed lemma sd_2_at_least_sd_3: assumes "\ \ other.t_stop" shows "safe_distance_3r \ safe_distance_2r" proof - from assms have "\ = other.t_stop \ \ < other.t_stop" by auto then have "safe_distance_3r = safe_distance_2r \ safe_distance_3r > safe_distance_2r" proof (rule Meson.disj_forward) assume "\ = other.t_stop" hence "\ = - v\<^sub>o / a\<^sub>o" unfolding other.t_stop_def by auto hence "- v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 = - v\<^sub>o * other.t_stop - 1/2 * a\<^sub>o * other.t_stop\<^sup>2" by (simp add: movement.t_stop_def) thus "safe_distance_3r = safe_distance_2r" using other.p_def other.p_max_def other.p_max_eq safe_distance_2r_def safe_distance_3r_def by auto next assume "\ < other.t_stop" hence "\ < - v\<^sub>o / a\<^sub>o" unfolding other.t_stop_def by auto hence "0 < v\<^sub>o + a\<^sub>o * \" using other.decel other.p'_def other.p'_pos_iff by auto hence "0 < v\<^sub>o + 1/2 * a\<^sub>o * (\ + other.t_stop)" by (auto simp add:field_simps other.t_stop_def) hence "0 > v\<^sub>o * (\ - other.t_stop) + 1/2 * a\<^sub>o * (\ + other.t_stop) * (\ - other.t_stop)" using \\ < other.t_stop\ mult_less_cancel_left_neg[where c="\ - other.t_stop" and a ="v\<^sub>o + 1 / 2 * a\<^sub>o * (\ + other.t_stop)" and b="0"] by (auto simp add: field_simps) hence " (\ + other.t_stop) * (\ - other.t_stop) = (\\<^sup>2 - other.t_stop\<^sup>2)" by (simp add: power2_eq_square square_diff_square_factored) hence "0 > v\<^sub>o * (\ - other.t_stop) + 1/2 * a\<^sub>o * (\\<^sup>2 - other.t_stop\<^sup>2)" by (metis (no_types, opaque_lifting) \v\<^sub>o * (\ - other.t_stop) + 1 / 2 * a\<^sub>o * (\ + other.t_stop) * (\ - other.t_stop) < 0\ divide_divide_eq_left divide_divide_eq_right times_divide_eq_left) hence "0 > v\<^sub>o * \ - v\<^sub>o * other.t_stop + 1/2 * a\<^sub>o * \\<^sup>2 - 1/2 * a\<^sub>o * other.t_stop\<^sup>2 " by (simp add: right_diff_distrib) hence "- v\<^sub>o * \ - 1/2 * a\<^sub>o * \\<^sup>2 > - v\<^sub>o * (- v\<^sub>o / a\<^sub>o) - 1/2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2" unfolding movement.t_stop_def by argo thus "safe_distance_3r > safe_distance_2r" using other.p_def other.p_max_def other.p_max_eq other.t_stop_def safe_distance_2r_def safe_distance_3r_def by auto qed thus ?thesis by auto qed end subsection \Checker Design\ text \ We define two checkers for different cases: \<^item> one checker for the case that \\ \ other.t_stop (other.t_stop = - v\<^sub>o / a\<^sub>o)\ \<^item> a second checker for the case that \\ > other.t_stop\ \ definition check_precond_r1 :: "real \ real \ real \ real \ real \ real \ real \ bool" where "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ s\<^sub>o > s\<^sub>e \ 0 \ v\<^sub>e \ 0 \ v\<^sub>o \ a\<^sub>e < 0 \ a\<^sub>o < 0 \ 0 < \ \ \ \ - v\<^sub>o / a\<^sub>o" definition safe_distance0 :: "real \ real \ real \ real \ real" where "safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \ = v\<^sub>e * \ - v\<^sub>o * \ - a\<^sub>o * \\<^sup>2 / 2" definition safe_distance_1r :: "real \ real \ real \ real" where "safe_distance_1r a\<^sub>e v\<^sub>e \ = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" definition safe_distance_2r :: "real \ real \ real \ real \ real \ real" where "safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" definition safe_distance_4r :: "real \ real \ real \ real \ real \ real" where "safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = (v\<^sub>o + a\<^sub>o * \ - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \ - 1 / 2 * a\<^sub>o * \\<^sup>2 + v\<^sub>e * \" definition safe_distance_3r :: "real \ real \ real \ real \ real \ real" where "safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = v\<^sub>e * \ - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e - v\<^sub>o * \ - 1 / 2 * a\<^sub>o * \\<^sup>2" definition checker_r1 :: "real \ real \ real \ real \ real \ real \ real \ bool" where "checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ let distance = s\<^sub>o - s\<^sub>e; precond = check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \; vo_star = v\<^sub>o + a\<^sub>o * \; t_stop_o_star = - vo_star / a\<^sub>o; t_stop_e = - v\<^sub>e / a\<^sub>e; safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \; safe_dist1 = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \; safe_dist2 = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \; safe_dist3 = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ in if \ precond then False else if distance > safe_dist0 \ distance > safe_dist3 then True else if (a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star) then distance > safe_dist2 else distance > safe_dist1" theorem checker_r1_correctness: "(checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..})" proof assume asm: "checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" have pre: "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" proof (rule ccontr) assume "\ check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" with asm show "False" unfolding checker_r1_def Let_def by auto qed from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by (unfold_locales) (auto simp add: check_precond_r1_def) interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ rewrites "sdn.distance0 = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \" and "sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_4r = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_3r = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" proof - from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by auto next show "safe_distance_normal.distance0 v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \ " unfolding safe_distance_normal.distance0_def[OF sdn'] safe_distance0_def by auto next show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \ = safe_distance_1r a\<^sub>e v\<^sub>e \" unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto next show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto next show "safe_distance_normal.safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ " unfolding safe_distance_normal.safe_distance_4r_def[OF sdn'] safe_distance_4r_def by auto next show "safe_distance_normal.safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.safe_distance_3r_def[OF sdn'] safe_distance_3r_def by auto qed have "0 < \" and "\ \ - v\<^sub>o / a\<^sub>o" using pre unfolding check_precond_r1_def by auto define so_delta where "so_delta = s\<^sub>o + v\<^sub>o * \ + a\<^sub>o * \\<^sup>2 / 2" define q_e_delta where "q_e_delta \ s\<^sub>e + v\<^sub>e * \" define u_stop_e where "u_stop_e \ q_e_delta - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)" define vo_star where "vo_star = v\<^sub>o + a\<^sub>o * \" define t_stop_o_star where "t_stop_o_star \ - vo_star / a\<^sub>o" define t_stop_e where "t_stop_e = - v\<^sub>e / a\<^sub>e" define distance where "distance \ s\<^sub>o - s\<^sub>e" define distance0 where "distance0 = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist0 where "safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \" define safe_dist2 where "safe_dist2 \ safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist1 where "safe_dist1 \ safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist3 where "safe_dist3 = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" note abb = so_delta_def q_e_delta_def u_stop_e_def vo_star_def t_stop_o_star_def t_stop_e_def distance_def safe_dist2_def safe_dist1_def safe_dist0_def safe_dist3_def distance0_def consider "distance > safe_dist0" | "distance > safe_dist3" | "distance \ safe_dist0 \ distance \ safe_dist3" by linarith hence "sdn.no_collision_react {0..}" proof (cases) case 1 then show ?thesis using sdn.sd_1r_correct unfolding abb by auto next case 2 hence pre2: "distance > distance0" using sdn.distance0_at_most_sd3r unfolding abb by auto hence "sdn.u \ < sdn.other.s \" using pre unfolding sdn.u_def sdn.ego.q_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def by auto from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ by (unfold_locales) (auto simp add:check_precond_r1_def \sdn.u \ < sdn.other.s \\) show ?thesis using sdr.sd_3r_correct 2 pre unfolding check_precond_r1_def abb sdn.other.t_stop_def by auto next case 3 hence "distance \ safe_dist3" by auto hence "sdn.other.s \ \ sdn.u_max" using pre unfolding check_precond_r1_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def sdn.u_max_eq sdn.ego.q_def abb sdn.safe_distance_3r_def by auto have " (a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star) \ \ (a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star) " by auto moreover { assume cond: "(a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star)" with 3 pre have "distance > safe_dist2" using asm unfolding checker_r1_def Let_def abb by auto with sdn.distance0_at_most_sd4r have "distance > distance0" unfolding abb using cond by auto hence "sdn.u \ < sdn.other.s \" using pre unfolding sdn.u_def sdn.ego.q_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def by auto from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ by (unfold_locales) (auto simp add:check_precond_r1_def \sdn.u \ < sdn.other.s \\) from sdr.sd_4r_correct[OF _ \sdn.other.s \ \ sdn.u_max\] \distance > safe_dist2\ have ?thesis using pre cond unfolding check_precond_r1_def sdn.other.t_stop_def abb by auto } moreover { assume not_cond: "\ (a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star)" with 3 pre have "distance > safe_dist1" using asm unfolding checker_r1_def Let_def abb by auto with sdn.dist0_sd2r_1 have "distance > distance0" using pre not_cond unfolding check_precond_r1_def sdn.other.t_stop_def sdn.other.s'_def sdn.other.p'_def abb by (auto simp add:field_simps) hence "sdn.u \ < sdn.other.s \" using pre unfolding sdn.u_def sdn.ego.q_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def by auto from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ by (unfold_locales) (auto simp add:check_precond_r1_def \sdn.u \ < sdn.other.s \\) from sdr.sd_2r_correct_for_3r_2[OF _ \sdn.other.s \ \ sdn.u_max\] not_cond \distance > safe_dist1\ have ?thesis using pre unfolding abb sdn.other.s'_def check_precond_r1_def sdn.other.t_stop_def sdn.other.p'_def by (auto simp add:field_simps) } ultimately show ?thesis by auto qed with pre show " check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ sdn.no_collision_react {0..}" by auto next assume "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}" hence pre: "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" and as2: "safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}" by auto show "checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ " proof (rule ccontr) assume as1: "\ checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" from pre have "0 < \" and "\ \ - v\<^sub>o / a\<^sub>o" unfolding check_precond_r1_def by auto define so_delta where "so_delta = s\<^sub>o + v\<^sub>o * \ + a\<^sub>o * \\<^sup>2 / 2" define q_e_delta where "q_e_delta \ s\<^sub>e + v\<^sub>e * \" define u_stop_e where "u_stop_e \ q_e_delta - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)" define vo_star where "vo_star \ v\<^sub>o + a\<^sub>o * \" define t_stop_o_star where "t_stop_o_star \ - vo_star / a\<^sub>o" define t_stop_e where "t_stop_e \ - v\<^sub>e / a\<^sub>e" define distance where "distance \ s\<^sub>o - s\<^sub>e" define distance0 where "distance0 \ safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist0 where "safe_dist0 \ safe_distance_1r a\<^sub>e v\<^sub>e \" define safe_dist2 where "safe_dist2 \ safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist1 where "safe_dist1 \ safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist3 where "safe_dist3 \ safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" note abb = so_delta_def q_e_delta_def u_stop_e_def vo_star_def t_stop_o_star_def t_stop_e_def distance_def safe_dist2_def safe_dist1_def safe_dist0_def safe_dist3_def distance0_def from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by (unfold_locales) (auto simp add: check_precond_r1_def) interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ rewrites "sdn.distance0 = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \" and "sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_4r = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_3r = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" proof - from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by auto next show "safe_distance_normal.distance0 v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \ " unfolding safe_distance_normal.distance0_def[OF sdn'] safe_distance0_def by auto next show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \ = safe_distance_1r a\<^sub>e v\<^sub>e \" unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto next show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto next show "safe_distance_normal.safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ " unfolding safe_distance_normal.safe_distance_4r_def[OF sdn'] safe_distance_4r_def by auto next show "safe_distance_normal.safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.safe_distance_3r_def[OF sdn'] safe_distance_3r_def by auto qed have "\ distance > distance0 \ distance > distance0" by auto moreover { assume "\ distance > distance0" hence "distance \ distance0" by auto with sdn.cond_3r_1' have "sdn.collision_react {0..\}" using pre unfolding check_precond_r1_def abb sdn.other.t_stop_def by auto with sdn.collision_react_subset have "sdn.collision_react {0..}" by auto with as2 have "False" by auto } moreover { assume if2: "distance > distance0" have "\ (distance > safe_dist0 \ distance > safe_dist3)" proof (rule ccontr) assume "\ \ (safe_dist0 < distance \ safe_dist3 < distance)" hence "(safe_dist0 < distance \ safe_dist3 < distance)" by auto with as1 show "False" using pre if2 unfolding checker_r1_def Let_def abb by auto qed hence if31: "distance \ safe_dist0" and if32: "distance \ safe_dist3" by auto have "sdn.u \ < sdn.other.s \" using if2 pre unfolding sdn.u_def sdn.ego.q_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def by auto from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ by (unfold_locales) (auto simp add:check_precond_r1_def \sdn.u \ < sdn.other.s \\) have " s\<^sub>o \ sdn.u_max" using if31 unfolding sdn.u_max_eq sdn.ego.q_def abb sdn.safe_distance_1r_def by auto have "sdn.other.s \ \ sdn.u_max" using if32 pre unfolding sdn.other.s_def check_precond_r1_def sdn.other.t_stop_def sdn.other.p_def sdn.u_max_eq sdn.ego.q_def abb sdn.safe_distance_3r_def by auto consider "(a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star)" | "\ (a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star)" by auto hence "False" proof (cases) case 1 hence rest_conjunct:"(a\<^sub>e < a\<^sub>o \ sdn.other.s' \ < v\<^sub>e \ v\<^sub>e - a\<^sub>e / a\<^sub>o * sdn.other.s' \ < 0)" using pre unfolding check_precond_r1_def unfolding sdn.other.s'_def sdn.other.t_stop_def sdn.other.p'_def abb by (auto simp add:field_simps) from 1 have "distance \ safe_dist2" using as1 pre if2 if31 if32 unfolding checker_r1_def Let_def abb by auto hence cond_f: "sdn.other.s \ - sdn.ego.q \ \ sdr.delayed_safe_distance.snd_safe_distance" using pre unfolding check_precond_r1_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def sdn.ego.q_def sdr.delayed_safe_distance.snd_safe_distance_def using sdn.other.s'_def[of "\"] unfolding sdn.other.t_stop_def sdn.other.p'_def abb sdn.safe_distance_4r_def by auto have "distance > safe_dist1 \ distance \ safe_dist1" by auto moreover { assume "distance > safe_dist1" hence "sdn.u_max < sdn.other.s_stop" unfolding sdn.u_max_eq sdn.ego.q_def sdn.other.s_t_stop sdn.other.p_max_eq abb sdn.safe_distance_2r_def by (auto simp add:field_simps) from sdr.cond_3r_2[OF \s\<^sub>o \ sdn.u_max\ this \sdn.other.s \ \ sdn.u_max\] have "sdn.collision_react {0..}" using cond_f rest_conjunct by auto with as2 have "False" by auto } moreover { assume "distance \ safe_dist1" hence "sdn.u_max \ sdn.other.s_stop" unfolding sdn.u_max_eq sdn.ego.q_def sdn.other.s_t_stop sdn.other.p_max_eq abb sdn.safe_distance_2r_def by (auto simp add:field_simps) with sdn.cond_2r[OF this] have "sdn.collision_react {0..}" by auto with as2 have "False" by auto } ultimately show ?thesis by auto next case 2 hence "distance \ safe_dist1" using as1 pre if2 if31 if32 unfolding checker_r1_def Let_def abb by auto hence "sdn.u_max \ sdn.other.s_stop" unfolding sdn.u_max_eq sdn.ego.q_def sdn.other.s_t_stop sdn.other.p_max_eq abb sdn.safe_distance_2r_def by (auto simp add:field_simps) with sdn.cond_2r[OF this] have "sdn.collision_react {0..}" by auto with as2 show "False" by auto qed } ultimately show "False" by auto qed qed definition check_precond_r2 :: "real \ real \ real \ real \ real \ real \ real \ bool" where "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ s\<^sub>o > s\<^sub>e \ 0 \ v\<^sub>e \ 0 \ v\<^sub>o \ a\<^sub>e < 0 \ a\<^sub>o < 0 \ 0 < \ \ \ > - v\<^sub>o / a\<^sub>o" definition safe_distance0_2 :: "real \ real \ real \ real \ real" where "safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \ = v\<^sub>e * \ + 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o" definition checker_r2 :: "real \ real \ real \ real \ real \ real \ real \ bool" where "checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ let distance = s\<^sub>o - s\<^sub>e; precond = check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \; safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \; safe_dist1 = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ in if \ precond then False else if distance > safe_dist0 then True else distance > safe_dist1" theorem checker_r2_correctness: "(checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..})" proof assume asm: "checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" have pre: "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" proof (rule ccontr) assume "\ check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" with asm show "False" unfolding checker_r2_def Let_def by auto qed from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by (unfold_locales) (auto simp add: check_precond_r2_def) interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ rewrites "sdn.distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \" and "sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" proof - from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by auto next show "safe_distance_normal.distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.distance0_2_def[OF sdn'] safe_distance0_2_def by auto next show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \ = safe_distance_1r a\<^sub>e v\<^sub>e \" unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto next show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto qed have "0 < \" and "\ > - v\<^sub>o / a\<^sub>o" using pre unfolding check_precond_r2_def by auto define distance where "distance \ s\<^sub>o - s\<^sub>e" define distance0_2 where "distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist0 where "safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \" define safe_dist1 where "safe_dist1 \ safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" note abb = distance_def safe_dist1_def safe_dist0_def distance0_2_def consider "distance > safe_dist0" | "distance \ safe_dist0" by linarith hence "sdn.no_collision_react {0..}" proof (cases) case 1 then show ?thesis using sdn.sd_1r_correct unfolding abb by auto next case 2 hence "(s\<^sub>o \ sdn.u_max)" using distance_def safe_dist0_def sdn.sd_1r_eq by linarith with 2 pre have "distance > safe_dist1" using asm unfolding checker_r2_def Let_def abb by auto with sdn.dist0_sd2r_2 have "distance > distance0_2" using abb \- v\<^sub>o / a\<^sub>o < \\ by auto hence "sdn.u \ < sdn.other.s \" using abb sdn.distance0_2_eq \\ > - v\<^sub>o / a\<^sub>o\ sdn.other.t_stop_def by auto have "sdn.u_max < sdn.other.s \" using abb sdn.sd2r_eq \\ > - v\<^sub>o / a\<^sub>o\ sdn.other.t_stop_def \distance > safe_dist1\ by auto from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ by (unfold_locales) (auto simp add:check_precond_r2_def \sdn.u \ < sdn.other.s \\) from sdr.sd_2r_correct_for_3r_3[OF] \distance > safe_dist1\ \sdn.u \ < sdn.other.s \\ \sdn.u_max < sdn.other.s \\ show ?thesis using pre unfolding abb sdn.other.s'_def check_precond_r2_def sdn.other.t_stop_def sdn.other.p'_def by (auto simp add:field_simps) qed with pre show " check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ sdn.no_collision_react {0..}" by auto next assume "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}" hence pre: "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" and as2: "safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}" by auto show "checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" proof (rule ccontr) assume as1: "\ checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" from pre have "0 < \" and "\ > - v\<^sub>o / a\<^sub>o" unfolding check_precond_r2_def by auto define distance where "distance \ s\<^sub>o - s\<^sub>e" define distance0_2 where "distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \" define safe_dist0 where "safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \" define safe_dist1 where "safe_dist1 \ safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" note abb = distance_def safe_dist1_def safe_dist0_def distance0_2_def from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by (unfold_locales) (auto simp add: check_precond_r2_def) interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ rewrites "sdn.distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \" and "sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \" and "sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" proof - from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \" by auto next show "safe_distance_normal.distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.distance0_2_def[OF sdn'] safe_distance0_2_def by auto next show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \ = safe_distance_1r a\<^sub>e v\<^sub>e \" unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto next show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \" unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto qed have "\ distance > distance0_2 \ distance > distance0_2" by auto moreover { assume "\ distance > distance0_2" hence "distance \ distance0_2" by auto with sdn.cond_3r_1'_2 have "sdn.collision_react {0..\}" using pre unfolding check_precond_r2_def abb sdn.other.t_stop_def by auto with sdn.collision_react_subset have "sdn.collision_react {0..}" by auto with as2 have "False" by auto } moreover { assume if2: "distance > distance0_2" have "\ (distance > safe_dist0)" proof (rule ccontr) assume "\ \ (safe_dist0 < distance)" hence "(safe_dist0 < distance)" by auto with as1 show "False" using pre if2 unfolding checker_r2_def Let_def abb by auto qed hence if3: "distance \ safe_dist0" by auto with pre have "distance \ safe_dist1" using as1 unfolding checker_r2_def Let_def abb by auto have "sdn.u \ < sdn.other.s \" using abb if2 sdn.distance0_2_eq \\ > - v\<^sub>o / a\<^sub>o\ sdn.other.t_stop_def by auto from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ by (unfold_locales) (auto simp add:check_precond_r2_def \sdn.u \ < sdn.other.s \\) have "sdn.u_max \ sdn.other.s \" using abb sdn.sd2r_eq \\ > - v\<^sub>o / a\<^sub>o\ sdn.other.t_stop_def \distance \ safe_dist1\ by auto with \\ > - v\<^sub>o / a\<^sub>o\ have "sdn.u_max \ sdn.other.s_stop" using sdn.other.s_mono sdn.other.t_stop_nonneg sdn.other.p_t_stop sdn.other.p_zero sdn.other.t_stop_def apply (auto simp: sdn.other.s_def movement.t_stop_def split: if_splits) using sdn.other.p_zero by auto hence "sdn.collision_react {0..}" using sdn.cond_2r by auto with as2 have "False" by auto } ultimately show "False" by auto qed qed text \Combine the two checkers into one.\ definition check_precond_r :: "real \ real \ real \ real \ real \ real \ real \ bool" where "check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ s\<^sub>o > s\<^sub>e \ 0 \ v\<^sub>e \ 0 \ v\<^sub>o \ a\<^sub>e < 0 \ a\<^sub>o < 0 \ 0 < \" definition checker_r :: "real \ real \ real \ real \ real \ real \ real \ bool" where "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ let distance = s\<^sub>o - s\<^sub>e; precond = check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \; vo_star = v\<^sub>o + a\<^sub>o * \; t_stop_o_star = -vo_star / a\<^sub>o; t_stop_e = -v\<^sub>e / a\<^sub>e; t_stop_o = -v\<^sub>o / a\<^sub>o; safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \; safe_dist1 = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \; safe_dist2 = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \; safe_dist3 = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ in if \ precond then False else if distance > safe_dist0 then True else if \ \ t_stop_o \ distance > safe_dist3 then True else if \ \ t_stop_o \ (a\<^sub>o > a\<^sub>e \ vo_star < v\<^sub>e \ t_stop_e < t_stop_o_star) then distance > safe_dist2 else distance > safe_dist1" theorem checker_eq_1: "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ \ \ - v\<^sub>o / a\<^sub>o \ checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" proof - have "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ \ \ - v\<^sub>o / a\<^sub>o \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \ \ s\<^sub>o - s\<^sub>e > safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ \ (((a\<^sub>o > a\<^sub>e \ v\<^sub>o + a\<^sub>o * \ < v\<^sub>e \ - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o) \ s\<^sub>o - s\<^sub>e > safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \) \ (\ (a\<^sub>o > a\<^sub>e \ v\<^sub>o + a\<^sub>o * \ < v\<^sub>e \ - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o) \ s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \))) \ \ \ - v\<^sub>o / a\<^sub>o" using checker_r_def by metis also have "... \ check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \ \ s\<^sub>o - s\<^sub>e > safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \ \ (((a\<^sub>o > a\<^sub>e \ v\<^sub>o + a\<^sub>o * \ < v\<^sub>e \ - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o) \ s\<^sub>o - s\<^sub>e > safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \) \ (\ (a\<^sub>o > a\<^sub>e \ v\<^sub>o + a\<^sub>o * \ < v\<^sub>e \ - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o) \ s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \)))" by (auto simp add:check_precond_r_def check_precond_r1_def) also have "... \ checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" by (metis checker_r1_def) finally show ?thesis by auto qed theorem checker_eq_2: "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ \ > - v\<^sub>o / a\<^sub>o \ checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" proof - have "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ \ > - v\<^sub>o / a\<^sub>o \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ (\ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \ \ (\ \ - v\<^sub>o / a\<^sub>o \ s\<^sub>o - s\<^sub>e > safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \) \ (\ \ - v\<^sub>o / a\<^sub>o \ a\<^sub>o > a\<^sub>e \ v\<^sub>o + a\<^sub>o * \ < v\<^sub>e \ - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \) / a\<^sub>o \ s\<^sub>o - s\<^sub>e > safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \) \ s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \) \ \ > - v\<^sub>o / a\<^sub>o" unfolding checker_r_def Let_def if_splits by auto also have "... \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \ \ s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \) \ \ > - v\<^sub>o / a\<^sub>o" by (auto simp add:HOL.disjE) also have "... \ check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \ \ s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \)" by (auto simp add:check_precond_r_def check_precond_r2_def) also have "... \ checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" by (auto simp add:checker_r2_def Let_def if_splits) thus ?thesis using calculation by auto qed theorem checker_r_correctness: "(checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..})" proof - have "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ (checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ \ \ - v\<^sub>o / a\<^sub>o) \ (checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ \ > - v\<^sub>o / a\<^sub>o)" by auto also have "... \ checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \" using checker_eq_1 checker_eq_2 by auto also have "... \ (check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}) \ (check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..})" using checker_r1_correctness checker_r2_correctness by auto also have "... \ (\ \ - v\<^sub>o / a\<^sub>o \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}) \ (\ > - v\<^sub>o / a\<^sub>o \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..})" by (auto simp add:check_precond_r_def check_precond_r1_def check_precond_r2_def) also have "... \ check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \ \ safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \ {0..}" by auto finally show ?thesis by auto qed end \ No newline at end of file diff --git a/thys/Stirling_Formula/Gamma_Asymptotics.thy b/thys/Stirling_Formula/Gamma_Asymptotics.thy --- a/thys/Stirling_Formula/Gamma_Asymptotics.thy +++ b/thys/Stirling_Formula/Gamma_Asymptotics.thy @@ -1,1893 +1,1893 @@ (* File: Gamma_Asymptotics.thy Author: Manuel Eberl The complete asymptotics of the real and complex logarithmic Gamma functions. Also of the real Polygamma functions (could be extended to the complex ones fairly easily if needed). *) section \Complete asymptotics of the logarithmic Gamma function\ theory Gamma_Asymptotics imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" Bernoulli.Bernoulli_FPS Bernoulli.Periodic_Bernpoly Stirling_Formula begin subsection \Auxiliary Facts\ (* TODO: could be automated with Laurent series expansions in the future *) lemma stirling_limit_aux1: "((\y. Ln (1 + z * of_real y) / of_real y) \ z) (at_right 0)" for z :: complex proof (cases "z = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + z * of_real y)) has_vector_derivative 1 * z) (at 0)" by (rule has_vector_derivative_real_field) (auto intro!: derivative_eq_intros) then have "(\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \0\ 0" by (auto simp add: has_vector_derivative_def has_derivative_def netlimit_at scaleR_conv_of_real field_simps) then have "((\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \ 0) (at_right 0)" by (rule filterlim_mono[OF _ _ at_le]) simp_all also have "?this \ ((\y. Ln (1 + z * of_real y) / (of_real y) - z) \ 0) (at_right 0)" using eventually_at_right_less[of "0::real"] by (intro filterlim_cong refl) (auto elim!: eventually_mono simp: field_simps) finally show ?thesis by (simp only: LIM_zero_iff) qed lemma stirling_limit_aux2: "((\y. y * Ln (1 + z / of_real y)) \ z) at_top" for z :: complex using stirling_limit_aux1[of z] by (subst filterlim_at_top_to_right) (simp add: field_simps) lemma Union_atLeastAtMost: assumes "N > 0" shows "(\n\{0.. {0..real N}" thus "x \ (\n\{0.. 0" "x < real N" by simp_all hence "x \ real (nat \x\)" "x \ real (nat \x\ + 1)" by linarith+ moreover from x have "nat \x\ < N" by linarith ultimately have "\n\{0.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto subsection \Cones in the complex plane\ definition complex_cone :: "real \ real \ complex set" where "complex_cone a b = {z. \y\{a..b}. z = rcis (norm z) y}" abbreviation complex_cone' :: "real \ complex set" where "complex_cone' a \ complex_cone (-a) a" lemma zero_in_complex_cone [simp, intro]: "a \ b \ 0 \ complex_cone a b" by (auto simp: complex_cone_def) lemma complex_coneE: assumes "z \ complex_cone a b" obtains r \ where "r \ 0" "\ \ {a..b}" "z = rcis r \" proof - from assms obtain y where "y \ {a..b}" "z = rcis (norm z) y" unfolding complex_cone_def by auto thus ?thesis using that[of "norm z" y] by auto qed lemma arg_cis [simp]: assumes "x \ {-pi<..pi}" shows "Arg (cis x) = x" using assms by (intro cis_Arg_unique) auto lemma arg_mult_of_real_left [simp]: assumes "r > 0" shows "Arg (of_real r * z) = Arg z" proof (cases "z = 0") case False thus ?thesis using Arg_bounded[of z] assms by (intro cis_Arg_unique) (auto simp: sgn_mult sgn_of_real cis_Arg) qed auto lemma arg_mult_of_real_right [simp]: assumes "r > 0" shows "Arg (z * of_real r) = Arg z" by (subst mult.commute, subst arg_mult_of_real_left) (simp_all add: assms) lemma arg_rcis [simp]: assumes "x \ {-pi<..pi}" "r > 0" shows "Arg (rcis r x) = x" using assms by (simp add: rcis_def) lemma rcis_in_complex_cone [intro]: assumes "\ \ {a..b}" "r \ 0" shows "rcis r \ \ complex_cone a b" using assms by (auto simp: complex_cone_def) lemma arg_imp_in_complex_cone: assumes "Arg z \ {a..b}" shows "z \ complex_cone a b" proof - have "z = rcis (norm z) (Arg z)" by (simp add: rcis_cmod_Arg) also have "\ \ complex_cone a b" using assms by auto finally show ?thesis . qed lemma complex_cone_altdef: assumes "-pi < a" "a \ b" "b \ pi" shows "complex_cone a b = insert 0 {z. Arg z \ {a..b}}" proof (intro equalityI subsetI) fix z assume "z \ complex_cone a b" then obtain r \ where *: "r \ 0" "\ \ {a..b}" "z = rcis r \" by (auto elim: complex_coneE) have "Arg z \ {a..b}" if [simp]: "z \ 0" proof - have "r > 0" using that * by (subst (asm) *) auto hence "\ \ {a..b}" using *(1,2) assms by (auto simp: *(1)) moreover from assms *(2) have "\ \ {-pi<..pi}" by auto ultimately show ?thesis using *(3) \r > 0\ by (subst *) auto qed thus "z \ insert 0 {z. Arg z \ {a..b}}" by auto qed (use assms in \auto intro: arg_imp_in_complex_cone\) lemma nonneg_of_real_in_complex_cone [simp, intro]: assumes "x \ 0" "a \ 0" "0 \ b" shows "of_real x \ complex_cone a b" proof - from assms have "rcis x 0 \ complex_cone a b" by (intro rcis_in_complex_cone) auto thus ?thesis by simp qed lemma one_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ 1 \ complex_cone a b" using nonneg_of_real_in_complex_cone[of 1] by (simp del: nonneg_of_real_in_complex_cone) lemma of_nat_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ of_nat n \ complex_cone a b" using nonneg_of_real_in_complex_cone[of "real n"] by (simp del: nonneg_of_real_in_complex_cone) subsection \Another integral representation of the Beta function\ lemma complex_cone_inter_nonpos_Reals: assumes "-pi < a" "a \ b" "b < pi" shows "complex_cone a b \ \\<^sub>\\<^sub>0 = {0}" proof (safe elim!: nonpos_Reals_cases) fix x :: real assume "complex_of_real x \ complex_cone a b" "x \ 0" hence "\(x < 0)" using assms by (intro notI) (auto simp: complex_cone_altdef) with \x \ 0\ show "complex_of_real x = 0" by auto qed (use assms in auto) theorem assumes a: "a > 0" and b: "b > (0 :: real)" shows has_integral_Beta_real': "((\u. u powr (b - 1) / (1 + u) powr (a + b)) has_integral Beta a b) {0<..}" and Beta_conv_nn_integral: "Beta a b = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" proof - define I where "I = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" have "Gamma (a + b) > 0" "Beta a b > 0" using assms by (simp_all add: add_pos_pos Beta_def) from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add mult_ac) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \(density (distr lborel borel ((*) t)) (\x. ennreal \t\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(t::real). indicator {0<..} t * (\\<^sup>+u. indicator {0..} (u * t) * t powr a * (u * t) powr (b - 1) / exp (t + t * u) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_density nn_integral_distr algebra_simps powr_diff simp flip: ennreal_mult) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr a * (u * t) powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def zero_le_mult_iff algebra_simps) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: powr_add powr_diff indicator_def powr_mult field_simps) also have "\ = (\\<^sup>+(u::real). \\<^sup>+t. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (rule lborel_pair.Fubini') auto also have "\ = (\\<^sup>+(u::real). indicator {0..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \(density (distr lborel borel ((*) (1/(1+u)))) (\x. ennreal \1/(1+u)\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. ennreal (1 / (u + 1)) * ennreal (indicator {0<..} (t / (u + 1)) * (t / (1+u)) powr (a + b - 1) * u powr (b - 1) / exp t) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_distr nn_integral_density add_ac) also have "\ = (\\<^sup>+u. \\<^sup>+t. indicator ({0<..}\{0<..}) (u, t) * 1/(u+1) * (t / (u+1)) powr (a + b - 1) * u powr (b - 1) / exp t \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def field_simps divide_ennreal simp flip: ennreal_mult ennreal_mult') also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) * ennreal (indicator {0<..} t * t powr (a + b - 1) / exp t) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_diff powr_divide powr_minus divide_simps add_ac simp flip: ennreal_mult) also have "\ = I * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel)" by (simp add: nn_integral_cmult nn_integral_multc I_def) also have "(\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel) = ennreal (Gamma (a + b))" using assms by (subst Gamma_conv_nn_integral_real) (auto intro!: nn_integral_cong_AE[OF AE_I[of _ _ "{0}"]] simp: indicator_def split: if_splits split_of_bool_asm) finally have "ennreal (Gamma a * Gamma b) = I * ennreal (Gamma (a + b))" . hence "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = I * ennreal (Gamma (a + b)) / ennreal (Gamma (a + b))" by simp also have "\ = I" using \Gamma (a + b) > 0\ by (intro ennreal_mult_divide_eq) auto also have "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = ennreal (Gamma a * Gamma b / Gamma (a + b))" using assms by (intro divide_ennreal) auto also have "\ = ennreal (Beta a b)" by (simp add: Beta_def) finally show *: "ennreal (Beta a b) = I" . define f where "f = (\u. u powr (b - 1) / (1 + u) powr (a + b))" have "((\u. indicator {0<..} u * f u) has_integral Beta a b) UNIV" using * \Beta a b > 0\ by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: f_def measurable_completion nn_integral_completion I_def mult_ac) also have "(\u. indicator {0<..} u * f u) = (\u. if u \ {0<..} then f u else 0)" by (auto simp: fun_eq_iff) also have "(\ has_integral Beta a b) UNIV \ (f has_integral Beta a b) {0<..}" by (rule has_integral_restrict_UNIV) finally show \ by (simp add: f_def) qed lemma has_integral_Beta2: fixes a :: real assumes "a < -1/2" shows "((\x. (1 + x ^ 2) powr a) has_integral Beta (- a - 1 / 2) (1 / 2) / 2) {0<..}" proof - define f where "f = (\u. u powr (-1/2) / (1 + u) powr (-a))" define C where "C = Beta (-a-1/2) (1/2)" have I: "(f has_integral C) {0<..}" using has_integral_Beta_real'[of "-a-1/2" "1/2"] assms by (simp_all add: diff_divide_distrib f_def C_def) define g where "g = (\x. x ^ 2 :: real)" have bij: "bij_betw g {0<..} {0<..}" by (intro bij_betwI[of _ _ _ sqrt]) (auto simp: g_def) have "(f absolutely_integrable_on g ` {0<..} \ integral (g ` {0<..}) f = C)" using I bij by (simp add: bij_betw_def has_integral_iff absolutely_integrable_on_def f_def) also have "?this \ ((\x. \2 * x\ *\<^sub>R f (g x)) absolutely_integrable_on {0<..} \ integral {0<..} (\x. \2 * x\ *\<^sub>R f (g x)) = C)" using bij by (intro has_absolute_integral_change_of_variables_1' [symmetric]) (auto intro!: derivative_eq_intros simp: g_def bij_betw_def) finally have "((\x. \2 * x\ * f (g x)) has_integral C) {0<..}" by (simp add: absolutely_integrable_on_def f_def has_integral_iff) also have "?this \ ((\x::real. 2 * (1 + x\<^sup>2) powr a) has_integral C) {0<..}" by (intro has_integral_cong) (auto simp: f_def g_def powr_def exp_minus ln_realpow field_simps) finally have "((\x::real. 1/2 * (2 * (1 + x\<^sup>2) powr a)) has_integral 1/2 * C) {0<..}" by (intro has_integral_mult_right) thus ?thesis by (simp add: C_def) qed lemma has_integral_Beta3: fixes a b :: real assumes "a < -1/2" and "b > 0" shows "((\x. (b + x ^ 2) powr a) has_integral Beta (-a - 1/2) (1/2) / 2 * b powr (a + 1/2)) {0<..}" proof - define C where "C = Beta (- a - 1 / 2) (1 / 2) / 2" have int: "nn_integral lborel (\x. indicator {0<..} x * (1 + x ^ 2) powr a) = C" using nn_integral_has_integral_lebesgue[OF _ has_integral_Beta2[OF assms(1)]] by (auto simp: C_def) have "nn_integral lborel (\x. indicator {0<..} x * (b + x ^ 2) powr a) = (\\<^sup>+x. ennreal (indicat_real {0<..} (x * sqrt b) * (b + (x * sqrt b)\<^sup>2) powr a * sqrt b) \lborel)" using assms by (subst lborel_distr_mult'[of "sqrt b"]) (auto simp: nn_integral_density nn_integral_distr mult_ac simp flip: ennreal_mult) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * (b * (1 + x ^ 2)) powr a * sqrt b) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def field_simps zero_less_mult_iff) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * b powr (a + 1/2) * (1 + x ^ 2) powr a) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_half_sqrt powr_mult) also have "\ = b powr (a + 1/2) * (\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel)" using assms by (subst nn_integral_cmult [symmetric]) (simp_all add: mult_ac flip: ennreal_mult) also have "(\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel) = C" using int by simp also have "ennreal (b powr (a + 1/2)) * ennreal C = ennreal (C * b powr (a + 1/2))" using assms by (subst ennreal_mult) (auto simp: C_def mult_ac Beta_def) finally have *: "(\\<^sup>+ x. ennreal (indicat_real {0<..} x * (b + x\<^sup>2) powr a) \lborel) = \" . hence "((\x. indicator {0<..} x * (b + x^2) powr a) has_integral C * b powr (a + 1/2)) UNIV" using assms by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: C_def measurable_completion nn_integral_completion Beta_def) also have "(\x. indicator {0<..} x * (b + x^2) powr a) = (\x. if x \ {0<..} then (b + x^2) powr a else 0)" by (auto simp: fun_eq_iff) finally show ?thesis by (subst (asm) has_integral_restrict_UNIV) (auto simp: C_def) qed subsection \Asymptotics of the real $\log\Gamma$ function and its derivatives\ text \ This is the error term that occurs in the expansion of @{term ln_Gamma}. It can be shown to be of order $O(s^{-n})$. \ definition stirling_integral :: "nat \ 'a :: {real_normed_div_algebra, banach} \ 'a" where "stirling_integral n s = lim (\N. integral {0..N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" context fixes s :: complex assumes s: "s \ \\<^sub>\\<^sub>0" fixes approx :: "nat \ complex" defines "approx \ (\N. (\n = 1.. \\\ ln_Gamma s\\ (ln_Gamma (of_nat N) - ln (2 * pi / of_nat N) / 2 - of_nat N * ln (of_nat N) + of_nat N) - \ \\\ 0\\ s * (harm (N - 1) - ln (of_nat (N - 1)) - euler_mascheroni) + \ \\\ 0\\ s * (ln (of_nat N + s) - ln (of_nat (N - 1))) - \ \\\ 0\\ (1/2) * (ln (of_nat N + s) - ln (of_nat N)) + \ \\\ 0\\ of_nat N * (ln (of_nat N + s) - ln (of_nat N)) - \ \\\ s\\ (s - 1/2) * ln s - ln (2 * pi) / 2)" begin qualified lemma assumes N: "N > 0" shows integrable_pbernpoly_1: "(\x. of_real (-pbernpoly 1 x) / (of_real x + s)) integrable_on {0..real N}" and integral_pbernpoly_1_aux: "integral {0..real N} (\x. -of_real (pbernpoly 1 x) / (of_real x + s)) = approx N" and has_integral_pbernpoly_1: "((\x. pbernpoly 1 x /(x + s)) has_integral (\mx. -pbernpoly 1 x / (x + s)) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" for n proof (rule has_integral_spike) have "((\x. (of_nat n + 1/2 + s) * (1 / (of_real x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_real (real (n + 1)) + s) - ln (of_real (real n) + s)) - 1) {of_nat n..of_nat (n + 1)}" using s has_integral_const_real[of 1 "of_nat n" "of_nat (n + 1)"] by (intro has_integral_diff has_integral_mult_right fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field - simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps + simp: has_real_derivative_iff_has_vector_derivative [symmetric] field_simps complex_nonpos_Reals_iff) thus "((\x. (of_nat n + 1/2 + s) * (1 / (of_real x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" by simp show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" if "x \ {of_nat n..of_nat (n + 1)} - {of_nat (n + 1)}" for x proof - have x: "x \ real n" "x < real (n + 1)" using that by simp_all hence "floor x = int n" by linarith moreover from s x have "complex_of_real x \ -s" by (auto simp add: complex_eq_iff complex_nonpos_Reals_iff simp del: of_nat_Suc) ultimately show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" by (auto simp: pbernpoly_def bernpoly_def frac_def divide_simps add_eq_0_iff2) qed qed simp_all hence *: "\I. I\?A \ ((\x. -pbernpoly 1 x / (x + s)) has_integral (Inf I + 1/2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1) I" by (auto simp: add_ac) have "((\x. - pbernpoly 1 x / (x + s)) has_integral (\I\?A. (Inf I + 1 / 2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1)) (\n\{0..x. - pbernpoly 1 x / (x + s)) has_integral ?i) {0..real N}" by (subst has_integral_spike_set_eq) (use Union_atLeastAtMost assms in \auto simp: intro!: empty_imp_negligible\) hence "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" and integral: "integral {0..real N} (\x. - pbernpoly 1 x / (x + s)) = ?i" by (simp_all add: has_integral_iff) show "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" by fact note has_integral_neg[OF has_integral] also have "-?i = (\xx. of_real (pbernpoly 1 x) / (of_real x + s)) has_integral \) {0..real N}" by simp note integral also have "?i = (\nnnn=1.. = (\n=1..n=1..n=1..n=1.. - (\n = 1..n=1.. 0" for x by (auto simp: complex_nonpos_Reals_iff complex_eq_iff) hence "(\n=1..n=1.. = ln (fact (N - 1)) + (\n=1..n=1..n=1..n=1..n = 1..n = 1..x. -of_real (pbernpoly 1 x) / (of_real x + s)) = \" by simp qed lemma integrable_ln_Gamma_aux: shows "(\x. of_real (pbernpoly n x) / (of_real x + s) ^ n) integrable_on {0..real N}" proof (cases "n = 1") case True with s show ?thesis using integrable_neg[OF integrable_pbernpoly_1[of N]] by (cases "N = 0") (simp_all add: integrable_negligible) next case False from s have "of_real x + s \ 0" if "x \ 0" for x using that by (auto simp: complex_eq_iff add_eq_0_iff2 complex_nonpos_Reals_iff) with False s show ?thesis by (auto intro!: integrable_continuous_real continuous_intros) qed text \ This following proof is based on ``Rudiments of the theory of the gamma function'' by Bruce Berndt~\cite{berndt}. \ lemma tendsto_of_real_0_I: "(f \ 0) G \ ((\x. (of_real (f x))) \ (0 ::'a::real_normed_div_algebra)) G" using tendsto_of_real_iff by force qualified lemma integral_pbernpoly_1: "(\N. integral {0..real N} (\x. pbernpoly 1 x / (x + s))) \ -ln_Gamma s - s + (s - 1 / 2) * ln s + ln (2 * pi) / 2" proof - have neq: "s + of_real x \ 0" if "x \ 0" for x :: real using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) have "(approx \ ln_Gamma s - 0 - 0 + 0 - 0 + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" unfolding approx_def proof (intro tendsto_add tendsto_diff) from s have s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff elim!: nonpos_Ints_cases) have "(\n. \i=1.. ln_Gamma s + euler_mascheroni * s + ln s" (is "?f \ _") using ln_Gamma_series'_aux[OF s'] unfolding sums_def by (subst filterlim_sequentially_Suc [symmetric], subst (asm) sum.atLeast1_atMost_eq [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) thus "((\n. ?f n - (euler_mascheroni * s + ln s)) \ ln_Gamma s) at_top" by (auto intro: tendsto_eq_intros) next show "(\x. complex_of_real (ln_Gamma (real x) - ln (2 * pi / real x) / 2 - real x * ln (real x) + real x)) \ 0" proof (intro tendsto_of_real_0_I filterlim_compose[OF tendsto_sandwich filterlim_real_sequentially]) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 0) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(1), simp add: algebra_simps) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 1 / 12 * inverse x) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(2), simp add: field_simps) show "((\x::real. 1 / 12 * inverse x) \ 0) at_top" by (intro tendsto_mult_right_zero tendsto_inverse_0_at_top filterlim_ident) qed simp_all next have "(\x. s * of_real (harm (x - 1) - ln (real (x - 1)) - euler_mascheroni)) \ s * of_real (euler_mascheroni - euler_mascheroni)" by (subst filterlim_sequentially_Suc [symmetric], intro tendsto_intros) (insert euler_mascheroni_LIMSEQ, simp_all) also have "?this \ (\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "1::nat"]]) (auto simp: Ln_of_nat of_real_harm) finally show "(\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" . next have "((\x. ln (1 + (s + 1) / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real (x + 1) + s) - ln (complex_of_real x) = ln (1 + (s + 1) / of_real x)" if "x > 1" for x using that s using Ln_divide_of_real[of x "of_real (x + 1) + s", symmetric] neq[of "x+1"] by (simp add: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real (x + 1) + s) - ln (of_real x)) \ 0) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "((\n. ln (of_real (real n + 1) + s) - ln (of_real (real n))) \ 0) at_top" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) hence "((\n. ln (of_nat n + s) - ln (of_nat (n - 1))) \ 0) at_top" by (subst filterlim_sequentially_Suc [symmetric]) (simp add: add_ac) thus "(\x. s * (ln (of_nat x + s) - ln (of_nat (x - 1)))) \ 0" by (rule tendsto_mult_right_zero) next have "((\x. ln (1 + s / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real x + s) - ln (of_real x) = ln (1 + s / of_real x)" if "x > 0" for x using Ln_divide_of_real[of x "of_real x + s"] neq[of x] that by (auto simp: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real x + s) - ln (of_real x)) \ 0) at_top" using s by (intro filterlim_cong refl) (auto intro: eventually_mono [OF eventually_gt_at_top[of "1::real"]]) finally have "(\x. (1/2) * (ln (of_real (real x) + s) - ln (of_real (real x)))) \ 0" by (rule tendsto_mult_right_zero[OF filterlim_compose[OF _ filterlim_real_sequentially]]) thus "(\x. (1/2) * (ln (of_nat x + s) - ln (of_nat x))) \ 0" by simp next have "((\x. x * (ln (1 + s / of_real x))) \ s) at_top" (is ?P) by (rule stirling_limit_aux2) also have "ln (1 + s / of_real x) = ln (of_real x + s) - ln (of_real x)" if "x > 1" for x using that s Ln_divide_of_real [of x "of_real x + s", symmetric] neq[of x] by (auto simp: Ln_of_real field_simps) hence "?P \ ((\x. of_real x * (ln (of_real x + s) - ln (of_real x))) \ s) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "(\n. of_real (real n) * (ln (of_real (real n) + s) - ln (of_real (real n)))) \ s" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) thus "(\n. of_nat n * (ln (of_nat n + s) - ln (of_nat n))) \ s" by simp qed simp_all also have "?this \ ((\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) \ ln_Gamma s + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" using integral_pbernpoly_1_aux by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) also have "(\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) = (\N. -integral {0..real N} (\x. pbernpoly 1 x / (x + s)))" by (simp add: fun_eq_iff) finally show ?thesis by (simp add: tendsto_minus_cancel_left [symmetric] algebra_simps) qed qualified lemma pbernpoly_integral_conv_pbernpoly_integral_Suc: assumes "n \ 1" shows "integral {0..real N} (\x. pbernpoly n x / (x + s) ^ n) = of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * integral {0..real N} (\x. of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)" proof - note [derivative_intros] = has_field_derivative_pbernpoly_Suc' define I where "I = -of_real (pbernpoly (Suc n) (of_nat N)) / (of_nat (Suc n) * (of_nat N + s) ^ n) + of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n + integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)" have "((\x. (-of_nat n * inverse (of_real x + s) ^ Suc n) * (of_real (pbernpoly (Suc n) x) / (of_nat (Suc n)))) has_integral -I) {0..real N}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_mult]) fix x :: real assume x: "x \ {0<.. \" proof assume "x \ \" then obtain n where "x = of_int n" by (auto elim!: Ints_cases) with x have x': "x = of_nat (nat n)" by simp from x show False by (auto simp: x') qed hence "((\x. of_real (pbernpoly (Suc n) x / of_nat (Suc n))) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by (intro has_vector_derivative_of_real) (auto intro!: derivative_eq_intros) thus "((\x. of_real (pbernpoly (Suc n) x) / of_nat (Suc n)) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by simp from x s have "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "((\x. inverse (of_real x + s) ^ n) has_vector_derivative - of_nat n * inverse (of_real x + s) ^ Suc n) (at x)" using x s assms by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: divide_simps power_add [symmetric] simp del: power_Suc) next have "complex_of_real x + s \ 0" if "x \ 0" for x using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "continuous_on {0..real N} (\x. inverse (of_real x + s) ^ n)" "continuous_on {0..real N} (\x. complex_of_real (pbernpoly (Suc n) x) / of_nat (Suc n))" using assms s by (auto intro!: continuous_intros simp del: of_nat_Suc) next have "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral pbernpoly (Suc n) (of_nat N) / (of_nat (Suc n) * (of_nat N + s) ^ n) - of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n - -I) {0..real N}" using integrable_ln_Gamma_aux[of n N] assms by (auto simp: I_def has_integral_integral divide_simps) thus "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral inverse (of_real (real N) + s) ^ n * (of_real (pbernpoly (Suc n) (real N)) / of_nat (Suc n)) - inverse (of_real 0 + s) ^ n * (of_real (pbernpoly (Suc n) 0) / of_nat (Suc n)) - - I) {0..real N}" by (simp_all add: field_simps) qed simp_all also have "(\x. - of_nat n * inverse (of_real x + s) ^ Suc n * (of_real (pbernpoly (Suc n) x) / of_nat (Suc n))) = (\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))" by (simp add: divide_simps fun_eq_iff) finally have "((\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)) has_integral - I) {0..real N}" . from has_integral_neg[OF this] show ?thesis by (auto simp add: I_def has_integral_iff algebra_simps integral_mult_right [symmetric] simp del: power_Suc of_nat_Suc ) qed lemma pbernpoly_over_power_tendsto_0: assumes "n > 0" shows "(\x. of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ 0" proof - from s have neq: "s + of_nat n \ 0" for n by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) obtain c where c: "\x. norm (pbernpoly (Suc n) x) \ c" using bounded_pbernpoly by auto have "eventually (\x. real x + Re s > 0) at_top" by real_asymp hence "eventually (\x. norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / (real x + Re s) ^ n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim x) have "norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / norm (s + of_nat x) ^ n" (is "_ \ ?rhs") using c[of x] by (auto simp: norm_divide norm_mult norm_power neq field_simps simp del: of_nat_Suc) also have "(real x + Re s) \ cmod (s + of_nat x)" using complex_Re_le_cmod[of "s + of_nat x"] s by (auto simp add: complex_nonpos_Reals_iff) hence "?rhs \ (c / real (Suc n)) / (real x + Re s) ^ n" using s elim c[of 0] neq[of x] by (intro divide_left_mono power_mono mult_pos_pos divide_nonneg_pos zero_less_power) auto finally show ?case . qed moreover have "(\x. (c / real (Suc n)) / (real x + Re s) ^ n) \ 0" using \n > 0\ by real_asymp ultimately show ?thesis by (rule Lim_null_comparison) qed lemma convergent_stirling_integral: assumes "n > 0" shows "convergent (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" (is "convergent (?f n)") proof - have "convergent (?f (Suc n))" for n proof (induction n) case 0 thus ?case using integral_pbernpoly_1 by (auto intro!: convergentI) next case (Suc n) have "convergent (\N. ?f (Suc n) N - of_real (pbernpoly (Suc (Suc n)) (real N)) / (of_nat (Suc (Suc n)) * (s + of_nat N) ^ Suc n) + of_real (bernoulli (Suc (Suc n)) / (real (Suc (Suc n)))) / s ^ Suc n)" (is "convergent ?g") by (intro convergent_add convergent_diff Suc convergent_const convergentI[OF pbernpoly_over_power_tendsto_0]) simp_all also have "?g = (\N. of_nat (Suc n) / of_nat (Suc (Suc n)) * ?f (Suc (Suc n)) N)" using s by (subst pbernpoly_integral_conv_pbernpoly_integral_Suc) (auto simp: fun_eq_iff field_simps simp del: of_nat_Suc power_Suc) also have "convergent \ \ convergent (?f (Suc (Suc n)))" by (intro convergent_mult_const_iff) (simp_all del: of_nat_Suc) finally show ?case . qed from this[of "n - 1"] assms show ?thesis by simp qed lemma stirling_integral_conv_stirling_integral_Suc: assumes "n > 0" shows "stirling_integral n s = of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" proof - have "(\N. of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (real (Suc n) * s ^ n) + integral {0..real N} (\x. of_nat n / of_nat (Suc n) * (of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))) \ 0 - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s" (is "?f \ _") unfolding stirling_integral_def integral_mult_right using convergent_stirling_integral[of "Suc n"] assms s by (intro tendsto_intros pbernpoly_over_power_tendsto_0) (auto simp: convergent_LIMSEQ_iff simp del: of_nat_Suc) also have "?this \ (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" using eventually_gt_at_top[of "0::nat"] pbernpoly_integral_conv_pbernpoly_integral_Suc[of n] assms unfolding integral_mult_right by (intro filterlim_cong refl) (auto elim!: eventually_mono simp del: power_Suc) finally show ?thesis unfolding stirling_integral_def[of n] by (rule limI) qed lemma stirling_integral_1_unfold: assumes "m > 0" shows "stirling_integral 1 s = stirling_integral m s / of_nat m - (\k=1..k=1..k = 1.. 0" shows "ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..k = 1.. 0 \ (\x. integral {0..real x} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ stirling_integral n s" unfolding stirling_integral_def using convergent_stirling_integral[of n] by (simp only: convergent_LIMSEQ_iff) end lemmas has_integral_of_real = has_integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemmas integral_of_real = integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemma integrable_ln_Gamma_aux_real: assumes "0 < s" shows "(\x. pbernpoly n x / (x + s) ^ n) integrable_on {0..real N}" proof - have "(\x. complex_of_real (pbernpoly n x / (x + s) ^ n)) integrable_on {0..real N}" using integrable_ln_Gamma_aux[of "of_real s" n N] assms by simp from integrable_linear[OF this bounded_linear_Re] show ?thesis by (simp only: o_def Re_complex_of_real) qed lemma assumes "x > 0" "n > 0" shows stirling_integral_complex_of_real: "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" and LIMSEQ_stirling_integral_real: "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" and stirling_integral_real_convergent: "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" proof - have "(\N. integral {0..real N} (\t. of_real (pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using LIMSEQ_stirling_integral[of "complex_of_real x" n] assms by simp hence "(\N. of_real (integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using integrable_ln_Gamma_aux_real[OF assms(1), of n] by (subst (asm) integral_of_real) simp from tendsto_Re[OF this] have "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ Re (stirling_integral n (complex_of_real x))" by simp thus "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" by (rule convergentI) thus "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" unfolding stirling_integral_def by (simp add: convergent_LIMSEQ_iff) from tendsto_of_real[OF this, where 'a = complex] integrable_ln_Gamma_aux_real[OF assms(1), of n] have "(\xa. integral {0..real xa} (\xa. complex_of_real (pbernpoly n xa) / (complex_of_real xa + x) ^ n)) \ complex_of_real (stirling_integral n x)" by (subst (asm) integral_of_real [symmetric]) simp_all from LIMSEQ_unique[OF this LIMSEQ_stirling_integral[of "complex_of_real x" n]] assms show "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" by simp qed lemma ln_Gamma_stirling_real: assumes "x > (0 :: real)" "m > (0::nat)" shows "ln_Gamma x = (x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1..k = 1.. (1::nat)" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp have "norm (stirling_integral n s) \ c / (real n - 1) / Re s ^ (n - 1)" if s: "Re s > 0" for s proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) have pos: "x + norm s > 0" if "x \ 0" for x using s that by (intro add_nonneg_pos) auto have nz: "of_real x + s \ 0" if "x \ 0" for x using s that by (auto simp: complex_eq_iff) let ?bound = "\N. c / (Re s ^ (n - 1) * (real n - 1)) - c / ((real N + Re s) ^ (n - 1) * (real n - 1))" show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) let ?F = "\x. -c / ((x + Re s) ^ (n - 1) * (real n - 1))" from n s have "((\x. c / (x + Re s) ^ n) has_integral (?F (real N) - ?F 0)) {0..real N}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 - has_field_derivative_iff_has_vector_derivative [symmetric]) + has_real_derivative_iff_has_vector_derivative [symmetric]) also have "?F (real N) - ?F 0 = ?bound N" by simp finally have *: "((\x. c / (x + Re s) ^ n) has_integral ?bound N) {0..real N}" . have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / (x + Re s) ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) fix x assume x: "x \ {0..real N}" have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / (x + Re s) ^ n" using x c c_nonneg s nz[of x] complex_Re_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power add_nonneg_pos) auto finally show "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ \" . qed (insert n s * pos nz c, auto simp: complex_nonpos_Reals_iff) also have "\ = ?bound N" using * by (simp add: has_integral_iff) also have "\ \ c / (Re s ^ (n - 1) * (real n - 1))" using c_nonneg elim s n by simp also have "\ = c / (real n - 1) / (Re s ^ (n - 1))" by simp finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)" . qed qed (insert s n, simp_all add: complex_nonpos_Reals_iff) thus ?thesis by (rule that) qed lemma stirling_integral_bound_aux_integral1: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" "l < a - b" "r > a + b" shows "((\x. c / max b \x - a\ ^ n) has_integral 2*c*(n / (n - 1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))) {l..r}" proof - define x1 x2 where "x1 = a - b" and "x2 = a + b" define F1 where "F1 = (\x::real. c / (a - x) ^ (n - 1) / (n - 1))" define F3 where "F3 = (\x::real. -c / (x - a) ^ (n - 1) / (n - 1))" have deriv: "(F1 has_vector_derivative (c / (a - x) ^ n)) (at x within A)" "(F3 has_vector_derivative (c / (x - a) ^ n)) (at x within A)" if "x \ a" for x :: real and A unfolding F1_def F3_def using assms that by (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 - simp flip: has_field_derivative_iff_has_vector_derivative) + simp flip: has_real_derivative_iff_has_vector_derivative) from assms have "((\x. c / (a - x) ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x1_def max_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" using assms by (intro has_integral_spike_finite_eq[of "{l}"]) (auto simp: x1_def max_def split: if_splits) finally have I1: "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" . have "((\x. c / b ^ n) has_integral (x2 - x1) * c / b ^ n) {x1..x2}" using has_integral_const_real[of "c / b ^ n" x1 x2] assms by (simp add: x1_def x2_def) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" by (intro has_integral_spike_finite_eq[of "{x1, x2}"]) (auto simp: x1_def x2_def split: if_splits) finally have I2: "((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" . from assms have I3: "((\x. c / (x - a) ^ n) has_integral (F3 r - F3 x2)) {x2..r}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x2_def min_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" using assms by (intro has_integral_spike_finite_eq[of "{r}"]) (auto simp: x2_def min_def split: if_splits) finally have I3: "((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" . have "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2)) {l..r}" using assms by (intro has_integral_combine[OF _ _ has_integral_combine[OF _ _ I1 I2] I3]) (auto simp: x1_def x2_def) also have "(F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2) = F1 x1 - F1 l + F3 r - F3 x2 + (x2 - x1) * c / b ^ n" by (simp add: algebra_simps) also have "x2 - x1 = 2 * b" using assms by (simp add: x2_def x1_def min_def max_def) also have "2 * b * c / b ^ n = 2 * c / b ^ (n - 1)" using assms by (simp add: power_diff field_simps) also have "F1 x1 - F1 l + F3 r - F3 x2 = c/(n-1) * (2/b^(n-1) - 1/(a-l)^(n-1) - 1/(r-a)^(n-1))" using assms by (simp add: x1_def x2_def F1_def F3_def field_simps) also have "\ + 2 * c / b ^ (n - 1) = 2*c*(1 + 1/(n-1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))" using assms by (simp add: field_simps) also have "1 + 1 / (n - 1) = n / (n - 1)" using assms by (simp add: field_simps) finally show ?thesis . qed lemma stirling_integral_bound_aux_integral2: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" obtains I where "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" "I \ 2 * c * (n / (n - 1)) / b ^ (n-1)" proof - define l' where "l' = min l (a - b - 1)" define r' where "r' = max r (a + b + 1)" define A where "A = 2 * c * (n / (n - 1)) / b ^ (n - 1)" define B where "B = c / real (n - 1) * (1 / (a - l') ^ (n - 1) + 1 / (r' - a) ^ (n - 1))" have has_int: "((\x. c / max b \x - a\ ^ n) has_integral (A - B)) {l'..r'}" using assms unfolding A_def B_def by (intro stirling_integral_bound_aux_integral1) (auto simp: l'_def r'_def) have "(\x. c / max b \x - a\ ^ n) integrable_on {l..r}" by (rule integrable_on_subinterval[OF has_integral_integrable[OF has_int]]) (auto simp: l'_def r'_def) then obtain I where has_int': "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" by (auto simp: integrable_on_def) from assms have "I \ A - B" by (intro has_integral_subset_le[OF _ has_int' has_int]) (auto simp: l'_def r'_def) also have "\ \ A" using assms by (simp add: B_def l'_def r'_def) finally show ?thesis using that[of I] has_int' unfolding A_def by blast qed lemma stirling_integral_bound_aux': assumes n: "n > (1::nat)" and \: "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp define D where "D = c * Beta (- (real_of_int (- int n) / 2) - 1 / 2) (1 / 2) / 2" define C where "C = max D (2*c*(n/(n-1))/sin \^(n-1))" have *: "norm (stirling_integral n s) \ C / norm s ^ (n - 1)" if s: "s \ complex_cone' \ - {0}" for s :: complex proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) from s \ have Arg: "\Arg s\ \ \" by (auto simp: complex_cone_altdef) have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ s by auto from s have [simp]: "s \ 0" by auto show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ C / norm s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) show ?case proof (cases "Re s > 0") case True have int: "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0<..}" using has_integral_mult_left[OF has_integral_Beta3[of "-n/2" "norm s ^ 2"], of c] assms True unfolding D_def by (simp add: algebra_simps) hence int': "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0..}" by (subst has_integral_interior [symmetric]) simp_all hence integrable: "(\x. c * (x^2 + norm s^2) powr (-n / 2)) integrable_on {0..}" by (simp add: has_integral_iff) have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c * (x^2 + norm s^2) powr (-n / 2))" proof (intro integral_norm_bound_integral s ballI integrable_ln_Gamma_aux) have [simp]: "{0<..} - {0::real..} = {}" "{0..} - {0<..} = {0::real}" by auto have "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0<..}" using int by (simp add: has_integral_iff) also have "?this \ (\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..}" by (intro integrable_spike_set_eq) auto finally show "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..real N}" by (rule integrable_on_subinterval) auto next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" using True x by (auto simp: complex_eq_iff) have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / sqrt (x ^ 2 + norm s ^ 2) ^ n" proof (intro divide_left_mono mult_pos_pos zero_less_power power_mono) show "sqrt (x\<^sup>2 + (cmod s)\<^sup>2) \ cmod (complex_of_real x + s)" using x True by (simp add: cmod_def algebra_simps power2_eq_square) qed (use x True c_nonneg assms nz in \auto simp: add_nonneg_pos\) also have "sqrt (x ^ 2 + norm s ^ 2) ^ n = (x ^ 2 + norm s ^ 2) powr (1/2 * n)" by (subst powr_powr [symmetric], subst powr_realpow) (auto simp: powr_half_sqrt add_nonneg_pos) also have "c / \ = c * (x^2 + norm s^2) powr (-n / 2)" by (simp add: powr_minus field_simps) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed fact+ also have "\ \ integral {0..} (\x. c * (x^2 + norm s^2) powr (-n / 2))" using c_nonneg by (intro integral_subset_le integrable integrable_on_subinterval[OF integrable]) auto also have "\ = D * (norm s ^ 2) powr (-n / 2 + 1 / 2)" using int' by (simp add: has_integral_iff) also have "(norm s ^ 2) powr (-n / 2 + 1 / 2) = norm s powr (2 * (-n / 2 + 1 / 2))" by (subst powr_powr [symmetric]) auto also have "\ = norm s powr (-real (n - 1))" using assms by (simp add: of_nat_diff) also have "D * \ = D / norm s ^ (n - 1)" by (auto simp: powr_minus powr_realpow field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ \" . next case False have "cos \Arg s\ = cos (Arg s)" by (simp add: abs_if) also have "cos (Arg s) = Re (rcis (norm s) (Arg s)) / norm s" by (subst Re_rcis) auto also have "\ = Re s / norm s" by (subst rcis_cmod_Arg) auto also have "\ \ cos (pi / 2)" using False by (auto simp: field_simps) finally have "\Arg s\ \ pi / 2" using Arg \ by (subst (asm) cos_mono_le_eq) auto have "sin \ * norm s = sin (pi - \) * norm s" by simp also have "\ \ sin (pi - \Arg s\) * norm s" using \ Arg \\Arg s\ \ pi / 2\ by (intro mult_right_mono sin_monotone_2pi_le) auto also have "sin \Arg s\ \ 0" using Arg_bounded[of s] by (intro sin_ge_zero) auto hence "sin (pi - \Arg s\) = \sin \Arg s\\" by simp also have "\ = \sin (Arg s)\" by (simp add: abs_if) also have "\ * norm s = \Im (rcis (norm s) (Arg s))\" by (simp add: abs_mult) also have "\ = \Im s\" by (subst rcis_cmod_Arg) auto finally have abs_Im_ge: "\Im s\ \ sin \ * norm s" . have [simp]: "Im s \ 0" "s \ 0" using s \s \ \\<^sub>\\<^sub>0\ False by (auto simp: cmod_def zero_le_mult_iff complex_nonpos_Reals_iff) have "sin \ > 0" using assms by (intro sin_gt_zero) auto obtain I where I: "((\x. c / max \Im s\ \x + Re s\ ^ n) has_integral I) {0..real N}" "I \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using s c_nonneg assms False stirling_integral_bound_aux_integral2[of "-Re s" "\Im s\" c n 0 "real N"] by auto have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / max \Im s\ \x + Re s\ ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) show "(\x. c / max \Im s\ \x + Re s\ ^ n) integrable_on {0..real N}" using I(1) by (simp add: has_integral_iff) next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff) have "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ c / norm (complex_of_real x + s) ^ n" unfolding norm_divide norm_power using c[of x] by (intro divide_right_mono) simp_all also have "\ \ c / max \Im s\ \x + Re s\ ^ n" using c_nonneg nz abs_Re_le_cmod[of "of_real x + s"] abs_Im_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power) (auto simp: less_max_iff_disj) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed (auto simp: complex_nonpos_Reals_iff) also have "\ \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using I by (simp add: has_integral_iff) also have "\ \ 2*c*(n/(n-1)) / (sin \ * norm s) ^ (n - 1)" using \sin \ > 0\ s c_nonneg abs_Im_ge by (intro divide_left_mono mult_pos_pos zero_less_power power_mono mult_nonneg_nonneg) auto also have "\ = 2*c*(n/(n-1))/sin \^(n-1) / norm s ^ (n - 1)" by (simp add: field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show ?thesis . qed qed qed (use that assms complex_cone_inter_nonpos_Reals[of "-\" \] \ in auto) thus ?thesis by (rule that) qed lemma stirling_integral_bound: assumes "n > 0" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux[of "Suc n"] assms obtain c where c: "\s. Re s > 0 \ norm (stirling_integral (Suc n) s) \ c / Re s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "Re s > 0" hence s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff) have "stirling_integral n s = ?f s" using s' assms by (rule stirling_integral_conv_stirling_integral_Suc) also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / Re s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / Re s ^ n" by (simp add: c1_def) also have "c2 / norm s ^ n \ c2 / Re s ^ n" using s c2_nonneg by (intro divide_left_mono power_mono complex_Re_le_cmod mult_pos_pos zero_less_power) auto also have "c1 / Re s ^ n + c2 / Re s ^ n = (c1 + c2) / Re s ^ n" using s by (simp add: field_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / Re s ^ n" by - simp_all qed qed lemma stirling_integral_bound': assumes "n > 0" and "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux'[of "Suc n"] assms obtain c where c: "\s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral (Suc n) s) \ c / norm s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "s \ complex_cone' \ - {0}" have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] assms s by auto have "stirling_integral n s = ?f s" using s' assms by (intro stirling_integral_conv_stirling_integral_Suc) auto also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / norm s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / norm s ^ n" by (simp add: c1_def) also have "c1 / norm s ^ n + c2 / norm s ^ n = (c1 + c2) / norm s ^ n" using s by (simp add: divide_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / norm s ^ n" by - simp_all qed qed lemma stirling_integral_holomorphic [holomorphic_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "stirling_integral m holomorphic_on A" proof - from assms have [simp]: "z \ \\<^sub>\\<^sub>0" if "z \ A" for z using that by auto let ?f = "\s::complex. of_nat m * ((s - 1 / 2) * Ln s - s + of_real (ln (2 * pi) / 2) + (\k=1.. stirling_integral m holomorphic_on A" using assms by (intro holomorphic_cong refl) (simp_all add: field_simps ln_Gamma_stirling_complex) finally show "stirling_integral m holomorphic_on A" . qed lemma stirling_integral_continuous_on_complex [continuous_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "continuous_on A (stirling_integral m :: _ \ complex)" by (intro holomorphic_on_imp_continuous_on stirling_integral_holomorphic assms) lemma has_field_derivative_stirling_integral_complex: fixes x :: complex assumes "x \ \\<^sub>\\<^sub>0" "n > 0" shows "(stirling_integral n has_field_derivative deriv (stirling_integral n) x) (at x)" using assms by (intro holomorphic_derivI[OF stirling_integral_holomorphic, of n "-\\<^sub>\\<^sub>0"]) auto lemma assumes n: "n > 0" and "x > 0" shows deriv_stirling_integral_complex_of_real: "(deriv ^^ j) (stirling_integral n) (complex_of_real x) = complex_of_real ((deriv ^^ j) (stirling_integral n) x)" (is "?lhs x = ?rhs x") and differentiable_stirling_integral_real: "(deriv ^^ j) (stirling_integral n) field_differentiable at x" (is ?thesis2) proof - let ?A = "{s. Re s > 0}" let ?f = "\j x. (deriv ^^ j) (stirling_integral n) (complex_of_real x)" let ?f' = "\j x. complex_of_real ((deriv ^^ j) (stirling_integral n) x)" have [simp]: "open ?A" by (simp add: open_halfspace_Re_gt) have "?lhs x = ?rhs x \ (deriv ^^ j) (stirling_integral n) field_differentiable at x" if "x > 0" for x using that proof (induction j arbitrary: x) case 0 have "((\x. Re (stirling_integral n (of_real x))) has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using 0 n by (auto intro!: derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic simp: complex_nonpos_Reals_iff) also have "?this \ (stirling_integral n has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] 0 n by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono simp: stirling_integral_complex_of_real) finally have "stirling_integral n field_differentiable at x" by (auto simp: field_differentiable_def) with 0 n show ?case by (auto simp: stirling_integral_complex_of_real) next case (Suc j x) note IH = conjunct1[OF Suc.IH] conjunct2[OF Suc.IH] have *: "(deriv ^^ Suc j) (stirling_integral n) (complex_of_real x) = of_real ((deriv ^^ Suc j) (stirling_integral n) x)" if x: "x > 0" for x proof - have "deriv ((deriv ^^ j) (stirling_integral n)) (complex_of_real x) = vector_derivative (\x. (deriv ^^ j) (stirling_integral n) (of_real x)) (at x)" using n x by (intro vector_derivative_of_real_right [symmetric] holomorphic_on_imp_differentiable_at[of _ ?A] holomorphic_higher_deriv stirling_integral_holomorphic) (auto simp: complex_nonpos_Reals_iff) also have "\ = vector_derivative (\x. of_real ((deriv ^^ j) (stirling_integral n) x)) (at x)" using eventually_nhds_in_open[of "{0<..}" x] x by (intro vector_derivative_cong_eq) (auto elim!: eventually_mono simp: IH(1)) also have "\ = of_real (deriv ((deriv ^^ j) (stirling_integral n)) x)" by (intro vector_derivative_of_real_left holomorphic_on_imp_differentiable_at[of _ ?A] field_differentiable_imp_differentiable IH(2) x) finally show ?thesis by simp qed have "((\x. Re ((deriv ^^ Suc j) (stirling_integral n) (of_real x))) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using Suc.prems n by (intro derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic holomorphic_higher_deriv) (auto simp: complex_nonpos_Reals_iff) also have "?this \ ((deriv ^^ Suc j) (stirling_integral n) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems * by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono) finally have "(deriv ^^ Suc j) (stirling_integral n) field_differentiable at x" by (auto simp: field_differentiable_def) with *[OF Suc.prems] show ?case by blast qed from this[OF assms(2)] show "?lhs x = ?rhs x" ?thesis2 by blast+ qed text \ Unfortunately, asymptotic power series cannot, in general, be differentiated. However, since @{term ln_Gamma} is holomorphic on the entire positive real half-space, we can differentiate its asymptotic expansion after all. To do this, we use an ad-hoc version of the more general approach outlined in Erdelyi's ``Asymptotic Expansions'' for holomorphic functions: We bound the value of the $j$-th derivative of the remainder term at some value $x$ by applying Cauchy's integral formula along a circle centred at $x$ with radius $\frac{1}{2} x$. \ lemma deriv_stirling_integral_real_bound: assumes m: "m > 0" shows "(deriv ^^ j) (stirling_integral m) \ O(\x::real. 1 / x ^ (m + j))" proof - obtain c where c: "\s. 0 < Re s \ cmod (stirling_integral m s) \ c / Re s ^ m" using stirling_integral_bound[OF m] by auto have "0 \ cmod (stirling_integral m 1)" by simp also have "\ \ c" using c[of 1] by simp finally have c_nonneg: "c \ 0" . define B where "B = c * 2 ^ (m + Suc j)" define B' where "B' = B * fact j / 2" have "eventually (\x::real. norm ((deriv ^^ j) (stirling_integral m) x) \ B' * norm (1 / x ^ (m+ j))) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) have "s \ \\<^sub>\\<^sub>0" if "s \ cball (of_real x) (x/2)" for s :: complex proof - have "x - Re s \ norm (of_real x - s)" using complex_Re_le_cmod[of "of_real x - s"] by simp also from that have "\ \ x/2" by (simp add: dist_complex_def) finally show ?thesis using elim by (auto simp: complex_nonpos_Reals_iff) qed hence "((\u. stirling_integral m u / (u - of_real x) ^ Suc j) has_contour_integral complex_of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) (circlepath (of_real x) (x/2))" using m elim by (intro Cauchy_has_contour_integral_higher_derivative_circlepath stirling_integral_continuous_on_complex stirling_integral_holomorphic) auto hence "norm (of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) \ B / x ^ (m + Suc j) * (2 * pi * (x / 2))" proof (rule has_contour_integral_bound_circlepath) fix u :: complex assume dist: "norm (u - of_real x) = x / 2" have "Re (of_real x - u) \ norm (of_real x - u)" by (rule complex_Re_le_cmod) also have "\ = x / 2" using dist by (simp add: norm_minus_commute) finally have Re_u: "Re u \ x/2" using elim by simp have "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ c / Re u ^ m / (x / 2) ^ Suc j" using Re_u elim unfolding norm_divide norm_power dist by (intro divide_right_mono zero_le_power c) simp_all also have "\ \ c / (x/2) ^ m / (x / 2) ^ Suc j" using c_nonneg elim Re_u by (intro divide_right_mono divide_left_mono power_mono) simp_all also have "\ = B / x ^ (m + Suc j)" using elim by (simp add: B_def field_simps power_add) finally show "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ B / x ^ (m + Suc j)" . qed (insert elim c_nonneg, auto simp: B_def simp del: power_Suc) hence "cmod ((deriv ^^ j) (stirling_integral m) (of_real x)) \ B' / x ^ (j + m)" using elim by (simp add: field_simps norm_divide norm_mult norm_power B'_def) with elim m show ?case by (simp_all add: add_ac deriv_stirling_integral_complex_of_real) qed thus ?thesis by (rule bigoI) qed definition stirling_sum where "stirling_sum j m x = (-1) ^ j * (\k = 1..k\m. (of_real (bernoulli' k) * pochhammer (of_nat (Suc k)) (j - 1) * inverse x ^ (k + j)))" lemma stirling_sum_complex_of_real: "stirling_sum j m (complex_of_real x) = complex_of_real (stirling_sum j m x)" by (simp add: stirling_sum_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma stirling_sum'_complex_of_real: "stirling_sum' j m (complex_of_real x) = complex_of_real (stirling_sum' j m x)" by (simp add: stirling_sum'_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma has_field_derivative_stirling_sum_complex [derivative_intros]: "Re x > 0 \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum_real [derivative_intros]: "x > (0::real) \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum'_complex [derivative_intros]: assumes "j > 0" "Re x > 0" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * complex_of_real (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma has_field_derivative_stirling_sum'_real [derivative_intros]: assumes "j > 0" "x > (0::real)" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma higher_deriv_stirling_sum_complex: "Re x > 0 \ (deriv ^^ i) (stirling_sum j m) x = stirling_sum (i + j) m x" proof (induction i arbitrary: x) case (Suc i) have "deriv ((deriv ^^ i) (stirling_sum j m)) x = deriv (stirling_sum (i + j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_halfspace_Re_gt Suc.IH) also from Suc.prems have "\ = stirling_sum (Suc (i + j)) m x" by (intro DERIV_imp_deriv has_field_derivative_stirling_sum_complex) finally show ?case by simp qed simp_all definition Polygamma_approx :: "nat \ nat \ 'a \ 'a :: {real_normed_field, ln}" where "Polygamma_approx j m = (deriv ^^ j) (\x. (x - 1 / 2) * ln x - x + of_real (ln (2 * pi)) / 2 + stirling_sum 0 m x)" lemma Polygamma_approx_Suc: "Polygamma_approx (Suc j) m = deriv (Polygamma_approx j m)" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_0: "Polygamma_approx 0 m x = (x - 1/2) * ln x - x + of_real (ln (2*pi)) / 2 + stirling_sum 0 m x" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_1_complex: "Re x > 0 \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma Polygamma_approx_1_real: "x > (0 :: real) \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma stirling_sum_2_conv_stirling_sum'_1: fixes x :: "'a :: {real_div_algebra, field_char_0}" assumes "m > 0" "x \ 0" shows "stirling_sum' 1 m x = 1 / x + 1 / (2 * x^2) + stirling_sum 2 m x" proof - have pochhammer_2: "pochhammer (of_nat k) 2 = of_nat k * of_nat (Suc k)" for k by (simp add: pochhammer_Suc eval_nat_numeral add_ac) have "stirling_sum 2 m x = (\k = Suc 0.. = (\k=0.. = (\k=0.. = (\k\m. of_real (bernoulli' k) * inverse x ^ Suc k)" by (intro sum.cong) auto also have "\ = stirling_sum' 1 m x" by (simp add: stirling_sum'_def) finally show ?thesis by (simp add: add_ac) qed lemma Polygamma_approx_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro deriv_cong_ev) (auto elim!: eventually_mono simp: Polygamma_approx_1_real) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (simp add: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{s. Re s > 0}" x] assms by (intro deriv_cong_ev) (auto simp: open_halfspace_Re_gt elim!: eventually_mono simp: Polygamma_approx_1_complex) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (subst stirling_sum_2_conv_stirling_sum'_1) (auto simp: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_ge_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_real) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_ge_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_complex) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH open_halfspace_Re_gt) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_complex_of_real: assumes "x > 0" "m > 0" shows "Polygamma_approx j m (complex_of_real x) = of_real (Polygamma_approx j m x)" proof (cases j) case 0 with assms show ?thesis by (simp add: Polygamma_approx_0 Ln_of_real stirling_sum_complex_of_real) next case [simp]: (Suc j') thus ?thesis proof (cases j') case 0 with assms show ?thesis by (simp add: Polygamma_approx_1_complex Polygamma_approx_1_real stirling_sum_complex_of_real Ln_of_real) next case (Suc j'') with assms show ?thesis by (simp add: Polygamma_approx_ge_2_complex Polygamma_approx_ge_2_real stirling_sum'_complex_of_real) qed qed lemma higher_deriv_Polygamma_approx [simp]: "(deriv ^^ j) (Polygamma_approx i m) = Polygamma_approx (j + i) m" by (simp add: Polygamma_approx_def funpow_add) lemma stirling_sum_holomorphic [holomorphic_intros]: "0 \ A \ stirling_sum j m holomorphic_on A" unfolding stirling_sum_def by (intro holomorphic_intros) auto lemma Polygamma_approx_holomorphic [holomorphic_intros]: "Polygamma_approx j m holomorphic_on {s. Re s > 0}" unfolding Polygamma_approx_def by (intro holomorphic_intros) (auto simp: open_halfspace_Re_gt elim!: nonpos_Reals_cases) lemma higher_deriv_lnGamma_stirling: assumes m: "m > 0" shows "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ O(\x. 1 / x ^ (m + j))" proof - have "eventually (\x. \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\ = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) note x = this have "\\<^sub>F y in nhds (complex_of_real x). y \ - \\<^sub>\\<^sub>0" using elim by (intro eventually_nhds_in_open) auto hence "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) (\x. (-inverse (of_nat m)) * stirling_integral m x) (complex_of_real x)" using x m by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: ln_Gamma_stirling_complex Polygamma_approx_def field_simps open_halfspace_Re_gt stirling_sum_def) also have "\ = - inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (of_real x)" using x m by (intro higher_deriv_cmult[of _ "-\\<^sub>\\<^sub>0"] stirling_integral_holomorphic) (auto simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) ln_Gamma (of_real x) - (deriv ^^ j) (Polygamma_approx 0 m) (of_real x)" using x by (intro higher_deriv_diff[of _ "{s. Re s > 0}"]) (auto intro!: holomorphic_intros elim!: nonpos_Reals_cases simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (Polygamma_approx 0 m) (complex_of_real x) = of_real (Polygamma_approx j m x)" using x m by (simp add: Polygamma_approx_complex_of_real) also have "norm (- inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (complex_of_real x)) = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\" using x m by (simp add: norm_mult norm_inverse deriv_stirling_integral_complex_of_real) also have "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using x by (simp add: higher_deriv_ln_Gamma_complex_of_real) also have "norm (\ - of_real (Polygamma_approx j m x)) = \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\" by (simp only: of_real_diff [symmetric] norm_of_real) finally show ?case . qed from bigthetaI_cong[OF this] m have "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ \(\x. (deriv ^^ j) (stirling_integral m) x)" by simp also have "(\x::real. (deriv ^^ j) (stirling_integral m) x) \ O(\x. 1 / x ^ (m + j))" using m by (rule deriv_stirling_integral_real_bound) finally show ?thesis . qed lemma Polygamma_approx_1_real': assumes x: "(x::real) > 0" and m: "m > 0" shows "Polygamma_approx 1 m x = ln x - (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" proof - have "Polygamma_approx 1 m x = ln x - (1 / (2 * x) + (\k=Suc 0..k=Suc 0.. = (\k=0.. = (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" using assms by (subst sum.shift_bounds_Suc_ivl [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) finally show ?thesis . qed theorem assumes m: "m > 0" shows ln_Gamma_real_asymptotics: "(\x. ln_Gamma x - ((x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1.. O(\x. 1 / x ^ m)" (is ?th1) and Digamma_real_asymptotics: "(\x. Digamma x - (ln x - (\k=1..m. bernoulli' k / real k / x ^ k))) \ O(\x. 1 / (x ^ Suc m))" (is ?th2) and Polygamma_real_asymptotics: "j > 0 \ (\x. Polygamma j x - (- 1) ^ Suc j * (\k\m. bernoulli' k * pochhammer (real (Suc k)) (j - 1) / x ^ (k + j))) \ O(\x. 1 / x ^ (m+j+1))" (is "_ \ ?th3") proof - define G :: "nat \ real \ real" where "G = (\m. if m = 0 then ln_Gamma else Polygamma (m - 1))" have *: "(\x. G j x - h x) \ O(\x. 1 / x ^ (m + j))" if "\x::real. x > 0 \ Polygamma_approx j m x = h x" for j h proof - have "(\x. G j x - h x) \ \(\x. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x)" (is "_ \ \(?f)") using that by (intro bigthetaI_cong) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::real"]] simp del: funpow.simps simp: higher_deriv_ln_Gamma_real G_def) also have "?f \ O(\x::real. 1 / x ^ (m + j))" using m by (rule higher_deriv_lnGamma_stirling) finally show ?thesis . qed note [[simproc del: simplify_landau_sum]] from *[OF Polygamma_approx_0] assms show ?th1 by (simp add: G_def Polygamma_approx_0 stirling_sum_def field_simps) from *[OF Polygamma_approx_1_real'] assms show ?th2 by (simp add: G_def field_simps) assume j: "j > 0" from *[OF Polygamma_approx_ge_2_real, of "j - 1"] assms j show ?th3 by (simp add: G_def stirling_sum'_def power_add power_diff field_simps) qed subsection \Asymptotics of the complex Gamma function\ text \ The \m\-th order remainder of Stirling's formula for $\log\Gamma$ is $O(s^{-m})$ uniformly over any complex cone $\text{Arg}(z) \leq \alpha$, $z\neq 0$ for any angle $\alpha\in(0, \pi)$. This means that there is bounded by $c z^{-m}$ for some constant $c$ for all $z$ in this cone. \ context fixes F and \ assumes \: "\ \ {0<.. principal (complex_cone' \ - {0})" begin lemma stirling_integral_bigo: fixes m :: nat assumes m: "m > 0" shows "stirling_integral m \ O[F](\s. 1 / s ^ m)" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF \m > 0\ \] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] \ by simp finally have "c \ 0" . have "eventually (\s. s \ complex_cone' \ - {0}) F" unfolding F_def by (auto simp: eventually_principal) hence "eventually (\s. norm (stirling_integral m s) \ c * norm (1 / s ^ m)) F" by eventually_elim (use c in \simp add: norm_divide norm_power\) thus "stirling_integral m \ O[F](\s. 1 / s ^ m)" by (intro bigoI[of _ c]) auto qed end text \ The following is a more explicit statement of this: \ theorem ln_Gamma_complex_asymptotics_explicit: fixes m :: nat and \ :: real assumes "m > 0" and "\ \ {0<.. complex" where "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1/2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \Arg s\ \ \ \ norm (R s) \ C / norm s ^ m" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF assms] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] assms by simp finally have "c \ 0" . define R where "R = (\s::complex. stirling_integral m s / of_nat m)" show ?thesis proof (rule that) from ln_Gamma_stirling_complex[of _ m] assms show "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \Arg s\ \ \ \ cmod (R s) \ c / real m / cmod s ^ m" proof (safe, goal_cases) case (1 s) show ?case using 1 c[of s] assms by (auto simp: complex_cone_altdef abs_le_iff R_def norm_divide field_simps) qed qed qed text \ Lastly, we can also derive the asymptotics of $\Gamma$ itself: \[\Gamma(z) \sim \sqrt{2\pi / z} \left(\frac{z}{e}\right)^z\] uniformly for $|z|\to\infty$ within the cone $\text{Arg}(z) \leq \alpha$ for $\alpha\in(0,\pi)$: \ context fixes F and \ assumes \: "\ \ {0<.. inf at_infinity (principal (complex_cone' \))" begin lemma Gamma_complex_asymp_equiv: "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2))" proof - define I :: "complex \ complex" where "I = stirling_integral 1" have "eventually (\s. s \ complex_cone' \) F" by (auto simp: eventually_inf_principal F_def) moreover have "eventually (\s. s \ 0) F" unfolding F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto ultimately have "eventually (\s. Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)) F" proof eventually_elim case (elim s) from elim have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ by auto from elim have [simp]: "s \ 0" by auto from s' have "Gamma s = exp (ln_Gamma s)" unfolding Gamma_complex_altdef using nonpos_Ints_subset_nonpos_Reals by auto also from s' have "ln_Gamma s = (s-1/2) * Ln s - s + complex_of_real (ln (2 * pi) / 2) - I s" by (subst ln_Gamma_stirling_complex[of _ 1]) (simp_all add: exp_add exp_diff I_def) also have "exp \ = exp ((s - 1 / 2) * Ln s) / exp s * exp (complex_of_real (ln (2 * pi) / 2)) / exp (I s)" unfolding exp_diff exp_add by (simp add: exp_diff exp_add) also have "exp ((s - 1 / 2) * Ln s) = s powr (s - 1 / 2)" by (simp add: powr_def) also have "exp (complex_of_real (ln (2 * pi) / 2)) = sqrt (2 * pi)" by (subst exp_of_real) (auto simp: powr_def simp flip: powr_half_sqrt) also have "exp s = exp 1 powr s" by (simp add: powr_def) also have "s powr (s - 1 / 2) / exp 1 powr s = (s powr s / exp 1 powr s) / s powr (1/2)" by (subst powr_diff) auto also have *: "Ln (s / exp 1) = Ln s - 1" using Ln_divide_of_real[of "exp 1" s] by (simp flip: exp_of_real) hence "s powr s / exp 1 powr s = (s / exp 1) powr s" unfolding powr_def by (subst *) (auto simp: exp_diff field_simps) finally show "Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)" by (simp add: algebra_simps) qed hence "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s))" by (rule asymp_equiv_refl_ev) also have "\ \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / 1)" proof (intro asymp_equiv_intros) have "F \ principal (complex_cone' \ - {0})" unfolding le_principal F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto moreover have "I \ O[principal (complex_cone' \ - {0})](\s. 1 / s)" using stirling_integral_bigo[of \ 1] \ unfolding F_def by (simp add: I_def) ultimately have "I \ O[F](\s. 1 / s)" by (rule landau_o.big.filter_mono) also have "(\s. 1 / s) \ o[F](\s. 1)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" hence "eventually (\z::complex. norm z \ 1 / c) at_infinity" by (auto simp: eventually_at_infinity) moreover have "eventually (\z::complex. z \ 0) at_infinity" by (rule eventually_not_equal_at_infinity) ultimately show "eventually (\z::complex. norm (1 / z) \ c * norm (1 :: complex)) F" unfolding F_def eventually_inf_principal by eventually_elim (use \c > 0\ in \auto simp: norm_divide field_simps\) qed finally have "I \ o[F](\s. 1)" . from smalloD_tendsto[OF this] have [tendsto_intros]: "(I \ 0) F" by simp show "(\x. exp (I x)) \[F] (\x. 1)" by (rule asymp_equivI' tendsto_eq_intros refl | simp)+ qed finally show ?thesis by simp qed end end diff --git a/thys/Stirling_Formula/Stirling_Formula.thy b/thys/Stirling_Formula/Stirling_Formula.thy --- a/thys/Stirling_Formula/Stirling_Formula.thy +++ b/thys/Stirling_Formula/Stirling_Formula.thy @@ -1,722 +1,722 @@ (* File: Stirling_Formula.thy Author: Manuel Eberl A proof of Stirling's formula, i.e. an asymptotic approximation of the Gamma function and the factorial. *) section \Stirling's Formula\ theory Stirling_Formula imports "HOL-Analysis.Analysis" "Landau_Symbols.Landau_More" begin context begin text \ First, we define the $S_n^*$ from Jameson's article: \ private definition S' :: "nat \ real \ real" where "S' n x = 1/(2*x) + (\r=1.. Next, the trapezium (also called $T$ in Jameson's article): \ private definition T :: "real \ real" where "T x = 1/(2*x) + 1/(2*(x+1))" text \ Now we define The difference $\Delta(x)$: \ private definition D :: "real \ real" where "D x = T x - ln (x + 1) + ln x" private lemma S'_telescope_trapezium: assumes "n > 0" shows "S' n x = (\rrrrrrrr=0..r=0..r=Suc 0.. = (\r=1.. + ?a + ?b = S' n x" by (simp add: S'_def Suc) finally show ?thesis .. qed (insert assms, simp_all) private lemma stirling_trapezium: assumes x: "(x::real) > 0" shows "D x \ {0 .. 1/(12*x^2) - 1/(12 * (x+1)^2)}" proof - define y where "y = 1 / (2*x + 1)" from x have y: "y > 0" "y < 1" by (simp_all add: divide_simps y_def) from x have "D x = T x - ln ((x + 1) / x)" by (subst ln_div) (simp_all add: D_def) also from x have "(x + 1) / x = 1 + 1 / x" by (simp add: field_simps) finally have D: "D x = T x - ln (1 + 1/x)" . from y have "(\n. y * y^n) sums (y * (1 / (1 - y)))" by (intro geometric_sums sums_mult) simp_all hence "(\n. y ^ Suc n) sums (y / (1 - y))" by simp also from x have "y / (1 - y) = 1 / (2*x)" by (simp add: y_def divide_simps) finally have *: "(\n. y ^ Suc n) sums (1 / (2*x))" . from y have "(\n. (-y) * (-y)^n) sums ((-y) * (1 / (1 - (-y))))" by (intro geometric_sums sums_mult) simp_all hence "(\n. (-y) ^ Suc n) sums (-(y / (1 + y)))" by simp also from x have "y / (1 + y) = 1 / (2*(x+1))" by (simp add: y_def divide_simps) finally have **: "(\n. (-y) ^ Suc n) sums (-(1 / (2*(x+1))))" . from sums_diff[OF * **] have sum1: "(\n. y ^ Suc n - (-y) ^ Suc n) sums T x" by (simp add: T_def) from y have "abs y < 1" "abs (-y) < 1" by simp_all from sums_diff[OF this[THEN ln_series']] have "(\n. y ^ n / real n - (-y) ^ n / real n) sums (ln (1 + y) - ln (1 - y))" by simp also from y have "ln (1 + y) - ln (1 - y) = ln ((1 + y) / (1 - y))" by (simp add: ln_div) also from x have "(1 + y) / (1 - y) = 1 + 1/x" by (simp add: divide_simps y_def) finally have "(\n. y^n / real n - (-y)^n / real n) sums ln (1 + 1/x)" . hence sum2: "(\n. y^Suc n / real (Suc n) - (-y)^Suc n / real (Suc n)) sums ln (1 + 1/x)" by (subst sums_Suc_iff) simp have "ln (1 + 1/x) \ T x" proof (rule sums_le [OF _ sum2 sum1]) fix n :: nat show "y ^ Suc n / real (Suc n) - (-y) ^ Suc n / real (Suc n) \ y^Suc n - (-y) ^ Suc n" proof (cases "even n") case True hence eq: "A - (-y) ^ Suc n / B = A + y ^ Suc n / B" "A - (-y) ^ Suc n = A + y ^ Suc n" for A B by simp_all from y show ?thesis unfolding eq by (intro add_mono) (auto simp: divide_simps) qed simp_all qed hence "D x \ 0" by (simp add: D) define c where "c = (\n. if even n then 2 * (1 - 1 / real (Suc n)) else 0)" note sums_diff[OF sum1 sum2] also have "(\n. y ^ Suc n - (-y) ^ Suc n - (y ^ Suc n / real (Suc n) - (-y) ^ Suc n / real (Suc n))) = (\n. c n * y ^ Suc n)" by (intro ext) (simp add: c_def algebra_simps) finally have sum3: "(\n. c n * y ^ Suc n) sums D x" by (simp add: D) from y have "(\n. y^2 * (of_nat (Suc n) * y^n)) sums (y^2 * (1 / (1 - y)^2))" by (intro sums_mult geometric_deriv_sums) simp_all hence "(\n. of_nat (Suc n) * y^(n+2)) sums (y^2 / (1 - y)^2)" by (simp add: algebra_simps power2_eq_square) also from x have "y^2 / (1 - y)^2 = 1 / (4*x^2)" by (simp add: y_def divide_simps) finally have *: "(\n. real (Suc n) * y ^ (Suc (Suc n))) sums (1 / (4 * x\<^sup>2))" by simp from y have "(\n. y^2 * (of_nat (Suc n) * (-y)^n)) sums (y^2 * (1 / (1 - -y)^2))" by (intro sums_mult geometric_deriv_sums) simp_all hence "(\n. of_nat (Suc n) * (-y)^(n+2)) sums (y^2 / (1 + y)^2)" by (simp add: algebra_simps power2_eq_square) also from x have "y^2 / (1 + y)^2 = 1 / (2^2*(x+1)^2)" unfolding power_mult_distrib [symmetric] by (simp add: y_def divide_simps add_ac) finally have **: "(\n. real (Suc n) * (- y) ^ (Suc (Suc n))) sums (1 / (4 * (x + 1)\<^sup>2))" by simp define d where "d = (\n. if even n then 2 * real n else 0)" note sums_diff[OF * **] also have "(\n. real (Suc n) * y^(Suc (Suc n)) - real (Suc n) * (-y)^(Suc (Suc n))) = (\n. d (Suc n) * y ^ Suc (Suc n))" by (intro ext) (simp_all add: d_def) finally have "(\n. d n * y ^ Suc n) sums (1 / (4 * x\<^sup>2) - 1 / (4 * (x + 1)\<^sup>2))" by (subst (asm) sums_Suc_iff) (simp add: d_def) from sums_mult[OF this, of "1/3"] x have sum4: "(\n. d n / 3 * y ^ Suc n) sums (1 / (12 * x^2) - 1 / (12 * (x + 1)^2))" by (simp add: field_simps) have "D x \ (1 / (12 * x^2) - 1 / (12 * (x + 1)^2))" proof (intro sums_le [OF _ sum3 sum4] allI) fix n :: nat define c' :: "nat \ real" where "c' = (\n. if odd n \ n = 0 then 0 else if n = 2 then 4/3 else 2)" show "c n * y ^ Suc n \ d n / 3 * y ^ Suc n" proof (intro mult_right_mono) have "c n \ c' n" by (simp add: c_def c'_def) also consider "n = 0" | "n = 1" | "n = 2" | "n \ 3" by force hence "c' n \ d n / 3" by cases (simp_all add: c'_def d_def) finally show "c n \ d n / 3" . qed (insert y, simp) qed with \D x \ 0\ show ?thesis by simp qed text \ The following functions correspond to the $p_n(x)$ from the article. The special case $n = 0$ would not, strictly speaking, be necessary, but it allows some theorems to work even without the precondition $n \neq 0$: \ private definition p :: "nat \ real \ real" where "p n x = (if n = 0 then 1/x else (\r We can write the Digamma function in terms of @{term S'}: \ private lemma S'_LIMSEQ_Digamma: assumes x: "x \ 0" shows "(\n. ln (real n) - S' n x - 1/(2*x)) \ Digamma x" proof - define c where "c = (\n. ln (real n) - (\rn. 1 / (2 * (x + real n)) = c n - (ln (real n) - S' n x - 1/(2*x))) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have "c n - (ln (real n) - S' n x - 1/(2*x)) = -(\rr=1..r=1..rn. 1 / (2 * (x + real n))) \ 0" by (rule real_tendsto_divide_at_top tendsto_const filterlim_tendsto_pos_mult_at_top filterlim_tendsto_add_at_top filterlim_real_sequentially | simp)+ ultimately have "(\n. c n - (ln (real n) - S' n x - 1/(2*x))) \ 0" by (blast intro: Lim_transform_eventually) from tendsto_minus[OF this] have "(\n. (ln (real n) - S' n x - 1/(2*x)) - c n) \ 0" by simp moreover from Digamma_LIMSEQ[OF x] have "c \ Digamma x" by (simp add: c_def) ultimately show "(\n. ln (real n) - S' n x - 1/(2*x)) \ Digamma x" by (rule Lim_transform [rotated]) qed text \ Moreover, we can give an expansion of @{term S'} with the @{term p} as variation terms. \ private lemma S'_approx: "S' n x = ln (real n + x) - ln x + p n x" proof (cases "n = 0") case True thus ?thesis by (simp add: p_def S'_def) next case False hence "S' n x = (\r = (\r = (\rr We define the limit of the @{term p} (simply called $p(x)$ in Jameson's article): \ private definition P :: "real \ real" where "P x = (\n. D (real n + x))" private lemma D_summable: assumes x: "x > 0" shows "summable (\n. D (real n + x))" proof - have *: "summable (\n. 1 / (12 * (x + real n)\<^sup>2) - 1 / (12 * (x + real (Suc n))\<^sup>2))" by (rule telescope_summable' real_tendsto_divide_at_top tendsto_const filterlim_tendsto_pos_mult_at_top filterlim_pow_at_top filterlim_tendsto_add_at_top filterlim_real_sequentially | simp)+ show "summable (\n. D (real n + x))" proof (rule summable_comparison_test[OF _ *], rule exI[of _ 2], safe) fix n :: nat assume "n \ 2" show "norm (D (real n + x)) \ 1 / (12 * (x + real n)^2) - 1 / (12 * (x + real (Suc n))^2)" using stirling_trapezium[of "real n + x"] x by (auto simp: algebra_simps) qed qed private lemma p_LIMSEQ: assumes x: "x > 0" shows "(\n. p n x) \ P x" proof (rule Lim_transform_eventually) from D_summable[OF x] have "(\n. D (real n + x)) sums P x" unfolding P_def by (simp add: sums_iff) then show "(\n. \r P x" by (simp add: sums_def) moreover from eventually_gt_at_top[of 1] show "eventually (\n. (\r This gives us an expansion of the Digamma function: \ lemma Digamma_approx: assumes x: "(x :: real) > 0" shows "Digamma x = ln x - 1 / (2 * x) - P x" proof - have "eventually (\n. ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x = ln (real n) - S' n x - 1/(2*x)) at_top" using eventually_gt_at_top[of "1::nat"] proof eventually_elim fix n :: nat assume n: "n > 1" have "ln (real n) - S' n x = ln ((real n) / (real n + x)) + ln x - p n x" using assms n unfolding S'_approx by (subst ln_div) (auto simp: algebra_simps) also from n have "real n / (real n + x) = inverse (1 + x / real n)" by (simp add: field_simps) finally show "ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x = ln (real n) - S' n x - 1/(2*x)" by simp qed moreover have "(\n. ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x) \ ln (inverse (1 + 0)) + ln x - 1/(2*x) - P x" by (rule tendsto_intros p_LIMSEQ x real_tendsto_divide_at_top filterlim_real_sequentially | simp)+ hence "(\n. ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x) \ ln x - 1/(2*x) - P x" by simp ultimately have "(\n. ln (real n) - S' n x - 1 / (2 * x)) \ ln x - 1/(2*x) - P x" by (blast intro: Lim_transform_eventually) moreover from x have "(\n. ln (real n) - S' n x - 1 / (2 * x)) \ Digamma x" by (intro S'_LIMSEQ_Digamma) simp_all ultimately show "Digamma x = ln x - 1 / (2 * x) - P x" by (rule LIMSEQ_unique [rotated]) qed text \ Next, we derive some bounds on @{term "P"}: \ private lemma p_ge_0: "x > 0 \ p n x \ 0" using stirling_trapezium[of "real n + x" for n] by (auto simp add: p_def intro!: sum_nonneg) private lemma P_ge_0: "x > 0 \ P x \ 0" by (rule tendsto_lowerbound[OF p_LIMSEQ]) (insert p_ge_0[of x], simp_all) private lemma p_upper_bound: assumes "x > 0" "n > 0" shows "p n x \ 1/(12*x^2)" proof - from assms have "p n x = (\r \ (\r = 1 / (12 * x\<^sup>2) - 1 / (12 * (real n + x)\<^sup>2)" by (subst sum_lessThan_telescope') simp also have "\ \ 1 / (12 * x^2)" by simp finally show ?thesis . qed private lemma P_upper_bound: assumes "x > 0" shows "P x \ 1/(12*x^2)" proof (rule tendsto_upperbound) show "eventually (\n. p n x \ 1 / (12 * x^2)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (use p_upper_bound[of x] assms in auto) show "(\n. p n x) \ P x" by (simp add: assms p_LIMSEQ) qed auto text \ We can now use this approximation of the Digamma function to approximate @{term ln_Gamma} (since the former is the derivative of the latter). We therefore define the function $g$ from Jameson's article, which measures the difference between @{term ln_Gamma} and its approximation: \ private definition g :: "real \ real" where "g x = ln_Gamma x - (x - 1/2) * ln x + x" private lemma DERIV_g: "x > 0 \ (g has_field_derivative -P x) (at x)" unfolding g_def [abs_def] using Digamma_approx[of x] by (auto intro!: derivative_eq_intros simp: field_simps) private lemma isCont_P: assumes "x > 0" shows "isCont P x" proof - define D' :: "real \ real" where "D' = (\x. - 1 / (2 * x^2 * (x+1)^2))" have DERIV_D: "(D has_field_derivative D' x) (at x)" if "x > 0" for x unfolding D_def [abs_def] D'_def T_def by (insert that, (rule derivative_eq_intros refl | simp)+) (simp add: power2_eq_square divide_simps, (simp add: algebra_simps)?) note this [THEN DERIV_chain2, derivative_intros] have "(P has_field_derivative (\n. D' (real n + x))) (at x)" unfolding P_def [abs_def] proof (rule has_field_derivative_series') show "convex {x/2<..}" by simp next fix n :: nat and y :: real assume y: "y \ {x/2<..}" with assms have "y > 0" by simp thus "((\a. D (real n + a)) has_real_derivative D' (real n + y)) (at y within {x/2<..})" by (auto intro!: derivative_eq_intros) next from assms D_summable[of x] show "summable (\n. D (real n + x))" by simp next show "uniformly_convergent_on {x/2<..} (\n x. \i {x/2<..}" with assms have "y > 0" by auto have "norm (D' (real n + y)) = (1 / (2 * (y + real n)^2)) * (1 / (y + real (Suc n))^2)" by (simp add: D'_def add_ac) also from y assms have "\ \ (1 / (2 * (x/2)^2)) * (1 / (real (Suc n))^2)" by (intro mult_mono divide_left_mono power_mono) simp_all also have "1 / (real (Suc n))^2 = inverse ((real (Suc n))^2)" by (simp add: field_simps) finally show "norm (D' (real n + y)) \ (1 / (2 * (x/2)^2)) * inverse ((real (Suc n))^2)" . next show "summable (\n. (1 / (2 * (x/2)^2)) * inverse ((real (Suc n))^2))" by (subst summable_Suc_iff, intro summable_mult inverse_power_summable) simp_all qed qed (insert assms, simp_all add: interior_open) thus ?thesis by (rule DERIV_isCont) qed private lemma P_continuous_on [THEN continuous_on_subset]: "continuous_on {0<..} P" by (intro continuous_at_imp_continuous_on ballI isCont_P) auto private lemma P_integrable: assumes a: "a > 0" shows "P integrable_on {a..}" proof - define f where "f = (\n x. if x \ {a..real n} then P x else 0)" show "P integrable_on {a..}" proof (rule dominated_convergence) fix n :: nat from a have "P integrable_on {a..real n}" by (intro integrable_continuous_real P_continuous_on) auto hence "f n integrable_on {a..real n}" by (rule integrable_eq) (simp add: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat show "norm (f n x) \ of_real (1/12) * (1 / x^2)" if "x \ {a..}" for x using a P_ge_0 P_upper_bound by (auto simp: f_def) next show "(\x::real. of_real (1/12) * (1 / x^2)) integrable_on {a..}" using has_integral_inverse_power_to_inf[of 2 a] a by (intro integrable_on_cmult_left) auto next show "(\n. f n x) \ P x" if "x\{a..}" for x proof - have "eventually (\n. real n \ x) at_top" using filterlim_real_sequentially by (simp add: filterlim_at_top) with that have "eventually (\n. f n x = P x) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ P x" by (simp add: tendsto_eventually) qed qed qed private definition c :: real where "c = integral {1..} (\x. -P x) + g 1" text \ We can now give bounds on @{term g}: \ private lemma g_bounds: assumes x: "x \ 1" shows "g x \ {c..c + 1/(12*x)}" proof - from assms have int_nonneg: "integral {x..} P \ 0" by (intro Henstock_Kurzweil_Integration.integral_nonneg P_integrable) (auto simp: P_ge_0) have int_upper_bound: "integral {x..} P \ 1/(12*x)" proof (rule has_integral_le) from x show "(P has_integral integral {x..} P) {x..}" by (intro integrable_integral P_integrable) simp_all from x has_integral_mult_right[OF has_integral_inverse_power_to_inf[of 2 x], of "1/12"] show "((\x. 1/(12*x^2)) has_integral (1/(12*x))) {x..}" by (simp add: field_simps) qed (insert P_upper_bound x, simp_all) note DERIV_g [THEN DERIV_chain2, derivative_intros] from assms have int1: "((\x. -P x) has_integral (g x - g 1)) {1..x}" by (intro fundamental_theorem_of_calculus) - (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] + (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric] intro!: derivative_eq_intros) from x have int2: "((\x. -P x) has_integral integral {x..} (\x. -P x)) {x..}" by (intro integrable_integral integrable_neg P_integrable) simp_all from has_integral_Un[OF int1 int2] x have "((\x. - P x) has_integral g x - g 1 - integral {x..} P) ({1..x} \ {x..})" by (simp add: max_def) also from x have "{1..x} \ {x..} = {1..}" by auto finally have "((\x. -P x) has_integral g x - g 1 - integral {x..} P) {1..}" . moreover have "((\x. -P x) has_integral integral {1..} (\x. -P x)) {1..}" by (intro integrable_integral integrable_neg P_integrable) simp_all ultimately have "g x - g 1 - integral {x..} P = integral {1..} (\x. -P x)" by (simp add: has_integral_unique) hence "g x = c + integral {x..} P" by (simp add: c_def algebra_simps) with int_upper_bound int_nonneg show "g x \ {c..c + 1/(12*x)}" by simp qed text \ Finally, we have bounds on @{term ln_Gamma}, @{term Gamma}, and @{term fact}. \ private lemma ln_Gamma_bounds_aux: "x \ 1 \ ln_Gamma x \ c + (x - 1/2) * ln x - x" "x \ 1 \ ln_Gamma x \ c + (x - 1/2) * ln x - x + 1/(12*x)" using g_bounds[of x] by (simp_all add: g_def) private lemma Gamma_bounds_aux: assumes x: "x \ 1" shows "Gamma x \ exp c * x powr (x - 1/2) / exp x" "Gamma x \ exp c * x powr (x - 1/2) / exp x * exp (1/(12*x))" proof - have "exp (ln_Gamma x) \ exp (c + (x - 1/2) * ln x - x)" by (subst exp_le_cancel_iff, rule ln_Gamma_bounds_aux) (simp add: x) with x show "Gamma x \ exp c * x powr (x - 1/2) / exp x" by (simp add: Gamma_real_pos_exp exp_add exp_diff powr_def del: exp_le_cancel_iff) next have "exp (ln_Gamma x) \ exp (c + (x - 1/2) * ln x - x + 1/(12*x))" by (subst exp_le_cancel_iff, rule ln_Gamma_bounds_aux) (simp add: x) with x show "Gamma x \ exp c * x powr (x - 1/2) / exp x * exp (1/(12*x))" by (simp add: Gamma_real_pos_exp exp_add exp_diff powr_def del: exp_le_cancel_iff) qed private lemma Gamma_asymp_equiv_aux: "Gamma \[at_top] (\x. exp c * x powr (x - 1/2) / exp x)" proof (rule asymp_equiv_sandwich) include asymp_equiv_notation show "eventually (\x. exp c * x powr (x - 1/2) / exp x \ Gamma x) at_top" "eventually (\x. exp c * x powr (x - 1/2) / exp x * exp (1/(12*x)) \ Gamma x) at_top" using eventually_ge_at_top[of "1::real"] by (eventually_elim; use Gamma_bounds_aux in force)+ have "((\x::real. exp (1 / (12 * x))) \ exp 0) at_top" by (rule tendsto_intros real_tendsto_divide_at_top filterlim_tendsto_pos_mult_at_top)+ (simp_all add: filterlim_ident) hence "(\x. exp (1 / (12 * x))) \ (\x. 1 :: real)" by (intro asymp_equivI') simp_all hence "(\x. exp c * x powr (x - 1 / 2) / exp x * 1) \ (\x. exp c * x powr (x - 1 / 2) / exp x * exp (1 / (12 * x)))" by (intro asymp_equiv_mult asymp_equiv_refl) (simp add: asymp_equiv_sym) thus "(\x. exp c * x powr (x - 1 / 2) / exp x) \ (\x. exp c * x powr (x - 1 / 2) / exp x * exp (1 / (12 * x)))" by simp qed simp_all private lemma exp_1_powr_real [simp]: "exp (1::real) powr x = exp x" by (simp add: powr_def) private lemma fact_asymp_equiv_aux: "fact \[at_top] (\x. exp c * sqrt (real x) * (real x / exp 1) powr real x)" proof - include asymp_equiv_notation have "fact \ (\n. Gamma (real (Suc n)))" by (simp add: Gamma_fact) also have "eventually (\n. Gamma (real (Suc n)) = real n * Gamma (real n)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (insert Gamma_plus1[of "real n" for n], auto simp: add_ac of_nat_in_nonpos_Ints_iff) also have "(\n. Gamma (real n)) \ (\n. exp c * real n powr (real n - 1/2) / exp (real n))" by (rule asymp_equiv_compose'[OF Gamma_asymp_equiv_aux] filterlim_real_sequentially)+ also have "eventually (\n. real n * (exp c * real n powr (real n - 1 / 2) / exp (real n)) = exp c * sqrt (real n) * (real n / exp 1) powr real n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" thus "real n * (exp c * real n powr (real n - 1 / 2) / exp (real n)) = exp c * sqrt (real n) * (real n / exp 1) powr real n" by (subst powr_diff) (simp_all add: powr_divide powr_half_sqrt field_simps) qed finally show ?thesis by - (simp_all add: asymp_equiv_mult) qed text \ We still need to determine the constant term @{term c}, which we do using Wallis' product formula: $$\prod_{n=1}^\infty \frac{4n^2}{4n^2-1} = \frac{\pi}{2}$$ \ private lemma powr_mult_2: "(x::real) > 0 \ x powr (y * 2) = (x^2) powr y" by (subst mult.commute, subst powr_powr [symmetric]) (simp add: powr_numeral) private lemma exp_mult_2: "exp (y * 2 :: real) = exp y * exp y" by (subst exp_add [symmetric]) simp private lemma exp_c: "exp c = sqrt (2*pi)" proof - include asymp_equiv_notation define p where "p = (\n. \k=1..n. (4*real k^2) / (4*real k^2 - 1))" have p_0 [simp]: "p 0 = 1" by (simp add: p_def) have p_Suc: "p (Suc n) = p n * (4 * real (Suc n)^2) / (4 * real (Suc n)^2 - 1)" for n unfolding p_def by (subst prod.nat_ivl_Suc') simp_all have p: "p = (\n. 16 ^ n * fact n ^ 4 / (fact (2 * n))\<^sup>2 / (2 * real n + 1))" proof fix n :: nat have "p n = (\k=1..n. (2*real k)^2 / (2*real k - 1) / (2 * real k + 1))" unfolding p_def by (intro prod.cong refl) (simp add: field_simps power2_eq_square) also have "\ = (\k=1..n. (2*real k)^2 / (2*real k - 1)) / (\k=1..n. (2 * real (Suc k) - 1))" by (simp add: prod_dividef prod.distrib add_ac) also have "(\k=1..n. (2 * real (Suc k) - 1)) = (\k=Suc 1..Suc n. (2 * real k - 1))" by (subst prod.atLeast_Suc_atMost_Suc_shift) simp_all also have "\ = (\k=1..n. (2 * real k - 1)) * (2 * real n + 1)" by (induction n) (simp_all add: prod.nat_ivl_Suc') also have "(\k = 1..n. (2 * real k)\<^sup>2 / (2 * real k - 1)) / \ = (\k = 1..n. (2 * real k)^2 / (2 * real k - 1)^2) / (2 * real n + 1)" unfolding power2_eq_square by (simp add: prod.distrib prod_dividef) also have "(\k = 1..n. (2 * real k)^2 / (2 * real k - 1)^2) = (\k = 1..n. (2 * real k)^4 / ((2*real k)*(2 * real k - 1))^2)" by (rule prod.cong) (simp_all add: power2_eq_square eval_nat_numeral) also have "\ = 16^n * fact n^4 / (\k=1..n. (2*real k) * (2*real k - 1))^2" by (simp add: prod.distrib prod_dividef fact_prod prod_power_distrib [symmetric] prod_constant) also have "(\k=1..n. (2*real k) * (2*real k - 1)) = fact (2*n)" by (induction n) (simp_all add: algebra_simps prod.nat_ivl_Suc') finally show "p n = 16 ^ n * fact n ^ 4 / (fact (2 * n))\<^sup>2 / (2 * real n + 1)" . qed have "p \ (\n. 16 ^ n * fact n ^ 4 / (fact (2 * n))\<^sup>2 / (2 * real n + 1))" by (simp add: p) also have "\ \ (\n. 16^n * (exp c * sqrt (real n) * (real n / exp 1) powr real n)^4 / (exp c * sqrt (real (2*n)) * (real (2*n) / exp 1) powr real (2*n))^2 / (2 * real n + 1))" (is "_ \ ?f") by (intro asymp_equiv_mult asymp_equiv_divide asymp_equiv_refl mult_nat_left_at_top fact_asymp_equiv_aux asymp_equiv_power asymp_equiv_compose'[OF fact_asymp_equiv_aux]) simp_all also have "eventually (\n. \ n = exp c ^ 2 / (4 + 2/n)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have [simp]: "16^n = 4^n * (4^n :: real)" by (simp add: power_mult_distrib [symmetric]) from n have "?f n = exp c ^ 2 * (n / (2*(2*n+1)))" by (simp add: power_mult_distrib divide_simps powr_mult real_sqrt_power_even) (simp add: field_simps power2_eq_square eval_nat_numeral powr_mult_2 exp_mult_2 powr_realpow) also from n have "\ = exp c ^ 2 / (4 + 2/n)" by (simp add: field_simps) finally show "?f n = \" . qed also have "(\x. 4 + 2 / real x) \ (\x. 4)" by (subst asymp_equiv_add_right) auto finally have "p \ exp c ^ 2 / 4" by (rule asymp_equivD_const) (simp_all add: asymp_equiv_divide) moreover have "p \ pi / 2" unfolding p_def by (rule wallis) ultimately have "exp c ^ 2 / 4 = pi / 2" by (rule LIMSEQ_unique) hence "2 * pi = exp c ^ 2" by simp also have "sqrt (exp c ^ 2) = exp c" by simp finally show "exp c = sqrt (2 * pi)" .. qed private lemma c: "c = ln (2*pi) / 2" proof - note exp_c [symmetric] also have "ln (exp c) = c" by simp finally show ?thesis by (simp add: ln_sqrt) qed text \ This gives us the final bounds: \ theorem Gamma_bounds: assumes "x \ 1" shows "Gamma x \ sqrt (2*pi/x) * (x / exp 1) powr x" (is ?th1) "Gamma x \ sqrt (2*pi/x) * (x / exp 1) powr x * exp (1 / (12 * x))" (is ?th2) proof - from assms have "exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x" by (subst powr_diff) (simp add: exp_c real_sqrt_divide powr_divide powr_half_sqrt field_simps) with Gamma_bounds_aux[OF assms] show ?th1 ?th2 by simp_all qed theorem ln_Gamma_bounds: assumes "x \ 1" shows "ln_Gamma x \ ln (2*pi/x) / 2 + x * ln x - x" (is ?th1) "ln_Gamma x \ ln (2*pi/x) / 2 + x * ln x - x + 1/(12*x)" (is ?th2) proof - from ln_Gamma_bounds_aux[OF assms] assms show ?th1 ?th2 by (simp_all add: c field_simps ln_div) from assms have "exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x" by (subst powr_diff) (simp add: exp_c real_sqrt_divide powr_divide powr_half_sqrt field_simps) qed theorem fact_bounds: assumes "n > 0" shows "(fact n :: real) \ sqrt (2*pi*n) * (n / exp 1) ^ n" (is ?th1) "(fact n :: real) \ sqrt (2*pi*n) * (n / exp 1) ^ n * exp (1 / (12 * n))" (is ?th2) proof - from assms have n: "real n \ 1" by simp from assms Gamma_plus1[of "real n"] have "real n * Gamma (real n) = Gamma (real (Suc n))" by (simp add: of_nat_in_nonpos_Ints_iff add_ac) also have "Gamma (real (Suc n)) = fact n" by (subst Gamma_fact [symmetric]) simp finally have *: "fact n = real n * Gamma (real n)" by simp have "2*pi/n = 2*pi*n / n^2" by (simp add: power2_eq_square) also have "sqrt \ = sqrt (2*pi*n) / n" by (subst real_sqrt_divide) simp_all also have "real n * \ = sqrt (2*pi*n)" by simp finally have **: "real n * sqrt (2*pi/real n) = sqrt (2*pi*real n)" . note * also note Gamma_bounds(2)[OF n] also have "real n * (sqrt (2 * pi / real n) * (real n / exp 1) powr real n * exp (1 / (12 * real n))) = (real n * sqrt (2*pi/n)) * (n / exp 1) powr n * exp (1 / (12 * n))" by (simp add: algebra_simps) also from n have "(real n / exp 1) powr real n = (real n / exp 1) ^ n" by (subst powr_realpow) simp_all also note ** finally show ?th2 by - (insert assms, simp_all) have "sqrt (2*pi*n) * (n / exp 1) powr n = n * (sqrt (2*pi/n) * (n / exp 1) powr n)" by (subst ** [symmetric]) (simp add: field_simps) also from assms have "\ \ real n * Gamma (real n)" by (intro mult_left_mono Gamma_bounds(1)) simp_all also from n have "(real n / exp 1) powr real n = (real n / exp 1) ^ n" by (subst powr_realpow) simp_all also note * [symmetric] finally show ?th1 . qed theorem ln_fact_bounds: assumes "n > 0" shows "ln (fact n :: real) \ ln (2*pi*n)/2 + n * ln n - n" (is ?th1) "ln (fact n :: real) \ ln (2*pi*n)/2 + n * ln n - n + 1/(12*real n)" (is ?th2) proof - have "ln (fact n :: real) \ ln (sqrt (2*pi*real n)*(real n/exp 1)^n)" using fact_bounds(1)[OF assms] assms by (subst ln_le_cancel_iff) auto also from assms have "ln (sqrt (2*pi*real n)*(real n/exp 1)^n) = ln (2*pi*n)/2 + n * ln n - n" by (simp add: ln_mult ln_div ln_realpow ln_sqrt field_simps) finally show ?th1 . next have "ln (fact n :: real) \ ln (sqrt (2*pi*real n) * (real n/exp 1)^n * exp (1/(12*real n)))" using fact_bounds(2)[OF assms] assms by (subst ln_le_cancel_iff) auto also from assms have "\ = ln (2*pi*n)/2 + n * ln n - n + 1/(12*real n)" by (simp add: ln_mult ln_div ln_realpow ln_sqrt field_simps) finally show ?th2 . qed theorem Gamma_asymp_equiv: "Gamma \[at_top] (\x. sqrt (2*pi/x) * (x / exp 1) powr x :: real)" proof - note Gamma_asymp_equiv_aux also have "eventually (\x. exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim fix x :: real assume x: "x > 0" thus "exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x" by (subst powr_diff) (simp add: exp_c powr_half_sqrt powr_divide field_simps real_sqrt_divide) qed finally show ?thesis . qed theorem fact_asymp_equiv: "fact \[at_top] (\n. sqrt (2*pi*n) * (n / exp 1) ^ n :: real)" proof - note fact_asymp_equiv_aux also have "eventually (\n. exp c * sqrt (real n) = sqrt (2 * pi * real n)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: exp_c real_sqrt_mult) also have "eventually (\n. (n / exp 1) powr n = (n / exp 1) ^ n) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: powr_realpow) finally show ?thesis . qed end end diff --git a/thys/UpDown_Scheme/Triangular_Function.thy b/thys/UpDown_Scheme/Triangular_Function.thy --- a/thys/UpDown_Scheme/Triangular_Function.thy +++ b/thys/UpDown_Scheme/Triangular_Function.thy @@ -1,422 +1,422 @@ section \ Hat Functions \ theory Triangular_Function imports "HOL-Analysis.Equivalence_Lebesgue_Henstock_Integration" Grid begin lemma continuous_on_max[continuous_intros]: fixes f :: "_ \ 'a::linorder_topology" shows "continuous_on S f \ continuous_on S g \ continuous_on S (\x. max (f x) (g x))" by (auto simp: continuous_on_def intro: tendsto_max) definition \ :: "(nat \ int) \ real \ real" where "\ \ (\(l,i) x. max 0 (1 - \ x * 2^(l + 1) - real_of_int i \))" definition \ :: "(nat \ int) list \ (nat \ real) \ real" where "\ p x = (\d (p ! d) (x d))" definition l2_\ where "l2_\ p1 p2 = (\x. \ p1 x * \ p2 x \lborel)" definition l2 where "l2 a b = (\x. \ a x * \ b x \(\\<^sub>M d\{..[measurable]: "\ p \ borel_measurable borel" by (cases p) (simp add: \_def) lemma \_nonneg: "0 \ \ p x" by (simp add: \_def split: prod.split) lemma \_zero_iff: "\ (l,i) x = 0 \ x \ {real_of_int (i - 1) / 2^(l + 1) <..< real_of_int (i + 1) / 2^(l + 1)}" by (auto simp: \_def field_simps split: split_max) lemma \_zero: "x \ {real_of_int (i - 1) / 2^(l + 1) <..< real_of_int (i + 1) / 2^(l + 1)} \ \ (l,i) x = 0" unfolding \_zero_iff by simp lemma \_eq_0: assumes x: "x < 0 \ 1 < x" and i: "0 < i" "i < 2^Suc l" shows "\ (l, i) x = 0" using x proof assume "x < 0" also have "0 \ real_of_int (i - 1) / 2^(l + 1)" using i by (auto simp: field_simps) finally show ?thesis by (auto intro!: \_zero simp: field_simps) next have "real_of_int (i + 1) / 2^(l + 1) \ 1" using i by (subst divide_le_eq_1_pos) (auto simp del: of_int_add power_Suc) also assume "1 < x" finally show ?thesis by (auto intro!: \_zero simp: field_simps) qed lemma ix_lt: "p \ sparsegrid dm lm \ d < dm \ ix p d < 2^(lv p d + 1)" unfolding sparsegrid_def lgrid_def using grid_estimate[of d "start dm" p "{0 ..< dm}"] by auto lemma ix_gt: "p \ sparsegrid dm lm \ d < dm \ 0 < ix p d" unfolding sparsegrid_def lgrid_def using grid_estimate[of d "start dm" p "{0 ..< dm}"] by auto lemma \_eq_0: assumes x: "\d 1 < x d" and p: "p \ sparsegrid dm lm" shows "\ p x = 0" unfolding \_def proof (rule prod_zero) from x obtain d where "d < length p \ (x d < 0 \ 1 < x d)" .. with p[THEN ix_lt, of d] p[THEN ix_gt, of d] p show "\a\{.. (p ! a) (x a) = 0" apply (cases "p!d") apply (intro bexI[of _ d]) apply (auto intro!: \_eq_0 simp: sparsegrid_length ix_def lv_def) done qed simp lemma \_left_support': "x \ {real_of_int (i - 1) / 2^(l + 1) .. real_of_int i / 2^(l + 1)} \ \ (l,i) x = 1 + x * 2^(l + 1) - real_of_int i" by (auto simp: \_def field_simps split: split_max) lemma \_left_support: "x \ {-1 .. 0::real} \ \ (l,i) ((x + real_of_int i) / 2^(l + 1)) = 1 + x" by (auto simp: \_def field_simps split: split_max) lemma \_right_support': "x \ {real_of_int i / 2^(l + 1) .. real_of_int (i + 1) / 2^(l + 1)} \ \ (l,i) x = 1 - x * 2^(l + 1) + real_of_int i" by (auto simp: \_def field_simps split: split_max) lemma \_right_support: "x \ {0 .. 1::real} \ \ (l,i) ((x + real i) / 2^(l + 1)) = 1 - x" by (auto simp: \_def field_simps split: split_max) lemma integrable_\: "integrable lborel (\ p)" proof (induct p) case (Pair l i) have "integrable lborel (\x. indicator {real_of_int (i - 1) / 2^(l + 1) .. real_of_int (i + 1) / 2^(l + 1)} x *\<^sub>R \ (l, i) x)" unfolding \_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?case - by (rule integrable_cong[THEN iffD1, rotated -1]) (auto simp: \_zero_iff) + by (rule Bochner_Integration.integrable_cong[THEN iffD1, rotated -1]) (auto simp: \_zero_iff) qed lemma integrable_\2: "integrable lborel (\x. \ p x * \ q x)" proof (cases p q rule: prod.exhaust[case_product prod.exhaust]) case (Pair_Pair l i l' i') have "integrable lborel (\x. indicator {real_of_int (i - 1) / 2^(l + 1) .. real_of_int (i + 1) / 2^(l + 1)} x *\<^sub>R (\ (l, i) x * \ (l', i') x))" unfolding \_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?thesis unfolding Pair_Pair - by (rule integrable_cong[THEN iffD1, rotated -1]) (auto simp: \_zero_iff) + by (rule Bochner_Integration.integrable_cong[THEN iffD1, rotated -1]) (auto simp: \_zero_iff) qed lemma l2_\I_DERIV: assumes n: "\ x. x \ { (real_of_int i' - 1) / 2^(l' + 1) .. real_of_int i' / 2^(l' + 1) } \ DERIV \_n x :> (\ (l', i') x * \ (l, i) x)" (is "\ x. x \ {?a..?b} \ DERIV _ _ :> ?P x") and p: "\ x. x \ { real_of_int i' / 2^(l' + 1) .. (real_of_int i' + 1) / 2^(l' + 1) } \ DERIV \_p x :> (\ (l', i') x * \ (l, i) x)" (is "\ x. x \ {?b..?c} \ _") shows "l2_\ (l', i') (l, i) = (\_n ?b - \_n ?a) + (\_p ?c - \_p ?b)" proof - have "has_bochner_integral lborel (\x. ?P x * indicator {?a..?b} x + ?P x * indicator {?b..?c} x) ((\_n ?b - \_n ?a) + (\_p ?c - \_p ?b))" by (intro has_bochner_integral_add has_bochner_integral_FTC_Icc_nonneg n p) (auto simp: \_nonneg field_simps) then have "has_bochner_integral lborel?P ((\_n ?b - \_n ?a) + (\_p ?c - \_p ?b))" by (rule has_bochner_integral_discrete_difference[where X="{?b}", THEN iffD1, rotated -1]) (auto simp: power_add intro!: \_zero integral_cong split: split_indicator) then show ?thesis by (simp add: has_bochner_integral_iff l2_\_def) qed lemma l2_eq: "length a = length b \ l2 a b = (\d (a!d) (b!d))" unfolding l2_def l2_\_def \_def apply (simp add: prod.distrib[symmetric]) proof (rule product_sigma_finite.product_integral_prod) show "product_sigma_finite (\d. lborel)" .. qed (auto intro: integrable_\2) lemma l2_when_disjoint: assumes "l \ l'" defines "d == l' - l" assumes "(i + 1) * 2^d < i' \ i' < (i - 1) * 2^d" (is "?right \ ?left") shows "l2_\ (l', i') (l, i) = 0" proof - let ?sup = "\l i. {real_of_int (i - 1) / 2^(l + 1) <..< real_of_int (i + 1) / 2^(l + 1)}" have l': "l' = l + d" using assms by simp have *: "\i l. 2 ^ l = real_of_int (2 ^ l::int)" by simp have [arith]: "0 < (2^d::int)" by simp from \?right \ ?left\ \l \ l'\ have empty_support: "?sup l i \ ?sup l' i' = {}" by (auto simp add: min_def max_def divide_simps l' power_add * of_int_mult[symmetric] simp del: of_int_diff of_int_add of_int_mult of_int_power) (simp_all add: field_simps) then have "\x. \ (l', i') x * \ (l, i) x = 0" unfolding \_zero_iff mult_eq_0_iff by blast then show ?thesis by (simp add: l2_\_def del: mult_eq_0_iff vector_space_over_itself.scale_eq_0_iff) qed lemma l2_commutative: "l2_\ p q = l2_\ q p" by (simp add: l2_\_def mult.commute) lemma l2_when_same: "l2_\ (l, i) (l, i) = 1/3 / 2^l" proof (subst l2_\I_DERIV) let ?l = "(2 :: real)^(l + 1)" let ?in = "real_of_int i - 1" let ?ip = "real_of_int i + 1" let ?\ = "\ (l,i)" let ?\2 = "\x. ?\ x * ?\ x" { fix x assume "x \ {?in / ?l .. real_of_int i / ?l}" hence \_eq: "?\ x = ?l * x - ?in" using \_left_support' by auto show "DERIV (\x. x^3 / 3 * ?l^2 + x * ?in^2 - x^2/2 * 2 * ?l * ?in) x :> ?\2 x" by (auto intro!: derivative_eq_intros simp add: power2_eq_square field_simps \_eq) } { fix x assume "x \ {real_of_int i / ?l .. ?ip / ?l}" hence \_eq: "?\ x = ?ip - ?l * x" using \_right_support' by auto show "DERIV (\x. x^3 / 3 * ?l^2 + x * ?ip^2 - x^2/2 * 2 * ?l * ?ip) x :> ?\2 x" by (auto intro!: derivative_eq_intros simp add: power2_eq_square field_simps \_eq) } qed (simp_all add: field_simps power_eq_if[of _ 2] power_eq_if[of _ 3]) lemma l2_when_left_child: assumes "l < l'" and i'_bot: "i' > (i - 1) * 2^(l' - l)" and i'_top: "i' < i * 2^(l' - l)" shows "l2_\ (l', i') (l, i) = (1 + real_of_int i' / 2^(l' - l) - real_of_int i) / 2^(l' + 1)" proof (subst l2_\I_DERIV) let ?l' = "(2 :: real)^(l' + 1)" let ?in' = "real_of_int i' - 1" let ?ip' = "real_of_int i' + 1" let ?l = "(2 :: real)^(l + 1)" let ?i = "real_of_int i - 1" let ?\' = "\ (l',i')" let ?\ = "\ (l, i)" let "?\2 x" = "?\' x * ?\ x" define \_n where "\_n x = x^3 / 3 * ?l' * ?l + x * ?i * ?in' - x^2 / 2 * (?in' * ?l + ?i * ?l')" for x define \_p where "\_p x = x^2 / 2 * (?ip' * ?l + ?i * ?l') - x^3 / 3 * ?l' * ?l - x * ?i * ?ip'" for x have level_diff: "2^(l' - l) = 2^l' / (2^l :: real)" using power_diff[of "2::real" l l'] \l < l'\ by auto { fix x assume x: "x \ {?in' / ?l' .. ?ip' / ?l'}" have "?i * 2^(l' - l) \ ?in'" using i'_bot int_less_real_le by auto hence "?i / ?l \ ?in' / ?l'" using level_diff by (auto simp: field_simps) hence "?i / ?l \ x" using x by auto moreover have "?ip' \ real_of_int i * 2^(l' - l)" using i'_top int_less_real_le by auto hence ip'_le_i: "?ip' / ?l' \ real_of_int i / ?l" using level_diff by (auto simp: field_simps) hence "x \ real_of_int i / ?l" using x by auto ultimately have "?\ x = ?l * x - ?i" using \_left_support' by auto } note \_eq = this { fix x assume x: "x \ {?in' / ?l' .. real_of_int i' / ?l'}" hence \'_eq: "?\' x = ?l' * x - ?in'" using \_left_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (auto simp add: field_simps) show "DERIV \_n x :> ?\2 x" unfolding \_eq[OF x'] \'_eq \_n_def by (auto intro!: derivative_eq_intros simp add: power2_eq_square algebra_simps) } { fix x assume x: "x \ {real_of_int i' / ?l' .. ?ip' / ?l'}" hence \'_eq: "?\' x = ?ip' - ?l' * x" using \_right_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (simp add: field_simps) show "DERIV \_p x :> ?\2 x" unfolding \_eq[OF x'] \'_eq \_p_def by (auto intro!: derivative_eq_intros simp add: power2_eq_square algebra_simps) } qed (simp_all add: field_simps power_eq_if[of _ 2] power_eq_if[of _ 3] power_diff[of "2::real", OF _ \l < l'\[THEN less_imp_le]] ) lemma l2_when_right_child: assumes "l < l'" and i'_bot: "i' > i * 2^(l' - l)" and i'_top: "i' < (i + 1) * 2^(l' - l)" shows "l2_\ (l', i') (l, i) = (1 - real_of_int i' / 2^(l' - l) + real_of_int i) / 2^(l' + 1)" proof (subst l2_\I_DERIV) let ?l' = "(2 :: real)^(l' + 1)" let ?in' = "real_of_int i' - 1" let ?ip' = "real_of_int i' + 1" let ?l = "(2 :: real)^(l + 1)" let ?i = "real_of_int i + 1" let ?\' = "\ (l',i')" let ?\ = "\ (l, i)" let ?\2 = "\x. ?\' x * ?\ x" define \_n where "\_n x = x^2 / 2 * (?in' * ?l + ?i * ?l') - x^3 / 3 * ?l' * ?l - x * ?i * ?in'" for x define \_p where "\_p x = x^3 / 3 * ?l' * ?l + x * ?i * ?ip' - x^2 / 2 * (?ip' * ?l + ?i * ?l')" for x have level_diff: "2^(l' - l) = 2^l' / (2^l :: real)" using power_diff[of "2::real" l l'] \l < l'\ by auto { fix x assume x: "x \ {?in' / ?l' .. ?ip' / ?l'}" have "real_of_int i * 2^(l' - l) \ ?in'" using i'_bot int_less_real_le by auto hence "real_of_int i / ?l \ ?in' / ?l'" using level_diff by (auto simp: field_simps) hence "real_of_int i / ?l \ x" using x by auto moreover have "?ip' \ ?i * 2^(l' - l)" using i'_top int_less_real_le by auto hence ip'_le_i: "?ip' / ?l' \ ?i / ?l" using level_diff by (auto simp: field_simps) hence "x \ ?i / ?l" using x by auto ultimately have "?\ x = ?i - ?l * x" using \_right_support' by auto } note \_eq = this { fix x assume x: "x \ {?in' / ?l' .. real_of_int i' / ?l'}" hence \'_eq: "?\' x = ?l' * x - ?in'" using \_left_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (simp add: field_simps) show "DERIV \_n x :> ?\2 x" unfolding \_n_def \_eq[OF x'] \'_eq by (auto intro!: derivative_eq_intros simp add: simp add: power2_eq_square algebra_simps) } { fix x assume x: "x \ {real_of_int i' / ?l' .. ?ip' / ?l'}" hence \'_eq: "?\' x = ?ip' - ?l' * x" using \_right_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (auto simp: field_simps) show "DERIV \_p x :> ?\2 x" unfolding \_eq[OF x'] \'_eq \_p_def by (auto intro!: derivative_eq_intros simp add: power2_eq_square algebra_simps) } qed (simp_all add: field_simps power_eq_if[of _ 2] power_eq_if[of _ 3] power_diff[of "2::real", OF _ \l < l'\[THEN less_imp_le]] ) lemma level_shift: "lc > l \ (x :: real) / 2 ^ (lc - Suc l) = x * 2 / 2 ^ (lc - l)" by (auto simp add: power_diff) lemma l2_child: assumes "d < length b" and p_grid: "p \ grid (child b dir d) ds" (is "p \ grid ?child ds") shows "l2_\ (p ! d) (b ! d) = (1 - real_of_int (sgn dir) * (real_of_int (ix p d) / 2^(lv p d - lv b d) - real_of_int (ix b d))) / 2^(lv p d + 1)" proof - have "lv ?child d \ lv p d" using \d < length b\ and p_grid using grid_single_level by auto hence "lv b d < lv p d" using \d < length b\ and p_grid using child_lv by auto let ?i_c = "ix ?child d" and ?l_c = "lv ?child d" let ?i_p = "ix p d" and ?l_p = "lv p d" let ?i_b = "ix b d" and ?l_b = "lv b d" have "(2::int) * 2^(?l_p - ?l_c) = 2^Suc (?l_p - ?l_c)" by auto also have "\ = 2^(Suc ?l_p - ?l_c)" proof - have "Suc (?l_p - ?l_c) = Suc ?l_p - ?l_c" using \lv ?child d \ lv p d\ by auto thus ?thesis by auto qed also have "\ = 2^(?l_p - ?l_b)" using \d < length b\ and \lv b d < lv p d\ by (auto simp add: child_def lv_def) finally have level: "2^(?l_p - ?l_b) = (2::int) * 2^(?l_p - ?l_c)" .. from \d < length b\ and p_grid have range_left: "?i_p > (?i_c - 1) * 2^(?l_p - ?l_c)" and range_right: "?i_p < (?i_c + 1) * 2^(?l_p - ?l_c)" using grid_estimate by auto show ?thesis proof (cases dir) case left with child_ix_left[OF \d < length b\] have "(?i_b - 1) * 2^(?l_p - ?l_b) = (?i_c - 1) * 2^(?l_p - ?l_c)" and "?i_b * 2^(?l_p - ?l_b) = (?i_c + 1) * 2^(?l_p - ?l_c)" using level by auto hence "?i_p > (?i_b - 1) * 2^(?l_p - ?l_b)" and "?i_p < ?i_b * 2^(?l_p - ?l_b)" using range_left and range_right by auto with \?l_b < ?l_p\ have "l2_\ (?l_p, ?i_p) (?l_b, ?i_b) = (1 + real_of_int ?i_p / 2^(?l_p - ?l_b) - real_of_int ?i_b) / 2^(?l_p + 1)" by (rule l2_when_left_child) thus ?thesis using left by (auto simp add: ix_def lv_def) next case right hence "?i_c = 2 * ?i_b + 1" using child_ix_right and \d < length b\ by auto hence "?i_b * 2^(?l_p - ?l_b) = (?i_c - 1) * 2^(?l_p - ?l_c)" and "(?i_b + 1) * 2^(?l_p - ?l_b) = (?i_c + 1) * 2^(?l_p - ?l_c)" using level by auto hence "?i_p > ?i_b * 2^(?l_p - ?l_b)" and "?i_p < (?i_b + 1) * 2^(?l_p - ?l_b)" using range_left and range_right by auto with \?l_b < ?l_p\ have "l2_\ (?l_p, ?i_p) (?l_b, ?i_b) = (1 - real_of_int ?i_p / 2^(?l_p - ?l_b) + real_of_int ?i_b) / 2^(?l_p + 1)" by (rule l2_when_right_child) thus ?thesis using right by (auto simp add: ix_def lv_def) qed qed lemma l2_same: "l2_\ (p!d) (p!d) = 1/3 / 2^(lv p d)" proof - have "l2_\ (p!d) (p!d) = l2_\ (lv p d, ix p d) (lv p d, ix p d)" by (auto simp add: lv_def ix_def) thus ?thesis using l2_when_same by auto qed lemma l2_disjoint: assumes "d < length b" and "p \ grid b {d}" and "p' \ grid b {d}" and "p' \ grid p {d}" and "lv p' d \ lv p d" shows "l2_\ (p' ! d) (p ! d) = 0" proof - have range: "ix p' d > (ix p d + 1) * 2^(lv p' d - lv p d) \ ix p' d < (ix p d - 1) * 2^(lv p' d - lv p d)" proof (rule ccontr) assume "\ ?thesis" hence "ix p' d \ (ix p d + 1) * 2^(lv p' d - lv p d)" and "ix p' d \ (ix p d - 1) * 2^(lv p' d - lv p d)" by auto with \p' \ grid b {d}\ and \p \ grid b {d}\ and \lv p' d \ lv p d\ and \d < length b\ have "p' \ grid p {d}" using grid_part[where p=p and b=b and d=d and p'=p'] by auto with \p' \ grid p {d}\ show False by auto qed have "l2_\ (p' ! d) (p ! d) = l2_\ (lv p' d, ix p' d) (lv p d, ix p d)" by (auto simp add: ix_def lv_def) also have "\ = 0" using range and \lv p' d \ lv p d\ and l2_when_disjoint by auto finally show ?thesis . qed lemma l2_down2: fixes pc pd p assumes "d < length pd" assumes pc_in_grid: "pc \ grid (child pd dir d) {d}" assumes pd_is_child: "pd = child p dir d" (is "pd = ?pd") shows "l2_\ (pc ! d) (pd ! d) / 2 = l2_\ (pc ! d) (p ! d)" proof - have "d < length p" using pd_is_child \d < length pd\ by auto moreover have "pc \ grid ?pd {d}" using pd_is_child and grid_child and pc_in_grid by auto hence "lv p d < lv pc d" using grid_child_level and \d < length pd\ and pd_is_child by auto moreover have "real_of_int (sgn dir) * real_of_int (sgn dir) = 1" by (cases dir, auto) ultimately show ?thesis unfolding l2_child[OF \d < length pd\ pc_in_grid] l2_child[OF \d < length p\ \pc \ grid ?pd {d}\] using child_lv and child_ix and pd_is_child and level_shift by (auto simp add: algebra_simps diff_divide_distrib add_divide_distrib) qed lemma l2_zigzag: assumes "d < length p" and p_child: "p = child p_p dir d" and p'_grid: "p' \ grid (child p (inv dir) d) {d}" and ps_intro: "child p (inv dir) d = child ps dir d" (is "?c_p = ?c_ps") shows "l2_\ (p' ! d) (p_p ! d) = l2_\ (p' ! d) (ps ! d) + l2_\ (p' ! d) (p ! d) / 2" proof - have "length p = length ?c_p" by auto also have "\ = length ?c_ps" using ps_intro by auto finally have "length p = length ps" using ps_intro by auto hence "d < length p_p" using p_child and \d < length p\ by auto moreover from ps_intro have "ps = p[d := (lv p d, ix p d - sgn dir)]" by (rule child_neighbour) hence "lv ps d = lv p d" and "real_of_int (ix ps d) = real_of_int (ix p d) - real_of_int (sgn dir)" using lv_def and ix_def and \length p = length ps\ and \d < length p\ by auto moreover have "d < length ps" and *: "p' \ grid (child ps dir d) {d}" using p'_grid ps_intro \length p = length ps\ \d < length p\ by auto have "p' \ grid p {d}" using p'_grid and grid_child by auto hence p_p_grid: "p' \ grid (child p_p dir d) {d}" using p_child by auto hence "lv p' d > lv p_p d" using grid_child_level and \d < length p_p\ by auto moreover have "real_of_int (sgn dir) * real_of_int (sgn dir) = 1" by (cases dir, auto) ultimately show ?thesis unfolding l2_child[OF \d < length p\ p'_grid] l2_child[OF \d < length ps\ *] l2_child[OF \d < length p_p\ p_p_grid] using child_lv and child_ix and p_child level_shift by (auto simp add: add_divide_distrib algebra_simps diff_divide_distrib) qed end diff --git a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy --- a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy +++ b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy @@ -1,2089 +1,2089 @@ (* File: Zeta_3_Irrational.thy Author: Manuel Eberl, TU München *) section \The Irrationality of $\zeta(3)$\ theory Zeta_3_Irrational imports "E_Transcendental.E_Transcendental" "Prime_Number_Theorem.Prime_Number_Theorem" "Prime_Distribution_Elementary.PNT_Consequences" begin text \ Ap\'{e}ry's original proof of the irrationality of $\zeta(3)$ contained several tricky identities of sums involving binomial coefficients that are difficult to prove. I instead follow Beukers's proof, which consists of several fairly straightforward manipulations of integrals with absolutely no caveats. Both Ap\'{e}ry and Beukers make use of an asymptotic upper bound on $\text{lcm}\{1\ldots n\}$ -- namely $\text{lcm}\{1\ldots n\} \in o(c^n)$ for any $c > e$, which is a consequence of the Prime Number Theorem (which, fortunately, is available in the \emph{Archive of Formal Proofs}~\cite{afp_primes1,afp_primes2}). I follow the concise presentation of Beukers's proof in Filaseta's lecture notes~\cite{filaseta}, which turned out to be very amenable to formalisation. There is another earlier formalisation of the irrationality of $\zeta(3)$ by Mahboubi\ \emph{et al.}~\cite{mahboubi}, who followed Ap\'{e}ry's original proof, but were ultimately forced to find a more elementary way to prove the asymptotics of $\text{lcm}\{1\ldots n\}$ than the Prime Number Theorem. First, we will require some auxiliary material before we get started with the actual proof. \ (* TODO: A lot of this (if not all of it) should really be in the library *) subsection \Auxiliary facts about polynomials\ lemma degree_higher_pderiv: "degree ((pderiv ^^ n) p) = degree p - n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" by (induction n arbitrary: p) (auto simp: degree_pderiv) lemma pcompose_power_left: "pcompose (p ^ n) q = pcompose p q ^ n" by (induction n) (auto simp: pcompose_mult pcompose_1) lemma pderiv_sum: "pderiv (\x\A. f x) = (\x\A. pderiv (f x))" by (induction A rule: infinite_finite_induct) (auto simp: pderiv_add) lemma higher_pderiv_minus: "(pderiv ^^ n) (-p :: 'a :: idom poly) = -(pderiv ^^ n) p" by (induction n) (auto simp: pderiv_minus) lemma pderiv_power: "pderiv (p ^ n) = smult (of_nat n) (p ^ (n - 1)) * pderiv p" by (cases n) (simp_all add: pderiv_power_Suc del: power_Suc) lemma pderiv_monom: "pderiv (monom c n) = monom (of_nat n * c) (n - 1)" by (simp add: monom_altdef pderiv_smult pderiv_power pderiv_pCons mult_ac) lemma higher_pderiv_monom: "k \ n \ (pderiv ^^ k) (monom c n) = monom (of_nat (pochhammer (n - k + 1) k) * c) (n - k)" by (induction k) (auto simp: pderiv_monom pochhammer_rec Suc_diff_le Suc_diff_Suc mult_ac) lemma higher_pderiv_mult: "(pderiv ^^ n) (p * q) = (\k\n. Polynomial.smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (n - k)) q))" proof (induction n) case (Suc n) have eq: "(Suc n choose k) = (n choose k) + (n choose (k-1))" if "k > 0" for k using that by (cases k) auto have "(pderiv ^^ Suc n) (p * q) = (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q))" by (simp add: Suc pderiv_sum pderiv_smult pderiv_mult sum.distrib smult_add_right algebra_simps Suc_diff_le) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) = (\k\insert 0 {1..n}. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q" by (subst sum.insert) (auto simp: add_ac) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q)) = (\k=1..n+1. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto also have "\ = (\k\insert (n+1) {1..n}. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (pderiv ^^ Suc n) p * q" by (subst sum.insert) (auto simp: add_ac) also have "(\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + \ = (\k=1..n. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + (pderiv ^^ Suc n) p * q" by (simp add: sum.distrib algebra_simps smult_add_right eq smult_add_left) also have "\ = (\k\{1..n}\{0,Suc n}. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (subst sum.union_disjoint) (auto simp: algebra_simps) also have "{1..n}\{0,Suc n} = {..Suc n}" by auto finally show ?case . qed auto subsection \Auxiliary facts about integrals\ theorem (in pair_sigma_finite) Fubini_set_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "set_borel_measurable (M1 \\<^sub>M M2) (A \ B) f" and integ1: "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2)" and integ2: "AE x\A in M1. set_integrable M2 B (\y. f (x, y))" shows "set_integrable (M1 \\<^sub>M M2) (A \ B) f" unfolding set_integrable_def proof (rule Fubini_integrable) note integ1 also have "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2) \ integrable M1 (\x. LINT y|M2. norm (indicat_real (A \ B) (x, y) *\<^sub>R f (x, y)))" unfolding set_integrable_def by (intro integrable_cong) (auto simp: indicator_def set_lebesgue_integral_def) finally show \ . next from integ2 show "AE x in M1. integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof eventually_elim case (elim x) show "integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof (cases "x \ A") case True with elim have "set_integrable M2 B (\y. f (x, y))" by simp also have "?this \ ?thesis" unfolding set_integrable_def using True by (intro integrable_cong) (auto simp: indicator_def) finally show ?thesis . qed auto qed qed (insert assms, auto simp: set_borel_measurable_def) lemma (in pair_sigma_finite) set_integral_fst': fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\x\A. (\y\B. f (x, y) \M2) \M1)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\x. \y. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M2 \M1)" using assms by (subst integral_fst' [symmetric]) (auto simp: set_integrable_def) also have "\ = (\x\A. (\y\B. f (x,y) \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) set_integral_snd: fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\y\B. (\x\A. f (x, y) \M1) \M2)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\y. \x. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M1 \M2)" using assms by (subst integral_snd) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\y\B. (\x\A. f (x,y) \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed proposition (in pair_sigma_finite) Fubini_set_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "set_integrable (M1 \\<^sub>M M2) (A \ B) (case_prod f)" shows "(\y\B. (\x\A. f x y \M1) \M2) = (\x\A. (\y\B. f x y \M2) \M1)" proof - have "(\y\B. (\x\A. f x y \M1) \M2) = (\y. (\x. indicator (A \ B) (x, y) *\<^sub>R f x y \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) also have "\ = (\x. (\y. indicator (A \ B) (x, y) *\<^sub>R f x y \M2) \M1)" using assms by (intro Fubini_integral) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\x\A. (\y\B. f x y \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) nn_integral_swap: assumes [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+x. f x \(M1 \\<^sub>M M2)) = (\\<^sup>+(y,x). f (x,y) \(M2 \\<^sub>M M1))" by (subst distr_pair_swap, subst nn_integral_distr) (auto simp: case_prod_unfold) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "set_integrable M A f \ set_borel_measurable M A g \ (AE x in M. x \ A \ norm (g x) \ norm (f x)) \ set_integrable M A g" unfolding set_integrable_def apply (erule Bochner_Integration.integrable_bound) apply (simp add: set_borel_measurable_def) apply (erule eventually_mono) apply (auto simp: indicator_def) done lemma set_integrableI_nonneg: fixes f :: "'a \ real" assumes "set_borel_measurable M A f" assumes "AE x in M. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \M) < \" shows "set_integrable M A f" unfolding set_integrable_def proof (rule integrableI_nonneg) from assms show "(\x. indicator A x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(2) show "AE x in M. 0 \ indicat_real A x *\<^sub>R f x" by eventually_elim (auto simp: indicator_def) have "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) = (\\<^sup>+x\A. f x \M)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ < \" by fact finally show "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) < \" . qed lemma set_integral_eq_nn_integral: assumes "set_borel_measurable M A f" assumes "set_nn_integral M A f = ennreal x" "x \ 0" assumes "AE x in M. x \ A \ f x \ 0" shows "set_integrable M A f" and "set_lebesgue_integral M A f = x" proof - have eq: "(\x. (if x \ A then 1 else 0) * f x) = (\x. if x \ A then f x else 0)" "(\x. if x \ A then ennreal (f x) else 0) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" "(\x. ennreal (f x * (if x \ A then 1 else 0))) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" by auto from assms show *: "set_integrable M A f" by (intro set_integrableI_nonneg) auto have "set_lebesgue_integral M A f = enn2real (set_nn_integral M A f)" unfolding set_lebesgue_integral_def using assms(1,4) * eq by (subst integral_eq_nn_integral) (auto intro!: nn_integral_cong simp: indicator_def of_bool_def set_integrable_def mult_ac) also have "\ = x" using assms by simp finally show "set_lebesgue_integral M A f = x" . qed lemma set_integral_0 [simp, intro]: "set_integrable M A (\y. 0)" by (simp add: set_integrable_def) lemma set_integrable_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_integrable M A (\y. \x\B. f x y)" using assms by (induction rule: finite_induct) (auto intro!: set_integral_add) lemma set_integral_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_lebesgue_integral M A (\y. \x\B. f x y) = (\x\B. set_lebesgue_integral M A (f x))" using assms apply (induction rule: finite_induct) apply simp apply simp apply (subst set_integral_add) apply (auto intro!: set_integrable_sum) done lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_nn_integral_mono: assumes "\x. x \ space M \ A \ f x \ g x" shows "set_nn_integral M A f \ set_nn_integral M A g" using assms by (intro nn_integral_mono) (auto simp: indicator_def) lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ (F has_field_derivative f x) (at x within {a..b})" assumes cont: "continuous_on {a..b} f" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" - unfolding has_field_derivative_iff_has_vector_derivative[symmetric] + unfolding has_real_derivative_iff_has_vector_derivative[symmetric] using deriv by auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 cont] integral_FTC_Icc[OF \a \ b\ 1 cont] by (auto simp: mult.commute) qed lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" proof - have "integrable lborel (\x. indicator {a..b} x *\<^sub>R ((F x) * (g x) + (f x) * (G x)))" by (intro borel_integrable_compact continuous_intros assms) (auto intro!: DERIV_continuous_on assms) thus ?thesis by (simp add: mult_ac) qed lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have 0: "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) (auto intro!: assms DERIV_continuous_on) have [continuous_intros]: "continuous_on {a..b} F" by (rule DERIV_continuous_on assms)+ have [continuous_intros]: "continuous_on {a..b} G" by (rule DERIV_continuous_on assms)+ have "(\x. indicator {a..b} x *\<^sub>R (F x * g x + f x * G x) \lborel) = (\x. indicator {a..b} x *\<^sub>R(F x * g x) \lborel) + \x. indicator {a..b} x *\<^sub>R (f x * G x) \lborel" apply (subst Bochner_Integration.integral_add[symmetric]) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (simp add: algebra_simps) done thus ?thesis using 0 by (simp add: algebra_simps) qed lemma interval_lebesgue_integral_by_parts: assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(LBINT x=a..b. F x * g x) = F b * G b - F a * G a - (LBINT x=a..b. f x * G x)" using integral_by_parts[of a b f g F G] assms by (simp add: interval_integral_Icc set_lebesgue_integral_def mult_ac) lemma interval_lebesgue_integral_by_parts_01: assumes cont_f[intro]: "continuous_on {0..1} f" assumes cont_g[intro]: "continuous_on {0..1} g" assumes [intro]: "\x. x \ {0..1} \ (F has_field_derivative f x) (at x within {0..1})" assumes [intro]: "\x. x \ {0..1} \ (G has_field_derivative g x) (at x within {0..1})" shows "(LBINT x=0..1. F x * g x) = F 1 * G 1 - F 0 * G 0 - (LBINT x=0..1. f x * G x)" using interval_lebesgue_integral_by_parts[of 0 1 f g F G] assms by (simp add: zero_ereal_def one_ereal_def) lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ real" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed subsection \Shifted Legendre polynomials\ text \ The first ingredient we need to show Apéry's theorem is the \<^emph>\shifted Legendre polynomials\ \[P_n(X) = \frac{1}{n!} \frac{\partial^n}{\partial X^n} (X^n(1-X)^n)\] and the auxiliary polynomials \[Q_{n,k}(X) = \frac{\partial^k}{\partial X^k} (X^n(1-X)^n)\ .\] Note that $P_n$ is in fact an \emph{integer} polynomial. Only some very basic properties of these will be proven, since that is all we will need. \ context fixes n :: nat begin definition gen_shleg_poly :: "nat \ int poly" where "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" definition shleg_poly where "shleg_poly = gen_shleg_poly n div [:fact n:]" text \ We can easily prove the following more explicit formula for $Q_{n,k}$: \[Q_{n,k}(X) = \sum_{i=0}^k (-1)^{k-1} {k \choose i} n^{\underline{i}}\, n^{\underline{k-i}}\, X^{n-i}\, (1-X)^{n-k+i}\] \ lemma gen_shleg_poly_altdef: assumes "k \ n" shows "gen_shleg_poly k = (\i\k. smult ((-1)^(k-i) * of_nat (k choose i) * pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n-k+i)))" proof - have *: "(pderiv ^^ i) (x \\<^sub>p [:1, -1:]) = smult ((-1) ^ i) ((pderiv ^^ i) x \\<^sub>p [:1, -1:])" for i and x :: "int poly" by (induction i arbitrary: x) (auto simp: pderiv_smult pderiv_pcompose funpow_Suc_right pderiv_pCons higher_pderiv_minus simp del: funpow.simps(2)) have "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" by (simp add: gen_shleg_poly_def) also have "[:0, 1, -1::int:] = [:0, 1:] * [:1, -1:]" by simp also have "\ ^ n = [:0, 1:] ^ n * [:1, -1:] ^ n" by (simp flip: power_mult_distrib) also have "(pderiv ^^ k) \ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) ([:0, 1:] ^ n) * (pderiv ^^ (k - i)) ([:1, -1:] ^ n)))" by (simp add: higher_pderiv_mult) also have "\ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * (pderiv ^^ (k - i)) (monom 1 n \\<^sub>p [:1, -1:])))" by (simp add: monom_altdef pcompose_pCons pcompose_power_left) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * ((pderiv ^^ (k - i)) (monom 1 n) \\<^sub>p [:1, -1:])))" by (simp add: * mult_ac) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) (monom (pochhammer (n - i + 1) i) (n - i) * monom (pochhammer (n - k + i + 1) (k - i)) (n - k + i) \\<^sub>p [:1, -1:]))" using assms by (simp add: higher_pderiv_monom) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i) * pochhammer (n - i + 1) i * pochhammer (n - k + i + 1) (k - i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n - k + i)))" by (simp add: monom_altdef algebra_simps pcompose_smult pcompose_power_left pcompose_pCons) finally show ?thesis . qed lemma degree_gen_shleg_poly [simp]: "degree (gen_shleg_poly k) = 2 * n - k" by (simp add: gen_shleg_poly_def degree_higher_pderiv degree_power_eq) lemma gen_shleg_poly_n: "gen_shleg_poly n = smult (fact n) shleg_poly" proof - obtain r where r: "gen_shleg_poly n = [:fact n:] * r" unfolding gen_shleg_poly_def using fact_dvd_higher_pderiv[of n "[:0,1,-1:]^n"] by blast have "smult (fact n) shleg_poly = smult (fact n) (gen_shleg_poly n div [:fact n:])" by (simp add: shleg_poly_def) also note r also have "[:fact n:] * r div [:fact n:] = r" by (rule nonzero_mult_div_cancel_left) auto finally show ?thesis by (simp add: r) qed lemma degree_shleg_poly [simp]: "degree shleg_poly = n" using degree_gen_shleg_poly[of n] by (simp add: gen_shleg_poly_n) lemma pderiv_gen_shleg_poly [simp]: "pderiv (gen_shleg_poly k) = gen_shleg_poly (Suc k)" by (simp add: gen_shleg_poly_def) text \ The following functions are the interpretation of the shifted Legendre polynomials and the auxiliary polynomials as a function from reals to reals. \ definition Gen_Shleg :: "nat \ real \ real" where "Gen_Shleg k x = poly (of_int_poly (gen_shleg_poly k)) x" definition Shleg :: "real \ real" where "Shleg = poly (of_int_poly shleg_poly)" lemma Gen_Shleg_altdef: assumes "k \ n" shows "Gen_Shleg k x = (\i\k. (-1)^(k-i) * of_nat (k choose i) * of_int (pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) * x ^ (n - i) * (1 - x) ^ (n-k+i))" using assms by (simp add: Gen_Shleg_def gen_shleg_poly_altdef poly_sum mult_ac) lemma Gen_Shleg_0 [simp]: "k < n \ Gen_Shleg k 0 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_1 [simp]: "k < n \ Gen_Shleg k 1 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_n_0 [simp]: "Gen_Shleg n 0 = fact n" proof - have "Gen_Shleg n 0 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{n}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (intro sum.mono_neutral_right) auto also have "\ = fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Gen_Shleg_n_1 [simp]: "Gen_Shleg n 1 = (-1) ^ n * fact n" proof - have "Gen_Shleg n 1 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{0}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (intro sum.mono_neutral_right) auto also have "\ = (-1) ^ n * fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Shleg_altdef: "Shleg x = Gen_Shleg n x / fact n" by (simp add: Shleg_def Gen_Shleg_def gen_shleg_poly_n) lemma Shleg_0 [simp]: "Shleg 0 = 1" and Shleg_1 [simp]: "Shleg 1 = (-1) ^ n" by (simp_all add: Shleg_altdef) lemma Gen_Shleg_0_left: "Gen_Shleg 0 x = x ^ n * (1 - x) ^ n" by (simp add: Gen_Shleg_def gen_shleg_poly_def power_mult_distrib) lemma has_field_derivative_Gen_Shleg: "(Gen_Shleg k has_field_derivative Gen_Shleg (Suc k) x) (at x)" proof - note [derivative_intros] = poly_DERIV show ?thesis unfolding Gen_Shleg_def by (rule derivative_eq_intros) auto qed lemma continuous_on_Gen_Shleg: "continuous_on A (Gen_Shleg k)" by (auto simp: Gen_Shleg_def intro!: continuous_intros) lemma continuous_on_Gen_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Gen_Shleg k (f x))" by (rule continuous_on_compose2[OF continuous_on_Gen_Shleg[of UNIV]]) auto lemma continuous_on_Shleg: "continuous_on A Shleg" by (auto simp: Shleg_def intro!: continuous_intros) lemma continuous_on_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Shleg (f x))" by (rule continuous_on_compose2[OF continuous_on_Shleg[of UNIV]]) auto lemma measurable_Gen_Shleg [measurable]: "Gen_Shleg n \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Gen_Shleg) lemma measurable_Shleg [measurable]: "Shleg \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Shleg) end subsection \Auxiliary facts about the \\\ function\ lemma Re_zeta_ge_1: assumes "x > 1" shows "Re (zeta (of_real x)) \ 1" proof - have *: "(\n. real (Suc n) powr -x) sums Re (zeta (complex_of_real x))" using sums_Re[OF sums_zeta[of "of_real x"]] assms by (simp add: powr_Reals_eq) show "Re (zeta (of_real x)) \ 1" proof (rule sums_le[OF _ _ *]) show "(\n. if n = 0 then 1 else 0) sums 1" by (rule sums_single) qed auto qed lemma sums_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. 1 / (k + 1) ^ n) sums zeta (of_nat n)" using sums_zeta[of "of_nat n"] n by (simp add: powr_minus field_simps flip: of_nat_Suc) from sums_split_initial_segment[OF this, of r] have "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\kkk=1..r. 1 / k ^ n)" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show ?thesis . qed lemma sums_Re_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (Re (zeta (of_nat n)) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. Re (1 / (r + k + 1) ^ n)) sums (Re (zeta (of_nat n) - (\k=1..r. 1 / k ^ n)))" by (intro sums_Re sums_zeta_of_nat_offset assms) thus ?thesis by simp qed subsection \Divisor of a sum of rationals\ text \ A finite sum of rationals of the form $\frac{a_1}{b_1} + \ldots + \frac{a_n}{b_n}$ can be brought into the form $\frac{c}{d}$, where $d$ is the LCM of the $b_i$ (or some integer multiple thereof). \ lemma sum_rationals_common_divisor: fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" using assms proof (induction rule: finite_induct) case empty thus ?case by auto next case (insert x A) define d where "d = (LCM x\A. g x)" from insert have [simp]: "d \ 0" by (auto simp: d_def Lcm_0_iff) from insert have [simp]: "g x \ 0" by auto from insert obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d" by (auto simp: d_def) define e1 where "e1 = lcm d (g x) div d" define e2 where "e2 = lcm d (g x) div g x" have "(\y\insert x A. f y / g y) = c / d + f x / g x" using insert c by simp also have "c / d = (c * e1) / lcm d (g x)" by (simp add: e1_def real_of_int_div) also have "f x / g x = (f x * e2) / lcm d (g x)" by (simp add: e2_def real_of_int_div) also have "(c * e1) / lcm d (g x) + \ = (c * e1 + f x * e2) / (LCM x\insert x A. g x)" using insert by (simp add: add_divide_distrib lcm.commute d_def) finally show ?case .. qed lemma sum_rationals_common_divisor': fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" and "(\x. x \ A \ g x dvd d)" and "d \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / real_of_int d" proof - define d' where "d' = (LCM x\A. g x)" have "d' dvd d" unfolding d'_def using assms(3) by (auto simp: Lcm_dvd_iff) then obtain e where e: "d = d' * e" by blast have "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" by (rule sum_rationals_common_divisor) fact+ then obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d'" unfolding d'_def .. also have "\ = real_of_int (c * e) / real_of_int d" using \d \ 0\ by (simp add: e) finally show ?thesis .. qed subsection \The first double integral\ text \ We shall now investigate the double integral \[I_1 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy}\,x^r y^s\, \text{d}x\,\text{d}y\ .\] Since everything is non-negative for now, we can work over the extended non-negative real numbers and the issues of integrability or summability do not arise at all. \ definition beukers_nn_integral1 :: "nat \ nat \ ennreal" where "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" definition beukers_integral1 :: "nat \ nat \ real" where "beukers_integral1 r s = (\(x,y)\{0<..<1}\{0<..<1}. (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" lemma fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0..1}" shows beukers_denom_ineq: "(1 - x * y) * z < 1" and beukers_denom_neq: "(1 - x * y) * z \ 1" proof - from xyz have *: "x * y < 1 * 1" by (intro mult_strict_mono) auto from * have "(1 - x * y) * z \ (1 - x * y) * 1" using xyz by (intro mult_left_mono) auto also have "\ < 1 * 1" using xyz by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * z < 1" "(1 - x * y) * z \ 1" by simp_all qed text \ We first evaluate the improper integral \[\int_0^1 -\ln x \cdot x^e\,\text{d}x = \frac{1}{(e+1)^2}\ .\] for any $e>-1$. \ lemma integral_0_1_ln_times_powr: assumes "e > -1" shows "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" and "interval_lebesgue_integrable lborel 0 1 (\x. -ln x * x powr e)" proof - define f where "f = (\x. -ln x * x powr e)" define F where "F = (\x. x powr (e + 1) * (1 - (e + 1) * ln x) / (e + 1) ^ 2)" have 0: "isCont f x" if "x \ {0<..<1}" for x using that by (auto intro!: continuous_intros simp: f_def) have 1: "(F has_real_derivative f x) (at x)" if "x \ {0<..<1}" for x proof - show "(F has_real_derivative f x) (at x)" unfolding F_def f_def using that assms apply (insert that assms) apply (rule derivative_eq_intros refl | simp)+ apply (simp add: divide_simps) apply (simp add: power2_eq_square algebra_simps powr_add power_numeral_reduce) done qed have 2: "AE x in lborel. ereal 0 < ereal x \ ereal x < ereal 1 \ 0 \ f x" by (intro AE_I2) (auto simp: f_def mult_nonpos_nonneg) have 3: "((F \ real_of_ereal) \ 0) (at_right (ereal 0))" unfolding ereal_tendsto_simps F_def using assms by real_asymp have 4: "((F \ real_of_ereal) \ F 1) (at_left (ereal 1))" unfolding ereal_tendsto_simps F_def using assms by real_asymp (simp add: field_simps) have "(LBINT x=ereal 0..ereal 1. f x) = F 1 - 0" by (rule interval_integral_FTC_nonneg[where F = F]) (use 0 1 2 3 4 in auto) thus "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" by (simp add: F_def zero_ereal_def one_ereal_def f_def) have "set_integrable lborel (einterval (ereal 0) (ereal 1)) f" by (rule interval_integral_FTC_nonneg) (use 0 1 2 3 4 in auto) thus "interval_lebesgue_integrable lborel 0 1 f" by (simp add: interval_lebesgue_integrable_def einterval_def) qed lemma interval_lebesgue_integral_lborel_01_cong: assumes "\x. x \ {0<..<1} \ f x = g x" shows "interval_lebesgue_integral lborel 0 1 f = interval_lebesgue_integral lborel 0 1 g" using assms by (subst (1 2) interval_integral_Ioo) (auto intro!: set_lebesgue_integral_cong assms) lemma nn_integral_0_1_ln_times_powr: assumes "e > -1" shows "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = ennreal (1 / (e + 1)\<^sup>2)" proof - have *: "(LBINT x=0..1. - ln x * x powr e = 1 / (e + 1)\<^sup>2)" "interval_lebesgue_integrable lborel 0 1 (\x. - ln x * x powr e)" using integral_0_1_ln_times_powr[OF assms] by simp_all have eq: "(\y. (if 0 < y \ y < 1 then 1 else 0) * ln y * y powr e) = (\y. if 0 < y \ y < 1 then ln y * y powr e else 0)" by auto have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = (\\<^sup>+y. ennreal (-ln y * y powr e * indicator {0<..<1} y) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal (1 / (e + 1)\<^sup>2)" using * eq by (subst nn_integral_eq_integral) (auto intro!: AE_I2 simp: indicator_def interval_lebesgue_integrable_def set_integrable_def one_ereal_def zero_ereal_def interval_integral_Ioo mult_ac mult_nonpos_nonneg set_lebesgue_integral_def) finally show ?thesis . qed lemma nn_integral_0_1_ln_times_power: "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = ennreal (1 / (n + 1)\<^sup>2)" proof - have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr real n) \lborel)" by (intro set_nn_integral_cong) (auto simp: powr_realpow) also have "\ = ennreal (1 / (n + 1)^2)" by (subst nn_integral_0_1_ln_times_powr) auto finally show ?thesis by simp qed text \ Next, we also evaluate the more trivial integral \[\int_0^1 x^n\ \text{d}x\ .\] \ lemma nn_integral_0_1_power: "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = ennreal (1 / (n + 1))" proof - have *: "((\a. a ^ (n + 1) / real (n + 1)) has_real_derivative x ^ n) (at x)" for x by (rule derivative_eq_intros refl | simp)+ have "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = (\\<^sup>+y\{0..1}. ennreal (y ^ n) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: indicator_def emeasure_lborel_countable) also have "\ = ennreal (1 ^ (n + 1) / (n + 1) - 0 ^ (n + 1) / (n + 1))" using * by (intro nn_integral_FTC_Icc) auto also have "\ = ennreal (1 / (n + 1))" by simp finally show ?thesis by simp qed text \ $I_1$ can alternatively be written as the triple integral \[\int_0^1\int_0^1\int_0^1 \frac{x^r y^s}{1-(1-xy)w}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ lemma beukers_nn_integral1_altdef: "beukers_nn_integral1 r s = (\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel)" proof - have "(\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\\<^sup>+w\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) \lborel)" by (subst lborel_prod [symmetric], subst lborel_pair.nn_integral_snd [symmetric]) (auto simp: case_prod_unfold indicator_def simp flip: lborel_prod intro!: nn_integral_cong) also have "\ = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y)/(1-x*y) * x^r * y^s) \lborel)" proof (intro nn_integral_cong, clarify) fix x y :: real have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (-ln (x*y)*x^r*y^s/(1-x*y))" if xy: "(x, y) \ {0<..<1} \ {0<..<1}" proof - from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have deriv: "((\w. -ln (1-(1-x*y)*w) / (1-x*y)) has_real_derivative 1/(1-(1-x*y)*w)) (at w)" if w: "w \ {0..1}" for w by (insert xy w \x*y<1\ beukers_denom_ineq[of x y w]) (rule derivative_eq_intros refl | simp add: divide_simps)+ have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (x^r*y^s) * (\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" using xy by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = (\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: emeasure_lborel_countable indicator_def) also have "(\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = ennreal (-ln (1-(1-x*y)*1)/(1-x*y) - (-ln (1-(1-x*y)*0)/(1-x*y)))" using xy deriv less_imp_le[OF beukers_denom_ineq[of x y]] by (intro nn_integral_FTC_Icc) auto finally show ?thesis using xy by (simp flip: ennreal_mult' ennreal_mult'' add: mult_ac) qed thus "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) * indicator ({0<..<1}\{0<..<1}) (x, y) = ennreal (-ln (x*y)/(1-x*y)*x^r*y^s) * indicator ({0<..<1}\{0<..<1}) (x, y)" by (auto simp: indicator_def) qed also have "\ = beukers_nn_integral1 r s" by (simp add: beukers_nn_integral1_def) finally show ?thesis .. qed context fixes r s :: nat and I1 I2' :: real and I2 :: ennreal and D :: "(real \ real \ real) set" assumes rs: "s \ r" defines "D \ {0<..<1}\{0<..<1}\{0<..<1}" begin text \ By unfolding the geometric series, pulling the summation out and evaluating the integrals, we find that \[I_1 = \sum_{k=0}^\infty \frac{1}{(k+r+1)^2(k+s+1)} + \frac{1}{(k+r+1)(k+s+1)^2}\ .\] \ lemma beukers_nn_integral1_series: "beukers_nn_integral1 r s = (\k. ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2)))" proof - have "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel)" unfolding beukers_nn_integral1_def proof (intro set_nn_integral_cong refl, clarify) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have "(\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) = ennreal (-ln (x*y) * x^r * y^s) * (\k. ennreal ((x*y)^k))" using xy by (subst ennreal_suminf_cmult [symmetric], subst ennreal_mult'' [symmetric]) (auto simp: power_add mult_ac power_mult_distrib) also have "(\k. ennreal ((x*y)^k)) = ennreal (1 / (1 - x*y))" using geometric_sums[of "x*y"] \x * y < 1\ xy by (intro suminf_ennreal_eq) auto also have "ennreal (-ln (x*y) * x^r * y^s) * \ = ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s)" using \x * y < 1\ by (subst ennreal_mult'' [symmetric]) auto finally show "ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) = (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)))" .. qed also have "\ = (\k. (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel))" unfolding case_prod_unfold by (subst nn_integral_suminf [symmetric]) (auto simp flip: borel_prod) also have "\ = (\k. ennreal (1 / ((k+r+1)^2*(k+s+1)) + 1 / ((k+r+1)*(k+s+1)^2)))" proof (rule suminf_cong) fix k :: nat define F where "F = (\x y::real. x + y)" have "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+x\{0<..<1}. (\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) \lborel)" unfolding case_prod_unfold lborel_prod [symmetric] by (subst lborel.nn_integral_fst [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2) \lborel)" proof (intro set_nn_integral_cong refl, clarify) fix x :: real assume x: "x \ {0<..<1}" have "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+y\{0<..<1}. (ennreal (-ln x * x^(k+r) * y^(k+s)) + ennreal (-ln y * x^(k+r) * y^(k+s))) \lborel)" by (intro set_nn_integral_cong) (use x in \auto simp: ln_mult ring_distribs mult_nonpos_nonneg simp flip: ennreal_plus\) also have "\ = (\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) + (\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult'') also have "(\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel) = ennreal (1/(k+s+1))" by (subst nn_integral_0_1_power) simp also have "ennreal (-ln x * x^(k+r)) * \ = ennreal (-ln x * x^(k+r) / (k+s+1))" by (subst ennreal_mult'' [symmetric]) auto also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel) = ennreal (x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (use x in \auto intro!: nn_integral_cong simp: indicator_def mult_ac simp flip: ennreal_mult'\) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel) = ennreal (1 / (k + s + 1)\<^sup>2)" by (subst nn_integral_0_1_ln_times_power) simp also have "ennreal (x ^ (k + r)) * \ = ennreal (x ^ (k + r) / (k + s + 1) ^ 2)" by (subst ennreal_mult'' [symmetric]) auto also have "ennreal (- ln x * x ^ (k + r) / (k + s + 1)) + \ = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" using x by (subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) finally show "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" . qed also have "\ = (\\<^sup>+x\{0<..<1}. (ennreal (-ln x * x^(k+r) / (k+s+1)) + ennreal (x^(k+r)/(k+s+1)^2)) \lborel)" by (intro set_nn_integral_cong refl, subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) + (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) = ennreal (1 / (k+s+1)) * (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1 / ((k+s+1) * (k+r+1)^2))" by (subst nn_integral_0_1_ln_times_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "(\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel) = ennreal (1/(k+s+1)^2) * (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1/((k+r+1)*(k+s+1)^2))" by (subst nn_integral_0_1_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "ennreal (1 / ((k+s+1) * (k+r+1)^2)) + \ = ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2))" by (subst ennreal_plus [symmetric]) (auto simp: algebra_simps) finally show "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = \" . qed finally show ?thesis . qed text \ Remembering that $\zeta(3) = \sum k^{-3}$, it is easy to see that if $r = s$, this sum is simply \[2\left(\zeta(3) - \sum_{k=1}^r \frac{1}{k^3}\right)\ .\] \ lemma beukers_nn_integral1_same: assumes "r = s" shows "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" and "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" proof - from assms have [simp]: "s = r" by simp have *: "Suc 2 = 3" by simp have "beukers_nn_integral1 r s = (\k. ennreal (2 / (r + k + 1) ^ 3))" unfolding beukers_nn_integral1_series by (simp only: assms power_Suc [symmetric] mult.commute[of "x ^ 2" for x] * times_divide_eq_right mult_1_right add_ac flip: mult_2) also have **: "(\k. 2 / (r + k + 1) ^ 3) sums (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" using sums_mult[OF sums_Re_zeta_of_nat_offset[of 3], of 2] by simp hence "(\k. ennreal (2 / (r + k + 1) ^ 3)) = ennreal \" by (intro suminf_ennreal_eq) auto finally show "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" . show "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" by (rule sums_le[OF _ sums_zero **]) auto qed lemma beukers_integral1_same: assumes "r = s" shows "beukers_integral1 r s = 2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3))" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_same[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed text \ In contrast, for \r > s\, we find that \[I_1 = \frac{1}{r-s} \sum_{k=s+1}^r \frac{1}{k^2}\ .\] \ lemma beukers_nn_integral1_different: assumes "r > s" shows "beukers_nn_integral1 r s = ennreal ((\k\{s<..r}. 1 / k ^ 2) / (r - s))" proof - have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) sums (1 / (r - s) * ((Re (zeta (of_nat 2)) - (\k=1..s. 1/k^2)) - (Re (zeta (of_nat 2)) - (\k=1..r. 1/k^2))))" (is "_ sums ?S") by (intro sums_mult sums_diff sums_Re_zeta_of_nat_offset) auto also have "?S = ((\k=1..r. 1/k^2) - (\k=1..s. 1/k^2)) / (r - s)" by (simp add: algebra_simps diff_divide_distrib) also have "(\k=1..r. 1/k^2) - (\k=1..s. 1/k^2) = (\k\{1..r}-{1..s}. 1/k^2)" using assms by (subst Groups_Big.sum_diff) auto also have "{1..r} - {1..s} = {s<..r}" by auto also have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) = (\k. 1 / ((k+r+1) * (k+s+1)^2) + 1 / ((k+r+1)^2 * (k+s+1)))" proof (intro ext, goal_cases) case (1 k) define x where "x = real (k + r + 1)" define y where "y = real (k + s + 1)" have [simp]: "x \ 0" "y \ 0" by (auto simp: x_def y_def) have "(x\<^sup>2 * y + x * y\<^sup>2) * (real r - real s) = x * y * (x\<^sup>2 - y\<^sup>2)" by (simp add: algebra_simps power2_eq_square x_def y_def) hence "1 / (x*y^2) + 1 / (x^2*y) = 1 / (r - s) * (1 / y^2 - 1 / x^2)" using assms by (simp add: divide_simps of_nat_diff) thus ?case by (simp add: x_def y_def algebra_simps) qed finally show ?thesis unfolding beukers_nn_integral1_series by (intro suminf_ennreal_eq) (auto simp: add_ac) qed lemma beukers_integral1_different: assumes "r > s" shows "beukers_integral1 r s = (\k\{s<..r}. 1 / k ^ 2) / (r - s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_different[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg intro!: sum_nonneg divide_nonneg_nonneg) qed end text \ It is also easy to see that if we exchange \r\ and \s\, nothing changes. \ lemma beukers_nn_integral1_swap: "beukers_nn_integral1 r s = beukers_nn_integral1 s r" unfolding beukers_nn_integral1_def lborel_prod [symmetric] by (subst lborel_pair.nn_integral_swap, simp) (intro nn_integral_cong, auto simp: indicator_def algebra_simps split: if_splits) lemma beukers_nn_integral1_finite: "beukers_nn_integral1 r s < \" using beukers_nn_integral1_different[of r s] beukers_nn_integral1_different[of s r] by (cases r s rule: linorder_cases) (simp_all add: beukers_nn_integral1_same beukers_nn_integral1_swap) lemma beukers_integral1_integrable: "set_integrable lborel ({0<..<1}\{0<..<1}) (\(x,y). (-ln (x*y) / (1 - x*y) * x^r * y^s :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" have "0 \ ln (x * y) / (1 - x * y) * x ^ r * y ^ s" using mult_strict_mono[of x 1 y 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) (use xy in auto) thus "0 \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s" by simp next show "\\<^sup>+x\{0<..<1}\{0<..<1}. ennreal (case x of (x, y) \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s) \lborel < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_def case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_integrable': "set_integrable lborel ({0<..<1}\{0<..<1}\{0<..<1}) (\(z,x,y). (x^r * y^s / (1 - (1 - x*y) * z) :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" show "0 \ x^r * y^s / (1 - (1 - x*y) * z)" using mult_strict_mono[of x 1 y 1] xyz beukers_denom_ineq[of x y z] by (intro mult_nonneg_nonneg divide_nonneg_nonneg) auto next show "\\<^sup>+x\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (case x of (z,x,y) \ x ^ r * y ^ s / (1-(1-x*y)*z)) \lborel < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_altdef case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_conv_nn_integral: "beukers_integral1 r s = enn2real (beukers_nn_integral1 r s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using mult_strict_mono[of a 1 b 1] that by (intro divide_nonpos_nonneg mult_nonpos_nonneg) auto thus ?thesis unfolding beukers_integral1_def using beukers_nn_integral1_finite[of r s] by (intro set_integral_eq_nn_integral) (auto simp: case_prod_unfold beukers_nn_integral1_def set_borel_measurable_def simp flip: borel_prod intro!: AE_I2 intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed lemma beukers_integral1_swap: "beukers_integral1 r s = beukers_integral1 s r" by (simp add: beukers_integral1_conv_nn_integral beukers_nn_integral1_swap) subsection \The second double integral\ context fixes n :: nat fixes D :: "(real \ real) set" and D' :: "(real \ real \ real) set" fixes P :: "real \ real" and Q :: "nat \ real \ real" defines "D \ {0<..<1} \ {0<..<1}" and "D' \ {0<..<1} \ {0<..<1} \ {0<..<1}" defines "Q \ Gen_Shleg n" and "P \ Shleg n" begin text \ The next integral to consider is the following variant of $I_1$: \[I_2 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy} P_n(x) P_n(y)\,\text{d}x\,\text{d}y\ .\] \ definition beukers_integral2 :: real where "beukers_integral2 = (\(x,y)\D. (-ln (x*y) / (1-x*y) * P x * P y) \lborel)" text \ $I_2$ is simply a sum of integrals of type $I_1$, so using our results for $I_1$, we can write $I_2$ in the form $A \zeta(3) + \frac{B}{\text{lcm}\{1\ldots n\}^3}$ where $A$ and $B$ are integers and $A > 0$: \ lemma beukers_integral2_conv_int_combination: obtains A B :: int where "A > 0" and "beukers_integral2 = of_int A * Re (zeta 3) + of_int B / of_nat (Lcm {1..n} ^ 3)" proof - let ?I1 = "(\i. (i, i)) ` {..n}" let ?I2 = "Set.filter (\(i,j). i \ j) ({..n}\{..n})" let ?I3 = "Set.filter (\(i,j). i < j) ({..n}\{..n})" let ?I4 = "Set.filter (\(i,j). i > j) ({..n}\{..n})" define p where "p = shleg_poly n" define I where "I = (SIGMA i:{..n}. {1..i})" define J where "J = (SIGMA (i,j):?I4. {j<..i})" define h where "h = beukers_integral1" define A :: int where "A = (\i\n. 2 * poly.coeff p i ^ 2)" define B1 where "B1 = (\(i,k)\I. real_of_int (-2 * coeff p i ^ 2) / real_of_int (k ^ 3))" define B2 where "B2 = (\((i,j),k)\J. real_of_int (2 * coeff p i * coeff p j) / real_of_int (k^2*(i-j)))" define d where "d = Lcm {1..n} ^ 3" have [simp]: "h i j = h j i" for i j by (simp add: h_def beukers_integral1_swap) have "beukers_integral2 = (\(x,y)\D. (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * -ln (x*y) / (1-x*y) * x ^ i * y ^ j) \lborel)" unfolding beukers_integral2_def by (subst sum.cartesian_product [symmetric]) (simp add: poly_altdef P_def Shleg_def mult_ac case_prod_unfold p_def sum_distrib_left sum_distrib_right sum_negf sum_divide_distrib) also have "\ = (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * h i j)" unfolding case_prod_unfold proof (subst set_integral_sum) fix ij :: "nat \ nat" have "set_integrable lborel D (\(x,y). real_of_int (coeff p (fst ij) * coeff p (snd ij)) * (-ln (x*y) / (1-x*y) * x ^ fst ij * y ^ snd ij))" unfolding case_prod_unfold using beukers_integral1_integrable[of "fst ij" "snd ij"] by (intro set_integrable_mult_right) (auto simp: D_def case_prod_unfold) thus "set_integrable lborel D (\pa. real_of_int (coeff p (fst ij) * coeff p (snd ij)) * -ln (fst pa * snd pa) / (1 - fst pa * snd pa) * fst pa ^ fst ij * snd pa ^ snd ij)" by (simp add: mult_ac case_prod_unfold) qed (auto simp: beukers_integral1_def h_def case_prod_unfold mult.assoc D_def simp flip: set_integral_mult_right) also have "\ = (\(i,j)\?I1\?I2. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I1. coeff p i * coeff p j * h i j) + (\(i,j)\?I2. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I1. coeff p i * coeff p j * h i j) = (\i\n. coeff p i ^ 2 * h i i)" by (subst sum.reindex) (auto intro: inj_onI simp: case_prod_unfold power2_eq_square) also have "\ = (\i\n. coeff p i ^ 2 * 2 * (Re (zeta 3) - (\k=1..i. 1 / k ^ 3)))" unfolding h_def D_def by (intro sum.cong refl, subst beukers_integral1_same) auto also have "\ = of_int A * Re (zeta 3) - (\i\n. 2 * coeff p i ^ 2 * (\k=1..i. 1 / k ^ 3))" by (simp add: sum_subtractf sum_distrib_left sum_distrib_right algebra_simps A_def) also have "\ = of_int A * Re (zeta 3) + B1" unfolding I_def B1_def by (subst sum.Sigma [symmetric]) (auto simp: sum_distrib_left sum_negf) also have "(\(i,j)\?I2. coeff p i * coeff p j * h i j) = (\(i,j)\?I3\?I4. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I3. coeff p i * coeff p j * h i j) + (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I3. coeff p i * coeff p j * h i j) = (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.reindex_bij_witness[of _ "\(i,j). (j,i)" "\(i,j). (j,i)"]) auto also have "\ + \ = 2 * \" by simp also have "\ = (\(i,j)\?I4. \k\{j<..i}. 2 * coeff p i * coeff p j / k ^ 2 / (i - j))" unfolding sum_distrib_left by (intro sum.cong refl) (auto simp: h_def beukers_integral1_different sum_divide_distrib sum_distrib_left mult_ac) also have "\ = B2" unfolding J_def B2_def by (subst sum.Sigma [symmetric]) (auto simp: case_prod_unfold) also have "\B1'. B1 = real_of_int B1' / real_of_int d" unfolding B1_def case_prod_unfold by (rule sum_rationals_common_divisor') (auto simp: d_def I_def) then obtain B1' where "B1 = real_of_int B1' / real_of_int d" .. also have "\B2'. B2 = real_of_int B2' / real_of_int d" unfolding B2_def case_prod_unfold J_def proof (rule sum_rationals_common_divisor'; clarsimp?) fix i j k :: nat assume ijk: "i \ n" "j < k" "k \ i" have "int (k ^ 2 * (i - j)) dvd int (Lcm {1..n} ^ 2 * Lcm {1..n})" unfolding int_dvd_int_iff using ijk by (intro mult_dvd_mono dvd_power_same dvd_Lcm) auto also have "\ = d" by (simp add: d_def power_numeral_reduce) finally show "int k ^ 2 * int (i - j) dvd int d" by simp qed(auto simp: d_def J_def intro!: Nat.gr0I) then obtain B2' where "B2 = real_of_int B2' / real_of_int d" .. finally have "beukers_integral2 = of_int A * Re (zeta 3) + of_int (B1' + B2') / of_nat (Lcm {1..n} ^ 3)" by (simp add: add_divide_distrib d_def) moreover have "coeff p 0 = P 0" unfolding P_def p_def Shleg_def by (simp add: poly_0_coeff_0) hence "coeff p 0 = 1" by (simp add: P_def) hence "A > 0" unfolding A_def by (intro sum_pos2[of _ 0]) auto ultimately show ?thesis by (intro that[of A "B1' + B2'"]) auto qed lemma beukers_integral2_integrable: "set_integrable lborel D (\(x,y). -ln (x*y) / (1 - x*y) * P x * P y)" proof - have "bounded (P ` {0..1})" unfolding P_def Shleg_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast have [measurable]: "P \ borel_measurable borel" by (simp add: P_def) from C[of 0] have "C \ 0" by simp show ?thesis proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C ^ 2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y :: real assume xy: "(x, y) \ D" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by (simp add: D_def) have "norm (-ln (x*y) / (1 - x*y) * P x * P y) = (-ln (x*y)) / (1 - x*y) * norm (P x) * norm (P y)" using xy \x * y < 1\ by (simp add: abs_mult abs_divide D_def) also have "\ \ (-ln (x*y)) / (1-x*y) * C * C" using xy C[of x] C[of y] \x * y < 1\ \C \ 0\ by (intro mult_mono divide_left_mono) (auto simp: D_def divide_nonpos_nonneg mult_nonpos_nonneg) also have "\ = norm ((-ln (x*y)) / (1-x*y) * C * C)" using xy \x * y < 1\ \C \ 0\ by (simp add: abs_divide abs_mult D_def) finally show "norm (-ln (x*y) / (1 - x*y) * P x * P y) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (auto simp: algebra_simps power2_eq_square abs_mult abs_divide) qed (auto simp: D_def set_borel_measurable_def case_prod_unfold simp flip: lborel_prod) qed subsection \The triple integral\ text \ Lastly, we turn to the triple integral \[I_3 := \int_0^1 \int_0^1 \int_0^1 \frac{(x(1-x)y(1-y)w(1-w))^n}{(1-(1-xy)w)^{n+1}}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ definition beukers_nn_integral3 :: ennreal where "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" definition beukers_integral3 :: real where "beukers_integral3 = (\(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" text \ We first prove the following bound (which is a consequence of the arithmetic--geometric mean inequality) that will help us to bound the triple integral. \ lemma beukers_integral3_integrand_bound: fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" shows "(x*(1-x)*y*(1-y)*z*(1-z)) / (1-(1-x*y)*z) \ 1 / 27" (is "?lhs \ _") proof - have ineq1: "x * (1 - x) \ 1 / 4" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) - 1 / 4 = -((x - 1 / 2) ^ 2)" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by simp finally show ?thesis by simp qed have ineq2: "x * (1 - x) ^ 2 \ 4 / 27" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) ^ 2 - 4 / 27 = (x - 4 / 3) * (x - 1 / 3) ^ 2" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by (rule mult_nonpos_nonneg) (use x in auto) finally show ?thesis by simp qed have "1 - (1-x*y)*z = (1 - z) + x * y * z" by (simp add: algebra_simps) also have "\ \ 2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z" using arith_geo_mean_sqrt[of "1 - z" "x * y * z"] xyz by (auto simp: real_sqrt_mult) finally have *: "?lhs \ (x*(1-x)*y*(1-y)*z*(1-z)) / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z)" using xyz beukers_denom_ineq[of x y z] by (intro divide_left_mono mult_nonneg_nonneg mult_pos_pos) auto have "(x*(1-x)*y*(1-y)*z*(1-z)) = (sqrt x * sqrt x * (1-x) * sqrt y * sqrt y * (1-y) * sqrt z * sqrt z * sqrt (1-z) * sqrt (1-z))" using xyz by simp also have "\ / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z) = sqrt (x * (1 - x) ^ 2) * sqrt (y * (1 - y) ^ 2) * sqrt (z * (1 - z)) / 2" using xyz by (simp add: divide_simps real_sqrt_mult del: real_sqrt_mult_self) also have "\ \ sqrt (4 / 27) * sqrt (4 / 27) * sqrt (1 / 4) / 2" using xyz by (intro divide_right_mono mult_mono real_sqrt_le_mono ineq1 ineq2) auto also have "\ = 1 / 27" by (simp add: real_sqrt_divide) finally show ?thesis using * by argo qed text \ Connecting the above bound with our results of $I_1$, it is easy to see that $I_3 \leq 2 \cdot 27^{-n} \cdot \zeta(3)$: \ lemma beukers_nn_integral3_le: "beukers_nn_integral3 \ ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" by (simp add: beukers_nn_integral3_def) also have "\ \ (\\<^sup>+(w,x,y)\D'. ((1 / 27) ^ n / (1-(1-x*y)*w)) \lborel)" proof (intro set_nn_integral_mono ennreal_leI, clarify, goal_cases) case (1 w x y) have "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) = ((x*(1-x)*y*(1-y)*w*(1-w)) / (1-(1-x*y)*w)) ^ n / (1-(1-x*y)*w)" by (simp add: divide_simps) also have "\ \ (1 / 27) ^ n / (1 - (1 - x * y) * w)" using beukers_denom_ineq[of x y w] 1 by (intro divide_right_mono power_mono beukers_integral3_integrand_bound) (auto simp: D'_def) finally show ?case . qed also have "\ = ennreal ((1 / 27) ^ n) * (\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel)" unfolding lborel_prod [symmetric] case_prod_unfold by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (- (ln (x * y) / (1 - x * y))) \lborel)" using beukers_nn_integral1_altdef[of 0 0] by (simp add: beukers_nn_integral1_def D'_def case_prod_unfold) also have "\ = ennreal (2 * Re (zeta 3))" using beukers_nn_integral1_same[of 0 0] by (simp add: D_def beukers_nn_integral1_def) also have "ennreal ((1 / 27) ^ n) * \ = ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" by (subst ennreal_mult' [symmetric]) (simp_all add: mult_ac) finally show ?thesis . qed lemma beukers_nn_integral3_finite: "beukers_nn_integral3 < \" by (rule le_less_trans, rule beukers_nn_integral3_le) simp_all lemma beukers_integral3_integrable: "set_integrable lborel D' (\(w,x,y). (x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1))" unfolding case_prod_unfold using less_imp_le[OF beukers_denom_ineq] beukers_nn_integral3_finite by (intro set_integrableI_nonneg AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod intro!: divide_nonneg_nonneg mult_nonneg_nonneg) lemma beukers_integral3_conv_nn_integral: "beukers_integral3 = enn2real beukers_nn_integral3" unfolding beukers_integral3_def using beukers_nn_integral3_finite less_imp_le[OF beukers_denom_ineq] by (intro set_integral_eq_nn_integral AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod) lemma beukers_integral3_le: "beukers_integral3 \ 2 * (1 / 27) ^ n * Re (zeta 3)" proof - have "beukers_integral3 = enn2real beukers_nn_integral3" by (rule beukers_integral3_conv_nn_integral) also have "\ \ enn2real (ennreal (2 * (1 / 27) ^ n * Re (zeta 3)))" by (intro enn2real_mono beukers_nn_integral3_le) auto also have "\ = 2 * (1 / 27) ^ n * Re (zeta 3)" using Re_zeta_ge_1[of 3] by (intro enn2real_ennreal mult_nonneg_nonneg) auto finally show ?thesis . qed text \ It is also easy to see that $I_3 > 0$. \ lemma beukers_nn_integral3_pos: "beukers_nn_integral3 > 0" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have *: "\(AE (w,x,y) in lborel. ennreal ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) * indicator D' (w,x,y) \ 0)" (is "\(AE z in lborel. ?P z)") proof - { fix w x y :: real assume xyw: "(w,x,y) \ D'" hence "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) > 0" using beukers_denom_ineq[of x y w] by (intro divide_pos_pos mult_pos_pos zero_less_power) (auto simp: D'_def) with xyw have "\?P (w,x,y)" by (auto simp: indicator_def D'_def) } hence *: "\?P z" if "z \ D'" for z using that by blast hence "{z\space lborel. \?P z} = D'" by auto moreover have "emeasure lborel D' = 1" proof - have "D' = box (0,0,0) (1,1,1)" by (auto simp: D'_def box_def Basis_prod_def) also have "emeasure lborel \ = 1" by (subst emeasure_lborel_box) (auto simp: Basis_prod_def) finally show ?thesis by simp qed ultimately show ?thesis by (subst AE_iff_measurable[of D']) (simp_all flip: borel_prod) qed hence "nn_integral lborel (\_::real\real\real. 0) < beukers_nn_integral3" unfolding beukers_nn_integral3_def by (intro nn_integral_less) (simp_all add: case_prod_unfold flip: lborel_prod) thus ?thesis by simp qed lemma beukers_integral3_pos: "beukers_integral3 > 0" proof - have "0 < enn2real beukers_nn_integral3" using beukers_nn_integral3_pos beukers_nn_integral3_finite by (subst enn2real_positive_iff) auto also have "\ = beukers_integral3" by (rule beukers_integral3_conv_nn_integral [symmetric]) finally show ?thesis . qed subsection \Connecting the double and triple integral\ text \ In this section, we will prove the most technically involved part, namely that $I_2 = I_3$. I will not go into detail about how this works -- the reader is advised to simply look at Filaseta's presentation of the proof. The basic idea is to integrate by parts \n\ times with respect to \y\ to eliminate the factor $P(y)$, then change variables $z = \frac{1-w}{1-(1-xy)w}$, and then apply the same integration by parts \n\ times to \x\ to eliminate $P(x)$. \ text \ The first expand \[-\frac{\ln (xy)}{1-xy} = \int_0^1 \frac{1}{1-(1-xy)z}\,\text{d}z\ .\] \ lemma beukers_aux_ln_conv_integral: fixes x y :: real assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "-ln (x*y) / (1-x*y) = (LBINT z=0..1. 1 / (1-(1-x*y)*z))" proof - have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - x * y) * u < 1" if u: "u \ {0..1}" for u proof - from u \x * y < 1\ have "(1 - x * y) * u \ (1 - x * y) * 1" by (intro mult_left_mono) auto also have "\ < 1 * 1" using xy by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * u < 1" by simp qed have neq: "(1 - x * y) * u \ 1" if "u \ {0..1}" for u using less[of u] that by simp let ?F = "\z. ln (1-(1-x*y)*z)/(x*y-1)" have "(LBINT z=ereal 0..ereal 1. 1 / (1-(1-x*y)*z)) = ?F 1 - ?F 0" proof (rule interval_integral_FTC_finite, goal_cases cont deriv) case cont show ?case using neq by (intro continuous_intros) auto next case (deriv z) show ?case - unfolding has_field_derivative_iff_has_vector_derivative [symmetric] + unfolding has_real_derivative_iff_has_vector_derivative [symmetric] by (insert less[of z] xy \x * y < 1\ deriv) (rule derivative_eq_intros refl | simp)+ qed also have "\ = -ln (x*y) / (1-x*y)" using \x * y < 1\ by (simp add: field_simps) finally show ?thesis by (simp add: zero_ereal_def one_ereal_def) qed text \ The first part we shall show is the integration by parts. \ lemma beukers_aux_by_parts_aux: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" and "k \ n" shows "(LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z))) = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" using assms(3) proof (induction k) case (Suc k) note [derivative_intros] = DERIV_chain2[OF has_field_derivative_Gen_Shleg] define G where "G = (\y. -fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1))" define g where "g = (\y. fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))" have less: "(1 - x * y) * z < 1" and neq: "(1 - x * y) * z \ 1" if y: "y \ {0..1}" for y proof - from y xz have "x * y \ x * 1" by (intro mult_left_mono) auto also have "\ < 1" using xz by simp finally have "(1 - x * y) * z \ 1 * z" using xz y by (intro mult_right_mono) auto also have "\ < 1" using xz by simp finally show "(1 - x * y) * z < 1" by simp thus "(1 - x * y) * z \ 1" by simp qed have cont: "continuous_on {0..1} g" using neq by (auto simp: g_def intro!: continuous_intros) have deriv: "(G has_real_derivative g y) (at y within {0..1})" if y: "y \ {0..1}" for y unfolding G_def by (insert neq xz y, (rule derivative_eq_intros refl power_not_zero)+) (auto simp: divide_simps g_def) have deriv2: "(Q (n - Suc k) has_real_derivative Q (n - k) y) (at y within {0..1})" for y using Suc.prems by (auto intro!: derivative_eq_intros simp: Suc_diff_Suc Q_def) have "(LBINT y=0..1. Q (n-Suc k) y * (fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))) = (LBINT y=0..1. Q (n-Suc k) y * g y)" by (simp add: g_def) also have "(LBINT y=0..1. Q (n-Suc k) y * g y) = -(LBINT y=0..1. Q (n-k) y * G y)" using Suc.prems deriv deriv2 cont by (subst interval_lebesgue_integral_by_parts_01[where f = "Q (n-k)" and G = G]) (auto intro!: continuous_intros simp: Q_def) also have "\ = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" by (simp add: G_def flip: interval_lebesgue_integral_uminus) finally show ?case using Suc by simp qed auto lemma beukers_aux_by_parts: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" shows "(LBINT y=0..1. P y / (1-(1-x*y)*z)) = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have "(LBINT y=0..1. P y * (1/(1-(1-x*y)*z))) = 1 / fact n * (LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z)))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: P_def Q_def Shleg_altdef) also have "\ = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" by (subst beukers_aux_by_parts_aux[OF assms, of n], simp, subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: Q_def mult_ac Gen_Shleg_0_left power_mult_distrib) finally show ?thesis by simp qed text \ We then get the following by applying the integration by parts to \y\: \ lemma beukers_aux_integral_transform1: fixes z :: real assumes z: "z \ {0<..<1}" shows "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "1 - (1 - fst p * snd p) * z \ 0" by simp qed hence integrable: "set_integrable lborel (box (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" by (rule set_integrable_subset) (auto simp: box simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "(1 - (1 - fst p * snd p) * z) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (rule set_integrable_subset) (auto simp: box D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT x=0..1. (LBINT y=0..1. P x * P y / (1-(1-x*y)*z)))" unfolding D_def lborel_prod [symmetric] using box integrable by (subst lborel_pair.set_integral_fst') (simp_all add: interval_integral_Ioo lborel_prod) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using z by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) auto also have "\ = (LBINT x=0..1. (LBINT y=0..1. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding D_def lborel_prod [symmetric] using box integrable' by (subst lborel_pair.set_integral_fst') (simp_all add: D_def interval_integral_Ioo lborel_prod) finally show "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = \" . qed text \ We then change variables for \z\: \ lemma beukers_aux_integral_transform2: assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "(LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - define g where "g = (\z. (1 - z) / (1-(1-x*y)*z))" define g' where "g' = (\z. -x*y / (1-(1-x*y)*z)^2)" have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - (x * y)) * w < 1" and neq: "(1 - (x * y)) * w \ 1" if w: "w \ {0..1}" for w proof - have "(1 - (x * y)) * w \ (1 - (x * y)) * 1" using w \x * y < 1\ by (intro mult_left_mono) auto also have "\ < 1" using xy by simp finally show "(1 - (x * y)) * w < 1" . thus "(1 - (x * y)) * w \ 1" by simp qed have deriv: "(g has_real_derivative g' w) (at w within {0..1})" if "w \ {0..1}" for w unfolding g_def g'_def apply (insert that xy neq) apply (rule derivative_eq_intros refl)+ apply (simp_all add: divide_simps power2_eq_square) apply (auto simp: algebra_simps) done have "continuous_on {0..1} (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" using neq by (auto intro!: continuous_intros) moreover have "g ` {0..1} \ {0..1}" proof clarify fix w :: real assume w: "w \ {0..1}" have "(1 - x * y) * w \ 1 * w" using \x * y < 1\ xy w by (intro mult_right_mono) auto thus "g w \ {0..1}" unfolding g_def using less[of w] w by (auto simp: divide_simps) qed ultimately have cont: "continuous_on (g ` {0..1}) (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" by (rule continuous_on_subset) have cont': "continuous_on {0..1} g'" using neq by (auto simp: g'_def intro!: continuous_intros) have "(LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = (1-y)^n * (LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) also have "(LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w)) = -(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w))" by (subst interval_integral_endpoints_reverse)(simp add: g_def zero_ereal_def one_ereal_def) also have "(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w)) = (LBINT w=0..1. g' w * ((1 - g w) ^ n / (1 - (1-x*y) * g w)))" using deriv cont cont' by (subst interval_integral_substitution_finite [symmetric, where g = g and g' = g']) (simp_all add: zero_ereal_def one_ereal_def) also have "-\ = (LBINT z=0..1. ((x*y)^n * z^n / (1-(1-x*y)*z)^(n+1)))" unfolding interval_lebesgue_integral_uminus [symmetric] using xy apply (intro interval_lebesgue_integral_lborel_01_cong) apply (simp add: divide_simps g_def g'_def) apply (auto simp: algebra_simps power_mult_distrib power2_eq_square) done also have "(1-y)^n * \ = (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) finally show "\ = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" .. qed text \ Lastly, we apply the same integration by parts to \x\: \ lemma beukers_aux_integral_transform3: assumes w: "w \ {0<..<1}" shows "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "1 - (1 - fst p * snd p) * w \ 0" by simp qed hence integrable: "set_integrable lborel D (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "(1 - (1 - fst p * snd p) * w) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT y=0..1. (LBINT x=0..1. P x * (1-y)^n / (1-(1-x*y)*w)))" using integrable unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. P x / (1-(1-y*x)*w)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" using w by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (LBINT x=0..1. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" using integrable' unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) finally show ?thesis . qed text \ We need to show the existence of some of these triple integrals. \ lemma beukers_aux_integrable1: "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x,y),z). P x * P y / (1-(1-x*y)*z))" proof - have D [measurable]: "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp flip: borel_prod) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def case_prod_unfold proof (subst lborel_prod [symmetric], intro lborel_pair.Fubini_set_integrable AE_I2 impI; clarsimp?) fix x y :: real assume xy: "x > 0" "x < 1" "y > 0" "y < 1" have "x * y < 1" using xy mult_strict_mono[of x 1 y 1] by simp show "set_integrable lborel {0<..<1} (\z. P x * P y / (1-(1-x*y)*z))" by (rule set_integrable_subset[of _ "{0..1}"], rule borel_integrable_atLeastAtMost') (use \x*y<1\ beukers_denom_neq[of x y] xy in \auto intro!: continuous_intros simp: P_def\) next have "set_integrable lborel D (\(x,y). (\z\{0<..<1}. norm (P x * P y / (1-(1-x*y)*z)) \lborel))" proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C\<^sup>2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y assume xy: "(x, y) \ D" have "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) = norm (LBINT z:{0<..<1}. \P x\ * \P y\ * (1 / (1-(1-x*y)*z)))" proof (intro arg_cong[where f = norm] set_lebesgue_integral_cong allI impI, goal_cases) case (2 z) with beukers_denom_ineq[of x y z] xy show ?case by (auto simp: abs_mult D_def) qed (auto simp: abs_mult D_def) also have "\ = norm (\P x\ * \P y\ * (LBINT z=0..1. (1 / (1-(1-x*y)*z))))" by (subst set_integral_mult_right) (auto simp: interval_integral_Ioo) also have "\ = norm (norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)))" using beukers_aux_ln_conv_integral[of x y] xy by (simp add: D_def) also have "\ = norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y))" using xy mult_strict_mono[of x 1 y 1] by (auto simp: D_def divide_nonpos_nonneg abs_mult) also have "norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)) \ norm (C * C * (- ln (x * y) / (1 - x * y)))" using xy C[of x] C[of y] mult_strict_mono[of x 1 y 1] unfolding norm_mult norm_divide by (intro mult_mono C) (auto simp: D_def divide_nonpos_nonneg) finally show "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (simp add: power2_eq_square mult_ac) next show "set_borel_measurable lborel D (\(x, y). LBINT z:{0<..<1}. norm (P x * P y / (1 - (1 - x * y) * z)))" unfolding lborel_prod [symmetric] set_borel_measurable_def case_prod_unfold set_lebesgue_integral_def P_def by measurable qed thus "set_integrable lborel ({0<..<1} \ {0<..<1}) (\x. LBINT y:{0<..<1}. \P (fst x) * P (snd x)\ / \1 - (1 - fst x * snd x) * y\)" by (simp add: case_prod_unfold D_def) qed (auto simp: case_prod_unfold lborel_prod [symmetric] set_borel_measurable_def P_def) qed lemma beukers_aux_integrable2: "set_integrable lborel D' (\(z,x,y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" have "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = norm (P x) * (1-y)^n * ((x*y*z) / (1-(1-x*y)*z))^n / (1-(1-x*y)*z)" using xyz beukers_denom_ineq[of x y z] by (simp add: abs_mult power_divide mult_ac) also have "(x*y*z) / (1-(1-x*y)*z) = 1/((1-z)/(z*x*y)+1)" using xyz by (simp add: field_simps) also have "norm (P x) * (1-y)^n * \^n / (1-(1-x*y)*z) \ C * 1^n * 1^n / (1-(1-x*y)*z)" using xyz C[of x] beukers_denom_ineq[of x y z] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*z)\" by linarith finally show "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) \ norm (case (z,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed lemma beukers_aux_integrable3: "set_integrable lborel D' (\(w,x,y). P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y w :: real assume xyw: "x \ {0<..<1}" "y \ {0<..<1}" "w \ {0<..<1}" have "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = norm (P x) * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)" using xyw beukers_denom_ineq[of x y w] by (simp add: abs_mult power_divide mult_ac) also have "\ \ C * 1^n * 1^n / (1-(1-x*y)*w)" using xyw C[of x] beukers_denom_ineq[of x y w] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*w)\" by linarith finally show "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) \ norm (case (w,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed text \ Now we only need to put all of these results together: \ lemma beukers_integral2_conv_3: "beukers_integral2 = beukers_integral3" proof - have cont_P: "continuous_on {0..1} P" by (auto simp: P_def intro: continuous_intros) have D [measurable]: "D \ sets borel" "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp_all flip: borel_prod) have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "beukers_integral2 = (LBINT (x,y):D. P x * P y * (LBINT z=0..1. 1 / (1-(1-x*y)*z)))" unfolding beukers_integral2_def case_prod_unfold by (intro set_lebesgue_integral_cong allI impI, measurable) (subst beukers_aux_ln_conv_integral, auto simp: D_def) also have "\ = (LBINT (x,y):D. (LBINT z=0..1. P x * P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) auto also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * P y / (1-(1-x*y)*z)))" by (simp add: interval_integral_Ioo) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)))" proof (subst lborel_pair.Fubini_set_integral [symmetric]) have "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x, y), z). P x * P y / (1 - (1 - x * y) * z))" using beukers_aux_integrable1 by simp also have "?this \ set_integrable (lborel \\<^sub>M lborel) ({0<..<1} \ D) (\(z,x,y). P x * P y / (1 - (1 - x * y) * z))" unfolding set_integrable_def by (subst lborel_pair.integrable_product_swap_iff [symmetric], intro integrable_cong) (auto simp: indicator_def case_prod_unfold lborel_prod D_def) finally show \ . qed (auto simp: case_prod_unfold) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (rule set_lebesgue_integral_cong) (use beukers_aux_integral_transform1 in simp_all) also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using beukers_aux_integrable2 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT (x,y):D. (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" proof (intro set_lebesgue_integral_cong allI impI; clarify?) fix x y :: real assume xy: "(x, y) \ D" have "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = P x * (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = P x * (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" using xy by (subst beukers_aux_integral_transform2) (auto simp: D_def) also have "\ = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) finally show "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" . qed (auto simp: D_def simp flip: borel_prod) also have "\ = (LBINT w:{0<..<1}. (LBINT (x,y):D. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" using beukers_aux_integrable3 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)))" unfolding case_prod_unfold by (subst set_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. (x*y*w*(1-x)*(1-y))^n / (1-(1-x*y)*w)^(n+1)))" by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_integral_transform3) (auto simp: mult_ac power_mult_distrib) also have "\ = (LBINT w=0..1. (LBINT (x,y):D. (x*y*w*(1-x)*(1-y)*(1-w))^n / (1-(1-x*y)*w)^(n+1)))" by (subst set_integral_mult_right [symmetric]) (auto simp: case_prod_unfold mult_ac power_mult_distrib) also have "\ = beukers_integral3" using beukers_integral3_integrable unfolding D'_def D_def beukers_integral3_def by (subst (2) lborel_prod [symmetric], subst lborel_pair.set_integral_fst') (auto simp: case_prod_unfold interval_integral_Ioo lborel_prod algebra_simps) finally show ?thesis . qed subsection \The main result\ text \ Combining all of the results so far, we can derive the key inequalities \[0 < A\zeta(3) + B < 2 \zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\] for integers $A$, $B$ with $A > 0$. \ lemma zeta_3_linear_combination_bounds: obtains A B :: int where "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" proof - define I where "I = beukers_integral2" define d where "d = Lcm {1..n} ^ 3" have "d > 0" by (auto simp: d_def intro!: Nat.gr0I) from beukers_integral2_conv_int_combination obtain A' B :: int where *: "A' > 0" "I = A' * Re (zeta 3) + B / d" unfolding I_def d_def . define A where "A = A' * d" from * have A: "A > 0" "I = (A * Re (zeta 3) + B) / d" using \d > 0\ by (simp_all add: A_def field_simps) have "0 < I" using beukers_integral3_pos by (simp add: I_def beukers_integral2_conv_3) with \d > 0\ have "A * Re (zeta 3) + B > 0" by (simp add: field_simps A) moreover have "I \ 2 * (1 / 27) ^ n * Re (zeta 3)" using beukers_integral2_conv_3 beukers_integral3_le by (simp add: I_def) hence "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * d / 27 ^ n" using \d > 0\ by (simp add: A field_simps) ultimately show ?thesis using A by (intro that[of A B]) (auto simp: d_def) qed text \ If $\zeta(3) = \frac{a}{b}$ for some integers $a$ and $b$, we can thus derive the inequality $2b\zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\geq 1$ for any natural number $n$. \ lemma beukers_key_inequality: fixes a :: int and b :: nat assumes "b > 0" and ab: "Re (zeta 3) = a / b" shows "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" proof - from zeta_3_linear_combination_bounds obtain A B :: int where AB: "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" . from AB have "0 < (A * Re (zeta 3) + B) * b" using \b > 0\ by (intro mult_pos_pos) auto also have "\ = A * (Re (zeta 3) * b) + B * b" using \b > 0\ by (simp add: algebra_simps) also have "Re (zeta 3) * b = a" using \b > 0\ by (simp add: ab) also have "of_int A * of_int a + of_int (B * b) = of_int (A * a + B * b)" by simp finally have "1 \ A * a + B * b" by linarith hence "1 \ real_of_int (A * a + B * b)" by linarith also have "\ = (A * (a / b) + B) * b" using \b > 0\ by (simp add: ring_distribs) also have "a / b = Re (zeta 3)" by (simp add: ab) also have "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n" using AB by simp finally show "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" using \b > 0\ by (simp add: mult_ac) qed end (* TODO: Move *) lemma smallo_power: "n > 0 \ f \ o[F](g) \ (\x. f x ^ n) \ o[F](\x. g x ^ n)" by (induction n rule: nat_induct_non_zero) (auto intro: landau_o.small.mult) text \ This is now a contradiction, since $\text{lcm}\{1\ldots n\} \in o(3^n)$ by the Prime Number Theorem -- hence the main result. \ theorem zeta_3_irrational: "zeta 3 \ \" proof assume "zeta 3 \ \" obtain a :: int and b :: nat where "b > 0" and ab': "zeta 3 = a / b" proof - from \zeta 3 \ \\ obtain r where r: "zeta 3 = of_rat r" by (elim Rats_cases) also have "r = rat_of_int (fst (quotient_of r)) / rat_of_int (snd (quotient_of r))" by (intro quotient_of_div) auto also have "of_rat \ = (of_int (fst (quotient_of r)) / of_int (snd (quotient_of r)) :: complex)" by (simp add: of_rat_divide) also have "of_int (snd (quotient_of r)) = of_nat (nat (snd (quotient_of r)))" using quotient_of_denom_pos'[of r] by auto finally have "zeta 3 = of_int (fst (quotient_of r)) / of_nat (nat (snd (quotient_of r)))" . thus ?thesis using quotient_of_denom_pos'[of r] by (intro that[of "nat (snd (quotient_of r))" "fst (quotient_of r)"]) auto qed hence ab: "Re (zeta 3) = a / b" by simp interpret prime_number_theorem by standard (rule prime_number_theorem) have Lcm_upto_smallo: "(\n. real (Lcm {1..n})) \ o(\n. c ^ n)" if c: "c > exp 1" for c proof - have "0 < exp (1::real)" by simp also note c finally have "c > 0" . have "(\n. real (Lcm {1..n})) = (\n. real (Lcm {1..nat \real n\}))" by simp also have "\ \ o(\n. c powr real n)" using Lcm_upto.smallo' by (rule landau_o.small.compose) (simp_all add: c filterlim_real_sequentially) also have "(\n. c powr real n) = (\n. c ^ n)" using c \c > 0\ by (subst powr_realpow) auto finally show ?thesis . qed have "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ O(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n)" using \b > 0\ Re_zeta_ge_1[of 3] by simp also have "exp 1 < (3 :: real)" using e_approx_32 by (simp add: abs_if split: if_splits) hence "(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\n. (3 ^ n) ^ 3 / 27 ^ n)" unfolding of_nat_power by (intro landau_o.small.divide_right smallo_power Lcm_upto_smallo) auto also have "(\n. (3 ^ n) ^ 3 / 27 ^ n :: real) = (\_. 1)" by (simp add: power_mult [of 3, symmetric] mult.commute[of _ 3] power_mult[of _ 3]) finally have *: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\_. 1)" . have lim: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ 0" using smalloD_tendsto[OF *] by simp moreover have "1 \ real (2 * b) * Re (zeta 3) * real (Lcm {1..n} ^ 3) / 27 ^ n" for n using beukers_key_inequality[of b a] ab \b > 0\ by simp ultimately have "1 \ (0 :: real)" by (intro tendsto_lowerbound[OF lim] always_eventually allI) auto thus False by simp qed end \ No newline at end of file diff --git a/thys/Zeta_Function/Zeta_Library.thy b/thys/Zeta_Function/Zeta_Library.thy --- a/thys/Zeta_Function/Zeta_Library.thy +++ b/thys/Zeta_Function/Zeta_Library.thy @@ -1,898 +1,898 @@ section \Various preliminary material\ theory Zeta_Library imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" "Dirichlet_Series.Dirichlet_Series_Analysis" begin subsection \Facts about limits\ lemma at_within_altdef: "at x within A = (INF S\{S. open S \ x \ S}. principal (S \ (A - {x})))" unfolding at_within_def nhds_def inf_principal [symmetric] by (subst INF_inf_distrib [symmetric]) (auto simp: INF_constant) lemma tendsto_at_left_realI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_left c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_left c)" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \dx\{d<.. A m" proof (rule ccontr) assume "\ (\m. \dx\{d<.. A m)" then obtain m where **: "\d. d < c \ \x\{d<.. A m" by auto have "\X. \n. (f (X n) \ A m \ X n < c) \ X (Suc n) > c - max 0 ((c - X n) / 2)" proof (intro dependent_nat_choice, goal_cases) case 1 from **[of "c - 1"] show ?case by auto next case (2 x n) with **[of "c - max 0 (c - x) / 2"] show ?case by force qed then obtain X where X: "\n. f (X n) \ A m" "\n. X n < c" "\n. X (Suc n) > c - max 0 ((c - X n) / 2)" by auto (metis diff_gt_0_iff_gt half_gt_zero_iff max.absorb3 max.commute) have X_ge: "X n \ c - (c - X 0) / 2 ^ n" for n proof (induction n) case (Suc n) have "c - (c - X 0) / 2 ^ Suc n = c - (c - (c - (c - X 0) / 2 ^ n)) / 2" by simp also have "c - (c - (c - (c - X 0) / 2 ^ n)) / 2 \ c - (c - X n) / 2" by (intro diff_left_mono divide_right_mono Suc diff_right_mono) auto also have "\ = c - max 0 ((c - X n) / 2)" using X[of n] by (simp add: max_def) also have "\ < X (Suc n)" using X[of n] by simp finally show ?case by linarith qed auto have "X \ c" proof (rule tendsto_sandwich) show "eventually (\n. X n \ c) sequentially" using X by (intro always_eventually) (auto intro!: less_imp_le) show "eventually (\n. X n \ c - (c - X 0) / 2 ^ n) sequentially" using X_ge by (intro always_eventually) auto qed real_asymp+ hence "filterlim X (at_left c) sequentially" by (rule tendsto_imp_filterlim_at_left) (use X in \auto intro!: always_eventually less_imp_le\) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain d where d: "d m < c" "x \ {d m<.. f x \ A m" for m x by metis have ***: "at_left c = (INF S\{S. open S \ c \ S}. principal (S \ {..m. {d m<..}"]) auto qed lemma shows at_right_PInf [simp]: "at_right (\ :: ereal) = bot" and at_left_MInf [simp]: "at_left (-\ :: ereal) = bot" proof - have "{(\::ereal)<..} = {}" "{..<-(\::ereal)} = {}" by auto thus "at_right (\ :: ereal) = bot" "at_left (-\ :: ereal) = bot" by (simp_all add: at_within_def) qed lemma tendsto_at_left_erealI_sequentially: fixes f :: "ereal \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_left c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_left c)" proof (cases c) case [simp]: PInf have "((\x. f (ereal x)) \ y) at_top" using assms by (intro tendsto_at_topI_sequentially assms) (simp_all flip: ereal_tendsto_simps add: o_def filterlim_at) thus ?thesis by (simp add: at_left_PInf filterlim_filtermap) next case [simp]: MInf thus ?thesis by auto next case [simp]: (real c') have "((\x. f (ereal x)) \ y) (at_left c')" proof (intro tendsto_at_left_realI_sequentially assms) fix X assume *: "filterlim X (at_left c') sequentially" show "filterlim (\n. ereal (X n)) (at_left c) sequentially" by (rule filterlim_compose[OF _ *]) (simp add: sequentially_imp_eventually_within tendsto_imp_filterlim_at_left) qed thus ?thesis by (simp add: at_left_ereal filterlim_filtermap) qed lemma tendsto_at_right_realI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_right c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_right c)" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \d>c. \x\{c<.. A m" proof (rule ccontr) assume "\ (\m. \d>c. \x\{c<.. A m)" then obtain m where **: "\d. d > c \ \x\{c<.. A m" by auto have "\X. \n. (f (X n) \ A m \ X n > c) \ X (Suc n) < c + max 0 ((X n - c) / 2)" proof (intro dependent_nat_choice, goal_cases) case 1 from **[of "c + 1"] show ?case by auto next case (2 x n) with **[of "c + max 0 (x - c) / 2"] show ?case by force qed then obtain X where X: "\n. f (X n) \ A m" "\n. X n > c" "\n. X (Suc n) < c + max 0 ((X n - c) / 2)" by auto (metis add.left_neutral half_gt_zero_iff less_diff_eq max.absorb4) have X_le: "X n \ c + (X 0 - c) / 2 ^ n" for n proof (induction n) case (Suc n) have "X (Suc n) < c + max 0 ((X n - c) / 2)" by (intro X) also have "\ = c + (X n - c) / 2" using X[of n] by (simp add: field_simps max_def) also have "\ \ c + (c + (X 0 - c) / 2 ^ n - c) / 2" by (intro add_left_mono divide_right_mono Suc diff_right_mono) auto also have "\ = c + (X 0 - c) / 2 ^ Suc n" by simp finally show ?case by linarith qed auto have "X \ c" proof (rule tendsto_sandwich) show "eventually (\n. X n \ c) sequentially" using X by (intro always_eventually) (auto intro!: less_imp_le) show "eventually (\n. X n \ c + (X 0 - c) / 2 ^ n) sequentially" using X_le by (intro always_eventually) auto qed real_asymp+ hence "filterlim X (at_right c) sequentially" by (rule tendsto_imp_filterlim_at_right) (use X in \auto intro!: always_eventually less_imp_le\) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain d where d: "d m > c" "x \ {c<.. f x \ A m" for m x by metis have ***: "at_right c = (INF S\{S. open S \ c \ S}. principal (S \ {c<..}))" by (simp add: at_within_altdef) from d show ?thesis unfolding *** A using A(1,2) by (intro filterlim_base[of _ "\m. {.. 'b::first_countable_topology" assumes *: "\X. filterlim X (at_right c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_right c)" proof (cases c) case [simp]: MInf have "((\x. f (-ereal x)) \ y) at_top" using assms by (intro tendsto_at_topI_sequentially assms) (simp_all flip: uminus_ereal.simps ereal_tendsto_simps add: o_def filterlim_at) thus ?thesis by (simp add: at_right_MInf filterlim_filtermap at_top_mirror) next case [simp]: PInf thus ?thesis by auto next case [simp]: (real c') have "((\x. f (ereal x)) \ y) (at_right c')" proof (intro tendsto_at_right_realI_sequentially assms) fix X assume *: "filterlim X (at_right c') sequentially" show "filterlim (\n. ereal (X n)) (at_right c) sequentially" by (rule filterlim_compose[OF _ *]) (simp add: sequentially_imp_eventually_within tendsto_imp_filterlim_at_right) qed thus ?thesis by (simp add: at_right_ereal filterlim_filtermap) qed proposition analytic_continuation': assumes hol: "f holomorphic_on S" "g holomorphic_on S" and "open S" and "connected S" and "U \ S" and "\ \ S" and "\ islimpt U" and fU0 [simp]: "\z. z \ U \ f z = g z" and "w \ S" shows "f w = g w" using analytic_continuation[OF holomorphic_on_diff[OF hol] assms(3-7) _ assms(9)] assms(8) by simp subsection \Various facts about integrals\ lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" assumes "set_integrable M A f" "set_borel_measurable M A g" assumes "AE x in M. x \ A \ norm (g x) \ norm (f x)" shows "set_integrable M A g" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) from assms(1) show "integrable M (\x. indicator A x *\<^sub>R f x)" by (simp add: set_integrable_def) from assms(2) show "(\x. indicat_real A x *\<^sub>R g x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(3) show "AE x in M. norm (indicat_real A x *\<^sub>R g x) \ norm (indicat_real A x *\<^sub>R f x)" by eventually_elim (simp add: indicator_def) qed (* TODO replace version in library. Better name? *) 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 of_bool_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 subsection \Uniform convergence of integrals\ lemma has_absolute_integral_change_of_variables_1': fixes f :: "real \ real" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "(\x. \g' x\ *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1) \ (\x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \ integral (g ` S) (\x. vec (f x)) = (vec b :: real ^ 1)" - using assms unfolding has_field_derivative_iff_has_vector_derivative + using assms unfolding has_real_derivative_iff_has_vector_derivative by (intro has_absolute_integral_change_of_variables_1 assms) auto thus ?thesis by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq) qed lemma set_nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes "set_borel_measurable borel A f" assumes "\x. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \lborel) < \" shows "(\\<^sup>+x\A. f x \lborel) = integral A f" proof - have eq: "(\\<^sup>+x\A. f x \lborel) = (\\<^sup>+x. ennreal (indicator A x * f x) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = integral UNIV (\x. indicator A x * f x)" using assms eq by (intro nn_integral_lborel_eq_integral) (auto simp: indicator_def set_borel_measurable_def) also have "integral UNIV (\x. indicator A x * f x) = integral A (\x. indicator A x * f x)" by (rule integral_spike_set) (auto intro: empty_imp_negligible) also have "\ = integral A f" by (rule integral_cong) (auto simp: indicator_def) finally show ?thesis . qed 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. ennreal (f x) * indicator \ x) = ennreal I" proof - have "integral\<^sup>N lborel (\x. ennreal (f x) * indicator \ x) = integral\<^sup>N lborel (\x. ennreal (indicator \ x * f x))" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal I" using assms by (intro nn_integral_has_integral_lebesgue) finally show ?thesis . qed lemma uniform_limit_set_lebesgue_integral: fixes f :: "'a \ 'b :: euclidean_space \ 'c :: {banach, second_countable_topology}" assumes "set_integrable lborel X' g" assumes [measurable]: "X' \ sets borel" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel X' (f y)" assumes "\y. y \ Y \ (AE t\X' in lborel. norm (f y t) \ g t)" assumes "eventually (\x. X x \ sets borel \ X x \ X') F" assumes "filterlim (\x. set_lebesgue_integral lborel (X x) g) (nhds (set_lebesgue_integral lborel X' g)) F" shows "uniform_limit Y (\x y. set_lebesgue_integral lborel (X x) (f y)) (\y. set_lebesgue_integral lborel X' (f y)) F" proof (rule uniform_limitI, goal_cases) case (1 \) have integrable_g: "set_integrable lborel U g" if "U \ sets borel" "U \ X'" for U by (rule set_integrable_subset[OF assms(1)]) (use that in auto) have "eventually (\x. dist (set_lebesgue_integral lborel (X x) g) (set_lebesgue_integral lborel X' g) < \) F" using \\ > 0\ assms by (auto simp: tendsto_iff) from this show ?case using \eventually (\_. _ \ _) F\ proof eventually_elim case (elim x) hence [measurable]:"X x \ sets borel" and "X x \ X'" by auto have integrable: "set_integrable lborel U (f y)" if "y \ Y" "U \ sets borel" "U \ X'" for y U apply (rule set_integrable_subset) apply (rule set_integrable_bound[OF assms(1)]) apply (use assms(3) that in \simp add: set_borel_measurable_def\) using assms(4)[OF \y \ Y\] apply eventually_elim apply force using that apply simp_all done show ?case proof fix y assume "y \ Y" have "dist (set_lebesgue_integral lborel (X x) (f y)) (set_lebesgue_integral lborel X' (f y)) = norm (set_lebesgue_integral lborel X' (f y) - set_lebesgue_integral lborel (X x) (f y))" by (simp add: dist_norm norm_minus_commute) also have "set_lebesgue_integral lborel X' (f y) - set_lebesgue_integral lborel (X x) (f y) = set_lebesgue_integral lborel (X' - X x) (f y)" unfolding set_lebesgue_integral_def apply (subst Bochner_Integration.integral_diff [symmetric]) unfolding set_integrable_def [symmetric] apply (rule integrable; (fact | simp)) apply (rule integrable; fact) apply (intro Bochner_Integration.integral_cong) apply (use \X x \ X'\ in \auto simp: indicator_def\) done also have "norm \ \ (\t\X'-X x. norm (f y t) \lborel)" by (intro set_integral_norm_bound integrable) (fact | simp)+ also have "AE t\X' - X x in lborel. norm (f y t) \ g t" using assms(4)[OF \y \ Y\] by eventually_elim auto with \y \ Y\ have "(\t\X'-X x. norm (f y t) \lborel) \ (\t\X'-X x. g t \lborel)" by (intro set_integral_mono_AE set_integrable_norm integrable integrable_g) auto also have "\ = (\t\X'. g t \lborel) - (\t\X x. g t \lborel)" unfolding set_lebesgue_integral_def apply (subst Bochner_Integration.integral_diff [symmetric]) unfolding set_integrable_def [symmetric] apply (rule integrable_g; (fact | simp)) apply (rule integrable_g; fact) apply (intro Bochner_Integration.integral_cong) apply (use \X x \ X'\ in \auto simp: indicator_def\) done also have "\ \ dist (\t\X x. g t \lborel) (\t\X'. g t \lborel)" by (simp add: dist_norm) also have "\ < \" by fact finally show "dist (set_lebesgue_integral lborel (X x) (f y)) (set_lebesgue_integral lborel X' (f y)) < \" . qed qed qed lemma integral_dominated_convergence_at_right: fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M and c :: real assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) (at_right c)" assumes bound: "\\<^sub>F i in at_right c. AE x in M. norm (s i x) \ w x" shows "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) (at_right c)" proof (rule tendsto_at_right_realI_sequentially) fix X :: "nat \ real" assume X: "filterlim X (at_right c) sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) (at_right c)" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma integral_dominated_convergence_at_left: fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M and c :: real assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) (at_left c)" assumes bound: "\\<^sub>F i in at_left c. AE x in M. norm (s i x) \ w x" shows "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) (at_left c)" proof (rule tendsto_at_left_realI_sequentially) fix X :: "nat \ real" assume X: "filterlim X (at_left c) sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) (at_left c)" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma uniform_limit_interval_integral_right: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes "a < b" shows "uniform_limit Y (\b' y. LBINT t=a..b'. f y t) (\y. LBINT t=a..b. f y t) (at_left b)" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\b'. b' \ {a<..a < b\ by (intro eventually_at_leftI) with \a < b\ have "?thesis \ uniform_limit Y (\b' y. \t\einterval a (min b b'). f y t \lborel) (\y. \t\einterval a b. f y t \lborel) (at_left b)" by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff min_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F b' in at_left b. einterval a (min b b') \ sets borel \ einterval a (min b b') \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\b'. set_lebesgue_integral lborel (einterval a (min b b')) g) \ set_lebesgue_integral lborel (einterval a b) g) (at_left b)" unfolding set_lebesgue_integral_def proof (intro tendsto_at_left_erealI_sequentially integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix X :: "nat \ ereal" and n :: nat have "set_borel_measurable borel (einterval a (min b (X n))) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next fix X :: "nat \ ereal" assume X: "filterlim X (at_left b) sequentially" show "AE x in lborel. (\n. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval a (min b (X t))) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval a (min b (X t))"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {max a x<..t. X t \ {max a x<..t. indicator (einterval a (min b (X t))) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval a (min b (X n))) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma uniform_limit_interval_integral_left: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes "a < b" shows "uniform_limit Y (\a' y. LBINT t=a'..b. f y t) (\y. LBINT t=a..b. f y t) (at_right a)" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\b'. b' \ {a<..a < b\ by (intro eventually_at_rightI) with \a < b\ have "?thesis \ uniform_limit Y (\a' y. \t\einterval (max a a') b. f y t \lborel) (\y. \t\einterval a b. f y t \lborel) (at_right a)" by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff max_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F a' in at_right a. einterval (max a a') b \ sets borel \ einterval (max a a') b \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\a'. set_lebesgue_integral lborel (einterval (max a a') b) g) \ set_lebesgue_integral lborel (einterval a b) g) (at_right a)" unfolding set_lebesgue_integral_def proof (intro tendsto_at_right_erealI_sequentially integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix X :: "nat \ ereal" and n :: nat have "set_borel_measurable borel (einterval (max a (X n)) b) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next fix X :: "nat \ ereal" assume X: "filterlim X (at_right a) sequentially" show "AE x in lborel. (\n. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval (max a (X t)) b) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval (max a (X t)) b"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {a<..t. X t \ {a<..t. indicator (einterval (max a (X t)) b) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma uniform_limit_interval_integral_sequentially: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes a': "filterlim a' (at_right a) sequentially" assumes b': "filterlim b' (at_left b) sequentially" assumes "a < b" shows "uniform_limit Y (\n y. LBINT t=a' n..b' n. f y t) (\y. LBINT t=a..b. f y t) sequentially" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\n. a < a' n \ a' n < b' n \ b' n < b) sequentially" proof - from ereal_dense2[OF \a < b\] obtain t where t: "a < ereal t" "ereal t < b" by blast from t have "eventually (\n. a' n \ {a<..n. b' n \ {t<..n. a < a' n \ a' n < b' n \ b' n < b) sequentially" by eventually_elim auto qed have "?thesis \ uniform_limit Y (\n y. \t\einterval (max a (a' n)) (min b (b' n)). f y t \lborel) (\y. \t\einterval a b. f y t \lborel) sequentially" using \a < b\ by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff min_def max_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F n in sequentially. einterval (max a (a' n)) (min b (b' n)) \ sets borel \ einterval (max a (a' n)) (min b (b' n)) \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\n. set_lebesgue_integral lborel (einterval (max a (a' n)) (min b (b' n))) g) \ set_lebesgue_integral lborel (einterval a b) g) sequentially" unfolding set_lebesgue_integral_def proof (intro integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix n :: nat have "set_borel_measurable borel (einterval (max a (a' n)) (min b (b' n))) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next show "AE x in lborel. (\n. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval (max a (a' t)) (min b (b' t))) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval (max a (a' t)) (min b (b' t))"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {a<..n. x \ {a' n<..n. a' n \ {a<..n. b' n \ {x<..n. x \ {a' n<..t. indicator (einterval (max a (a' t)) (min b (b' t))) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma interval_lebesgue_integrable_combine: assumes "interval_lebesgue_integrable lborel A B f" assumes "interval_lebesgue_integrable lborel B C f" assumes "set_borel_measurable borel (einterval A C) f" assumes "A \ B" "B \ C" shows "interval_lebesgue_integrable lborel A C f" proof - have meas: "set_borel_measurable borel (einterval A B \ einterval B C) f" by (rule set_borel_measurable_subset[OF assms(3)]) (use assms in \auto simp: einterval_def\) have "set_integrable lborel (einterval A B \ einterval B C) f" using assms by (intro set_integrable_Un) (auto simp: interval_lebesgue_integrable_def) also have "?this \ set_integrable lborel (einterval A C) f" proof (cases "B \ {\, -\}") case True with assms have "einterval A B \ einterval B C = einterval A C" by (auto simp: einterval_def) thus ?thesis by simp next case False then obtain B' where [simp]: "B = ereal B'" by (cases B) auto have "indicator (einterval A C) x = (indicator (einterval A B \ einterval B C) x :: real)" if "x \ B'" for x using assms(4,5) that by (cases A; cases C) (auto simp: einterval_def indicator_def) hence "{x \ space lborel. indicat_real (einterval A B \ einterval B C) x *\<^sub>R f x \ indicat_real (einterval A C) x *\<^sub>R f x} \ {B'}" by force thus ?thesis unfolding set_integrable_def using meas assms by (intro integrable_cong_AE AE_I[of _ _ "{B'}"]) (simp_all add: set_borel_measurable_def) qed also have "\ \ ?thesis" using order.trans[OF assms(4,5)] by (simp add: interval_lebesgue_integrable_def) finally show ?thesis . qed lemma interval_lebesgue_integrable_bigo_right: fixes A B :: real fixes f :: "real \ real" assumes "f \ O[at_left B](g)" assumes cont: "continuous_on {A.. 0" "eventually (\x. norm (f x) \ c * norm (g x)) (at_left B)" by (elim landau_o.bigE) then obtain B' where B': "B' < B" "\x. x \ {B'<.. norm (f x) \ c * norm (g x)" using \A < B\ by (auto simp: Topological_Spaces.eventually_at_left[of A]) show ?thesis proof (rule interval_lebesgue_integrable_combine) show "interval_lebesgue_integrable lborel A (max A B') f" using B' assms by (intro interval_integrable_continuous_on continuous_on_subset[OF cont]) auto show "set_borel_measurable borel (einterval (ereal A) (ereal B)) f" using assms by simp have meas': "set_borel_measurable borel {max A B'<..x. c * g x)" using assms by (simp add: interval_lebesgue_integrable_def) thus "set_integrable lborel {max A B'<..x. c * g x)" by (rule set_integrable_subset) auto next fix x assume "x \ {max A B'<.. c * norm (g x)" by (intro B') auto also have "\ \ norm (c * g x)" unfolding norm_mult by (intro mult_right_mono) auto finally show "norm (f x) \ norm (c * g x)" . qed (use meas' in \simp_all add: set_borel_measurable_def\) thus "interval_lebesgue_integrable lborel (ereal (max A B')) (ereal B) f" unfolding interval_lebesgue_integrable_def einterval_eq_Icc using \B' < B\ assms by simp qed (use B' assms in auto) qed lemma interval_lebesgue_integrable_bigo_left: fixes A B :: real fixes f :: "real \ real" assumes "f \ O[at_right A](g)" assumes cont: "continuous_on {A<..B} f" assumes meas: "set_borel_measurable borel {A<.. 0" "eventually (\x. norm (f x) \ c * norm (g x)) (at_right A)" by (elim landau_o.bigE) then obtain A' where A': "A' > A" "\x. x \ {A<.. norm (f x) \ c * norm (g x)" using \A < B\ by (auto simp: Topological_Spaces.eventually_at_right[of A]) show ?thesis proof (rule interval_lebesgue_integrable_combine) show "interval_lebesgue_integrable lborel (min B A') B f" using A' assms by (intro interval_integrable_continuous_on continuous_on_subset[OF cont]) auto show "set_borel_measurable borel (einterval (ereal A) (ereal B)) f" using assms by simp have meas': "set_borel_measurable borel {A<..x. c * g x)" using assms by (simp add: interval_lebesgue_integrable_def) thus "set_integrable lborel {A<..x. c * g x)" by (rule set_integrable_subset) auto next fix x assume "x \ {A<.. c * norm (g x)" by (intro A') auto also have "\ \ norm (c * g x)" unfolding norm_mult by (intro mult_right_mono) auto finally show "norm (f x) \ norm (c * g x)" . qed (use meas' in \simp_all add: set_borel_measurable_def\) thus "interval_lebesgue_integrable lborel (ereal A) (ereal (min B A')) f" unfolding interval_lebesgue_integrable_def einterval_eq_Icc using \A' > A\ assms by simp qed (use A' assms in auto) qed subsection \Other material\ (* TODO: Library *) 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 fps_expansion_cong: assumes "eventually (\x. g x = h x) (nhds x)" shows "fps_expansion g x = fps_expansion h x" proof - have "(deriv ^^ n) g x = (deriv ^^ n) h x" for n by (intro higher_deriv_cong_ev assms refl) thus ?thesis by (simp add: fps_expansion_def) qed lemma fps_expansion_eq_zero_iff: assumes "g holomorphic_on ball z r" "r > 0" shows "fps_expansion g z = 0 \ (\z\ball z r. g z = 0)" proof assume *: "\z\ball z r. g z = 0" have "eventually (\w. w \ ball z r) (nhds z)" using assms by (intro eventually_nhds_in_open) auto hence "eventually (\z. g z = 0) (nhds z)" by eventually_elim (use * in auto) hence "fps_expansion g z = fps_expansion (\_. 0) z" by (intro fps_expansion_cong) thus "fps_expansion g z = 0" by (simp add: fps_expansion_def fps_zero_def) next assume *: "fps_expansion g z = 0" have "g w = 0" if "w \ ball z r" for w by (rule holomorphic_fun_eq_0_on_ball[OF assms(1) that]) (use * in \auto simp: fps_expansion_def fps_eq_iff\) thus "\w\ball z r. g w = 0" by blast qed lemma fds_nth_higher_deriv: "fds_nth ((fds_deriv ^^ k) F) = (\n. (-1) ^ k * of_real (ln n) ^ k * fds_nth F n)" by (induction k) (auto simp: fds_nth_deriv fun_eq_iff simp flip: scaleR_conv_of_real) lemma binomial_n_n_minus_one [simp]: "n > 0 \ n choose (n - Suc 0) = n" by (cases n) auto lemma has_field_derivative_complex_powr_right: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z within A)" by (rule DERIV_subset, rule has_field_derivative_powr_right) auto lemmas has_field_derivative_complex_powr_right' = has_field_derivative_complex_powr_right[THEN DERIV_chain2] end \ No newline at end of file