diff --git a/thys/Derangements/Derangements.thy b/thys/Derangements/Derangements.thy --- a/thys/Derangements/Derangements.thy +++ b/thys/Derangements/Derangements.thy @@ -1,515 +1,515 @@ (* Author: Lukas Bulwahn *) section \Derangements\ theory Derangements imports Complex_Main "HOL-Library.Permutations" begin subsection \Preliminaries\ subsubsection \Additions to @{theory HOL.Finite_Set} Theory\ lemma card_product_dependent: assumes "finite S" "\x \ S. finite (T x)" shows "card {(x, y). x \ S \ y \ T x} = (\x \ S. card (T x))" using card_SigmaI[OF assms, symmetric] by (auto intro!: arg_cong[where f=card] simp add: Sigma_def) subsubsection \Additions to @{theory "HOL-Library.Permutations"} Theory\ lemma permutes_imp_bij': assumes "p permutes S" shows "bij p" using assms by (meson bij_def permutes_inj permutes_surj) lemma permutesE: assumes "p permutes S" obtains "bij p" "\x. x \ S \ p x = x" using assms by (simp add: permutes_def permutes_imp_bij') lemma bij_imp_permutes': assumes "bij p" "\x. x \ A \ p x = x" shows "p permutes A" using assms bij_imp_permutes permutes_superset by force lemma permutes_swap: assumes "p permutes S" shows "Fun.swap x y p permutes (insert x (insert y S))" proof - from assms have "p permutes (insert x (insert y S))" by (meson permutes_subset subset_insertI) moreover have "Fun.swap x y id permutes (insert x (insert y S))" by (simp add: permutes_swap_id) ultimately show "Fun.swap x y p permutes (insert x (insert y S))" by (metis comp_id comp_swap permutes_compose) qed lemma bij_extends: "bij p \ p x = x \ bij (p(x := y, inv p y := x))" unfolding bij_def proof (rule conjI; erule conjE) assume a: "inj p" "p x = x" show "inj (p(x := y, inv p y := x))" proof (intro injI) fix z z' assume "(p(x := y, inv p y := x)) z = (p(x := y, inv p y := x)) z'" from this a show "z = z'" by (auto split: if_split_asm simp add: inv_f_eq inj_eq) qed next assume a: "inj p" "surj p" "p x = x" { fix x' from a have "(p(x := y, inv p y := x)) (((inv p)(y := x, x := inv p y)) x') = x'" by (auto split: if_split_asm) (metis surj_f_inv_f)+ } from this show "surj (p(x := y, inv p y := x))" by (metis surjI) qed lemma permutes_add_one: assumes "p permutes S" "x \ S" "y \ S" shows "p(x := y, inv p y := x) permutes (insert x S)" proof (rule bij_imp_permutes') from assms show "bij (p(x := y, inv p y := x))" by (meson bij_extends permutes_def permutes_imp_bij') from assms show "\z. z \ insert x S \ (p(x := y, inv p y := x)) z = z" by (metis fun_upd_apply insertCI permutes_def permutes_inverses(1)) qed lemma permutations_skip_one: assumes "p permutes S" "x : S" shows "p(x := x, inv p x := p x) permutes (S - {x})" proof (rule bij_imp_permutes') from assms show "\z. z \ S - {x} \ (p(x := x, inv p x := p x)) z = z" by (auto elim: permutesE simp add: bij_inv_eq_iff) (simp add: assms(1) permutes_in_image permutes_inv) from assms have "inj (p(x := x, inv p x := p x))" by (intro injI) (auto split: if_split_asm; metis permutes_inverses(2))+ from assms this show "bij (p(x := x, inv p x := p x))" by (metis UNIV_I bij_betw_imageI bij_betw_swap_iff permutes_inj permutes_surj surj_f_inv_f swap_def) qed lemma permutes_drop_cycle_size_two: assumes "p permutes S" "p (p x) = x" shows "Fun.swap x (p x) p permutes (S - {x, p x})" using assms by (auto intro!: bij_imp_permutes' elim!: permutesE) (metis swap_apply(1,3)) subsection \Fixpoint-Free Permutations\ definition derangements :: "nat set \ (nat \ nat) set" where "derangements S = {p. p permutes S \ (\x \ S. p x \ x)}" lemma derangementsI: assumes "p permutes S" "\x. x \ S \ p x \ x" shows "p \ derangements S" using assms unfolding derangements_def by auto lemma derangementsE: assumes "d : derangements S" obtains "d permutes S" "\x\S. d x \ x" using assms unfolding derangements_def by auto subsection \Properties of Derangements\ lemma derangements_inv: assumes d: "d \ derangements S" shows "inv d \ derangements S" using assms by (auto intro!: derangementsI elim!: derangementsE simp add: permutes_inv permutes_inv_eq) lemma derangements_in_image: assumes "d \ derangements A" "x \ A" shows "d x \ A" using assms by (auto elim: derangementsE simp add: permutes_in_image) lemma derangements_in_image_strong: assumes "d \ derangements A" "x \ A" shows "d x \ A - {x}" using assms by (auto elim: derangementsE simp add: permutes_in_image) lemma derangements_inverse_in_image: assumes "d \ derangements A" "x \ A" shows "inv d x \ A" using assms by (auto intro: derangements_in_image derangements_inv) lemma derangements_fixpoint: assumes "d \ derangements A" "x \ A" shows "d x = x" using assms by (auto elim!: derangementsE simp add: permutes_def) lemma derangements_no_fixpoint: assumes "d \ derangements A" "x \ A" shows "d x \ x" using assms by (auto elim: derangementsE) lemma finite_derangements: assumes "finite A" shows "finite (derangements A)" using assms unfolding derangements_def by (auto simp add: finite_permutations) subsection \Construction of Derangements\ lemma derangements_empty[simp]: "derangements {} = {id}" unfolding derangements_def by auto lemma derangements_singleton[simp]: "derangements {x} = {}" unfolding derangements_def by auto lemma derangements_swap: assumes "d \ derangements S" "x \ S" "y \ S" "x \ y" shows "Fun.swap x y d \ derangements (insert x (insert y S))" proof (rule derangementsI) from assms show "Fun.swap x y d permutes (insert x (insert y S))" by (auto intro: permutes_swap elim: derangementsE) from assms have s: "d x = x" "d y = y" by (auto intro: derangements_fixpoint) { fix x' assume "x' : insert x (insert y S)" from s assms \x \ y\ this show "Fun.swap x y d x' \ x'" by (cases "x' = x"; cases "x' = y") (auto dest: derangements_no_fixpoint) } qed lemma derangements_skip_one: assumes d: "d \ derangements S" and "x \ S" "d (d x) \ x" shows "d(x := x, inv d x := d x) \ derangements (S - {x})" proof - from d have bij: "bij d" by (auto elim: derangementsE simp add: permutes_imp_bij') from d \x : S\ have that: "d x : S - {x}" by (auto dest: derangements_in_image derangements_no_fixpoint) from d \d (d x) \ x\ bij have "\x'\S - {x}. (d(x := x, inv d x := d x)) x' \ x'" by (auto elim!: derangementsE simp add: bij_inv_eq_iff) from d \x : S\ this show derangements: "d(x:=x, inv d x:= d x) : derangements (S - {x})" by (meson derangementsE derangementsI permutations_skip_one) qed lemma derangements_add_one: assumes "d \ derangements S" "x \ S" "y \ S" shows "d(x := y, inv d y := x) \ derangements (insert x S)" proof (rule derangementsI) from assms show "d(x := y, inv d y := x) permutes (insert x S)" by (auto intro: permutes_add_one elim: derangementsE) next fix z assume "z : insert x S" from assms this derangements_inverse_in_image[OF assms(1), of y] show "(d(x := y, inv d y := x)) z \ z" by (auto elim: derangementsE) qed lemma derangements_drop_minimal_cycle: assumes "d \ derangements S" "d (d x) = x" shows "Fun.swap x (d x) d \ derangements (S - {x, d x})" proof (rule derangementsI) from assms show "Fun.swap x (d x) d permutes (S - {x, d x})" by (meson derangementsE permutes_drop_cycle_size_two) next fix y assume "y \ S - {x, d x}" from assms this show "Fun.swap x (d x) d y \ y" by (auto elim: derangementsE) qed subsection \Cardinality of Derangements\ subsubsection \Recursive Characterization\ fun count_derangements :: "nat \ nat" where "count_derangements 0 = 1" | "count_derangements (Suc 0) = 0" | "count_derangements (Suc (Suc n)) = (n + 1) * (count_derangements (Suc n) + count_derangements n)" lemma card_derangements: assumes "finite S" "card S = n" shows "card (derangements S) = count_derangements n" using assms proof (induct n arbitrary: S rule: count_derangements.induct) case 1 from this show ?case by auto next case 2 from this derangements_singleton finite_derangements show ?case using Finite_Set.card_0_eq card_eq_SucD count_derangements.simps(2) by fastforce next case (3 n) from 3(4) obtain x where "x \ S" using card_eq_SucD insertI1 by auto let ?D1 = "(\(y, d). Fun.swap x y d) ` {(y, d). y \ S & y \ x & d : derangements (S - {x, y})}" let ?D2 = "(\(y, f). f(x:=y, inv f y := x)) ` ((S - {x}) \ derangements (S - {x}))" from \x \ S\ have subset1: "?D1 \ derangements S" proof (auto) fix y d assume "y \ S" "y \ x" assume d: "d \ derangements (S - {x, y})" from \x : S\ \y : S\ have S: "S = insert x (insert y (S - {x, y}))" by auto from d \x : S\ \y : S\ \y \ x\ show "Fun.swap x y d \ derangements S" by (subst S) (auto intro!: derangements_swap) qed have subset2: "?D2 \ derangements S" proof (rule subsetI, erule imageE, simp split: prod.split_asm, (erule conjE)+) fix d y assume "d : derangements (S - {x})" "y : S" "y \ x" from this have "d(x := y, inv d y := x) \ derangements (insert x (S - {x}))" by (intro derangements_add_one) auto from this \x : S\ show "d(x := y, inv d y := x) \ derangements S" using insert_Diff by fastforce qed have split: "derangements S = ?D1 \ ?D2" proof from subset1 subset2 show "?D1 \ ?D2 \ derangements S" by simp next show "derangements S \ ?D1 \ ?D2" proof fix d assume d: "d : derangements S" show "d : ?D1 \ ?D2" proof (cases "d (d x) = x") case True from \x : S\ d have "d x \ S" "d x \ x" by (auto simp add: derangements_in_image derangements_no_fixpoint) from d True have "Fun.swap x (d x) d \ derangements (S - {x, d x})" by (rule derangements_drop_minimal_cycle) from \d x \ S\ \d x \ x\ this have "d : ?D1" by (auto intro: image_eqI[where x = "(d x, Fun.swap x (d x) d)"]) from this show ?thesis by auto next case False from d have bij: "bij d" by (auto elim: derangementsE simp add: permutes_imp_bij') from d \x : S\ have that: "d x : S - {x}" by (intro derangements_in_image_strong) from d \x : S\ False have derangements: "d(x:=x, inv d x:= d x) : derangements (S - {x})" by (auto intro: derangements_skip_one) from this have "bij (d(x := x, inv d x:= d x))" by (metis derangementsE permutes_imp_bij')+ from this have a: "inv (d(x := x, inv d x := d x)) (d x) = inv d x" by (metis bij_inv_eq_iff fun_upd_same) from bij have x: "d (inv d x) = x" by (meson bij_inv_eq_iff) from d derangements_inv[of d] \x : S\ have "inv d x \ x" "d x \ x" by (auto dest: derangements_no_fixpoint) from this a x have d_eq: "d = d(inv d x := d x, x := d x, inv (d(x := x, inv d x := d x)) (d x) := x)" by auto from derangements that have "(d x, d(x:=x, inv d x:=d x)) : ((S - {x}) \ derangements (S - {x}))" by auto from d_eq this have "d : ?D2" by (auto intro: image_eqI[where x = "(d x, d(x:=x, inv d x:=d x))"]) from this show ?thesis by auto qed qed qed have no_intersect: "?D1 \ ?D2 = {}" proof - have that: "\d. d \ ?D1 \ d (d x) = x" using Diff_iff Diff_insert2 derangements_fixpoint insertI1 swap_apply(2) by fastforce have "\d. d \ ?D2 \ d (d x) \ x" proof - fix d assume a: "d \ ?D2" from a obtain y d' where d: "d = d'(x := y, inv d' y := x)" "d' \ derangements (S - {x})" "y \ S - {x}" by auto from d(2) have inv: "inv d' \ derangements (S - {x})" by (rule derangements_inv) from d have inv_x: "inv d' y \ x" by (auto dest: derangements_inverse_in_image) from inv have inv_y: "inv d' y \ y" using d(3) derangements_no_fixpoint by blast from d inv_x have 1: "d x = y" by auto from d inv_y have 2: "d y = d' y" by auto from d(2, 3) have 3: "d' y \ S - {x}" by (auto dest: derangements_in_image) from 1 2 3 show "d (d x) \ x" by auto qed from this that show ?thesis by blast qed have inj: "inj_on (\(y, f). Fun.swap x y f) {(y, f). y \ S & y \ x & f : derangements (S - {x, y})}" unfolding inj_on_def by (clarify; metis DiffD2 derangements_fixpoint insertI1 insert_commute swap_apply(1) swap_nilpotent) have eq: "{(y, f). y \ S & y \ x & f : derangements (S - {x, y})} = {(y, f). y \ S - {x} & f : derangements (S - {x, y})}" by simp have eq': "{(y, f). y \ S & y \ x & f : derangements (S - {x, y})} = Sigma (S - {x}) (%y. derangements (S - {x, y}))" unfolding Sigma_def by auto have card: "\ y. y \ S - {x} \ card (derangements (S - {x, y})) = count_derangements n" proof - fix y assume "y \ S - {x}" from 3(3, 4) \x \ S\ this have "card (S - {x, y}) = n" by (metis Diff_insert2 card_Diff_singleton diff_Suc_1 finite_Diff) from 3(3) 3(2)[OF _ this] show "card (derangements (S - {x, y})) = count_derangements n" by auto qed from 3(3, 4) \x : S\ have card2: "card (S - {x}) = Suc n" by (simp add: card.insert_remove insert_absorb) from inj have "card ?D1 = card {(y, f). y \ S - {x} \ f \ derangements (S - {x, y})}" by (simp add: card_image) also from 3(3) have "... = (\y\S - {x}. card (derangements (S - {x, y})))" by (subst card_product_dependent) (simp add: finite_derangements)+ finally have card_1: "card ?D1 =(Suc n) * count_derangements n" using card card2 by auto have inj_2: "inj_on (\(y, f). f(x := y, inv f y := x)) ((S - {x}) \ derangements (S - {x}))" proof - { fix d d' y y' assume 1: "d \ derangements (S - {x})" "d' \ derangements (S - {x})" assume 2: "y \ S" "y \ x" "y' \ S" "y' \ x" assume eq: "d(x := y, inv d y := x) = d'(x := y', inv d' y' := x)" from 1 2 eq \x \ S\ have y: "y = y'" by (metis Diff_insert_absorb derangements_in_image derangements_inv fun_upd_same fun_upd_twist insert_iff mk_disjoint_insert) have "d = d'" proof fix z from 1 have d_x: "d x = d' x" by (auto dest!: derangements_fixpoint) from 1 have bij: "bij d" "bij d'" by (metis derangementsE permutes_imp_bij')+ from this have d_d: "d (inv d y) = y" "d' (inv d' y') = y'" by (simp add: bij_is_surj surj_f_inv_f)+ from 1 2 bij have neq: "d (inv d' y') \ x \ d' (inv d y) \ x" by (metis Diff_iff bij_inv_eq_iff derangements_fixpoint singletonI) from eq have "(d(x := y, inv d y := x)) z = (d'(x := y', inv d' y' := x)) z" by auto from y d_x d_d neq this show "d z = d' z" by (auto split: if_split_asm) qed from y this have "y = y' & d = d'" by auto } from this show ?thesis unfolding inj_on_def by auto qed from 3(3) 3(1)[OF _ card2] have card3: "card (derangements (S - {x})) = count_derangements (Suc n)" by auto from 3(3) inj_2 have card_2: "card ?D2 = (Suc n) * count_derangements (Suc n)" by (simp only: card_image) (auto simp add: card_cartesian_product card3 card2) from \finite S\have finite1: "finite ?D1" unfolding eq' by (auto intro: finite_derangements) from \finite S\ have finite2: "finite ?D2" by (auto intro: finite_cartesian_product finite_derangements) show ?case by (simp add: split card_Un_disjoint finite1 finite2 no_intersect card_1 card_2 field_simps) qed subsubsection \Closed-Form Characterization\ lemma count_derangements: "real (count_derangements n) = fact n * (\k \ {0..n}. (-1) ^ k / fact k)" proof (induct n rule: count_derangements.induct) case (3 n) let ?f = "\n. fact n * (\k = 0..n. (- 1) ^ k / fact k)" have "real (count_derangements (Suc (Suc n))) = (n + 1) * (count_derangements (n + 1) + count_derangements n)" unfolding count_derangements.simps by simp also have "... = real (n + 1) * (real (count_derangements (n + 1)) + real (count_derangements n))" by (simp only: of_nat_mult of_nat_add) also have "... = (n + 1) * (?f (n + 1) + ?f n)" unfolding 3(2) 3(1)[unfolded Suc_eq_plus1] .. also have "(n + 1) * (?f (n + 1) + ?f n) = ?f (n + 2)" proof - define f where "f n = ((fact n) :: real) * (\k = 0..n. (- 1) ^ k / fact k)" for n have f_eq: "\n. f (n + 1) = (n + 1) * f n + (- 1) ^ (n + 1)" proof - fix n have "?f (n + 1) = (n + 1) * fact n * (\k = 0..n. (- 1) ^ k / fact k) + fact (n + 1) * ((- 1) ^ (n + 1) / fact (n + 1))" by (simp add: field_simps) also have "... = (n + 1) * ?f n + (- 1) ^ (n + 1)" by (simp del: One_nat_def) finally show "?thesis n" unfolding f_def by simp qed from this have f_eq': "\ n. (n + 1) * f n = f (n + 1) + (- 1) ^ n" by auto from f_eq'[of n] have "(n + 1) * (f (n + 1) + f n) = ((n + 1) * f (n + 1)) + f (n + 1) + (- 1) ^ n" by (simp only: distrib_left of_nat_add of_nat_1) also have "... = (n + 2) * f (n + 1) + (- 1) ^ (n + 2)" by (simp del: One_nat_def add_2_eq_Suc' add: algebra_simps) simp also from f_eq[of "n + 1"] have "... = f (n + 2)" by simp finally show ?thesis unfolding f_def by simp qed finally show ?case by simp qed (auto) subsubsection \Approximation of Cardinality\ lemma two_power_fact_le_fact: assumes "n \ 1" shows "2^k * fact n \ (fact (n + k) :: 'a :: {semiring_char_0,linordered_semidom})" proof (induction k) case (Suc k) have "2 ^ Suc k * fact n = 2 * (2 ^ k * fact n)" by (simp add: algebra_simps) also note Suc.IH also from assms have "of_nat 1 + of_nat 1 \ of_nat n + (of_nat (Suc k) :: 'a)" by (intro add_mono) (unfold of_nat_le_iff, simp_all) hence "2 * (fact (n + k) :: 'a) \ of_nat (n + Suc k) * fact (n + k)" by (intro mult_right_mono) (simp_all add: add_ac) also have "\ = fact (n + Suc k)" by simp finally show ?case by - (simp add: mult_left_mono) qed simp_all lemma exp1_approx: assumes "n > 0" shows "exp (1::real) - (\k {0..2 / fact n}" proof (unfold atLeastAtMost_iff, safe) have "(\k. 1^k / fact k) sums exp (1::real)" using exp_converges[of "1::real"] by (simp add: field_simps) from sums_split_initial_segment[OF this, of n] have sums: "(\k. 1 / fact (n+k)) sums (exp 1 - (\kk 0" by (intro sums_le[OF _ sums_zero sums]) auto show "(exp 1 - (\k 2 / fact n" - proof (rule sums_le[OF allI]) + proof (rule sums_le) from assms have "(\k. (1 / fact n) * (1 / 2)^k :: real) sums ((1 / fact n) * (1 / (1 - 1 / 2)))" by (intro sums_mult geometric_sums) simp_all also have "\ = 2 / fact n" by simp finally show "(\k. 1 / fact n * (1 / 2) ^ k) sums (2 / fact n :: real)" . next fix k show "(1 / fact (n+k)) \ (1 / fact n) * (1 / 2 :: real)^k" for k using two_power_fact_le_fact[of n k] assms by (auto simp: field_simps) qed fact+ qed lemma exp1_bounds: "exp 1 \ {8 / 3..11 / 4 :: real}" using exp1_approx[of 4] by (simp add: eval_nat_numeral) lemma count_derangements_approximation: assumes "n \ 0" shows "abs(real (count_derangements n) - fact n / exp 1) < 1 / 2" proof (cases "n \ 5") case False from assms this have n: "n = 1 \ n = 2 \ n = 3 \ n = 4" by auto from exp1_bounds have 1: "abs(real (count_derangements 1) - fact 1 / exp 1) < 1 / 2" by simp from exp1_bounds have 2: "abs(real (count_derangements 2) - fact 2 / exp 1) < 1 / 2" by (auto simp add: eval_nat_numeral abs_real_def) from exp1_bounds have 3: "abs(real (count_derangements 3) - fact 3 / exp 1) < 1 / 2" by (auto simp add: eval_nat_numeral abs_if field_simps) from exp1_bounds have 4: "abs(real (count_derangements 4) - fact 4 / exp 1) < 1 / 2" by (auto simp: abs_if field_simps eval_nat_numeral) from 1 2 3 4 n show ?thesis by auto next case True from Maclaurin_exp_le[of "- 1" "n + 1"] obtain t where t: "abs (t :: real) \ abs (- 1)" and exp: "exp (- 1) = (\mk = 0..n. (- 1) ^ k / fact k) = exp (- 1) - exp t / fact (n + 1) * (- 1) ^ (n + 1)" by (simp add: atLeast0AtMost ivl_disj_un(2)[symmetric]) have fact_plus1: "fact (n + 1) = (n + 1) * fact n" by simp have eq: "\real (count_derangements n) - fact n / exp 1\ = exp t / (n + 1)" by (simp del: One_nat_def add: count_derangements sum_eq_exp exp_minus inverse_eq_divide algebra_simps abs_mult) (simp add: fact_plus1) from t have "exp t \ exp 1" by simp also from exp1_bounds have "\ < 3" by simp finally show ?thesis using \n \ 5\ by (simp add: eq) qed theorem derangements_formula: assumes "n \ 0" "finite S" "card S = n" shows "int (card (derangements S)) = round (fact n / exp 1 :: real)" using count_derangements_approximation[of n] assms by (intro round_unique' [symmetric]) (auto simp: card_derangements abs_minus_commute) theorem derangements_formula': assumes "n \ 0" "finite S" "card S = n" shows "card (derangements S) = nat (round (fact n / exp 1 :: real))" using derangements_formula[OF assms] by simp end diff --git a/thys/Dirichlet_Series/Arithmetic_Summatory_Asymptotics.thy b/thys/Dirichlet_Series/Arithmetic_Summatory_Asymptotics.thy --- a/thys/Dirichlet_Series/Arithmetic_Summatory_Asymptotics.thy +++ b/thys/Dirichlet_Series/Arithmetic_Summatory_Asymptotics.thy @@ -1,822 +1,822 @@ (* File: Arithmetic_Summatory_Asymptotics.thy Author: Manuel Eberl, TU München *) section \Asymptotics of summatory arithmetic functions\ theory Arithmetic_Summatory_Asymptotics imports Euler_MacLaurin.Euler_MacLaurin_Landau Arithmetic_Summatory Dirichlet_Series_Analysis Landau_Symbols.Landau_More begin subsection \Auxiliary bounds\ lemma sum_inverse_squares_tail_bound: assumes "d > 0" shows "summable (\n. 1 / (real (Suc n) + d) ^ 2)" "(\n. 1 / (real (Suc n) + d) ^ 2) \ 1 / d" proof - show *: "summable (\n. 1 / (real (Suc n) + d) ^ 2)" proof (rule summable_comparison_test, intro allI exI impI) fix n :: nat from assms show "norm (1 / (real (Suc n) + d) ^ 2) \ 1 / real (Suc n) ^ 2" unfolding norm_divide norm_one norm_power by (intro divide_left_mono power_mono) simp_all qed (insert inverse_squares_sums, simp add: sums_iff) show "(\n. 1 / (real (Suc n) + d) ^ 2) \ 1 / d" - proof (rule sums_le[OF allI]) + proof (rule sums_le) fix n have "1 / (real (Suc n) + d) ^ 2 \ 1 / ((real n + d) * (real (Suc n) + d))" unfolding power2_eq_square using assms by (intro divide_left_mono mult_mono mult_pos_pos add_nonneg_pos) simp_all also have "\ = 1 / (real n + d) - 1 / (real (Suc n) + d)" using assms by (simp add: divide_simps) finally show "1 / (real (Suc n) + d)\<^sup>2 \ 1 / (real n + d) - 1 / (real (Suc n) + d)" . next show "(\n. 1 / (real (Suc n) + d)\<^sup>2) sums (\n. 1 / (real (Suc n) + d)\<^sup>2)" using * by (simp add: sums_iff) next have "(\n. 1 / (real n + d) - 1 / (real (Suc n) + d)) sums (1 / (real 0 + d) - 0)" by (intro telescope_sums' real_tendsto_divide_at_top[OF tendsto_const], subst add.commute, rule filterlim_tendsto_add_at_top[OF tendsto_const filterlim_real_sequentially]) thus "(\n. 1 / (real n + d) - 1 / (real (Suc n) + d)) sums (1 / d)" by simp qed qed lemma moebius_sum_tail_bound: assumes "d > 0" shows "abs (\n. moebius_mu (Suc n + d) / real (Suc n + d)^2) \ 1 / d" (is "abs ?S \ _") proof - have *: "summable (\n. 1 / (real (Suc n + d))\<^sup>2)" by (insert sum_inverse_squares_tail_bound(1)[of "real d"] assms, simp_all add: add_ac) have **: "summable (\n. abs (moebius_mu (Suc n + d) / real (Suc n + d)^2))" proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (\moebius_mu (Suc n + d) / (real (Suc n + d))^2\) \ 1 / (real (Suc n + d))^2" unfolding real_norm_def abs_abs abs_divide power_abs abs_of_nat by (intro divide_right_mono abs_moebius_mu_le) simp_all qed (insert *) from ** have "abs ?S \ (\n. abs (moebius_mu (Suc n + d) / real (Suc n + d)^2))" by (rule summable_rabs) also have "\ \ (\n. 1 / (real (Suc n) + d) ^ 2)" proof (intro suminf_le allI) fix n :: nat show "abs (moebius_mu (Suc n + d) / (real (Suc n + d))^2) \ 1 / (real (Suc n) + real d)^2" unfolding abs_divide abs_of_nat power_abs of_nat_add [symmetric] by (intro divide_right_mono abs_moebius_mu_le) simp_all qed (insert * **, simp_all add: add_ac) also from assms have "\ \ 1 / d" by (intro sum_inverse_squares_tail_bound) simp_all finally show ?thesis . qed lemma sum_upto_inverse_bound: "sum_upto (\i. 1 / real i) x \ 0" "eventually (\x. sum_upto (\i. 1 / real i) x \ ln x + 13 / 22) at_top" proof - show "sum_upto (\i. 1 / real i) x \ 0" by (simp add: sum_upto_def sum_nonneg) from order_tendstoD(2)[OF euler_mascheroni_LIMSEQ euler_mascheroni_less_13_over_22] obtain N where N: "\n. n \ N \ harm n - ln (real n) < 13 / 22" unfolding eventually_at_top_linorder by blast show "eventually (\x. sum_upto (\i. 1 / real i) x \ ln x + 13 / 22) at_top" using eventually_ge_at_top[of "max (real N) 1"] proof eventually_elim case (elim x) have "sum_upto (\i. 1 / real i) x = (\i\{0<..nat \x\}. 1 / real i)" by (simp add: sum_upto_altdef) also have "\ = harm (nat \x\)" unfolding harm_def by (intro sum.cong refl) (auto simp: field_simps) also have "\ \ ln (real (nat \x\)) + 13 / 22" using N[of "nat \x\"] elim by (auto simp: le_nat_iff le_floor_iff) also have "ln (real (nat \x\)) \ ln x" using elim by (subst ln_le_cancel_iff) auto finally show ?case by - simp qed qed lemma sum_upto_inverse_bigo: "sum_upto (\i. 1 / real i) \ O(\x. ln x)" proof - have "eventually (\x. norm (sum_upto (\i. 1 / real i) x) \ 1 * norm (ln x + 13/22)) at_top" using eventually_ge_at_top[of "1::real"] sum_upto_inverse_bound(2) by eventually_elim (insert sum_upto_inverse_bound(1), simp_all) hence "sum_upto (\i. 1 / real i) \ O(\x. ln x + 13/22)" by (rule bigoI) also have "(\x::real. ln x + 13/22) \ O(\x. ln x)" by simp finally show ?thesis . qed lemma defines "G \ (\x::real. (\n. moebius_mu (n + Suc (nat \x\)) / (n + Suc (nat \x\))^2) :: real)" shows moebius_sum_tail_bound': "\t. t \ 2 \ \G t\ \ 1 / (t - 1)" and moebius_sum_tail_bigo: "G \ O(\t. 1 / t)" proof - show "\G t\ \ 1 / (t - 1)" if t: "t \ 2" for t proof - from t have "\G t\ \ 1 / real (nat \t\)" unfolding G_def using moebius_sum_tail_bound[of "nat \t\"] by simp also have "t \ 1 + real_of_int \t\" by linarith hence "1 / real (nat \t\) \ 1 / (t - 1)" using t by (simp add: field_simps) finally show ?thesis . qed hence "G \ O(\t. 1 / (t - 1))" by (intro bigoI[of _ 1] eventually_mono[OF eventually_ge_at_top[of "2::real"]]) auto also have "(\t::real. 1 / (t - 1)) \ \(\t. 1 / t)" by simp finally show "G \ O(\t. 1 / t)" . qed subsection \Summatory totient function\ theorem summatory_totient_asymptotics: "(\x. sum_upto (\n. real (totient n)) x - 3 / pi\<^sup>2 * x\<^sup>2) \ O(\x. x * ln x)" proof - define H where "H = (\x. of_int (floor x) * (of_int (floor x) + 1) / 2 - x ^ 2 / 2 :: real)" define H' where "H' = (\x. sum_upto (\n. moebius_mu n * H (x / real n)) x)" have H: "sum_upto real x = x^2/2 + H x" if "x \ 0" for x using that by (simp add: sum_upto_real H_def) define G where "G = (\x::real. (\n. moebius_mu (n + Suc (nat \x\)) / (n + Suc (nat \x\))^2))" have H_bound: "\H t\ \ t / 2" if "t \ 0" for t proof - have "H t - t / 2 = (-(t - of_int (floor t))) * (floor t + t + 1) / 2" by (simp add: H_def field_simps power2_eq_square) also have "\ \ 0" using that by (intro mult_nonpos_nonneg divide_nonpos_nonneg) simp_all finally have "H t \ t / 2" by simp have "-H t - t / 2 = (t - of_int (floor t) - 1) * (of_int (floor t) + t) / 2" by (simp add: H_def field_simps power2_eq_square) also have "\ \ 0" using that by (intro divide_nonpos_nonneg mult_nonpos_nonneg) ((simp; fail) | linarith)+ finally have "-H t \ t / 2" by simp with \H t \ t / 2\ show "\H t\ \ t / 2" by simp qed have H'_bound: "\H' t\ \ t / 2 * sum_upto (\i. 1 / real i) t" if "t \ 0" for t proof - have "\H' t\ \ (\i | 0 < i \ real i \ t. \moebius_mu i * H (t / real i)\)" unfolding H'_def sum_upto_def by (rule sum_abs) also have "\ \ (\i | 0 < i \ real i \ t. 1 * ((t / real i) / 2))" unfolding abs_mult using that by (intro sum_mono mult_mono abs_moebius_mu_le H_bound) simp_all also have "\ = t / 2 * sum_upto (\i. 1 / real i) t" by (simp add: sum_upto_def sum_distrib_left sum_distrib_right mult_ac) finally show ?thesis . qed hence "H' \ O(\t. t * sum_upto (\i. 1 / real i) t)" using sum_upto_inverse_bound(1) by (intro bigoI[of _ "1/2"] eventually_mono[OF eventually_ge_at_top[of "0::real"]]) (auto elim!: eventually_mono simp: abs_mult) also have "(\t. t * sum_upto (\i. 1 / real i) t) \ O(\t. t * ln t)" by (intro landau_o.big.mult sum_upto_inverse_bigo) simp_all finally have H'_bigo: "H' \ O(\x. x * ln x)" . { fix x :: real assume x: "x \ 0" have "sum_upto (\n. real (totient n)) x = sum_upto (\n. of_int (int (totient n))) x" by simp also have "\ = sum_upto (\n. moebius_mu n * sum_upto real (x / real n)) x" by (subst totient_conv_moebius_mu) (simp add: sum_upto_dirichlet_prod of_int_dirichlet_prod) also have "\ = sum_upto (\n. moebius_mu n * ((x / real n) ^ 2 / 2 + H (x / real n))) x" using x by (intro sum_upto_cong) (simp_all add: H) also have "\ = x^2 / 2 * sum_upto (\n. moebius_mu n / real n ^ 2) x + H' x" by (simp add: sum_upto_def H'_def sum.distrib ring_distribs sum_distrib_left sum_distrib_right power_divide mult_ac) also have "sum_upto (\n. moebius_mu n / real n ^ 2) x = (\n\{..x\)}. moebius_mu n / real n ^ 2)" unfolding sum_upto_altdef by (intro sum.mono_neutral_cong_left refl) auto also have "\ = 6 / pi ^ 2 - G x" using sums_split_initial_segment[OF moebius_over_square_sums, of "Suc (nat \x\)"] by (auto simp: sums_iff algebra_simps G_def) finally have "sum_upto (\n. real (totient n)) x = 3 / pi\<^sup>2 * x\<^sup>2 - x\<^sup>2 / 2 * G x + H' x" by (simp add: algebra_simps) } hence "(\x. sum_upto (\n. real (totient n)) x - 3 / pi^2 * x^2) \ \(\x. (-(x^2) / 2) * G x + H' x)" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of "0::real"]]) (auto elim!: eventually_mono) also have "(\x. (-(x^2) / 2) * G x + H' x) \ O(\x. x * ln x)" proof (intro sum_in_bigo H'_bigo) have "(\x. (- (x^2) / 2) * G x) \ O(\x. x^2 * (1 / x))" using moebius_sum_tail_bigo [folded G_def] by (intro landau_o.big.mult) simp_all also have "(\x::real. x^2 * (1 / x)) \ O(\x. x * ln x)" by simp finally show "(\x. (- (x^2) / 2) * G x) \ O(\x. x * ln x)" . qed finally show ?thesis . qed theorem summatory_totient_asymptotics': "(\x. sum_upto (\n. real (totient n)) x) =o (\x. 3 / pi\<^sup>2 * x\<^sup>2) +o O(\x. x * ln x)" using summatory_totient_asymptotics by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) theorem summatory_totient_asymptotics'': "sum_upto (\n. real (totient n)) \[at_top] (\x. 3 / pi\<^sup>2 * x\<^sup>2)" proof - have "(\x. sum_upto (\n. real (totient n)) x - 3 / pi\<^sup>2 * x\<^sup>2) \ O(\x. x * ln x)" by (rule summatory_totient_asymptotics) also have "(\x. x * ln x) \ o(\x. 3 / pi ^ 2 * x ^ 2)" by simp finally show ?thesis by (simp add: asymp_equiv_altdef) qed subsection \Asymptotic distribution of squarefree numbers\ lemma le_sqrt_iff: "x \ 0 \ x \ sqrt y \ x^2 \ y" using real_sqrt_le_iff[of "x^2" y] by (simp del: real_sqrt_le_iff) theorem squarefree_asymptotics: "(\x. card {n. real n \ x \ squarefree n} - 6 / pi\<^sup>2 * x) \ O(sqrt)" proof - define f :: "nat \ real" where "f = (\n. if n = 0 then 0 else 1)" define g :: "nat \ real" where "g = dirichlet_prod (ind squarefree) moebius_mu" interpret g: multiplicative_function g unfolding g_def by (intro multiplicative_dirichlet_prod squarefree.multiplicative_function_axioms moebius_mu.multiplicative_function_axioms) interpret g: multiplicative_function' g "\p k. if k = 2 then -1 else 0" "\_. 0" proof interpret g': multiplicative_dirichlet_prod' "ind squarefree" moebius_mu "\p k. if 1 < k then 0 else 1" "\p k. if k = 1 then - 1 else 0" "\_. 1" "\_. - 1" by (intro multiplicative_dirichlet_prod'.intro squarefree.multiplicative_function'_axioms moebius_mu.multiplicative_function'_axioms) fix p k :: nat assume "prime p" "k > 0" hence "g (p ^ k) = (\i\{0<.. = (\i\{0<..k > 0\ have "\ = (if k = 2 then -1 else 0)" by simp finally show "g (p ^ k) = \" . qed simp_all have mult_g_square: "multiplicative_function (\n. g (n ^ 2))" by standard (simp_all add: power_mult_distrib g.mult_coprime) have g_square: "g (m ^ 2) = moebius_mu m" for m using mult_g_square moebius_mu.multiplicative_function_axioms proof (rule multiplicative_function_eqI) fix p k :: nat assume *: "prime p" "k > 0" have "g ((p ^ k) ^ 2) = g (p ^ (2 * k))" by (simp add: power_mult [symmetric] mult_ac) also from * have "\ = (if k = 1 then -1 else 0)" by (simp add: g.prime_power) also from * have "\ = moebius_mu (p ^ k)" by (simp add: moebius_mu.prime_power) finally show "g ((p ^ k) ^ 2) = moebius_mu (p ^ k)" . qed have g_nonsquare: "g m = 0" if "\is_square m" for m proof (cases "m = 0") case False from that False obtain p where p: "prime p" "odd (multiplicity p m)" using is_nth_power_conv_multiplicity_nat[of 2 m] by auto from p have "multiplicity p m \ 2" by auto moreover from p have "p \ prime_factors m" by (auto simp: prime_factors_multiplicity intro!: Nat.gr0I) ultimately have "(\p\prime_factors m. if multiplicity p m = 2 then - 1 else 0 :: real) = 0" (is "?P = _") by auto also have "?P = g m" using False by (subst g.prod_prime_factors') auto finally show ?thesis . qed auto have abs_g_le: "abs (g m) \ 1" for m by (cases "is_square m") (auto simp: g_square g_nonsquare abs_moebius_mu_le elim!: is_nth_powerE) have fds_g: "fds g = fds_ind squarefree * fds moebius_mu" by (rule fds_eqI) (simp add: g_def fds_nth_mult) have "fds g * fds_zeta = fds_ind squarefree * (fds_zeta * fds moebius_mu)" by (simp add: fds_g mult_ac) also have "fds_zeta * fds moebius_mu = (1 :: real fds)" by (rule fds_zeta_times_moebius_mu) finally have *: "fds_ind squarefree = fds g * fds_zeta" by simp have ind_squarefree: "ind squarefree = dirichlet_prod g f" proof fix n :: nat from * show "ind squarefree n = dirichlet_prod g f n" by (cases "n = 0") (simp_all add: fds_eq_iff fds_nth_mult f_def) qed define H :: "real \ real" where "H = (\x. sum_upto (\m. g (m^2) * (real_of_int \x / real (m\<^sup>2)\ - x / real (m^2))) (sqrt x))" define J where "J = (\x::real. (\n. moebius_mu (n + Suc (nat \x\)) / (n + Suc (nat \x\))^2))" have "eventually (\x. norm (H x) \ 1 * norm (sqrt x)) at_top" using eventually_ge_at_top[of "0::real"] proof eventually_elim case (elim x) have "abs (H x) \ sum_upto (\m. abs (g (m^2) * (real_of_int \x / real (m\<^sup>2)\ - x / real (m^2)))) (sqrt x)" (is "_ \ ?S") unfolding H_def sum_upto_def by (rule sum_abs) also have "x / (real m)\<^sup>2 - real_of_int \x / (real m)\<^sup>2\ \ 1" for m by linarith hence "?S \ sum_upto (\m. 1 * 1) (sqrt x)" unfolding abs_mult sum_upto_def by (intro sum_mono mult_mono abs_g_le) simp_all also have "\ = of_int \sqrt x\" using elim by (simp add: sum_upto_altdef) also have "\ \ sqrt x" by linarith finally show ?case using elim by simp qed hence H_bigo: "H \ O(\x. sqrt x)" by (rule bigoI) let ?A = "\x. card {n. real n \ x \ squarefree n}" have "eventually (\x. ?A x - 6 / pi\<^sup>2 * x = (-x) * J (sqrt x) + H x) at_top" using eventually_ge_at_top[of "0::real"] proof eventually_elim fix x :: real assume x: "x \ 0" have "{n. real n \ x \ squarefree n} = {n. n > 0 \ real n \ x \ squarefree n}" by (auto intro!: Nat.gr0I) also have "card \ = sum_upto (ind squarefree :: nat \ real) x" by (rule sum_upto_ind [symmetric]) also have "\ = sum_upto (\d. g d * sum_upto f (x / real d)) x" (is "_ = ?S") unfolding ind_squarefree by (rule sum_upto_dirichlet_prod) also have "sum f {0<..nat \x / real i\} = of_int \x / real i\" if "i > 0" for i using x by (simp add: f_def) hence "?S = sum_upto (\d. g d * of_int \x / real d\) x" unfolding sum_upto_altdef by (intro sum.cong refl) simp_all also have "\ = sum_upto (\m. g (m ^ 2) * of_int \x / real (m ^ 2)\) (sqrt x)" unfolding sum_upto_def proof (intro sum.reindex_bij_betw_not_neutral [symmetric]) show "bij_betw power2 ({i. 0 < i \ real i \ sqrt x} - {}) ({i. 0 < i \ real i \ x} - {i\{0<..nat \x\}. \is_square i})" by (auto simp: bij_betw_def inj_on_def power_eq_iff_eq_base le_sqrt_iff is_nth_power_def le_nat_iff le_floor_iff) qed (auto simp: g_nonsquare) also have "\ = x * sum_upto (\m. g (m ^ 2) / real m ^ 2) (sqrt x) + H x" by (simp add: H_def sum_upto_def sum.distrib ring_distribs sum_subtractf sum_distrib_left sum_distrib_right mult_ac) also have "sum_upto (\m. g (m ^ 2) / real m ^ 2) (sqrt x) = sum_upto (\m. moebius_mu m / real m ^ 2) (sqrt x)" unfolding sum_upto_altdef by (intro sum.cong refl) (simp_all add: g_square) also have "sum_upto (\m. moebius_mu m / (real m)\<^sup>2) (sqrt x) = (\msqrt x\). moebius_mu m / (real m) ^ 2)" unfolding sum_upto_altdef by (intro sum.mono_neutral_cong_left) auto also have "\ = (6 / pi^2 - J (sqrt x))" using sums_split_initial_segment[OF moebius_over_square_sums, of "Suc (nat \sqrt x\)"] by (auto simp: sums_iff algebra_simps J_def sum_upto_altdef) finally show "?A x - 6 / pi\<^sup>2 * x = (-x) * J (sqrt x) + H x" by (simp add: algebra_simps) qed hence "(\x. ?A x - 6 / pi\<^sup>2 * x) \ \(\x. (-x) * J (sqrt x) + H x)" by (rule bigthetaI_cong) also have "(\x. (-x) * J (sqrt x) + H x) \ O(\x. sqrt x)" proof (intro sum_in_bigo H_bigo) have "(\x. J (sqrt x)) \ O(\x. 1 / sqrt x)" unfolding J_def using moebius_sum_tail_bigo sqrt_at_top by (rule landau_o.big.compose) hence "(\x. (-x) * J (sqrt x)) \ O(\x. x * (1 / sqrt x))" by (intro landau_o.big.mult) simp_all also have "(\x::real. x * (1 / sqrt x)) \ \(\x. sqrt x)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of "0::real"]]) (auto simp: field_simps) finally show "(\x. (-x) * J (sqrt x)) \ O(\x. sqrt x)" . qed finally show ?thesis . qed theorem squarefree_asymptotics': "(\x. card {n. real n \ x \ squarefree n}) =o (\x. 6 / pi\<^sup>2 * x) +o O(\x. sqrt x)" using squarefree_asymptotics by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) theorem squarefree_asymptotics'': "(\x. card {n. real n \ x \ squarefree n}) \[at_top] (\x. 6 / pi\<^sup>2 * x)" proof - have "(\x. card {n. real n \ x \ squarefree n} - 6 / pi\<^sup>2 * x) \ O(\x. sqrt x)" by (rule squarefree_asymptotics) also have "(sqrt :: real \ real) \ \(\x. x powr (1/2))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of "0::real"]]) (auto simp: powr_half_sqrt) also have "(\x::real. x powr (1/2)) \ o(\x. 6 / pi ^ 2 * x)" by simp finally show ?thesis by (simp add: asymp_equiv_altdef) qed subsection \The hyperbola method\ lemma hyperbola_method_bigo: fixes f g :: "nat \ 'a :: real_normed_field" assumes "(\x. sum_upto (\n. f n * sum_upto g (x / real n)) (sqrt x) - R x) \ O(b)" assumes "(\x. sum_upto (\n. sum_upto f (x / real n) * g n) (sqrt x) - S x) \ O(b)" assumes "(\x. sum_upto f (sqrt x) * sum_upto g (sqrt x) - T x) \ O(b)" shows "(\x. sum_upto (dirichlet_prod f g) x - (R x + S x - T x)) \ O(b)" proof - let ?A = "\x. (sum_upto (\n. f n * sum_upto g (x / real n)) (sqrt x) - R x) + (sum_upto (\n. sum_upto f (x / real n) * g n) (sqrt x) - S x) + (-(sum_upto f (sqrt x) * sum_upto g (sqrt x) - T x))" have "(\x. sum_upto (dirichlet_prod f g) x - (R x + S x - T x)) \ \(?A)" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of "0::real"]]) (auto simp: hyperbola_method_sqrt) also from assms have "?A \ O(b)" by (intro sum_in_bigo(1)) (simp_all only: landau_o.big.uminus_in_iff) finally show ?thesis . qed lemma frac_le_1: "frac x \ 1" unfolding frac_def by linarith lemma ln_minus_ln_floor_bound: assumes "x \ 2" shows "ln x - ln (floor x) \ {0..<1 / (x - 1)}" proof - from assms have "ln (floor x) \ ln (x - 1)" by (subst ln_le_cancel_iff) simp_all hence "ln x - ln (floor x) \ ln ((x - 1) + 1) - ln (x - 1)" by simp also from assms have "\ < 1 / (x - 1)" by (intro ln_diff_le_inverse) simp_all finally have "ln x - ln (floor x) < 1 / (x - 1)" by simp moreover from assms have "ln x \ ln (of_int \x\)" by (subst ln_le_cancel_iff) simp_all ultimately show ?thesis by simp qed lemma ln_minus_ln_floor_bigo: "(\x::real. ln x - ln (floor x)) \ O(\x. 1 / x)" proof - have "eventually (\x. norm (ln x - ln (floor x)) \ 1 * norm (1 / (x - 1))) at_top" using eventually_ge_at_top[of "2::real"] proof eventually_elim case (elim x) with ln_minus_ln_floor_bound[OF this] show ?case by auto qed hence "(\x::real. ln x - ln (floor x)) \ O(\x. 1 / (x - 1))" by (rule bigoI) also have "(\x::real. 1 / (x - 1)) \ O(\x. 1 / x)" by simp finally show ?thesis . qed lemma divisor_count_asymptotics_aux: "(\x. sum_upto (\n. sum_upto (\_. 1) (x / real n)) (sqrt x) - (x * ln x / 2 + euler_mascheroni * x)) \ O(sqrt)" proof - define R where "R = (\x. \i\{0<..nat \sqrt x\}. frac (x / real i))" define S where "S = (\x. ln (real (nat \sqrt x\)) - ln x / 2)" have R_bound: "R x \ {0..sqrt x}" if x: "x \ 0" for x proof - have "R x \ (\i\{0<..nat \sqrt x\}. 1)" unfolding R_def by (intro sum_mono frac_le_1) also from x have "\ = of_int \sqrt x\" by simp also have "\ \ sqrt x" by simp finally have "R x \ sqrt x" . moreover have "R x \ 0" unfolding R_def by (intro sum_nonneg) simp_all ultimately show ?thesis by simp qed have R_bound': "norm (R x) \ 1 * norm (sqrt x)" if "x \ 0" for x using R_bound[OF that] that by simp have R_bigo: "R \ O(sqrt)" using eventually_ge_at_top[of "0::real"] by (intro bigoI[of _ 1], elim eventually_mono) (rule R_bound') have "eventually (\x. sum_upto (\n. sum_upto (\_. 1 :: real) (x / real n)) (sqrt x) = x * harm (nat \sqrt x\) - R x) at_top" using eventually_ge_at_top[of "0 :: real"] proof eventually_elim case (elim x) have "sum_upto (\n. sum_upto (\_. 1 :: real) (x / real n)) (sqrt x) = (\i\{0<..nat \sqrt x\}. of_int \x / real i\)" using elim by (simp add: sum_upto_altdef) also have "\ = x * (\i\{0<..nat \sqrt x\}. 1 / real i) - R x" by (simp add: sum_subtractf frac_def R_def sum_distrib_left) also have "{0<..nat \sqrt x\} = {1..nat \sqrt x\}" by auto also have "(\i\\. 1 / real i) = harm (nat \sqrt x\)" by (simp add: harm_def divide_simps) finally show ?case . qed hence "(\x. sum_upto (\n. sum_upto (\_. 1 :: real) (x / real n)) (sqrt x) - (x * ln x / 2 + euler_mascheroni * x)) \ \(\x. x * (harm (nat \sqrt x\) - (ln (nat \sqrt x\) + euler_mascheroni)) - R x + x * S x)" (is "_ \ \(?A)") by (intro bigthetaI_cong) (elim eventually_mono, simp_all add: algebra_simps S_def) also have "?A \ O(sqrt)" proof (intro sum_in_bigo) have "(\x. - S x) \ \(\x. ln (sqrt x) - ln (of_int \sqrt x\))" by (intro bigthetaI_cong eventually_mono [OF eventually_ge_at_top[of "1::real"]]) (auto simp: S_def ln_sqrt) also have "(\x. ln (sqrt x) - ln (of_int \sqrt x\)) \ O(\x. 1 / sqrt x)" by (rule landau_o.big.compose[OF ln_minus_ln_floor_bigo sqrt_at_top]) finally have "(\x. x * S x) \ O(\x. x * (1 / sqrt x))" by (intro landau_o.big.mult) simp_all also have "(\x::real. x * (1 / sqrt x)) \ \(\x. sqrt x)" by (intro bigthetaI_cong eventually_mono [OF eventually_gt_at_top[of "0::real"]]) (auto simp: field_simps) finally show "(\x. x * S x) \ O(sqrt)" . next let ?f = "\x::real. harm (nat \sqrt x\) - (ln (real (nat \sqrt x\)) + euler_mascheroni)" have "?f \ O(\x. 1 / real (nat \sqrt x\))" proof (rule landau_o.big.compose[of _ _ _ "\x. nat \sqrt x\"]) show "filterlim (\x::real. nat \sqrt x\) at_top at_top" by (intro filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_floor_sequentially] sqrt_at_top) next show "(\a. harm a - (ln (real a) + euler_mascheroni)) \ O(\a. 1 / real a)" by (rule harm_expansion_bigo_simple2) qed also have "(\x. 1 / real (nat \sqrt x\)) \ O(\x. 1 / (sqrt x - 1))" proof (rule bigoI[of _ 1], use eventually_ge_at_top[of 2] in eventually_elim) case (elim x) have "sqrt x \ 1 + real_of_int \sqrt x\" by linarith with elim show ?case by (simp add: field_simps) qed also have "(\x::real. 1 / (sqrt x - 1)) \ O(\x. 1 / sqrt x)" by (rule landau_o.big.compose[OF _ sqrt_at_top]) simp_all finally have "(\x. x * ?f x) \ O(\x. x * (1 / sqrt x))" by (intro landau_o.big.mult landau_o.big_refl) also have "(\x::real. x * (1 / sqrt x)) \ \(\x. sqrt x)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of "0::real"]]) (auto elim!: eventually_mono simp: field_simps) finally show "(\x. x * ?f x) \ O(sqrt)" . qed fact+ finally show ?thesis . qed lemma sum_upto_sqrt_bound: assumes x: "x \ (0 :: real)" shows "norm ((sum_upto (\_. 1) (sqrt x))\<^sup>2 - x) \ 2 * norm (sqrt x)" proof - from x have "0 \ 2 * sqrt x * (1 - frac (sqrt x)) + frac (sqrt x) ^ 2" by (intro add_nonneg_nonneg mult_nonneg_nonneg) (simp_all add: frac_le_1) also from x have "\ = (sqrt x - frac (sqrt x)) ^ 2 - x + 2 * sqrt x" by (simp add: algebra_simps power2_eq_square) also have "sqrt x - frac (sqrt x) = of_int \sqrt x\" by (simp add: frac_def) finally have "(of_int \sqrt x\) ^ 2 - x \ -2 * sqrt x" by (simp add: algebra_simps) moreover from x have "of_int (\sqrt x\) ^ 2 \ sqrt x ^ 2" by (intro power_mono) simp_all with x have "of_int (\sqrt x\) ^ 2 - x \ 0" by simp ultimately have "sum_upto (\_. 1) (sqrt x) ^ 2 - x \ {-2 * sqrt x..0}" using x by (simp add: sum_upto_altdef) with x show ?thesis by simp qed lemma summatory_divisor_count_asymptotics: "(\x. sum_upto (\n. real (divisor_count n)) x - (x * ln x + (2 * euler_mascheroni - 1) * x)) \ O(sqrt)" proof - let ?f = "\x. x * ln x / 2 + euler_mascheroni * x" have "(\x. sum_upto (dirichlet_prod (\_. 1 :: real) (\_. 1)) x - (?f x + ?f x - x)) \ O(sqrt)" (is "?g \ _") proof (rule hyperbola_method_bigo) have "eventually (\x::real. norm (sum_upto (\_. 1) (sqrt x) ^ 2 - x) \ 2 * norm (sqrt x)) at_top" using eventually_ge_at_top[of "0::real"] by eventually_elim (rule sum_upto_sqrt_bound) thus "(\x::real. sum_upto (\_. 1) (sqrt x) * sum_upto (\_. 1) (sqrt x) - x) \ O(sqrt)" by (intro bigoI[of _ 2]) (simp_all add: power2_eq_square) next show "(\x. sum_upto (\n. 1 * sum_upto (\_. 1) (x / real n)) (sqrt x) - (x * ln x / 2 + euler_mascheroni * x)) \ O(sqrt)" using divisor_count_asymptotics_aux by simp next show "(\x. sum_upto (\n. sum_upto (\_. 1) (x / real n) * 1) (sqrt x) - (x * ln x / 2 + euler_mascheroni * x)) \ O(sqrt)" using divisor_count_asymptotics_aux by simp qed also have "divisor_count n = dirichlet_prod (\_. 1) (\_. 1) n" for n using fds_divisor_count by (cases "n = 0") (simp_all add: fds_eq_iff power2_eq_square fds_nth_mult) hence "?g = (\x. sum_upto (\n. real (divisor_count n)) x - (x * ln x + (2 * euler_mascheroni - 1) * x))" by (intro ext) (simp_all add: algebra_simps dirichlet_prod_def) finally show ?thesis . qed theorem summatory_divisor_count_asymptotics': "(\x. sum_upto (\n. real (divisor_count n)) x) =o (\x. x * ln x + (2 * euler_mascheroni - 1) * x) +o O(\x. sqrt x)" using summatory_divisor_count_asymptotics by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) theorem summatory_divisor_count_asymptotics'': "sum_upto (\n. real (divisor_count n)) \[at_top] (\x. x * ln x)" proof - have "(\x. sum_upto (\n. real (divisor_count n)) x - (x * ln x + (2 * euler_mascheroni - 1) * x)) \ O(sqrt)" by (rule summatory_divisor_count_asymptotics) also have "sqrt \ \(\x. x powr (1/2))" by (intro bigthetaI_cong eventually_mono [OF eventually_ge_at_top[of "0::real"]]) (auto elim!: eventually_mono simp: powr_half_sqrt) also have "(\x::real. x powr (1/2)) \ o(\x. x * ln x + (2 * euler_mascheroni - 1) * x)" by simp finally have "sum_upto (\n. real (divisor_count n)) \[at_top] (\x. x * ln x + (2 * euler_mascheroni - 1) * x)" by (simp add: asymp_equiv_altdef) also have "\ \[at_top] (\x. x * ln x)" by (subst asymp_equiv_add_right) simp_all finally show ?thesis . qed lemma summatory_divisor_eq: "sum_upto (\n. real (divisor_count n)) (real m) = card {(n,d). n \ {0<..m} \ d dvd n}" proof - have "sum_upto (\n. real (divisor_count n)) m = card (SIGMA n:{0<..m}. {d. d dvd n})" unfolding sum_upto_altdef divisor_count_def by (subst card_SigmaI) simp_all also have "(SIGMA n:{0<..m}. {d. d dvd n}) = {(n,d). n \ {0<..m} \ d dvd n}" by auto finally show ?thesis . qed context fixes M :: "nat \ real" defines "M \ \m. card {(n,d). n \ {0<..m} \ d dvd n} / card {0<..m}" begin lemma mean_divisor_count_asymptotics: "(\m. M m - (ln m + 2 * euler_mascheroni - 1)) \ O(\m. 1 / sqrt m)" proof - have "(\m. M m - (ln m + 2 * euler_mascheroni - 1)) \ \(\m. (sum_upto (\n. real (divisor_count n)) (real m) - (m * ln m + (2 * euler_mascheroni - 1) * m)) / m)" (is "_ \ \(?f)") unfolding M_def by (intro bigthetaI_cong eventually_mono [OF eventually_gt_at_top[of "0::nat"]]) (auto simp: summatory_divisor_eq field_simps) also have "?f \ O(\m. sqrt m / m)" by (intro landau_o.big.compose[OF _ filterlim_real_sequentially] landau_o.big.divide_right summatory_divisor_count_asymptotics eventually_at_top_not_equal) also have "(\m::nat. sqrt m / m) \ \(\m. 1 / sqrt m)" by (intro bigthetaI_cong eventually_mono [OF eventually_gt_at_top[of "0::nat"]]) (auto simp: field_simps) finally show ?thesis . qed theorem mean_divisor_count_asymptotics': "M =o (\x. ln x + 2 * euler_mascheroni - 1) +o O(\x. 1 / sqrt x)" using mean_divisor_count_asymptotics by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) theorem mean_divisor_count_asymptotics'': "M \[at_top] ln" proof - have "(\x. M x - (ln x + 2 * euler_mascheroni - 1)) \ O(\x. 1 / sqrt x)" by (rule mean_divisor_count_asymptotics) also have "(\x. 1 / sqrt (real x)) \ \(\x. x powr (-1/2))" using eventually_gt_at_top[of "0::nat"] by (intro bigthetaI_cong) (auto elim!: eventually_mono simp: powr_half_sqrt field_simps powr_minus) also have "(\x::nat. x powr (-1/2)) \ o(\x. ln x + 2 * euler_mascheroni - 1)" by (intro smallo_real_nat_transfer) simp_all finally have "M \[at_top] (\x. ln x + 2 * euler_mascheroni - 1)" by (simp add: asymp_equiv_altdef) also have "\ = (\x::nat. ln x + (2 * euler_mascheroni - 1))" by (simp add: algebra_simps) also have "\ \[at_top] (\x::nat. ln x)" by (subst asymp_equiv_add_right) auto finally show ?thesis . qed end subsection \The asymptotic ditribution of coprime pairs\ context fixes A :: "nat \ (nat \ nat) set" defines "A \ (\N. {(m,n) \ {1..N} \ {1..N}. coprime m n})" begin lemma coprime_pairs_asymptotics: "(\N. real (card (A N)) - 6 / pi\<^sup>2 * (real N)\<^sup>2) \ O(\N. real N * ln (real N))" proof - define C :: "nat \ (nat \ nat) set" where "C = (\N. (\m\{1..N}. (\n. (m,n)) ` totatives m))" define D :: "nat \ (nat \ nat) set" where "D = (\N. (\n\{1..N}. (\m. (m,n)) ` totatives n))" have fin: "finite (C N)" "finite (D N)" for N unfolding C_def D_def by (intro finite_UN_I finite_imageI; simp)+ have *: "card (A N) = 2 * (\m\{0<..N}. totient m) - 1" if N: "N > 0" for N proof - have "A N = C N \ D N" by (auto simp add: A_def C_def D_def totatives_def image_iff ac_simps) also have "card \ = card (C N) + card (D N) - card (C N \ D N)" using card_Un_Int[OF fin[of N]] by arith also have "C N \ D N = {(1, 1)}" using N by (auto simp: image_iff totatives_def C_def D_def) also have "D N = (\(x,y). (y,x)) ` C N" by (simp add: image_UN image_image C_def D_def) also have "card \ = card (C N)" by (rule card_image) (simp add: inj_on_def C_def) also have "card (C N) = (\m\{1..N}. card ((\n. (m,n)) ` totatives m))" unfolding C_def by (intro card_UN_disjoint) auto also have "\ = (\m\{1..N}. totient m)" unfolding totient_def by (subst card_image) (auto simp: inj_on_def) also have "\ = (\m\{0<..N}. totient m)" by (intro sum.cong refl) auto finally show "card (A N) = 2 * \ - 1" by simp qed have **: "(\m\{0<..N}. totient m) \ 1" if "N \ 1" for N proof - have "1 \ N" by fact also have "N = (\m\{0<..N}. 1)" by simp also have "(\m\{0<..N}. 1) \ (\m\{0<..N}. totient m)" by (intro sum_mono) (simp_all add: Suc_le_eq) finally show ?thesis . qed have "(\N. real (card (A N)) - 6 / pi\<^sup>2 * (real N)\<^sup>2) \ \(\N. 2 * (sum_upto (\m. real (totient m)) (real N) - (3 / pi\<^sup>2 * (real N)\<^sup>2)) - 1)" (is "_ \ \(?f)") using * ** by (intro bigthetaI_cong eventually_mono [OF eventually_gt_at_top[of "0::nat"]]) (auto simp: of_nat_diff sum_upto_altdef) also have "?f \ O(\N. real N * ln (real N))" proof (rule landau_o.big.compose[OF _ filterlim_real_sequentially], rule sum_in_bigo) show " (\x. 2 * (sum_upto (\m. real (totient m)) x - 3 / pi\<^sup>2 * x\<^sup>2)) \ O(\x. x * ln x)" by (subst landau_o.big.cmult_in_iff, simp, rule summatory_totient_asymptotics) qed simp_all finally show ?thesis . qed theorem coprime_pairs_asymptotics': "(\N. real (card (A N))) =o (\N. 6 / pi\<^sup>2 * (real N)\<^sup>2) +o O(\N. real N * ln (real N))" using coprime_pairs_asymptotics by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) theorem coprime_pairs_asymptotics'': "(\N. real (card (A N))) \[at_top] (\N. 6 / pi\<^sup>2 * (real N)\<^sup>2)" proof - have "(\N. real (card (A N)) - 6 / pi\<^sup>2 * (real N) ^ 2) \ O(\N. real N * ln (real N))" by (rule coprime_pairs_asymptotics) also have "(\N. real N * ln (real N)) \ o(\N. 6 / pi ^ 2 * real N ^ 2)" by (rule landau_o.small.compose[OF _ filterlim_real_sequentially]) simp finally show ?thesis by (simp add: asymp_equiv_altdef) qed theorem coprime_probability_tendsto: "(\N. card (A N) / card ({1..N} \ {1..N})) \ 6 / pi\<^sup>2" proof - have "(\N. 6 / pi ^ 2) \[at_top] (\N. 6 / pi ^ 2 * real N ^ 2 / real N ^ 2)" using eventually_gt_at_top[of "0::nat"] by (intro asymp_equiv_refl_ev) (auto elim!: eventually_mono) also have "\ \[at_top] (\N. real (card (A N)) / real N ^ 2)" by (intro asymp_equiv_intros asymp_equiv_symI[OF coprime_pairs_asymptotics'']) also have "\ \[at_top] (\N. real (card (A N)) / real (card ({1..N} \ {1..N})))" by (simp add: power2_eq_square) finally have "\ \[at_top] (\_. 6 / pi ^ 2)" by (simp add: asymp_equiv_sym) thus ?thesis by (rule asymp_equivD_const) qed end subsection \The asymptotics of the number of Farey fractions\ definition farey_fractions :: "nat \ rat set" where "farey_fractions N = {q :: rat \ {0<..1}. snd (quotient_of q) \ int N} " lemma Fract_eq_coprime: assumes "Rat.Fract a b = Rat.Fract c d" "b > 0" "d > 0" "coprime a b" "coprime c d" shows "a = c" "b = d" proof - from assms have "a * d = c * b" by (auto simp: eq_rat) hence "abs (a * d) = abs (c * b)" by (simp only:) hence "abs a * abs d = abs c * abs b" by (simp only: abs_mult) also have "?this \ abs a = abs c \ d = b" using assms by (subst coprime_crossproduct_int) simp_all finally show "b = d" by simp with \a * d = c * b\ and \b > 0\ show "a = c" by simp qed lemma quotient_of_split: "P (quotient_of q) = (\a b. b > 0 \ coprime a b \ q = Rat.Fract a b \ P (a, b))" by (cases q) (auto simp: quotient_of_Fract dest: Fract_eq_coprime) lemma quotient_of_split_asm: "P (Rat.quotient_of q) = (\(\a b. b > 0 \ coprime a b \ q = Rat.Fract a b \ \P (a, b)))" using quotient_of_split[of P q] by blast lemma farey_fractions_bij: "bij_betw (\(a,b). Rat.Fract (int a) (int b)) {(a,b)|a b. 0 < a \ a \ b \ b \ N \ coprime a b} (farey_fractions N)" proof (rule bij_betwI[of _ _ _ "\q. case quotient_of q of (a, b) \ (nat a, nat b)"], goal_cases) case 1 show ?case by (auto simp: farey_fractions_def Rat.zero_less_Fract_iff Rat.Fract_le_one_iff Rat.quotient_of_Fract Rat.normalize_def gcd_int_def Let_def) next case 2 show ?case by (auto simp add: farey_fractions_def Rat.Fract_le_one_iff Rat.zero_less_Fract_iff split: prod.splits quotient_of_split_asm) (simp add: coprime_int_iff [symmetric]) next case (3 x) thus ?case by (auto simp: Rat.quotient_of_Fract Rat.normalize_def Let_def gcd_int_def) next case (4 x) thus ?case unfolding farey_fractions_def by (split quotient_of_split) (auto simp: Rat.zero_less_Fract_iff) qed lemma card_farey_fractions: "card (farey_fractions N) = sum totient {0<..N}" proof - have "card (farey_fractions N) = card {(a,b)|a b. 0 < a \ a \ b \ b \ N \ coprime a b}" using farey_fractions_bij by (rule bij_betw_same_card [symmetric]) also have "{(a,b)|a b. 0 < a \ a \ b \ b \ N \ coprime a b} = (\b\{0<..N}. (\a. (a, b)) ` totatives b)" by (auto simp: totatives_def image_iff) also have "card \ = (\b\{0<..N}. card ((\a. (a, b)) ` totatives b))" by (intro card_UN_disjoint) auto also have "\ = (\b\{0<..N}. totient b)" unfolding totient_def by (intro sum.cong refl card_image) (auto simp: inj_on_def) finally show ?thesis . qed lemma card_farey_fractions_asymptotics: "(\N. real (card (farey_fractions N)) - 3 / pi\<^sup>2 * (real N)\<^sup>2) \ O(\N. real N * ln (real N))" proof - have "(\N. sum_upto (\n. real (totient n)) (real N) - 3 / pi\<^sup>2 * (real N)\<^sup>2) \ O(\N. real N * ln (real N))" (is "?f \ _") using summatory_totient_asymptotics filterlim_real_sequentially by (rule landau_o.big.compose) also have "?f = (\N. real (card (farey_fractions N)) - 3 / pi\<^sup>2 * (real N)\<^sup>2)" by (intro ext) (simp add: sum_upto_altdef card_farey_fractions) finally show ?thesis . qed theorem card_farey_fractions_asymptotics': "(\N. card (farey_fractions N)) =o (\N. 3 / pi\<^sup>2 * N^2) +o O(\N. N * ln N)" using card_farey_fractions_asymptotics by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) theorem card_farey_fractions_asymptotics'': "(\N. real (card (farey_fractions N))) \[at_top] (\N. 3 / pi\<^sup>2 * (real N)\<^sup>2)" proof - have "(\N. real (card (farey_fractions N)) - 3 / pi\<^sup>2 * (real N) ^ 2) \ O(\N. real N * ln (real N))" by (rule card_farey_fractions_asymptotics) also have "(\N. real N * ln (real N)) \ o(\N. 3 / pi ^ 2 * real N ^ 2)" by (rule landau_o.small.compose[OF _ filterlim_real_sequentially]) simp finally show ?thesis by (simp add: asymp_equiv_altdef) qed end diff --git a/thys/Furstenberg_Topology/Furstenberg_Topology.thy b/thys/Furstenberg_Topology/Furstenberg_Topology.thy --- a/thys/Furstenberg_Topology/Furstenberg_Topology.thy +++ b/thys/Furstenberg_Topology/Furstenberg_Topology.thy @@ -1,880 +1,878 @@ (* File: Furstenberg_Topology.thy Author: Manuel Eberl, TU München *) section \Furstenberg's topology and his proof of the infinitude of primes\ theory Furstenberg_Topology imports "HOL-Real_Asymp.Real_Asymp" "HOL-Analysis.Analysis" "HOL-Number_Theory.Number_Theory" begin text \ This article gives a formal version of Furstenberg's topological proof of the infinitude of primes~\cite{furstenberg}. He defines a topology on the integers based on arithmetic progressions (or, equivalently, residue classes). Apart from yielding a short proof of the infinitude of primes, this topology is also fairly `nice' in general: it is second countable, metrizable, and perfect. All of these (well-known) facts will be formally proven below. \ subsection \Arithmetic progressions of integers\ text \ We first define `bidirectional infinite arithmetic progressions' on \\\ in the sense that to an integer \a\ and a positive integer \b\, we associate all the integers \x\ such that $x \equiv a\ (\text{mod}\ b)$, or, equivalently, $\{a + nb\mid n\in\mathbb{Z}\}$. \ definition arith_prog :: "int \ nat \ int set" where "arith_prog a b = {x. [x = a] (mod int b)}" lemma arith_prog_0_right [simp]: "arith_prog a 0 = {a}" by (simp add: arith_prog_def) lemma arith_prog_Suc_0_right [simp]: "arith_prog a (Suc 0) = UNIV" by (auto simp: arith_prog_def) lemma in_arith_progI [intro]: "[x = a] (mod b) \ x \ arith_prog a b" by (auto simp: arith_prog_def) text \ Two arithmetic progressions with the same period and noncongruent starting points are disjoint. \ lemma arith_prog_disjoint: assumes "[a \ a'] (mod int b)" and "b > 0" shows "arith_prog a b \ arith_prog a' b = {}" using assms by (auto simp: arith_prog_def cong_def) text \ Multiplying the period gives us a subset of the original progression. \ lemma arith_prog_dvd_mono: "b dvd b' \ arith_prog a b' \ arith_prog a b" by (auto simp: arith_prog_def cong_dvd_modulus) text \ The following proves the alternative definition mentioned above. \ lemma bij_betw_arith_prog: assumes "b > 0" shows "bij_betw (\n. a + int b * n) UNIV (arith_prog a b)" proof (rule bij_betwI[of _ _ _ "\x. (x - a) div int b"], goal_cases) case 1 thus ?case by (auto simp: arith_prog_def cong_add_lcancel_0 cong_mult_self_right mult_of_nat_commute) next case 4 thus ?case by (auto simp: arith_prog_def cong_iff_lin) qed (use \b > 0\ in \auto simp: arith_prog_def\) lemma arith_prog_altdef: "arith_prog a b = range (\n. a + int b * n)" proof (cases "b = 0") case False thus ?thesis using bij_betw_arith_prog[of b] by (auto simp: bij_betw_def) qed auto text \ A simple corollary from this is also that any such arithmetic progression is infinite. \ lemma infinite_arith_prog: "b > 0 \ infinite (arith_prog a b)" using bij_betw_finite[OF bij_betw_arith_prog[of b]] by simp subsection \The Furstenberg topology on \\\\ text \ The typeclass-based topology is somewhat nicer to use in Isabelle/HOL, but the integers, of course, already have a topology associated to them. We therefore need to introduce a type copy of the integers and furnish them with the new topology. We can easily convert between them and the `proper' integers using Lifting and Transfer. \ typedef fbint = "UNIV :: int set" morphisms int_of_fbint fbint .. setup_lifting type_definition_fbint lift_definition arith_prog_fb :: "int \ nat \ fbint set" is "arith_prog" . instantiation fbint :: topological_space begin text \ Furstenberg defined the topology as the one generated by all arithmetic progressions. We use a slightly more explicit equivalent formulation that exploits the fact that the intersection of two arithmetic progressions is again an arithmetic progression (or empty). \ lift_definition open_fbint :: "fbint set \ bool" is "\U. (\x\U. \b>0. arith_prog x b \ U)" . text \ We now prove that this indeed forms a topology. \ instance proof show "open (UNIV :: fbint set)" by transfer auto next fix U V :: "fbint set" assume "open U" and "open V" show "open (U \ V)" proof (use \open U\ \open V\ in transfer, safe) fix U V :: "int set" and x :: int assume U: "\x\U. \b>0. arith_prog x b \ U" and V: "\x\V. \b>0. arith_prog x b \ V" assume x: "x \ U" "x \ V" from U x obtain b1 where b1: "b1 > 0" "arith_prog x b1 \ U" by auto from V x obtain b2 where b2: "b2 > 0" "arith_prog x b2 \ V" by auto from b1 b2 have "lcm b1 b2 > 0" "arith_prog x (lcm b1 b2) \ U \ V" using arith_prog_dvd_mono[of b1 "lcm b1 b2" x] arith_prog_dvd_mono[of b2 "lcm b1 b2" x] by (auto simp: lcm_pos_nat) thus "\b>0. arith_prog x b \ U \ V" by blast qed next fix F :: "fbint set set" assume *: "\U\F. open U" show "open (\F)" proof (use * in transfer, safe) fix F :: "int set set" and U :: "int set" and x :: int assume F: "\U\F. \x\U. \b>0. arith_prog x b \ U" assume "x \ U" "U \ F" with F obtain b where b: "b > 0" "arith_prog x b \ U" by blast with \U \ F\ show "\b>0. arith_prog x b \ \F" by blast qed qed end text \ Since any non-empty open set contains an arithmetic progression and arithmetic progressions are infinite, we obtain that all nonempty open sets are infinite. \ lemma open_fbint_imp_infinite: fixes U :: "fbint set" assumes "open U" and "U \ {}" shows "infinite U" using assms proof transfer fix U :: "int set" assume *: "\x\U. \b>0. arith_prog x b \ U" and "U \ {}" from \U \ {}\ obtain x where "x \ U" by auto with * obtain b where b: "b > 0" "arith_prog x b \ U" by auto from b have "infinite (arith_prog x b)" using infinite_arith_prog by blast with b show "infinite U" using finite_subset by blast qed lemma not_open_finite_fbint [simp]: assumes "finite (U :: fbint set)" "U \ {}" shows "\open U" using open_fbint_imp_infinite assms by blast text \ More or less by definition, any arithmetic progression is open. \ lemma open_arith_prog_fb [intro]: assumes "b > 0" shows "open (arith_prog_fb a b)" using assms proof transfer fix a :: int and b :: nat assume "b > 0" show "\x\arith_prog a b. \b'>0. arith_prog x b' \ arith_prog a b" proof (intro ballI exI[of _ b] conjI) fix x assume "x \ arith_prog a b" thus "arith_prog x b \ arith_prog a b" using cong_trans by (auto simp: arith_prog_def ) qed (use \b > 0\ in auto) qed text \ Slightly less obviously, any arithmetic progression is also closed. This can be seen by realising that for a period \b\, we can partition the integers into \b\ congruence classes and then the complement of each congruence class is the union of the other \b - 1\ classes, and unions of open sets are open. \ lemma closed_arith_prog_fb [intro]: assumes "b > 0" shows "closed (arith_prog_fb a b)" proof - have "open (-arith_prog_fb a b)" proof - have "-arith_prog_fb a b = (\i\{1.. arith_prog a b" if "x \ arith_prog (a + int i) b" "i \ {1.. a + int i] (mod int b)" proof assume "[a = a + int i] (mod int b)" hence "[a + 0 = a + int i] (mod int b)" by simp hence "[0 = int i] (mod int b)" by (subst (asm) cong_add_lcancel) auto with that show False by (auto simp: cong_def) qed thus ?thesis using arith_prog_disjoint[of a "a + int i" b] \b > 0\ that by auto qed have covering: "x \ arith_prog a b \ x \ (\i\{1..b > 0\ by simp also have "[a + (x - a) mod int b = a + (x - a)] (mod int b)" by (intro cong_add) auto finally have "[x = a + int i] (mod int b)" by (simp add: cong_sym_eq) hence "x \ arith_prog (a + int i) b" using \b > 0\ by (auto simp: arith_prog_def) moreover have "i < b" using \b > 0\ by (auto simp: i_def nat_less_iff) ultimately show ?thesis using \b > 0\ by (cases "i = 0") auto qed from disjoint and covering show "- arith_prog a b = (\i\{1..b > 0\ have "open \" by auto finally show ?thesis . qed thus ?thesis by (simp add: closed_def) qed subsection \The infinitude of primes\ text \ The infinite of the primes now follows quite obviously: The multiples of any prime form a closed set, so if there were only finitely many primes, the union of all of these would also be open. However, since any number other than \\1\ has a prime divisor, the union of all these sets is simply \\\{\1}\, which is obviously \<^emph>\not\ closed since the finite set \{\1}\ is not open. \ theorem "infinite {p::nat. prime p}" proof assume fin: "finite {p::nat. prime p}" define A where "A = (\p\{p::nat. prime p}. arith_prog_fb 0 p)" have "closed A" unfolding A_def using fin by (intro closed_Union) (auto simp: prime_gt_0_nat) hence "open (-A)" by (simp add: closed_def) also have "A = -{fbint 1, fbint (-1)}" unfolding A_def proof transfer show "(\p\{p::nat. prime p}. arith_prog 0 p) = - {1, - 1}" proof (intro equalityI subsetI) fix x :: int assume x: "x \ -{1, -1}" hence "\x\ \ 1" by auto show "x \ (\p\{p::nat. prime p}. arith_prog 0 p)" proof (cases "x = 0") case True thus ?thesis by (auto simp: A_def intro!: exI[of _ 2]) next case [simp]: False obtain p where p: "prime p" "p dvd x" using prime_divisor_exists[of x] and \\x\ \ 1\ by auto hence "x \ arith_prog 0 (nat p)" using prime_gt_0_int[of p] by (auto simp: arith_prog_def cong_0_iff) thus ?thesis using p by (auto simp: A_def intro!: exI[of _ "nat p"]) qed qed (auto simp: A_def arith_prog_def cong_0_iff) qed also have "-(-{fbint 1, fbint (-1)}) = {fbint 1, fbint (-1)}" by simp finally have "open {fbint 1, fbint (-1)}" . thus False by simp qed subsection \Additional topological properties\ text \ Just for fun, let us also show a few more properties of Furstenberg's topology. First, we show the equivalence to the above to Furstenberg's original definition (the topology generated by all arithmetic progressions). \ theorem topological_basis_fbint: "topological_basis {arith_prog_fb a b |a b. b > 0}" unfolding topological_basis_def proof safe fix a :: int and b :: nat assume "b > 0" thus "open (arith_prog_fb a b)" by auto next fix U :: "fbint set" assume "open U" hence "\x\U. \b. b > 0 \ arith_prog_fb (int_of_fbint x) b \ U" by transfer hence "\f. \x\U. f x > 0 \ arith_prog_fb (int_of_fbint x) (f x) \ U" by (subst (asm) bchoice_iff) then obtain f where f: "\x\U. f x > 0 \ arith_prog_fb (int_of_fbint x) (f x) \ U" .. define B where "B = (\x. arith_prog_fb (int_of_fbint x) (f x)) ` U" have "B \ {arith_prog_fb a b |a b. b > 0}" using f by (auto simp: B_def) moreover have "\B = U" proof safe fix x assume "x \ U" hence "x \ arith_prog_fb (int_of_fbint x) (f x)" using f by transfer auto with \x \ U\ show "x \ \B" by (auto simp: B_def) qed (use f in \auto simp: B_def\) ultimately show "\B'\{arith_prog_fb a b |a b. 0 < b}. \ B' = U" by auto qed lemma open_fbint_altdef: "open = generate_topology {arith_prog_fb a b |a b. b > 0}" using topological_basis_imp_subbasis[OF topological_basis_fbint] . text \ From this, we can immediately see that it is second countable: \ instance fbint :: second_countable_topology proof have "countable ((\(a,b). arith_prog_fb a b) ` (UNIV \ {b. b > 0}))" by (intro countable_image) auto also have "\ = {arith_prog_fb a b |a b. b > 0}" by auto ultimately show "\B::fbint set set. countable B \ open = generate_topology B" unfolding open_fbint_altdef by auto qed text \ A trivial consequence of the fact that nonempty open sets in this topology are infinite is that it is a perfect space: \ instance fbint :: perfect_space by standard auto text \ It is also Hausdorff, since given any two distinct integers, we can easily construct two non-overlapping arithmetic progressions that each contain one of them. We do not \<^emph>\really\ have to prove this since we will get it for free later on when we show that it is a metric space, but here is the proof anyway: \ instance fbint :: t2_space proof fix x y :: fbint assume "x \ y" define d where "d = nat \int_of_fbint x - int_of_fbint y\ + 1" from \x \ y\ have "d > 0" unfolding d_def by transfer auto define U where "U = arith_prog_fb (int_of_fbint x) d" define V where "V = arith_prog_fb (int_of_fbint y) d" have "U \ V = {}" unfolding U_def V_def d_def proof (use \x \ y\ in transfer, rule arith_prog_disjoint) fix x y :: int assume "x \ y" show "[x \ y] (mod int (nat \x - y\ + 1))" proof assume "[x = y] (mod int (nat \x - y\ + 1))" hence "\x - y\ + 1 dvd \x - y\" by (auto simp: cong_iff_dvd_diff algebra_simps) hence "\x - y\ + 1 \ \x - y\" by (rule zdvd_imp_le) (use \x \ y\ in auto) thus False by simp qed qed auto moreover have "x \ U" "y \ V" unfolding U_def V_def by (use \d > 0\ in transfer, fastforce)+ moreover have "open U" "open V" using \d > 0\ by (auto simp: U_def V_def) ultimately show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by blast qed (* TODO Move? *) text \ Next, we need a small lemma: Given an additional assumption, a $T_2$ space is also $T_3$: \ lemma t2_space_t3_spaceI: assumes "\(x :: 'a :: t2_space) U. x \ U \ open U \ \V. x \ V \ open V \ closure V \ U" shows "OFCLASS('a, t3_space_class)" proof fix X :: "'a set" and z :: 'a assume X: "closed X" "z \ X" with assms[of z "-X"] obtain V where V: "z \ V" "open V" "closure V \ -X" by auto show "\U V. open U \ open V \ z \ U \ X \ V \ U \ V = {}" by (rule exI[of _ V], rule exI[of _ "-closure V"]) (use X V closure_subset[of V] in auto) qed text \ Since the Furstenberg topology is $T_2$ and every arithmetic progression is also closed, we can now easily show that it is also $T_3$ (i.\,e.\ regular). Again, we do not really need this proof, but here it is: \ instance fbint :: t3_space proof (rule t2_space_t3_spaceI) fix x :: fbint and U :: "fbint set" assume "x \ U" and "open U" then obtain b where b: "b > 0" "arith_prog_fb (int_of_fbint x) b \ U" by transfer blast define V where "V = arith_prog_fb (int_of_fbint x) b" have "x \ V" unfolding V_def by transfer auto moreover have "open V" "closed V" using \b > 0\ by (auto simp: V_def) ultimately show "\V. x \ V \ open V \ closure V \ U" using b by (intro exI[of _ V]) (auto simp: V_def) qed subsection \Metrizability\ text \ The metrizability of Furstenberg's topology (i.\,e.\ that it is induced by some metric) can be shown from the fact that it is second countable and $T_3$ using Urysohn's Metrization Theorem, but this is not available in Isabelle yet. Let us therefore give an \<^emph>\explicit\ metric, as described by Zulfeqarr~\cite{zulfeqarr}. We follow the exposition by Dirmeier~\cite{dirmeier}. First, we define a kind of norm on the integers. The norm depends on a real parameter \q > 1\. The value of \q\ does not matter in the sense that all values induce the same topology (which we will show). For the final definition, we then simply pick \q = 2\. \ locale fbnorm = fixes q :: "real" assumes q_gt_1: "q > 1" begin definition N :: "int \ real" where "N n = (\k. if k = 0 \ int k dvd n then 0 else 1 / q ^ k)" lemma N_summable: "summable (\k. if k = 0 \ int k dvd n then 0 else 1 / q ^ k)" by (rule summable_comparison_test[OF _ summable_geometric[of "1/q"]]) (use q_gt_1 in \auto intro!: exI[of _ 0] simp: power_divide\) lemma N_sums: "(\k. if k = 0 \ int k dvd n then 0 else 1 / q ^ k) sums N n" using N_summable unfolding N_def by (rule summable_sums) lemma N_nonneg: "N n \ 0" by (rule sums_le[OF _ sums_zero N_sums]) (use q_gt_1 in auto) lemma N_uminus [simp]: "N (-n) = N n" by (simp add: N_def) lemma N_minus_commute: "N (x - y) = N (y - x)" using N_uminus[of "x - y"] by (simp del: N_uminus) lemma N_zero [simp]: "N 0 = 0" by (simp add: N_def) lemma not_dvd_imp_N_ge: assumes "\n dvd a" "n > 0" shows "N a \ 1 / q ^ n" by (rule sums_le[OF _ sums_single[of n] N_sums]) (use q_gt_1 assms in auto) lemma N_lt_imp_dvd: assumes "N a < 1 / q ^ n" and "n > 0" shows "n dvd a" using not_dvd_imp_N_ge[of n a] assms by auto lemma N_pos: assumes "n \ 0" shows "N n > 0" proof - have "0 < 1 / q ^ (nat \n\+1)" using q_gt_1 by simp also have "\1 + \n\ dvd \n\" using zdvd_imp_le[of "1 + \n\" "\n\"] assms by auto hence "1 / q ^ (nat \n\+1) \ N n" by (intro not_dvd_imp_N_ge) (use assms in auto) finally show ?thesis . qed lemma N_zero_iff [simp]: "N n = 0 \ n = 0" using N_pos[of n] by (cases "n = 0") auto lemma N_triangle_ineq: "N (n + m) \ N n + N m" proof (rule sums_le) let ?I = "\n k. if k = 0 \ int k dvd n then 0 else 1 / q ^ k" show "?I (n + m) sums N (n + m)" by (rule N_sums) show "(\k. ?I n k + ?I m k) sums (N n + N m)" by (intro sums_add N_sums) - show "\k. ?I (n + m) k \ ?I n k + ?I m k" - using q_gt_1 by auto -qed +qed (use q_gt_1 in auto) lemma N_1: "N 1 = 1 / (q * (q - 1))" proof (rule sums_unique2) have "(\k. if k = 0 \ int k dvd 1 then 0 else 1 / q ^ k) sums N 1" by (rule N_sums) also have "(\k. if k = 0 \ int k dvd 1 then 0 else 1 / q ^ k) = (\k. if k \ {0, 1} then 0 else (1 / q) ^ k)" by (simp add: power_divide cong: if_cong) finally show "(\k. if k \ {0, 1} then 0 else (1 / q) ^ k) sums N 1" . have "(\k. if k \ {0, 1} then 0 else (1 / q) ^ k) sums (1 / (1 - 1 / q) + (- (1 / q) - 1))" by (rule sums_If_finite_set'[OF geometric_sums]) (use q_gt_1 in auto) also have "\ = 1 / (q * (q - 1))" using q_gt_1 by (simp add: field_simps) finally show "(\k. if k \ {0, 1} then 0 else (1 / q) ^ k) sums \" . qed text \ It follows directly from the definition that norms fulfil a kind of monotonicity property with respect to divisibility: the norm of a number is at most as large as the norm of any of its factors: \ lemma N_dvd_mono: assumes "m dvd n" shows "N n \ N m" -proof (rule sums_le[OF allI N_sums N_sums]) +proof (rule sums_le[OF _ N_sums N_sums]) fix k :: nat show "(if k = 0 \ int k dvd n then 0 else 1 / q ^ k) \ (if k = 0 \ int k dvd m then 0 else 1 / q ^ k)" using q_gt_1 assms by auto qed text \ In particular, this means that 1 and -1 have the greatest norm. \ lemma N_le_N_1: "N n \ N 1" by (rule N_dvd_mono) auto text \ Primes have relatively large norms, almost reaching the norm of 1: \ lemma N_prime: assumes "prime p" shows "N p = N 1 - 1 / q ^ nat p" proof (rule sums_unique2) define p' where "p' = nat p" have p: "p = int p'" using assms by (auto simp: p'_def prime_ge_0_int) have "prime p'" using assms by (simp add: p) have "(\k. if k = 0 \ int k dvd p then 0 else 1 / q ^ k) sums N p" by (rule N_sums) also have "int k dvd p \ k \ {1, p'}" for k using assms by (auto simp: p prime_nat_iff) hence "(\k. if k = 0 \ int k dvd p then 0 else 1 / q ^ k) = (\k. if k \ {0, 1, p'} then 0 else (1 / q) ^ k)" using assms q_gt_1 by (simp add: power_divide cong: if_cong) finally show "\ sums N p" . have "(\k. if k \ {0, 1, p'} then 0 else (1 / q) ^ k) sums (1 / (1 - 1 / q) + (- (1 / q) - (1 / q) ^ p' - 1))" by (rule sums_If_finite_set'[OF geometric_sums]) (use \prime p'\ q_gt_1 prime_gt_Suc_0_nat[of p'] in \auto simp: \) also have "\ = N 1 - 1 / q ^ p'" using q_gt_1 by (simp add: field_simps N_1) finally show "(\k. if k \ {0, 1, p'} then 0 else (1 / q) ^ k) sums \" . qed lemma N_2: "N 2 = 1 / (q ^ 2 * (q - 1))" using q_gt_1 by (auto simp: N_prime N_1 field_simps power2_eq_square) lemma N_less_N_1: assumes "n \ 1" "n \ -1" shows "N n < N 1" proof (cases "n = 0") case False then obtain p where p: "prime p" "p dvd n" using prime_divisor_exists[of n] assms by force hence "N n \ N p" by (intro N_dvd_mono) also from p have "N p < N 1" using q_gt_1 by (simp add: N_prime) finally show ?thesis . qed (use q_gt_1 in \auto simp: N_1\) text \ Composites, on the other hand, do not achieve this: \ lemma nonprime_imp_N_lt: assumes "\prime_elem n" "\n\ \ 1" "n \ 0" shows "N n < N 1 - 1 / q ^ nat \n\" proof - obtain p where p: "prime p" "p dvd n" using prime_divisor_exists[of n] assms by auto define p' where "p' = nat p" have p': "p = int p'" using p by (auto simp: p'_def prime_ge_0_int) have "prime p'" using p by (simp add: p') define n' where "n' = nat \n\" have "n' > 1" using assms by (auto simp: n'_def) have "N n \ 1 / (q * (q - 1)) - 1 / q ^ p' - 1 / q ^ n'" proof (rule sums_le) show "(\k. if k = 0 \ int k dvd n then 0 else 1 / q ^ k) sums N n" by (rule N_sums) next from assms p have "n' \ p'" by (auto simp: n'_def p'_def nat_eq_iff) hence "(\k. if k \ {0, 1, p', n'} then 0 else (1 / q) ^ k) sums (1 / (1 - 1 / q) + (- (1 / q) - (1 / q) ^ p' - (1 / q) ^ n' - 1))" by (intro sums_If_finite_set'[OF geometric_sums]) (use \prime p'\ q_gt_1 prime_gt_Suc_0_nat[of p'] \n' > 1\ in \auto simp: \) also have "\ = 1 / (q * (q - 1)) - 1 / q ^ p' - 1 / q ^ n'" using q_gt_1 by (simp add: field_simps) finally show "(\k. if k \ {0, 1, p', n'} then 0 else (1 / q) ^ k) sums \" . next - show " \na. (if na = 0 \ int na dvd n then 0 else 1 / q ^ na) - \ (if na \ {0, 1, p', n'} then 0 else (1 / q) ^ na)" - using q_gt_1 assms p by (auto simp: p'_def n'_def power_divide) + show "\k. (if k = 0 \ int k dvd n then 0 else 1 / q ^ k) + \ (if k \ {0, 1, p', n'} then 0 else (1 / q) ^ k)" + using q_gt_1 p by (auto simp: p'_def n'_def power_divide) qed also have "\ < 1 / (q * (q - 1)) - 1 / q ^ n'" using q_gt_1 by simp finally show ?thesis by (simp add: n'_def N_1) qed text \ This implies that one can use the norm as a primality test: \ lemma prime_iff_N_eq: assumes "n \ 0" shows "prime_elem n \ N n = N 1 - 1 / q ^ nat \n\" proof - have *: "prime_elem n \ N n = N 1 - 1 / q ^ nat \n\" if "n > 0" for n proof - consider "n = 1" | "prime n" | "\prime n" "n > 1" using \n > 0\ by force thus ?thesis proof cases assume "n = 1" thus ?thesis using q_gt_1 by (auto simp: N_1) next assume n: "\prime n" "n > 1" with nonprime_imp_N_lt[of n] show ?thesis by simp qed (auto simp: N_prime prime_ge_0_int) qed show ?thesis proof (cases "n > 0") case True with * show ?thesis by blast next case False with *[of "-n"] assms show ?thesis by simp qed qed text \ Factorials, on the other hand, have very small norms: \ lemma N_fact_le: "N (fact m) \ 1 / (q - 1) * 1 / q ^ m" -proof (rule sums_le[OF allI N_sums]) +proof (rule sums_le[OF _ N_sums]) have "(\k. 1 / q ^ k / q ^ Suc m) sums (q / (q - 1) / q ^ Suc m)" using geometric_sums[of "1 / q"] q_gt_1 by (intro sums_divide) (auto simp: field_simps) also have "(q / (q - 1) / q ^ Suc m) = 1 / (q - 1) * 1 / q ^ m" using q_gt_1 by (simp add: field_simps) also have "(\k. 1 / q ^ k / q ^ Suc m) = (\k. 1 / q ^ (k + Suc m))" using q_gt_1 by (simp add: field_simps power_add) also have "\ = (\k. if k + Suc m \ m then 0 else 1 / q ^ (k + Suc m))" by auto finally have "\ sums (1 / (q - 1) * 1 / q ^ m)" . also have "?this \ (\k. if k \ m then 0 else 1 / q ^ k) sums (1 / (q - 1) * 1 / q ^ m)" by (rule sums_zero_iff_shift) auto finally show \ . next fix k :: nat have "int k dvd fact m" if "k > 0" "k \ m" proof - have "int k dvd int (fact m)" unfolding int_dvd_int_iff using that by (simp add: dvd_fact) thus "int k dvd fact m" unfolding of_nat_fact by simp qed thus "(if k = 0 \ int k dvd fact m then 0 else 1 / q ^ k) \ (if k \ m then 0 else 1 / q ^ k)" using q_gt_1 by auto qed lemma N_prime_mono: assumes "prime p" "prime p'" "p \ p'" shows "N p \ N p'" using assms q_gt_1 by (auto simp add: N_prime field_simps nat_le_iff prime_ge_0_int) lemma N_prime_ge: assumes "prime p" shows "N p \ 1 / (q\<^sup>2 * (q - 1))" proof - have "1 / (q ^ 2 * (q - 1)) = N 2" using q_gt_1 by (auto simp: N_prime N_1 field_simps power2_eq_square) also have "\ \ N p" using assms by (intro N_prime_mono) (auto simp: prime_ge_2_int) finally show ?thesis . qed lemma N_prime_elem_ge: assumes "prime_elem p" shows "N p \ 1 / (q\<^sup>2 * (q - 1))" proof (cases "p \ 0") case True with assms N_prime_ge show ?thesis by auto next case False with assms N_prime_ge[of "-p"] show ?thesis by auto qed text \ Next, we use this norm to derive a metric: \ lift_definition dist :: "fbint \ fbint \ real" is "\x y. N (x - y)" . lemma dist_self [simp]: "dist x x = 0" by transfer simp lemma dist_sym [simp]: "dist x y = dist y x" by transfer (simp add: N_minus_commute) lemma dist_pos: "x \ y \ dist x y > 0" by transfer (use N_pos in simp) lemma dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" using dist_pos[of x y] by (cases "x = y") auto lemma dist_triangle_ineq: "dist x z \ dist x y + dist y z" proof transfer fix x y z :: int show "N (x - z) \ N (x - y) + N (y - z)" using N_triangle_ineq[of "x - y" "y - z"] by simp qed text \ Lastly, we show that the metric we defined indeed induces the Furstenberg topology. \ theorem dist_induces_open: "open U \ (\x\U. \e>0. \y. dist x y < e \ y \ U)" proof (transfer, safe) fix U :: "int set" and x :: int assume *: "\x\U. \b>0. arith_prog x b \ U" assume "x \ U" with * obtain b where b: "b > 0" "arith_prog x b \ U" by blast define e where "e = 1 / q ^ b" show "\e>0. \y. N (x - y) < e \ y \ U" proof (rule exI; safe?) show "e > 0" using q_gt_1 by (simp add: e_def) next fix y assume "N (x - y) < e" also have "\ = 1 / q ^ b" by fact finally have "b dvd (x - y)" by (rule N_lt_imp_dvd) fact hence "y \ arith_prog x b" by (auto simp: arith_prog_def cong_iff_dvd_diff dvd_diff_commute) with b show "y \ U" by blast qed next fix U :: "int set" and x :: int assume *: "\x\U. \e>0. \y. N (x - y) < e \ y \ U" assume "x \ U" with * obtain e where e: "e > 0" "\y. N (x - y) < e \ y \ U" by blast have "eventually (\N. 1 / (q - 1) * 1 / q ^ N < e) at_top" using q_gt_1 \e > 0\ by real_asymp then obtain m where m: "1 / (q - 1) * 1 / q ^ m < e" by (auto simp: eventually_at_top_linorder) define b :: nat where "b = fact m" have "arith_prog x b \ U" proof fix y assume "y \ arith_prog x b" show "y \ U" proof (cases "y = x") case False from \y \ arith_prog x b\ obtain n where y: "y = x + int b * n" by (auto simp: arith_prog_altdef) from y and \y \ x\ have [simp]: "n \ 0" by auto have "N (x - y) = N (int b * n)" by (simp add: y) also have "\ \ N (int b)" by (rule N_dvd_mono) auto also have "\ \ 1 / (q - 1) * 1 / q ^ m" using N_fact_le by (simp add: b_def) also have "\ < e" by fact finally show "y \ U" using e by auto qed (use \x \ U\ in auto) qed moreover have "b > 0" by (auto simp: b_def) ultimately show "\b>0. arith_prog x b \ U" by blast qed end text \ We now show that the Furstenberg space is a metric space with this metric (with \q = 2\), which essentially only amounts to plugging together all the results from above. \ interpretation fb: fbnorm 2 by standard auto instantiation fbint :: dist begin definition dist_fbint where "dist_fbint = fb.dist" instance .. end instantiation fbint :: uniformity_dist begin definition uniformity_fbint :: "(fbint \ fbint) filter" where "uniformity_fbint = (INF e\{0 <..}. principal {(x, y). dist x y < e})" instance by standard (simp add: uniformity_fbint_def) end instance fbint :: open_uniformity proof fix U :: "fbint set" show "open U = (\x\U. eventually (\(x',y). x' = x \ y \ U) uniformity)" unfolding eventually_uniformity_metric dist_fbint_def using fb.dist_induces_open by simp qed instance fbint :: metric_space by standard (use fb.dist_triangle_ineq in \auto simp: dist_fbint_def\) text \ In particular, we can now show that the sequence \n!\ tends to 0 in the Furstenberg topology: \ lemma tendsto_fbint_fact: "(\n. fbint (fact n)) \ fbint 0" proof - have "(\n. dist (fbint (fact n)) (fbint 0)) \ 0" proof (rule tendsto_sandwich[OF always_eventually always_eventually]; safe?) fix n :: nat show "dist (fbint (fact n)) (fbint 0) \ 1 / 2 ^ n" unfolding dist_fbint_def by (transfer fixing: n) (use fb.N_fact_le[of n] in simp) show "dist (fbint (fact n)) (fbint 0) \ 0" by simp show "(\n. 1 / 2 ^ n :: real) \ 0" by real_asymp qed simp_all thus ?thesis using tendsto_dist_iff by metis qed end diff --git a/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy b/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy --- a/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy +++ b/thys/Irrational_Series_Erdos_Straus/Irrational_Series_Erdos_Straus.thy @@ -1,2020 +1,2017 @@ (* Title: Irrational_Series_Erdos_Straus.thy Author: Angeliki Koutsoukou-Argyraki and Wenda Li, University of Cambridge, UK. We formalise certain irrationality criteria for infinite series by P. Erdos and E.G. Straus. In particular, we formalise Theorem 2.1, Corollary 2.10 and Theorem 3.1 in [1]. The latter is an application of Theorem 2.1 involving the prime numbers. References: [1] P. Erdos and E.G. Straus, On the irrationality of certain series, Pacific Journal of Mathematics, Vol. 55, No 1, 1974 https://projecteuclid.org/euclid.pjm/1102911140 *) theory "Irrational_Series_Erdos_Straus" imports Prime_Number_Theorem.Prime_Number_Theorem Prime_Distribution_Elementary.PNT_Consequences begin section \Miscellaneous\ lemma suminf_comparison: - assumes "summable f" and "\n. norm (g n) \ f n" + assumes "summable f" and gf: "\n. norm (g n) \ f n" shows "suminf g \ suminf f" proof (rule suminf_le) - show "\n. g n \ f n" - apply rule - subgoal for n using assms(2)[rule_format,of n] by auto - done - show "summable g" - apply (rule summable_comparison_test'[OF \summable f\, of 0]) - using assms(2) by auto + show "g n \ f n" for n + using gf[of n] by auto + show "summable g" + using assms summable_comparison_test' by blast show "summable f" using assms(1) . qed lemma tendsto_of_int_diff_0: assumes "(\n. f n - of_int(g n)) \ (0::real)" "\\<^sub>F n in sequentially. f n > 0" shows "\\<^sub>F n in sequentially. 0 \ g n" proof - have "\\<^sub>F n in sequentially. \f n - of_int(g n)\ < 1 / 2" using assms(1)[unfolded tendsto_iff,rule_format,of "1/2"] by auto then show ?thesis using assms(2) apply eventually_elim by linarith qed lemma eventually_mono_sequentially: assumes "eventually P sequentially" assumes "\x. P (x+k) \ Q (x+k)" shows "eventually Q sequentially" using sequentially_offset[OF assms(1),of k] apply (subst eventually_sequentially_seg[symmetric,of _ k]) apply (elim eventually_mono) by fact lemma frequently_eventually_at_top: fixes P Q::"'a::linorder \ bool" assumes "frequently P at_top" "eventually Q at_top" shows "frequently (\x. P x \ (\y\x. Q y) ) at_top" using assms unfolding frequently_def eventually_at_top_linorder by (metis (mono_tags, hide_lams) le_cases order_trans) lemma eventually_at_top_mono: fixes P Q::"'a::linorder \ bool" assumes event_P:"eventually P at_top" assumes PQ_imp:"\x. x\z \ \y\x. P y \ Q x" shows "eventually Q at_top" proof - obtain N where N_P:"\n\N. P n" using event_P[unfolded eventually_at_top_linorder] by auto define N' where "N' = max N z" have "Q x" when "x\N'" for x apply (rule PQ_imp) using N_P that unfolding N'_def by auto then show ?thesis unfolding eventually_at_top_linorder by auto qed lemma frequently_at_top_elim: fixes P Q::"'a::linorder \ bool" assumes "\\<^sub>Fx in at_top. P x" assumes "\i. P i \ \j>i. Q j" shows "\\<^sub>Fx in at_top. Q x" using assms unfolding frequently_def eventually_at_top_linorder by (meson leD le_cases less_le_trans) lemma less_Liminf_iff: fixes X :: "_ \ _ :: complete_linorder" shows "Liminf F X < C \ (\yx. y \ X x) F)" apply (subst Not_eq_iff[symmetric]) apply (simp add:not_less not_frequently not_le le_Liminf_iff) by force lemma sequentially_even_odd_imp: assumes "\\<^sub>F N in sequentially. P (2*N)" "\\<^sub>F N in sequentially. P (2*N+1)" shows "\\<^sub>F n in sequentially. P n" proof - obtain N where N_P:"\x\N. P (2 * x) \ P (2 * x + 1)" using eventually_conj[OF assms] unfolding eventually_at_top_linorder by auto define N' where "N'=2*N " have "P n" when "n\2*N" for n proof - define n' where "n'= n div 2" then have "n'\N" using that by auto then have "P (2 * n') \ P (2 * n' + 1)" using N_P by auto then show ?thesis unfolding n'_def apply (cases "even n") by auto qed then show ?thesis unfolding eventually_at_top_linorder by auto qed section \Theorem 2.1 and Corollary 2.10\ context fixes a b ::"nat\int " assumes a_pos:"\ n. a n >0 " and a_large:"\\<^sub>F n in sequentially. a n > 1" and ab_tendsto: "(\n. \b n\ / (a (n-1)*a n)) \ 0" begin private lemma aux_series_summable: "summable (\n. b n / (\k\n. a k))" proof - have "\e>0. \\<^sub>F x in sequentially. \b x\ / (a (x-1) * a x) < e" using ab_tendsto[unfolded tendsto_iff] apply (simp add:of_int_abs[symmetric] abs_mult del:of_int_abs) by (subst (asm) (2) abs_of_pos,use \\ n. a n > 0\ in auto)+ from this[rule_format,of 1] have "\\<^sub>F x in sequentially. \real_of_int(b x)\ < (a (x-1) * a x)" using \\ n. a n >0\ by auto moreover have "\n. (\k\n. real_of_int (a k)) > 0" using a_pos by (auto intro!:linordered_semidom_class.prod_pos) ultimately have "\\<^sub>F n in sequentially. \b n\ / (\k\n. a k) < (a (n-1) * a n) / (\k\n. a k)" apply (elim eventually_mono) by (auto simp add:field_simps) moreover have "\b n\ / (\k\n. a k) = norm (b n / (\k\n. a k))" for n using \\n. (\k\n. real_of_int (a k)) > 0\[rule_format,of n] by auto ultimately have "\\<^sub>F n in sequentially. norm (b n / (\k\n. a k)) < (a (n-1) * a n) / (\k\n. a k)" by algebra moreover have "summable (\n. (a (n-1) * a n) / (\k\n. a k))" proof - obtain s where a_gt_1:"\ n\s. a n >1" using a_large[unfolded eventually_at_top_linorder] by auto define cc where "cc= (\k0" unfolding cc_def apply (rule linordered_semidom_class.prod_pos) using a_pos by auto have "(\k\n+s. a k) \ cc * 2^n" for n proof - have "prod a {s.. 2^n" proof (induct n) case 0 then show ?case using a_gt_1 by auto next case (Suc n) moreover have "a (s + Suc n) \ 2" using a_gt_1 by (smt le_add1) ultimately show ?case apply (subst prod.atLeastLessThan_Suc,simp) using mult_mono'[of 2 "a (Suc (s + n))" " 2 ^ n" "prod a {s..cc>0\ unfolding cc_def by (simp add: atLeast0AtMost) qed then have "1/(\k\n+s. a k) \ 1/(cc * 2^n)" for n proof - assume asm:"\n. cc * 2 ^ n \ prod a {..n + s}" then have "real_of_int (cc * 2 ^ n) \ prod a {..n + s}" using of_int_le_iff by blast moreover have "prod a {..n + s} >0" using \cc>0\ by (simp add: a_pos prod_pos) ultimately show ?thesis using \cc>0\ by (auto simp:field_simps simp del:of_int_prod) qed moreover have "summable (\n. 1/(cc * 2^n))" proof - have "summable (\n. 1/(2::int)^n)" using summable_geometric[of "1/(2::int)"] by (simp add:power_one_over) from summable_mult[OF this,of "1/cc"] show ?thesis by auto qed ultimately have "summable (\n. 1 / (\k\n+s. a k))" apply (elim summable_comparison_test'[where N=0]) apply (unfold real_norm_def, subst abs_of_pos) by (auto simp add: \\n. 0 < (\k\n. real_of_int (a k))\) then have "summable (\n. 1 / (\k\n. a k))" apply (subst summable_iff_shift[where k=s,symmetric]) by simp then have "summable (\n. (a (n+1) * a (n+2)) / (\k\n+2. a k))" proof - assume asm:"summable (\n. 1 / real_of_int (prod a {..n}))" have "1 / real_of_int (prod a {..n}) = (a (n+1) * a (n+2)) / (\k\n+2. a k)" for n proof - have "a (Suc (Suc n)) \ 0" "a (Suc n) \0" using a_pos by (metis less_irrefl)+ then show ?thesis by (simp add: atLeast0_atMost_Suc atMost_atLeast0) qed then show ?thesis using asm by auto qed then show "summable (\n. (a (n-1) * a n) / (\k\n. a k))" apply (subst summable_iff_shift[symmetric,of _ 2]) by auto qed ultimately show ?thesis apply (elim summable_comparison_test_ev[rotated]) by (simp add: eventually_mono) qed private fun get_c::"(nat \ int) \ (nat \ int) \ int \ nat \ (nat \ int)" where "get_c a' b' B N 0 = round (B * b' N / a' N)"| "get_c a' b' B N (Suc n) = get_c a' b' B N n * a' (n+N) - B * b' (n+N)" lemma ab_rationality_imp: assumes ab_rational:"(\n. (b n / (\i \ n. a i))) \ \" shows "\ (B::int)>0. \ c::nat\ int. (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\ (\n. c (Suc n) / a n) \ 0" proof - have [simp]:"a n \ 0" for n using a_pos by (metis less_numeral_extra(3)) obtain A::int and B::int where AB_eq:"(\n. real_of_int (b n) / real_of_int (prod a {..n})) = A / B" and "B>0" proof - obtain q::rat where "(\n. real_of_int (b n) / real_of_int (prod a {..n})) = real_of_rat q" using ab_rational by (rule Rats_cases) simp moreover obtain A::int and B::int where "q = Rat.Fract A B" "B > 0" "coprime A B" by (rule Rat_cases) auto ultimately show ?thesis by (auto intro!: that[of A B] simp:of_rat_rat) qed define f where "f = (\n. b n / real_of_int (prod a {..n}))" define R where "R = (\N. (\n. B*b (n+N+1) / prod a {N..n+N+1}))" have all_e_ubound:"\e>0. \\<^sub>F M in sequentially. \n. \B*b (n+M+1) / prod a {M..n+M+1}\ < e/4 * 1/2^n" proof safe fix e::real assume "e>0" obtain N where N_a2: "\n \ N. a n \ 2" and N_ba: "\n \ N. \b n\ / (a (n-1) * a n) < e/(4*B)" proof - have "\\<^sub>F n in sequentially. \b n\ / (a (n - 1) * a n) < e/(4*B)" using order_topology_class.order_tendstoD[OF ab_tendsto,of "e/(4*B)"] \B>0\ \e>0\ by auto moreover have "\\<^sub>F n in sequentially. a n \ 2" using a_large by (auto elim: eventually_mono) ultimately have "\\<^sub>F n in sequentially. \b n\ / (a (n - 1) * a n) < e/(4*B) \ a n \ 2" by eventually_elim auto then show ?thesis unfolding eventually_at_top_linorder using that by auto qed have geq_N_bound:"\B*b (n+M+1) / prod a {M..n+M+1}\ < e/4 * 1/2^n" when "M\N" for n M proof - define D where "D = B*b (n+M+1)/ (a (n+M) * a (n+M+1))" have "\B*b (n+M+1) / prod a {M..n+M+1}\ = \D / prod a {M.." proof - have "{M..n+M+1} = {M.. {n+M,n+M+1}" by auto then have "prod a {M..n+M+1} = a (n+M) * a (n+M+1)* prod a {M..e/4 * (1/prod a {M.." proof - have "\D\ < e/ 4" unfolding D_def using N_ba[rule_format, of "n+M+1"] \B>0\ \M \ N\ \e>0\ a_pos by (auto simp:field_simps abs_mult abs_of_pos) from mult_strict_right_mono[OF this,of "1/prod a {M..e>0\ show ?thesis apply (auto simp:abs_prod abs_mult prod_pos) by (subst (2) abs_of_pos,auto)+ qed also have "... \ e/4 * 1/2^n" proof - have "prod a {M.. 2^n" proof (induct n) case 0 then show ?case by simp next case (Suc n) then show ?case using \M\N\ by (simp add: N_a2 mult.commute mult_mono' prod.atLeastLessThan_Suc) qed then have "real_of_int (prod a {M.. 2^n" using numeral_power_le_of_int_cancel_iff by blast then show ?thesis using \e>0\ by (auto simp add:divide_simps) qed finally show ?thesis . qed show "\\<^sub>F M in sequentially. \n. \real_of_int (B * b (n + M + 1)) / real_of_int (prod a {M..n + M + 1})\ < e / 4 * 1 / 2 ^ n" apply (rule eventually_sequentiallyI[of N]) using geq_N_bound by blast qed have R_tendsto_0:"R \ 0" proof (rule tendstoI) fix e::real assume "e>0" show "\\<^sub>F x in sequentially. dist (R x) 0 < e" using all_e_ubound[rule_format,OF \e>0\] proof eventually_elim case (elim M) define g where "g = (\n. B*b (n+M+1) / prod a {M..n+M+1})" have g_lt:"\g n\ < e/4 * 1/2^n" for n using elim unfolding g_def by auto have g_abs_summable:"summable (\n. \g n\)" proof - have "summable (\n. e/4 * 1/2^n)" using summable_geometric[of "1/2",THEN summable_mult,of "e/4",simplified] by (auto simp add:algebra_simps power_divide) then show ?thesis apply (elim summable_comparison_test') using g_lt less_eq_real_def by auto qed have "\\n. g n\ \ (\n. \g n\)" by (rule summable_rabs[OF g_abs_summable]) also have "... \(\n. e/4 * 1/2^n)" proof (rule suminf_comparison) show "summable (\n. e/4 * 1/2^n)" using summable_geometric[of "1/2",THEN summable_mult,of "e/4",simplified] by (auto simp add:algebra_simps power_divide) - show "\n. norm \g n\ \ e / 4 * 1 / 2 ^ n" using g_lt less_eq_real_def by auto + show "\n. norm \g n\ \ e / 4 * 1 / 2 ^ n" using g_lt less_eq_real_def by auto qed also have "... = (e/4) * (\n. (1/2)^n)" apply (subst suminf_mult[symmetric]) subgoal apply (rule complete_algebra_summable_geometric) by simp subgoal by (auto simp:algebra_simps power_divide) done also have "... = e/2" by (simp add:suminf_geometric[of "1/2"]) finally have "\\n. g n\ \ e / 2" . then show "dist (R M) 0 < e" unfolding R_def g_def using \e>0\ by auto qed qed obtain N where R_N_bound:"\M \ N. \R M\ \ 1 / 4" and N_geometric:"\M\N. \n. \real_of_int (B * b (n + M + 1)) / (prod a {M..n + M + 1})\ < 1 / 2 ^ n" proof - obtain N1 where N1:"\M \ N1. \R M\ \ 1 / 4" using metric_LIMSEQ_D[OF R_tendsto_0,of "1/4"] all_e_ubound[rule_format,of 4,unfolded eventually_sequentially] by (auto simp:less_eq_real_def) obtain N2 where N2:"\M\N2. \n. \real_of_int (B * b (n + M + 1)) / (prod a {M..n + M + 1})\ < 1 / 2 ^ n" using all_e_ubound[rule_format,of 4,unfolded eventually_sequentially] by (auto simp:less_eq_real_def) define N where "N=max N1 N2" show ?thesis using that[of N] N1 N2 unfolding N_def by simp qed define C where "C = B * prod a {..nn. f n)" unfolding AB_eq f_def using \B>0\ by auto also have "... = B * prod a {..nn. f (n+N+1)))" using suminf_split_initial_segment[OF \summable f\, of "N+1"] by auto also have "... = B * prod a {..nn. f (n+N+1)))" using sum.atLeast0_lessThan_Suc by simp also have "... = C + B * b N / a N + (\n. B*b (n+N+1) / prod a {N..n+N+1})" proof - have "B * prod a {.. {N}" using ivl_disj_un_singleton(2) by blast then show ?thesis unfolding f_def by auto qed moreover have "B * prod a {..n. f (n+N+1)) = (\n. B*b (n+N+1) / prod a {N..n+N+1})" proof - have "summable (\n. f (n + N + 1))" using \summable f\ summable_iff_shift[of f "N+1"] by auto moreover have "prod a {.. {N..n + N + 1}" by auto then show ?thesis unfolding f_def apply simp apply (subst prod.union_disjoint) by auto qed ultimately show ?thesis apply (subst suminf_mult[symmetric]) by (auto simp add: mult.commute mult.left_commute) qed ultimately show ?thesis unfolding C_def by (auto simp:algebra_simps) qed also have "... = C +B * b N / a N + R N" unfolding R_def by simp finally show ?thesis . qed have R_bound:"\R M\ \ 1 / 4" and R_Suc:"R (Suc M) = a M * R M - B * b (Suc M) / a (Suc M)" when "M \ N" for M proof - define g where "g = (\n. B*b (n+M+1) / prod a {M..n+M+1})" have g_abs_summable:"summable (\n. \g n\)" proof - have "summable (\n.(1::real)/2^n)" using summable_geometric[of "(1::real)/2",simplified] by (auto elim!: back_subst[of "summable"] simp:field_simps) moreover have "\g n\ < 1/2^n" for n using N_geometric[rule_format,OF that] unfolding g_def by simp ultimately show ?thesis apply (elim summable_comparison_test') using less_eq_real_def by auto qed show "\R M\ \ 1 / 4" using R_N_bound[rule_format,OF that] . have "R M = (\n. g n)" unfolding R_def g_def by simp also have "... = g 0 + (\n. g (Suc n))" apply (subst suminf_split_head) using summable_rabs_cancel[OF g_abs_summable] by auto also have "... = g 0 + 1/a M * (\n. a M * g (Suc n))" apply (subst suminf_mult) by (auto simp add: g_abs_summable summable_Suc_iff summable_rabs_cancel) also have "... = g 0 + 1/a M * R (Suc M)" proof - have "a M * g (Suc n) = B * b (n + M + 2) / prod a {Suc M..n + M + 2}" for n proof - have "{M..Suc (Suc (M + n))} = {M} \ {Suc M..Suc (Suc (M + n))}" by auto then show ?thesis unfolding g_def using \B>0\ by (auto simp add:algebra_simps) qed then have "(\n. a M * g (Suc n)) = R (Suc M)" unfolding R_def by auto then show ?thesis by auto qed finally have "R M = g 0 + 1 / a M * R (Suc M)" . then have "R (Suc M) = a M * R M - g 0 * a M" by (auto simp add:algebra_simps) moreover have "{M..Suc M} = {M,Suc M}" by auto ultimately show "R (Suc M) = a M * R M - B * b (Suc M) / a (Suc M)" unfolding g_def by auto qed define c where "c = (\n. if n\N then get_c a b B N (n-N) else undefined)" have c_rec:"c (n+1) = c n * a n - B * b n" when "n \ N" for n unfolding c_def using that by (auto simp:Suc_diff_le) have c_R:"c (Suc n) / a n = R n" when "n \ N" for n using that proof (induct rule:nat_induct_at_least) case base have "\ c (N+1) / a N \ \ 1/2" proof - have "c N = round (B * b N / a N)" unfolding c_def by simp moreover have "c (N+1) / a N = c N - B * b N / a N" using a_pos[rule_format,of N] by (auto simp add:c_rec[of N,simplified] divide_simps) ultimately show ?thesis using of_int_round_abs_le by auto qed moreover have "\R N\ \ 1 / 4" using R_bound[of N] by simp ultimately have "\c (N+1) / a N - R N \ < 1" by linarith moreover have "c (N+1) / a N - R N \ \" proof - have "c (N+1) / a N = c N - B * b N / a N" using a_pos[rule_format,of N] by (auto simp add:c_rec[of N,simplified] divide_simps) moreover have " B * b N / a N + R N \ \" proof - have "C = B * (\nn {..n}" if "nB>0\ apply simp apply (subst prod.union_disjoint) by auto qed finally have "C = real_of_int (B * (\n \" using Ints_of_int by blast moreover note \A * prod a {.. ultimately show ?thesis by (metis Ints_diff Ints_of_int add.assoc add_diff_cancel_left') qed ultimately show ?thesis by (simp add: diff_diff_add) qed ultimately have "c (N+1) / a N - R N = 0" by (metis Ints_cases less_irrefl of_int_0 of_int_lessD) then show ?case by simp next case (Suc n) have "c (Suc (Suc n)) / a (Suc n) = c (Suc n) - B * b (Suc n) / a (Suc n)" apply (subst c_rec[of "Suc n",simplified]) using \N \ n\ by (auto simp add: divide_simps) also have "... = a n * R n - B * b (Suc n) / a (Suc n)" using Suc by (auto simp: divide_simps) also have "... = R (Suc n)" using R_Suc[OF \N \ n\] by simp finally show ?case . qed have ca_tendsto_zero:"(\n. c (Suc n) / a n) \ 0" using R_tendsto_0 apply (elim filterlim_mono_eventually) using c_R by (auto intro!:eventually_sequentiallyI[of N]) have ca_bound:"\c (n + 1)\ < a n / 2" when "n \ N" for n proof - have "\c (Suc n)\ / a n = \c (Suc n) / a n\" using a_pos[rule_format,of n] by auto also have "... = \R n\" using c_R[OF that] by auto also have "... < 1/2" using R_bound[OF that] by auto finally have "\c (Suc n)\ / a n < 1/2" . then show ?thesis using a_pos[rule_format,of n] by auto qed (* (* the following part corresponds to (2.7) (2.8) in the original paper, but turns out to be not necessary. *) have c_round:"c n = round (B * b n / a n)" when "n \ N" for n proof (cases "n=N") case True then show ?thesis unfolding c_def by simp next case False with \n\N\ obtain n' where n_Suc:"n=Suc n'" and "n' \ N" by (metis le_eq_less_or_eq lessE less_imp_le_nat) have "B * b n / a n = c n - R n" proof - have "R n = c n - B * b n / a n" using c_R[OF \n'\N\,symmetric,folded n_Suc] R_Suc[OF \n'\N\,folded n_Suc] by (auto simp:field_simps) then show ?thesis by (auto simp:field_simps) qed then have "\B * b n / a n - c n\ = \R n\" by auto then have "\B * b n / a n - c n\ < 1/2" using R_bound[OF \n \ N\] by auto from round_unique'[OF this] show ?thesis by auto qed *) show "\B>0. \c. (\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < a n / 2) \ (\n. c (Suc n) / a n) \ 0" unfolding eventually_at_top_linorder apply (rule exI[of _ B],use \B>0\ in simp) apply (intro exI[of _c] exI[of _ N]) using c_rec ca_bound ca_tendsto_zero by fastforce qed private lemma imp_ab_rational: assumes "\ (B::int)>0. \ c::nat\ int. (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\n. (b n / (\i \ n. a i))) \ \" proof - obtain B::int and c::"nat\int" and N::nat where "B>0" and large_n:"\n\N. B * b n = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < a n / 2 \ a n\2" proof - obtain B c where "B>0" and event1:"\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2" using assms by auto from eventually_conj[OF event1 a_large,unfolded eventually_at_top_linorder] obtain N where "\n\N. (B * b n = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2) \ 2 \ a n" by fastforce then show ?thesis using that[of B N c] \B>0\ by auto qed define f where "f=(\n. real_of_int (b n) / real_of_int (prod a {..n}))" define S where "S = (\n. f n)" have "summable f" unfolding f_def by (rule aux_series_summable) define C where "C=B*prod a {..nn. (c (n+N) * a (n+N)) / prod a {N..n+N})" define h2 where "h2 = (\n. c (n+N+1) / prod a {N..n+N})" have f_h12:"B*prod a {..n. B * b (n+N))" define g2 where "g2 = (\n. prod a {.. {N..n + N}) = prod a {.. {N..n + N}) = prod a {..n+N}" by (simp add: ivl_disj_un_one(4)) ultimately show ?thesis unfolding g2_def apply simp using a_pos by (metis less_irrefl) qed ultimately have "B*prod a {..nn. f (n+N)))" using suminf_split_initial_segment[OF \summable f\,of N] unfolding S_def by (auto simp:algebra_simps) also have "... = C + B*prod a {..n. f (n+N))" unfolding C_def by (auto simp:algebra_simps) also have "... = C + (\n. h1 n - h2 n)" apply (subst suminf_mult[symmetric]) subgoal using \summable f\ by (simp add: summable_iff_shift) subgoal using f_h12 by auto done also have "... = C + h1 0" proof - have "(\n. \i\n. h1 i - h2 i) \ (\i. h1 i - h2 i)" proof (rule summable_LIMSEQ') have "(\i. h1 i - h2 i) = (\i. real_of_int (B * prod a {..i. h1 i - h2 i)" using \summable f\ by (simp add: summable_iff_shift summable_mult) qed moreover have "(\i\n. h1 i - h2 i) = h1 0 - h2 n" for n proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(\i\Suc n. h1 i - h2 i) = (\i\n. h1 i - h2 i) + h1 (n+1) - h2 (n+1)" by auto also have "... = h1 0 - h2 n + h1 (n+1) - h2 (n+1)" using Suc by auto also have "... = h1 0 - h2 (n+1)" proof - have "h2 n = h1 (n+1)" unfolding h2_def h1_def apply (auto simp:prod.nat_ivl_Suc') using a_pos by (metis less_numeral_extra(3)) then show ?thesis by auto qed finally show ?case by simp qed ultimately have "(\n. h1 0 - h2 n) \ (\i. h1 i - h2 i)" by simp then have "h2 \ (h1 0 - (\i. h1 i - h2 i))" apply (elim metric_tendsto_imp_tendsto) by (auto intro!:eventuallyI simp add:dist_real_def) moreover have "h2 \ 0" proof - have h2_n:"\h2 n\ < (1 / 2)^(n+1)" for n proof - have "\h2 n\ = \c (n + N + 1)\ / prod a {N..n + N}" unfolding h2_def abs_divide using a_pos by (simp add: abs_of_pos prod_pos) also have "... < (a (N+n) / 2) / prod a {N..n + N}" unfolding h2_def apply (rule divide_strict_right_mono) subgoal using large_n[rule_format,of "N+n"] by (auto simp add:algebra_simps) subgoal using a_pos by (simp add: prod_pos) done also have "... = 1 / (2*prod a {N.. (1/2)^(n+1)" proof (induct n) case 0 then show ?case by auto next case (Suc n) define P where "P=1 / real_of_int (2 * prod a {N.. ( (1 / 2) ^ (n + 1) ) / a (n+N) " apply (rule divide_right_mono) subgoal unfolding P_def using Suc by auto subgoal by (simp add: a_pos less_imp_le) done also have "... \ ( (1 / 2) ^ (n + 1) ) / 2 " apply (rule divide_left_mono) using large_n[rule_format,of "n+N",simplified] by auto also have "... = (1 / 2) ^ (n + 2)" by auto finally show ?case by simp qed finally show ?thesis . qed have "(\n. (1 / 2)^(n+1)) \ (0::real)" using tendsto_mult_right_zero[OF LIMSEQ_abs_realpow_zero2[of "1/2",simplified],of "1/2"] by auto then show ?thesis apply (elim Lim_null_comparison[rotated]) using h2_n less_eq_real_def by (auto intro!:eventuallyI) qed ultimately have "(\i. h1 i - h2 i) = h1 0" using LIMSEQ_unique by fastforce then show ?thesis by simp qed also have "... = C + c N" unfolding h1_def using a_pos by auto (metis less_irrefl) finally show ?thesis . qed then have "S = (C + real_of_int (c N)) / (B*prod a {..0 < B\ a_pos less_irrefl mult.commute mult_pos_pos nonzero_mult_div_cancel_right of_int_eq_0_iff prod_pos) moreover have "... \ \" unfolding C_def f_def by (intro Rats_divide Rats_add Rats_mult Rats_of_int Rats_sum) ultimately show "S \ \" by auto qed theorem theorem_2_1_Erdos_Straus : "(\n. (b n / (\i \ n. a i))) \ \ \ (\ (B::int)>0. \ c::nat\ int. (\\<^sub>F n in sequentially. B*b n = c n * a n - c(n+1) \ \c(n+1)\The following is a Corollary to Theorem 2.1. \ corollary corollary_2_10_Erdos_Straus: assumes ab_event:"\\<^sub>F n in sequentially. b n > 0 \ a (n+1) \ a n" and ba_lim_leq:"lim (\n. (b(n+1) - b n )/a n) \ 0" and ba_lim_exist:"convergent (\n. (b(n+1) - b n )/a n)" and "liminf (\n. a n / b n) = 0 " shows "(\n. (b n / (\i \ n. a i))) \ \" proof assume "(\n. (b n / (\i \ n. a i))) \ \" then obtain B c where "B>0" and abc_event:"\\<^sub>F n in sequentially. B * b n = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" and ca_vanish: "(\n. c (Suc n) / a n) \ 0" using ab_rationality_imp by auto have bac_close:"(\n. B * b n / a n - c n) \ 0" proof - have "\\<^sub>F n in sequentially. B * b n - c n * a n + c (n + 1) = 0" using abc_event by (auto elim!:eventually_mono) then have "\\<^sub>F n in sequentially. (B * b n - c n * a n + c (n+1)) / a n = 0 " apply eventually_elim by auto then have "\\<^sub>F n in sequentially. B * b n / a n - c n + c (n + 1) / a n = 0" apply eventually_elim using a_pos by (auto simp:divide_simps) (metis less_irrefl) then have "(\n. B * b n / a n - c n + c (n + 1) / a n) \ 0" by (simp add: eventually_mono tendsto_iff) from tendsto_diff[OF this ca_vanish] show ?thesis by auto qed have c_pos:"\\<^sub>F n in sequentially. c n > 0" proof - from bac_close have *:"\\<^sub>F n in sequentially. c n \ 0" apply (elim tendsto_of_int_diff_0) using ab_event a_large apply (eventually_elim) using \B>0\ by auto show ?thesis proof (rule ccontr) assume "\ (\\<^sub>F n in sequentially. c n > 0)" moreover have "\\<^sub>F n in sequentially. c (Suc n) \ 0 \ c n\0" using * eventually_sequentially_Suc[of "\n. c n\0"] by (metis (mono_tags, lifting) eventually_at_top_linorder le_Suc_eq) ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0" using eventually_elim2 frequently_def by fastforce moreover have "\\<^sub>F n in sequentially. b n > 0 \ B * b n = c n * a n - c (n + 1)" using ab_event abc_event by eventually_elim auto ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0 \ b n > 0 \ B * b n = c n * a n - c (n + 1)" using frequently_eventually_frequently by fastforce from frequently_ex[OF this] obtain n where "c n = 0" "c (Suc n) \ 0" "b n > 0" "B * b n = c n * a n - c (n + 1)" by auto then have "B * b n \ 0" by auto then show False using \b n>0\ \B > 0\ using mult_pos_pos not_le by blast qed qed have bc_epsilon:"\\<^sub>F n in sequentially. b (n+1) / b n > (c (n+1) - \) / c n" when "\>0" "\<1" for \::real proof - have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \ / 2" using ca_vanish[unfolded tendsto_iff,rule_format, of "\/2"] \\>0\ by auto moreover then have "\\<^sub>F x in sequentially. \c (x+2) / a (x+1)\ < \ / 2" apply (subst (asm) eventually_sequentially_Suc[symmetric]) by simp moreover have "\\<^sub>F n in sequentially. B * b (n+1) = c (n+1) * a (n+1) - c (n + 2)" using abc_event apply (subst (asm) eventually_sequentially_Suc[symmetric]) by (auto elim:eventually_mono) moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0 \ c (n+2) > 0" proof - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" using c_pos by (subst eventually_sequentially_Suc) simp moreover then have "\\<^sub>F n in sequentially. 0 < c (Suc (Suc n))" using c_pos by (subst eventually_sequentially_Suc) simp ultimately show ?thesis using c_pos by eventually_elim auto qed ultimately show ?thesis using ab_event abc_event proof eventually_elim case (elim n) define \\<^sub>0 \\<^sub>1 where "\\<^sub>0 = c (n+1) / a n" and "\\<^sub>1 = c (n+2) / a (n+1)" have "\\<^sub>0 > 0" "\\<^sub>1 > 0" "\\<^sub>0 < \/2" "\\<^sub>1 < \/2" using a_pos elim by (auto simp add: \\<^sub>0_def \\<^sub>1_def) have "(\ - \\<^sub>1) * c n > 0" apply (rule mult_pos_pos) using \\\<^sub>1 > 0\ \\\<^sub>1 < \/2\ \\>0\ elim by auto moreover have "\\<^sub>0 * (c (n+1) - \) > 0" apply (rule mult_pos_pos[OF \\\<^sub>0 > 0\]) using elim(4) that(2) by linarith ultimately have "(\ - \\<^sub>1) * c n + \\<^sub>0 * (c (n+1) - \) > 0" by auto moreover have "c n - \\<^sub>0 > 0" using \\\<^sub>0 < \ / 2\ elim(4) that(2) by linarith moreover have "c n > 0" by (simp add: elim(4)) ultimately have "(c (n+1) - \) / c n < (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0)" by (auto simp add: field_simps) also have "... \ (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) * (a (n+1) / a n)" proof - have "(c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) > 0" by (smt \0 < (\ - \\<^sub>1) * real_of_int (c n)\ \0 < real_of_int (c n) - \\<^sub>0\ divide_pos_pos elim(4) mult_le_0_iff of_int_less_1_iff that(2)) moreover have "(a (n+1) / a n) \ 1" using a_pos elim(5) by auto ultimately show ?thesis by (metis mult_cancel_left1 real_mult_le_cancel_iff2) qed also have "... = (B * b (n+1)) / (B * b n)" proof - have "B * b n = c n * a n - c (n + 1)" using elim by auto also have "... = a n * (c n - \\<^sub>0)" using a_pos[rule_format,of n] unfolding \\<^sub>0_def by (auto simp:field_simps) finally have "B * b n = a n * (c n - \\<^sub>0)" . moreover have "B * b (n+1) = a (n+1) * (c (n+1) - \\<^sub>1)" unfolding \\<^sub>1_def using a_pos[rule_format,of "n+1"] apply (subst \B * b (n + 1) = c (n + 1) * a (n + 1) - c (n + 2)\) by (auto simp:field_simps) ultimately show ?thesis by (simp add: mult.commute) qed also have "... = b (n+1) / b n" using \B>0\ by auto finally show ?case . qed qed have eq_2_11:"\\<^sub>F n in sequentially. b (n+1) > b n + (1 - \)^2 * a n / B" when "\>0" "\<1" "\ (\\<^sub>F n in sequentially. c (n+1) \ c n)" for \::real proof - have "\\<^sub>F x in sequentially. c x < c (Suc x) " using that(3) by (simp add:not_eventually frequently_elim1) moreover have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \" using ca_vanish[unfolded tendsto_iff,rule_format, of \] \\>0\ by auto moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0" proof - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" using c_pos by (subst eventually_sequentially_Suc) simp then show ?thesis using c_pos by eventually_elim auto qed ultimately show ?thesis using ab_event abc_event bc_epsilon[OF \\>0\ \\<1\] proof (elim frequently_rev_mp,eventually_elim) case (elim n) then have "c (n+1) / a n < \" using a_pos[rule_format,of n] by auto also have "... \ \ * c n" using elim(7) that(1) by auto finally have "c (n+1) / a n < \ * c n" . then have "c (n+1) / c n < \ * a n" using a_pos[rule_format,of n] elim by (auto simp:field_simps) then have "(1 - \) * a n < a n - c (n+1) / c n" by (auto simp:algebra_simps) then have "(1 - \)^2 * a n / B < (1 - \) * (a n - c (n+1) / c n) / B" apply (subst (asm) real_mult_less_iff1[symmetric, of "(1-\)/B"]) using \\<1\ \B>0\ by (auto simp:divide_simps power2_eq_square) then have "b n + (1 - \)^2 * a n / B < b n + (1 - \) * (a n - c (n+1) / c n) / B" using \B>0\ by auto also have "... = b n + (1 - \) * ((c n *a n - c (n+1)) / c n) / B" using elim by (auto simp:field_simps) also have "... = b n + (1 - \) * (b n / c n)" proof - have "B * b n = c n * a n - c (n + 1)" using elim by auto from this[symmetric] show ?thesis using \B>0\ by simp qed also have "... = (1+(1-\)/c n) * b n" by (auto simp:algebra_simps) also have "... = ((c n+1-\)/c n) * b n" using elim by (auto simp:divide_simps) also have "... \ ((c (n+1) -\)/c n) * b n" proof - define cp where "cp = c n+1" have "c (n+1) \ cp" unfolding cp_def using \c n < c (Suc n)\ by auto moreover have "c n>0" "b n>0" using elim by auto ultimately show ?thesis apply (fold cp_def) by (auto simp:divide_simps) qed also have "... < b (n+1)" using elim by (auto simp:divide_simps) finally show ?case . qed qed have "\\<^sub>F n in sequentially. c (n+1) \ c n" proof (rule ccontr) assume "\ (\\<^sub>F n in sequentially. c (n + 1) \ c n)" from eq_2_11[OF _ _ this,of "1/2"] have "\\<^sub>F n in sequentially. b (n+1) > b n + 1/4 * a n / B" by (auto simp:algebra_simps power2_eq_square) then have *:"\\<^sub>F n in sequentially. (b (n+1) - b n) / a n > 1 / (B * 4)" apply (elim frequently_elim1) subgoal for n using a_pos[rule_format,of n] by (auto simp:field_simps) done define f where "f = (\n. (b (n+1) - b n) / a n)" have "f \ lim f" using convergent_LIMSEQ_iff ba_lim_exist unfolding f_def by auto from this[unfolded tendsto_iff,rule_format, of "1 / (B*4)"] have "\\<^sub>F x in sequentially. \f x - lim f\ < 1 / (B * 4)" using \B>0\ by (auto simp:dist_real_def) moreover have "\\<^sub>F n in sequentially. f n > 1 / (B * 4)" using * unfolding f_def by auto ultimately have "\\<^sub>F n in sequentially. f n > 1 / (B * 4) \ \f n - lim f\ < 1 / (B * 4)" by (auto elim:frequently_eventually_frequently[rotated]) from frequently_ex[OF this] obtain n where "f n > 1 / (B * 4)" "\f n - lim f\ < 1 / (B * 4)" by auto moreover have "lim f \ 0" using ba_lim_leq unfolding f_def by auto ultimately show False by linarith qed then obtain N where N_dec:"\n\N. c (n+1) \ c n" by (meson eventually_at_top_linorder) define max_c where "max_c = (MAX n \ {..N}. c n)" have max_c:"c n \ max_c" for n proof (cases "n\N") case True then show ?thesis unfolding max_c_def by simp next case False then have "n\N" by auto then have "c n\c N" proof (induct rule:nat_induct_at_least) case base then show ?case by simp next case (Suc n) then have "c (n+1) \ c n" using N_dec by auto then show ?case using \c n \ c N\ by auto qed moreover have "c N \ max_c" unfolding max_c_def by auto ultimately show ?thesis by auto qed have "max_c > 0 " proof - obtain N where "\n\N. 0 < c n" using c_pos[unfolded eventually_at_top_linorder] by auto then have "c N > 0" by auto then show ?thesis using max_c[of N] by simp qed have ba_limsup_bound:"1/(B*(B+1)) \ limsup (\n. b n/a n)" "limsup (\n. b n/a n) \ max_c / B + 1 / (B+1)" proof - define f where "f = (\n. b n/a n)" from tendsto_mult_right_zero[OF bac_close,of "1/B"] have "(\n. f n - c n / B) \ 0" unfolding f_def using \B>0\ by (auto simp:algebra_simps) from this[unfolded tendsto_iff,rule_format,of "1/(B+1)"] have "\\<^sub>F x in sequentially. \f x - c x / B\ < 1 / (B+1)" using \B>0\ by auto then have *:"\\<^sub>F n in sequentially. 1/(B*(B+1)) \ ereal (f n) \ ereal (f n) \ max_c / B + 1 / (B+1)" using c_pos proof eventually_elim case (elim n) then have "f n - c n / B < 1 / (B+1)" by auto then have "f n < c n / B + 1 / (B+1)" by simp also have "... \ max_c / B + 1 / (B+1)" using max_c[of n] using \B>0\ by (auto simp:divide_simps) finally have *:"f n < max_c / B + 1 / (B+1)" . have "1/(B*(B+1)) = 1/B - 1 / (B+1)" using \B>0\ by (auto simp:divide_simps) also have "... \ c n/B - 1 / (B+1)" using \0 < c n\ \B>0\ by (auto,auto simp:divide_simps) also have "... < f n" using elim by auto finally have "1/(B*(B+1)) < f n" . with * show ?case by simp qed show "limsup f \ max_c / B + 1 / (B+1)" apply (rule Limsup_bounded) using * by (auto elim:eventually_mono) have "1/(B*(B+1)) \ liminf f" apply (rule Liminf_bounded) using * by (auto elim:eventually_mono) also have "liminf f \ limsup f" by (simp add: Liminf_le_Limsup) finally show "1/(B*(B+1)) \ limsup f" . qed have "0 < inverse (ereal (max_c / B + 1 / (B+1)))" using \max_c > 0\ \B>0\ by (simp add: pos_add_strict) also have "... \ inverse (limsup (\n. b n/a n))" proof (rule ereal_inverse_antimono[OF _ ba_limsup_bound(2)]) have "0<1/(B*(B+1))" using \B>0\ by auto also have "... \ limsup (\n. b n/a n)" using ba_limsup_bound(1) . finally show "0\limsup (\n. b n/a n)" using zero_ereal_def by auto qed also have "... = liminf (\n. inverse (ereal ( b n/a n)))" apply (subst Liminf_inverse_ereal[symmetric]) using a_pos ab_event by (auto elim!:eventually_mono simp:divide_simps) also have "... = liminf (\n. ( a n/b n))" apply (rule Liminf_eq) using a_pos ab_event apply (auto elim!:eventually_mono) by (metis less_int_code(1)) finally have "liminf (\n. ( a n/b n)) > 0" . then show False using \liminf (\n. a n / b n) = 0\ by simp qed end section\Some auxiliary results on the prime numbers. \ lemma nth_prime_nonzero[simp]:"nth_prime n \ 0" by (simp add: prime_gt_0_nat prime_nth_prime) lemma nth_prime_gt_zero[simp]:"nth_prime n >0" by (simp add: prime_gt_0_nat prime_nth_prime) lemma ratio_of_consecutive_primes: "(\n. nth_prime (n+1)/nth_prime n) \1" proof - define f where "f=(\x. real (nth_prime (Suc x)) /real (nth_prime x))" define g where "g=(\x. (real x * ln (real x)) / (real (Suc x) * ln (real (Suc x))))" have p_n:"(\x. real (nth_prime x) / (real x * ln (real x))) \ 1" using nth_prime_asymptotics[unfolded asymp_equiv_def,simplified] . moreover have p_sn:"(\n. real (nth_prime (Suc n)) / (real (Suc n) * ln (real (Suc n)))) \ 1" using nth_prime_asymptotics[unfolded asymp_equiv_def,simplified ,THEN LIMSEQ_Suc] . ultimately have "(\x. f x * g x) \ 1" using tendsto_divide[OF p_sn p_n] unfolding f_def g_def by (auto simp:algebra_simps) moreover have "g \ 1" unfolding g_def by real_asymp ultimately have "(\x. if g x = 0 then 0 else f x) \ 1" apply (drule_tac tendsto_divide[OF _ \g \ 1\]) by auto then have "f \ 1" proof (elim filterlim_mono_eventually) have "\\<^sub>F x in sequentially. (if g (x+3) = 0 then 0 else f (x+3)) = f (x+3)" unfolding g_def by auto then show "\\<^sub>F x in sequentially. (if g x = 0 then 0 else f x) = f x" apply (subst (asm) eventually_sequentially_seg) by simp qed auto then show ?thesis unfolding f_def by auto qed lemma nth_prime_double_sqrt_less: assumes "\ > 0" shows "\\<^sub>F n in sequentially. (nth_prime (2*n) - nth_prime n) / sqrt (nth_prime n) < n powr (1/2+\)" proof - define pp ll where "pp=(\n. (nth_prime (2*n) - nth_prime n) / sqrt (nth_prime n))" and "ll=(\x::nat. x * ln x)" have pp_pos:"pp (n+1) > 0" for n unfolding pp_def by simp have "(\x. nth_prime (2 * x)) \[sequentially] (\x. (2 * x) * ln (2 * x))" using nth_prime_asymptotics[THEN asymp_equiv_compose ,of "(*) 2" sequentially,unfolded comp_def] using mult_nat_left_at_top pos2 by blast also have "... \[sequentially] (\x. 2 *x * ln x)" by real_asymp finally have "(\x. nth_prime (2 * x)) \[sequentially] (\x. 2 *x * ln x)" . from this[unfolded asymp_equiv_def, THEN tendsto_mult_left,of 2] have "(\x. nth_prime (2 * x) / (x * ln x)) \ 2" unfolding asymp_equiv_def by auto moreover have *:"(\x. nth_prime x / (x * ln x)) \ 1" using nth_prime_asymptotics unfolding asymp_equiv_def by auto ultimately have "(\x. (nth_prime (2 * x) - nth_prime x) / ll x) \ 1" unfolding ll_def apply - apply (drule (1) tendsto_diff) apply (subst of_nat_diff,simp) by (subst diff_divide_distrib,simp) moreover have "(\x. sqrt (nth_prime x) / sqrt (ll x)) \ 1" unfolding ll_def using tendsto_real_sqrt[OF *] by (auto simp: real_sqrt_divide) ultimately have "(\x. pp x * (sqrt (ll x) / (ll x))) \ 1" apply - apply (drule (1) tendsto_divide,simp) by (auto simp:field_simps of_nat_diff pp_def) moreover have "\\<^sub>F x in sequentially. sqrt (ll x) / ll x = 1/sqrt (ll x)" apply (subst eventually_sequentially_Suc[symmetric]) by (auto intro!:eventuallyI simp:ll_def divide_simps) ultimately have "(\x. pp x / sqrt (ll x)) \ 1" apply (elim filterlim_mono_eventually) by (auto elim!:eventually_mono) (metis mult.right_neutral times_divide_eq_right) moreover have "(\x. sqrt (ll x) / x powr (1/2+\)) \ 0" unfolding ll_def using \\>0\ by real_asymp ultimately have "(\x. pp x / x powr (1/2+\) * (sqrt (ll x) / sqrt (ll x))) \ 0" apply - apply (drule (1) tendsto_mult) by (auto elim:filterlim_mono_eventually) moreover have "\\<^sub>F x in sequentially. sqrt (ll x) / sqrt (ll x) = 1" apply (subst eventually_sequentially_Suc[symmetric]) by (auto intro!:eventuallyI simp:ll_def ) ultimately have "(\x. pp x / x powr (1/2+\)) \ 0" apply (elim filterlim_mono_eventually) by (auto elim:eventually_mono) from tendstoD[OF this, of 1,simplified] show "\\<^sub>F x in sequentially. pp x < x powr (1 / 2 + \)" apply (elim eventually_mono_sequentially[of _ 1]) using pp_pos by auto qed section \Theorem 3.1\ text\Theorem 3.1 is an application of Theorem 2.1 with the sequences considered involving the prime numbers.\ theorem theorem_3_10_Erdos_Straus: fixes a::"nat \ int" assumes a_pos:"\ n. a n >0" and "mono a" and nth_1:"(\n. nth_prime n / (a n)^2) \ 0" and nth_2:"liminf (\n. a n / nth_prime n) = 0" shows "(\n. (nth_prime n / (\i \ n. a i))) \ \" proof assume asm:"(\n. (nth_prime n / (\i \ n. a i))) \ \" have a2_omega:"(\n. (a n)^2) \ \(\x. x * ln x)" proof - have "(\n. real (nth_prime n)) \ o(\n. real_of_int ((a n)\<^sup>2))" apply (rule smalloI_tendsto[OF nth_1]) using a_pos by (metis (mono_tags, lifting) less_int_code(1) not_eventuallyD of_int_0_eq_iff zero_eq_power2) moreover have "(\x. real (nth_prime x)) \ \(\x. real x * ln (real x))" using nth_prime_bigtheta by blast ultimately show ?thesis using landau_omega.small_big_trans smallo_imp_smallomega by blast qed have a_gt_1:"\\<^sub>F n in sequentially. 1 < a n" proof - have "\\<^sub>F x in sequentially. \x * ln x\ \ (a x)\<^sup>2" using a2_omega[unfolded smallomega_def,simplified,rule_format,of 1] by auto then have "\\<^sub>F x in sequentially. \(x+3) * ln (x+3)\ \ (a (x+3))\<^sup>2" apply (subst (asm) eventually_sequentially_seg[symmetric, of _ 3]) by simp then have "\\<^sub>F n in sequentially. 1 < a ( n+3)" proof (elim eventually_mono) fix x assume "\real (x + 3) * ln (real (x + 3))\ \ real_of_int ((a (x + 3))\<^sup>2)" moreover have "\real (x + 3) * ln (real (x + 3))\ > 3" proof - have "ln (real (x + 3)) > 1" apply simp using ln3_gt_1 ln_gt_1 by force moreover have "real(x+3) \ 3" by simp ultimately have "(x+3)*ln (real (x + 3)) > 3*1 " apply (rule_tac mult_le_less_imp_less) by auto then show ?thesis by auto qed ultimately have "real_of_int ((a (x + 3))\<^sup>2) > 3" by auto then show "1 < a (x + 3)" by (smt Suc3_eq_add_3 a_pos add.commute of_int_1 one_power2) qed then show ?thesis apply (subst eventually_sequentially_seg[symmetric, of _ 3]) by auto qed obtain B::int and c where "B>0" and Bc_large:"\\<^sub>F n in sequentially. B * nth_prime n = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" and ca_vanish: "(\n. c (Suc n) / real_of_int (a n)) \ 0" proof - note a_gt_1 moreover have "(\n. real_of_int \int (nth_prime n)\ / real_of_int (a (n - 1) * a n)) \ 0" proof - define f where "f=(\n. nth_prime (n+1) / (a n * a (n+1)))" define g where "g=(\n. 2*nth_prime n / (a n)^2)" have "\\<^sub>F x in sequentially. norm (f x) \ g x" proof - have "\\<^sub>F n in sequentially. nth_prime (n+1) < 2*nth_prime n" using ratio_of_consecutive_primes[unfolded tendsto_iff ,rule_format,of 1,simplified] apply (elim eventually_mono) by (auto simp :divide_simps dist_norm) moreover have "\\<^sub>F n in sequentially. real_of_int (a n * a (n+1)) \ (a n)^2" apply (rule eventuallyI) using \mono a\ by (auto simp:power2_eq_square a_pos incseq_SucD) ultimately show ?thesis unfolding f_def g_def apply eventually_elim apply (subst norm_divide) apply (rule_tac linordered_field_class.frac_le) using a_pos[rule_format, THEN order.strict_implies_not_eq ] by auto qed moreover have "g \ 0 " using nth_1[THEN tendsto_mult_right_zero,of 2] unfolding g_def by auto ultimately have "f \ 0" using Lim_null_comparison[of f g sequentially] by auto then show ?thesis unfolding f_def by (rule_tac LIMSEQ_imp_Suc) auto qed moreover have "(\n. real_of_int (int (nth_prime n)) / real_of_int (prod a {..n})) \ \" using asm by simp ultimately have "\B>0. \c. (\\<^sub>F n in sequentially. B * int (nth_prime n) = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2) \ (\n. real_of_int (c (Suc n)) / real_of_int (a n)) \ 0" using ab_rationality_imp[OF a_pos,of nth_prime] by fast then show thesis apply clarify apply (rule_tac c=c and B=B in that) by auto qed have bac_close:"(\n. B * nth_prime n / a n - c n) \ 0" proof - have "\\<^sub>F n in sequentially. B * nth_prime n - c n * a n + c (n + 1) = 0" using Bc_large by (auto elim!:eventually_mono) then have "\\<^sub>F n in sequentially. (B * nth_prime n - c n * a n + c (n+1)) / a n = 0 " apply eventually_elim by auto then have "\\<^sub>F n in sequentially. B * nth_prime n / a n - c n + c (n + 1) / a n = 0" apply eventually_elim using a_pos by (auto simp:divide_simps) (metis less_irrefl) then have "(\n. B * nth_prime n / a n - c n + c (n + 1) / a n) \ 0" by (simp add: eventually_mono tendsto_iff) from tendsto_diff[OF this ca_vanish] show ?thesis by auto qed have c_pos:"\\<^sub>F n in sequentially. c n > 0" proof - from bac_close have *:"\\<^sub>F n in sequentially. c n \ 0" apply (elim tendsto_of_int_diff_0) using a_gt_1 apply (eventually_elim) using \B>0\ by auto show ?thesis proof (rule ccontr) assume "\ (\\<^sub>F n in sequentially. c n > 0)" moreover have "\\<^sub>F n in sequentially. c (Suc n) \ 0 \ c n\0" using * eventually_sequentially_Suc[of "\n. c n\0"] by (metis (mono_tags, lifting) eventually_at_top_linorder le_Suc_eq) ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0" using eventually_elim2 frequently_def by fastforce moreover have "\\<^sub>F n in sequentially. nth_prime n > 0 \ B * nth_prime n = c n * a n - c (n + 1)" using Bc_large by eventually_elim auto ultimately have "\\<^sub>F n in sequentially. c n = 0 \ c (Suc n) \ 0 \ B * nth_prime n = c n * a n - c (n + 1)" using frequently_eventually_frequently by fastforce from frequently_ex[OF this] obtain n where "c n = 0" "c (Suc n) \ 0" "B * nth_prime n = c n * a n - c (n + 1)" by auto then have "B * nth_prime n \ 0" by auto then show False using \B > 0\ by (simp add: mult_le_0_iff) qed qed have B_nth_prime:"\\<^sub>F n in sequentially. nth_prime n > B" proof - have "\\<^sub>F x in sequentially. B+1 \ nth_prime x" using nth_prime_at_top[unfolded filterlim_at_top_ge[where c="nat B+1"] ,rule_format,of "nat B + 1",simplified] apply (elim eventually_mono) using \B>0\ by auto then show ?thesis apply (elim eventually_mono) by auto qed have bc_epsilon:"\\<^sub>F n in sequentially. nth_prime (n+1) / nth_prime n > (c (n+1) - \) / c n" when "\>0" "\<1" for \::real proof - have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \ / 2" using ca_vanish[unfolded tendsto_iff,rule_format, of "\/2"] \\>0\ by auto moreover then have "\\<^sub>F x in sequentially. \c (x+2) / a (x+1)\ < \ / 2" apply (subst (asm) eventually_sequentially_Suc[symmetric]) by simp moreover have "\\<^sub>F n in sequentially. B * nth_prime (n+1) = c (n+1) * a (n+1) - c (n + 2)" using Bc_large apply (subst (asm) eventually_sequentially_Suc[symmetric]) by (auto elim:eventually_mono) moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0 \ c (n+2) > 0" proof - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" using c_pos by (subst eventually_sequentially_Suc) simp moreover then have "\\<^sub>F n in sequentially. 0 < c (Suc (Suc n))" using c_pos by (subst eventually_sequentially_Suc) simp ultimately show ?thesis using c_pos by eventually_elim auto qed ultimately show ?thesis using Bc_large proof eventually_elim case (elim n) define \\<^sub>0 \\<^sub>1 where "\\<^sub>0 = c (n+1) / a n" and "\\<^sub>1 = c (n+2) / a (n+1)" have "\\<^sub>0 > 0" "\\<^sub>1 > 0" "\\<^sub>0 < \/2" "\\<^sub>1 < \/2" using a_pos elim \mono a\ by (auto simp add: \\<^sub>0_def \\<^sub>1_def abs_of_pos) have "(\ - \\<^sub>1) * c n > 0" apply (rule mult_pos_pos) using \\\<^sub>1 > 0\ \\\<^sub>1 < \/2\ \\>0\ elim by auto moreover have "\\<^sub>0 * (c (n+1) - \) > 0" apply (rule mult_pos_pos[OF \\\<^sub>0 > 0\]) using elim(4) that(2) by linarith ultimately have "(\ - \\<^sub>1) * c n + \\<^sub>0 * (c (n+1) - \) > 0" by auto moreover have "c n - \\<^sub>0 > 0" using \\\<^sub>0 < \ / 2\ elim(4) that(2) by linarith moreover have "c n > 0" by (simp add: elim(4)) ultimately have "(c (n+1) - \) / c n < (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0)" by (auto simp add:field_simps) also have "... \ (c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) * (a (n+1) / a n)" proof - have "(c (n+1) - \\<^sub>1) / (c n - \\<^sub>0) > 0" by (smt \0 < (\ - \\<^sub>1) * real_of_int (c n)\ \0 < real_of_int (c n) - \\<^sub>0\ divide_pos_pos elim(4) mult_le_0_iff of_int_less_1_iff that(2)) moreover have "(a (n+1) / a n) \ 1" using a_pos \mono a\ by (simp add: mono_def) ultimately show ?thesis by (metis mult_cancel_left1 real_mult_le_cancel_iff2) qed also have "... = (B * nth_prime (n+1)) / (B * nth_prime n)" proof - have "B * nth_prime n = c n * a n - c (n + 1)" using elim by auto also have "... = a n * (c n - \\<^sub>0)" using a_pos[rule_format,of n] unfolding \\<^sub>0_def by (auto simp:field_simps) finally have "B * nth_prime n = a n * (c n - \\<^sub>0)" . moreover have "B * nth_prime (n+1) = a (n+1) * (c (n+1) - \\<^sub>1)" unfolding \\<^sub>1_def using a_pos[rule_format,of "n+1"] apply (subst \B * nth_prime (n + 1) = c (n + 1) * a (n + 1) - c (n + 2)\) by (auto simp:field_simps) ultimately show ?thesis by (simp add: mult.commute) qed also have "... = nth_prime (n+1) / nth_prime n" using \B>0\ by auto finally show ?case . qed qed have c_ubound:"\x. \n. c n > x" proof (rule ccontr) assume " \ (\x. \n. x < c n)" then obtain ub where "\n. c n \ ub" "ub > 0" by (meson dual_order.trans int_one_le_iff_zero_less le_cases not_le) define pa where "pa = (\n. nth_prime n / a n)" have pa_pos:"\n. pa n > 0" unfolding pa_def by (simp add: a_pos) have "liminf (\n. 1 / pa n) = 0" using nth_2 unfolding pa_def by auto then have "(\y\<^sub>F x in sequentially. ereal (1 / pa x) \ y)" apply (subst less_Liminf_iff[symmetric]) using \0 < B\ \0 < ub\ by auto then have "\\<^sub>F x in sequentially. 1 / pa x < B/(ub+1)" by (meson frequently_mono le_less_trans less_ereal.simps(1)) then have "\\<^sub>F x in sequentially. B*pa x > (ub+1)" apply (elim frequently_elim1) by (metis \0 < ub\ mult.left_neutral of_int_0_less_iff pa_pos pos_divide_less_eq pos_less_divide_eq times_divide_eq_left zless_add1_eq) moreover have "\\<^sub>F x in sequentially. c x \ ub" using \\n. c n \ ub\ by simp ultimately have "\\<^sub>F x in sequentially. B*pa x - c x > 1" apply (elim frequently_rev_mp eventually_mono) by linarith moreover have "(\n. B * pa n - c n) \0" unfolding pa_def using bac_close by auto from tendstoD[OF this,of 1] have "\\<^sub>F n in sequentially. \B * pa n - c n\ < 1" by auto ultimately have "\\<^sub>F x in sequentially. B*pa x - c x > 1 \ \B * pa x - c x\ < 1" using frequently_eventually_frequently by blast then show False by (simp add: frequently_def) qed have eq_2_11:"\\<^sub>F n in sequentially. c (n+1)>c n \ nth_prime (n+1) > nth_prime n + (1 - \)^2 * a n / B" when "\>0" "\<1" for \::real proof - have "\\<^sub>F x in sequentially. \c (Suc x) / a x\ < \" using ca_vanish[unfolded tendsto_iff,rule_format, of \] \\>0\ by auto moreover have "\\<^sub>F n in sequentially. c n > 0 \ c (n+1) > 0" proof - have "\\<^sub>F n in sequentially. 0 < c (Suc n)" using c_pos by (subst eventually_sequentially_Suc) simp then show ?thesis using c_pos by eventually_elim auto qed ultimately show ?thesis using Bc_large bc_epsilon[OF \\>0\ \\<1\] proof (eventually_elim, rule_tac impI) case (elim n) assume "c n < c (n + 1)" have "c (n+1) / a n < \" using a_pos[rule_format,of n] using elim(1,2) by auto also have "... \ \ * c n" using elim(2) that(1) by auto finally have "c (n+1) / a n < \ * c n" . then have "c (n+1) / c n < \ * a n" using a_pos[rule_format,of n] elim by (auto simp:field_simps) then have "(1 - \) * a n < a n - c (n+1) / c n" by (auto simp:algebra_simps) then have "(1 - \)^2 * a n / B < (1 - \) * (a n - c (n+1) / c n) / B" apply (subst (asm) real_mult_less_iff1[symmetric, of "(1-\)/B"]) using \\<1\ \B>0\ by (auto simp:divide_simps power2_eq_square) then have "nth_prime n + (1 - \)^2 * a n / B < nth_prime n + (1 - \) * (a n - c (n+1) / c n) / B" using \B>0\ by auto also have "... = nth_prime n + (1 - \) * ((c n *a n - c (n+1)) / c n) / B" using elim by (auto simp:field_simps) also have "... = nth_prime n + (1 - \) * (nth_prime n / c n)" proof - have "B * nth_prime n = c n * a n - c (n + 1)" using elim by auto from this[symmetric] show ?thesis using \B>0\ by simp qed also have "... = (1+(1-\)/c n) * nth_prime n" by (auto simp:algebra_simps) also have "... = ((c n+1-\)/c n) * nth_prime n" using elim by (auto simp:divide_simps) also have "... \ ((c (n+1) -\)/c n) * nth_prime n" proof - define cp where "cp = c n+1" have "c (n+1) \ cp" unfolding cp_def using \c n < c (n + 1)\ by auto moreover have "c n>0" "nth_prime n>0" using elim by auto ultimately show ?thesis apply (fold cp_def) by (auto simp:divide_simps) qed also have "... < nth_prime (n+1)" using elim by (auto simp:divide_simps) finally show "real (nth_prime n) + (1 - \)\<^sup>2 * real_of_int (a n) / real_of_int B < real (nth_prime (n + 1))" . qed qed have c_neq_large:"\\<^sub>F n in sequentially. c (n+1) \ c n" proof (rule ccontr) assume "\ (\\<^sub>F n in sequentially. c (n + 1) \ c n)" then have that:"\\<^sub>F n in sequentially. c (n + 1) = c n" unfolding frequently_def . have "\\<^sub>F x in sequentially. (B * int (nth_prime x) = c x * a x - c (x + 1) \ \real_of_int (c (x + 1))\ < real_of_int (a x) / 2) \ 0 < c x \ B < int (nth_prime x) \ (c (x+1)>c x \ nth_prime (x+1) > nth_prime x + a x / (2* B))" using Bc_large c_pos B_nth_prime eq_2_11[of "1-1/ sqrt 2",simplified] by eventually_elim (auto simp:divide_simps) then have "\\<^sub>F m in sequentially. nth_prime (m+1) > (1+1/(2*B))*nth_prime m" proof (elim frequently_eventually_at_top[OF that, THEN frequently_at_top_elim]) fix n assume "c (n + 1) = c n \ (\y\n. (B * int (nth_prime y) = c y * a y - c (y + 1) \ \real_of_int (c (y + 1))\ < real_of_int (a y) / 2) \ 0 < c y \ B < int (nth_prime y) \ (c y < c (y + 1) \ real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) < real (nth_prime (y + 1))))" then have "c (n + 1) = c n" and Bc_eq:"\y\n. B * int (nth_prime y) = c y * a y - c (y + 1) \ 0 < c y \ \real_of_int (c (y + 1))\ < real_of_int (a y) / 2 \ B < int (nth_prime y) \ (c y < c (y + 1) \ real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) < real (nth_prime (y + 1)))" by auto obtain m where "n c n" "c nN. N > n \ c N > c n" using c_ubound[rule_format, of "MAX x\{..n}. c x"] by (metis Max_ge atMost_iff dual_order.trans finite_atMost finite_imageI image_eqI linorder_not_le order_refl) then obtain N where "N>n" "c N>c n" by auto define A m where "A={m. n (m+1)\N \ c (m+1) > c n}" and "m = Min A" have "finite A" unfolding A_def by (metis (no_types, lifting) A_def add_leE finite_nat_set_iff_bounded_le mem_Collect_eq) moreover have "N-1\A" unfolding A_def using \c n < c N\ \n < N\ \c (n + 1) = c n\ by (smt Suc_diff_Suc Suc_eq_plus1 Suc_leI Suc_pred add.commute add_diff_inverse_nat add_leD1 diff_is_0_eq' mem_Collect_eq nat_add_left_cancel_less zero_less_one) ultimately have "m\A" using Min_in unfolding m_def by auto then have "n0" unfolding m_def A_def by auto moreover have "c m \ c n" proof (rule ccontr) assume " \ c m \ c n" then have "m-1\A" using \m\A\ \c (n + 1) = c n\ unfolding A_def by auto (smt One_nat_def Suc_eq_plus1 Suc_lessI less_diff_conv) from Min_le[OF \finite A\ this,folded m_def] \m>0\ show False by auto qed ultimately show ?thesis using that[of m] by auto qed have "(1 + 1 / (2 * B)) * nth_prime m < nth_prime m + a m / (2*B)" proof - have "nth_prime m < a m" proof - have "B * int (nth_prime m) < c m * (a m - 1)" using Bc_eq[rule_format,of m] \c m \ c n\ \c n < c (m + 1)\ \n < m\ by (auto simp:algebra_simps) also have "... \ c n * (a m - 1)" by (simp add: \c m \ c n\ a_pos mult_right_mono) finally have "B * int (nth_prime m) < c n * (a m - 1)" . moreover have "c n\B" proof - have " B * int (nth_prime n) = c n * (a n - 1)" "B < int (nth_prime n)" and c_a:"\real_of_int (c (n + 1))\ < real_of_int (a n) / 2" using Bc_eq[rule_format,of n] \c (n + 1) = c n\ by (auto simp:algebra_simps) from this(1) have " c n dvd (B * int (nth_prime n))" by simp moreover have "coprime (c n) (int (nth_prime n))" proof - have "c n < int (nth_prime n)" proof (rule ccontr) assume "\ c n < int (nth_prime n)" then have asm:"c n \ int (nth_prime n)" by auto then have "a n > 2 * nth_prime n" using c_a \c (n + 1) = c n\ by auto then have "a n -1 \ 2 * nth_prime n" by simp then have "a n - 1 > 2 * B" using \B < int (nth_prime n)\ by auto from mult_le_less_imp_less[OF asm this] \B>0\ have "int (nth_prime n) * (2 * B) < c n * (a n - 1)" by auto then show False using \B * int (nth_prime n) = c n * (a n - 1)\ by (smt \0 < B\ \B < int (nth_prime n)\ combine_common_factor mult.commute mult_pos_pos) qed then have "\ nth_prime n dvd c n" by (simp add: Bc_eq zdvd_not_zless) then have "coprime (int (nth_prime n)) (c n)" by (auto intro!:prime_imp_coprime_int) then show ?thesis using coprime_commute by blast qed ultimately have "c n dvd B" using coprime_dvd_mult_left_iff by auto then show ?thesis using \0 < B\ zdvd_imp_le by blast qed moreover have "c n > 0 " using Bc_eq by blast ultimately show ?thesis using \B>0\ by (smt a_pos mult_mono) qed then show ?thesis using \B>0\ by (auto simp:field_simps) qed also have "... < nth_prime (m+1)" using Bc_eq[rule_format, of m] \n \c m \ c n\ \c n < c (m+1)\ by linarith finally show "\j>n. (1 + 1 / real_of_int (2 * B)) * real (nth_prime j) < real (nth_prime (j + 1))" using \m>n\ by auto qed then have "\\<^sub>F m in sequentially. nth_prime (m+1)/nth_prime m > (1+1/(2*B))" by (auto elim:frequently_elim1 simp:field_simps) moreover have "\\<^sub>F m in sequentially. nth_prime (m+1)/nth_prime m < (1+1/(2*B))" using ratio_of_consecutive_primes[unfolded tendsto_iff,rule_format,of "1/(2*B)"] \B>0\ unfolding dist_real_def by (auto elim!:eventually_mono simp:algebra_simps) ultimately show False by (simp add: eventually_mono frequently_def) qed have c_gt_half:"\\<^sub>F N in sequentially. card {n\{N..<2*N}. c n > c (n+1)} > N / 2" proof - define h where "h=(\n. (nth_prime (2*n) - nth_prime n) / sqrt (nth_prime n))" have "\\<^sub>F n in sequentially. h n < n / 2" proof - have "\\<^sub>F n in sequentially. h n < n powr (5/6)" using nth_prime_double_sqrt_less[of "1/3"] unfolding h_def by auto moreover have "\\<^sub>F n in sequentially. n powr (5/6) < (n /2)" by real_asymp ultimately show ?thesis by eventually_elim auto qed moreover have "\\<^sub>F n in sequentially. sqrt (nth_prime n) / a n < 1 / (2*B)" using nth_1[THEN tendsto_real_sqrt,unfolded tendsto_iff ,rule_format,of "1/(2*B)"] \B>0\ a_pos by (auto simp:real_sqrt_divide abs_of_pos) ultimately have "\\<^sub>F x in sequentially. c (x+1) \ c x \ sqrt (nth_prime x) / a x < 1 / (2*B) \ h x < x / 2 \ (c (x+1)>c x \ nth_prime (x+1) > nth_prime x + a x / (2* B))" using c_neq_large B_nth_prime eq_2_11[of "1-1/ sqrt 2",simplified] by eventually_elim (auto simp:divide_simps) then show ?thesis proof (elim eventually_at_top_mono) fix N assume "N\1" and N_asm:"\y\N. c (y + 1) \ c y \ sqrt (real (nth_prime y)) / real_of_int (a y) < 1 / real_of_int (2 * B) \ h y < y / 2 \ (c y < c (y + 1) \ real (nth_prime y) + real_of_int (a y) / real_of_int (2 * B) < real (nth_prime (y + 1)))" define S where "S={n \ {N..<2 * N}. c n < c (n + 1)}" define g where "g=(\n. (nth_prime (n+1) - nth_prime n) / sqrt (nth_prime n))" define f where "f=(\n. nth_prime (n+1) - nth_prime n)" have g_gt_1:"g n>1" when "n\N" "c n < c (n + 1)" for n proof - have "nth_prime n + sqrt (nth_prime n) < nth_prime (n+1)" proof - have "nth_prime n + sqrt (nth_prime n) < nth_prime n + a n / (2*B)" using N_asm[rule_format,OF \n\N\] a_pos by (auto simp:field_simps) also have "... < nth_prime (n+1)" using N_asm[rule_format,OF \n\N\] \c n < c (n + 1)\ by auto finally show ?thesis . qed then show ?thesis unfolding g_def using \c n < c (n + 1)\ by auto qed have g_geq_0:"g n \ 0" for n unfolding g_def by auto have "finite S" "\x\S. x\N \ c x sum g S" proof (induct S) case empty then show ?case by auto next case (insert x F) moreover have "g x>1" proof - have "c x < c (x+1)" "x\N" using insert(4) by auto then show ?thesis using g_gt_1 by auto qed ultimately show ?case by simp qed also have "... \ sum g {N..<2*N}" apply (rule sum_mono2) unfolding S_def using g_geq_0 by auto also have "... \ sum (\n. f n/sqrt (nth_prime N)) {N..<2*N}" apply (rule sum_mono) unfolding f_def g_def by (auto intro!:divide_left_mono) also have "... = sum f {N..<2*N} / sqrt (nth_prime N)" unfolding sum_divide_distrib[symmetric] by auto also have "... = (nth_prime (2*N) - nth_prime N) / sqrt (nth_prime N)" proof - have "sum f {N..<2 * N} = nth_prime (2 * N) - nth_prime N" proof (induct N) case 0 then show ?case by simp next case (Suc N) have ?case if "N=0" proof - have "sum f {Suc N..<2 * Suc N} = sum f {1}" using that by (simp add: numeral_2_eq_2) also have "... = nth_prime 2 - nth_prime 1" unfolding f_def by (simp add:numeral_2_eq_2) also have "... = nth_prime (2 * Suc N) - nth_prime (Suc N)" using that by auto finally show ?thesis . qed moreover have ?case if "N\0" proof - have "sum f {Suc N..<2 * Suc N} = sum f {N..<2 * Suc N} - f N" apply (subst (2) sum.atLeast_Suc_lessThan) using that by auto also have "... = sum f {N..<2 * N}+ f (2*N) + f(2*N+1) - f N" by auto also have "... = nth_prime (2 * Suc N) - nth_prime (Suc N)" using Suc unfolding f_def by auto finally show ?thesis . qed ultimately show ?case by blast qed then show ?thesis by auto qed also have "... = h N" unfolding h_def by auto also have "... < N/2" using N_asm by auto finally have "card S < N/2" . define T where "T={n \ {N..<2 * N}. c n > c (n + 1)}" have "T \ S = {N..<2 * N}" "T \ S = {}" "finite T" unfolding T_def S_def using N_asm by fastforce+ then have "card T + card S = card {N..<2 * N}" using card_Un_disjoint \finite S\ by metis also have "... = N" by simp finally have "card T + card S = N" . with \card S < N/2\ show "card T > N/2" by linarith qed qed text\Inequality (3.5) in the original paper required a slight modification: \ have a_gt_plus:"\\<^sub>F n in sequentially. c n > c (n+1) \a (n+1) > a n + (a n - c(n+1) - 1) / c (n+1)" proof - note a_gt_1[THEN eventually_all_ge_at_top] c_pos[THEN eventually_all_ge_at_top] moreover have "\\<^sub>F n in sequentially. B * int (nth_prime (n+1)) = c (n+1) * a (n+1) - c (n + 2)" using Bc_large apply (subst (asm) eventually_sequentially_Suc[symmetric]) by (auto elim:eventually_mono) moreover have "\\<^sub>F n in sequentially. B * int (nth_prime n) = c n * a n - c (n + 1) \ \c (n + 1)\ < a n / 2" using Bc_large by (auto elim:eventually_mono) ultimately show ?thesis apply (eventually_elim) proof (rule impI) fix n assume "\y\n. 1 < a y" "\y\n. 0 < c y" and Suc_n_eq:"B * int (nth_prime (n + 1)) = c (n + 1) * a (n + 1) - c (n + 2)" and "B * int (nth_prime n) = c n * a n - c (n + 1) \ real_of_int \c (n + 1)\ < real_of_int (a n) / 2" and "c (n + 1) < c n" then have n_eq:"B * int (nth_prime n) = c n * a n - c (n + 1)" and c_less_a: "real_of_int \c (n + 1)\ < real_of_int (a n) / 2" by auto from \\y\n. 1 < a y\ \\y\n. 0 < c y\ have *:"a n>1" "a (n+1) > 1" "c n > 0" "c (n+1) > 0" "c (n+2) > 0" by auto then have "(1+1/c (n+1))* (a n - 1)/a (n+1) = (c (n+1)+1) * ((a n - 1) / (c (n+1) * a (n+1)))" by (auto simp:field_simps) also have "... \ c n * ((a n - 1) / (c (n+1) * a (n+1)))" apply (rule mult_right_mono) subgoal using \c (n + 1) < c n\ by auto subgoal by (smt \0 < c (n + 1)\ a_pos divide_nonneg_pos mult_pos_pos of_int_0_le_iff of_int_0_less_iff) done also have "... = (c n * (a n - 1)) / (c (n+1) * a (n+1))" by auto also have "... < (c n * (a n - 1)) / (c (n+1) * a (n+1) - c (n+2))" apply (rule divide_strict_left_mono) subgoal using \c (n+2) > 0\ by auto unfolding Suc_n_eq[symmetric] using * \B>0\ by auto also have "... < (c n * a n - c (n+1)) / (c (n+1) * a (n+1) - c (n+2))" apply (rule frac_less) unfolding Suc_n_eq[symmetric] using * \B>0\ \c (n + 1) < c n\ by (auto simp:algebra_simps) also have "... = nth_prime n / nth_prime (n+1)" unfolding Suc_n_eq[symmetric] n_eq[symmetric] using \B>0\ by auto also have "... < 1" by auto finally have "(1 + 1 / real_of_int (c (n + 1))) * real_of_int (a n - 1) / real_of_int (a (n + 1)) < 1 " . then show "a n + (a n - c (n + 1) - 1) / (c (n + 1)) < (a (n + 1))" using * by (auto simp:field_simps) qed qed have a_gt_1:"\\<^sub>F n in sequentially. c n > c (n+1) \ a (n+1) > a n + 1" using Bc_large a_gt_plus c_pos[THEN eventually_all_ge_at_top] apply eventually_elim proof (rule impI) fix n assume "c (n + 1) < c n \ a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" "c (n + 1) < c n" and B_eq:"B * int (nth_prime n) = c n * a n - c (n + 1) \ \real_of_int (c (n + 1))\ < real_of_int (a n) / 2" and c_pos:"\y\n. 0 < c y" from this(1,2) have "a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" by auto moreover have "a n - 2 * c (n+1) > 0" using B_eq c_pos[rule_format,of "n+1"] by auto then have "a n - 2 * c (n+1) \ 1" by simp then have "(a n - c (n + 1) - 1) / c (n + 1) \ 1" using c_pos[rule_format,of "n+1"] by (auto simp:field_simps) ultimately show "a n + 1 < a (n + 1)" by auto qed text\The following corresponds to inequality (3.6) in the paper, which had to be slightly corrected: \ have a_gt_sqrt:"\\<^sub>F n in sequentially. c n > c (n+1) \ a (n+1) > a n + (sqrt n - 2)" proof - have a_2N:"\\<^sub>F N in sequentially. a (2*N) \ N /2 +1" using c_gt_half a_gt_1[THEN eventually_all_ge_at_top] proof eventually_elim case (elim N) define S where "S={n \ {N..<2 * N}. c (n + 1) < c n}" define f where "f = (\n. a (Suc n) - a n)" have f_1:"\x\S. f x\1" and f_0:"\x. f x\0" subgoal using elim unfolding S_def f_def by auto subgoal using \mono a\[THEN incseq_SucD] unfolding f_def by auto done have "N / 2 < card S" using elim unfolding S_def by auto also have "... \ sum f S" unfolding of_int_sum apply (rule sum_bounded_below[of _ 1,simplified]) using f_1 by auto also have "... \ sum f {N..<2 * N}" unfolding of_int_sum apply (rule sum_mono2) unfolding S_def using f_0 by auto also have "... = a (2*N) - a N" unfolding of_int_sum f_def of_int_diff apply (rule sum_Suc_diff') by auto finally have "N / 2 < a (2*N) - a N" . then show ?case using a_pos[rule_format,of N] by linarith qed have a_n4:"\\<^sub>F n in sequentially. a n > n/4" proof - obtain N where a_N:"\n\N. a (2*n) \ n /2+1" using a_2N unfolding eventually_at_top_linorder by auto have "a n>n/4" when "n\2*N" for n proof - define n' where "n'=n div 2" have "n'\N" unfolding n'_def using that by auto have "n/4 < n' /2+1" unfolding n'_def by auto also have "... \ a (2*n')" using a_N \n'\N\ by auto also have "... \a n" unfolding n'_def apply (cases "even n") subgoal by simp subgoal by (simp add: assms(2) incseqD) done finally show ?thesis . qed then show ?thesis unfolding eventually_at_top_linorder by auto qed have c_sqrt:"\\<^sub>F n in sequentially. c n < sqrt n / 4" proof - have "\\<^sub>F x in sequentially. x>1" by simp moreover have "\\<^sub>F x in sequentially. real (nth_prime x) / (real x * ln (real x)) < 2" using nth_prime_asymptotics[unfolded asymp_equiv_def,THEN order_tendstoD(2),of 2] by simp ultimately have "\\<^sub>F n in sequentially. c n < B*8 *ln n + 1" using a_n4 Bc_large proof eventually_elim case (elim n) from this(4) have "c n=(B*nth_prime n+c (n+1))/a n" using a_pos[rule_format,of n] by (auto simp:divide_simps) also have "... = (B*nth_prime n)/a n+c (n+1)/a n" by (auto simp:divide_simps) also have "... < (B*nth_prime n)/a n + 1" proof - have "c (n+1)/a n < 1" using elim(4) by auto then show ?thesis by auto qed also have "... < B*8 * ln n + 1" proof - have "B*nth_prime n < 2*B*n*ln n" using \real (nth_prime n) / (real n * ln (real n)) < 2\ \B>0\ \ 1 < n\ by (auto simp:divide_simps) moreover have "real n / 4 < real_of_int (a n)" by fact ultimately have "(B*nth_prime n) / a n < (2*B*n*ln n) / (n/4)" apply (rule_tac frac_less) using \B>0\ \ 1 < n\ by auto also have "... = B*8 * ln n" using \ 1 < n\ by auto finally show ?thesis by auto qed finally show ?case . qed moreover have "\\<^sub>F n in sequentially. B*8 *ln n + 1 < sqrt n / 4" by real_asymp ultimately show ?thesis by eventually_elim auto qed have "\\<^sub>F n in sequentially. 0 < c (n+1)" "\\<^sub>F n in sequentially. c (n+1) < sqrt (n+1) / 4" "\\<^sub>F n in sequentially. n > 4" "\\<^sub>F n in sequentially. (n - 4) / sqrt (n + 1) + 1 > sqrt n" subgoal using c_pos[THEN eventually_all_ge_at_top] by eventually_elim auto subgoal using c_sqrt[THEN eventually_all_ge_at_top] by eventually_elim (use le_add1 in blast) subgoal by simp subgoal by real_asymp done then show ?thesis using a_gt_plus a_n4 apply eventually_elim proof (rule impI) fix n assume asm:"0 < c (n + 1)" "c (n + 1) < sqrt (real (n + 1)) / 4" and a_ineq:"c (n + 1) < c n \ a n + (a n - c (n + 1) - 1) / c (n + 1) < a (n + 1)" "c (n + 1) < c n" and "n / 4 < a n" "n > 4" and n_neq:" sqrt (real n) < real (n - 4) / sqrt (real (n + 1)) + 1" have "(n-4) / sqrt(n+1) = (n/4 - 1)/ (sqrt (real (n + 1)) / 4)" using \n>4\ by (auto simp:divide_simps) also have "... < (a n - 1) / c (n + 1)" apply (rule frac_less) using \n > 4\ \n / 4 < a n\ \0 < c (n + 1)\ \c (n + 1) < sqrt (real (n + 1)) / 4\ by auto also have "... - 1 = (a n - c (n + 1) - 1) / c (n + 1)" using \0 < c (n + 1)\ by (auto simp:field_simps) also have "a n + ... < a (n+1)" using a_ineq by auto finally have "a n + ((n - 4) / sqrt (n + 1) - 1) < a (n + 1)" by simp moreover have "(n - 4) / sqrt (n + 1) - 1 > sqrt n - 2" using n_neq[THEN diff_strict_right_mono,of 2] \n>4\ by (auto simp:algebra_simps of_nat_diff) ultimately show "real_of_int (a n) + (sqrt (real n) - 2) < real_of_int (a (n + 1))" by argo qed qed text\The following corresponds to inequality $ a_{2N} > N^{3/2}/2$ in the paper, which had to be slightly corrected: \ have a_2N_sqrt:"\\<^sub>F N in sequentially. a (2*N) > real N * (sqrt (real N)/2 - 1)" using c_gt_half a_gt_sqrt[THEN eventually_all_ge_at_top] eventually_gt_at_top[of 4] proof eventually_elim case (elim N) define S where "S={n \ {N..<2 * N}. c (n + 1) < c n}" define f where "f = (\n. a (Suc n) - a n)" have f_N:"\x\S. f x\sqrt N - 2" proof fix x assume "x\S" then have "sqrt (real x) - 2 < f x" "x\N" using elim unfolding S_def f_def by auto moreover have "sqrt x - 2 \ sqrt N - 2" using \x\N\ by simp ultimately show "sqrt (real N) - 2 \ real_of_int (f x)" by argo qed have f_0:"\x. f x\0" using \mono a\[THEN incseq_SucD] unfolding f_def by auto have "(N / 2) * (sqrt N - 2) < card S * (sqrt N - 2)" apply (rule mult_strict_right_mono) subgoal using elim unfolding S_def by auto subgoal using \N>4\ by (metis diff_gt_0_iff_gt numeral_less_real_of_nat_iff real_sqrt_four real_sqrt_less_iff) done also have "... \ sum f S" unfolding of_int_sum apply (rule sum_bounded_below) using f_N by auto also have "... \ sum f {N..<2 * N}" unfolding of_int_sum apply (rule sum_mono2) unfolding S_def using f_0 by auto also have "... = a (2*N) - a N" unfolding of_int_sum f_def of_int_diff apply (rule sum_Suc_diff') by auto finally have "real N / 2 * (sqrt (real N) - 2) < real_of_int (a (2 * N) - a N)" . then have "real N / 2 * (sqrt (real N) - 2) < a (2 * N)" using a_pos[rule_format,of N] by linarith then show ?case by (auto simp:field_simps) qed text\The following part is required to derive the final contradiction of the proof.\ have a_n_sqrt:"\\<^sub>F n in sequentially. a n > (((n-1)/2) powr (3/2) - (n-1)) /2" proof (rule sequentially_even_odd_imp) define f where "f=(\N. ((real (2 * N - 1) / 2) powr (3 / 2) - real (2 * N - 1)) / 2)" define g where "g=(\N. real N * (sqrt (real N) / 2 - 1))" have "\\<^sub>F N in sequentially. g N > f N" unfolding f_def g_def by real_asymp moreover have "\\<^sub>F N in sequentially. a (2 * N) > g N" unfolding g_def using a_2N_sqrt . ultimately show "\\<^sub>F N in sequentially. f N < a (2 * N)" by eventually_elim auto next define f where "f=(\N. ((real (2 * N + 1 - 1) / 2) powr (3 / 2) - real (2 * N + 1 - 1)) / 2)" define g where "g=(\N. real N * (sqrt (real N) / 2 - 1))" have "\\<^sub>F N in sequentially. g N = f N" using eventually_gt_at_top[of 0] apply eventually_elim unfolding f_def g_def by (auto simp:algebra_simps powr_half_sqrt[symmetric] powr_mult_base) moreover have "\\<^sub>F N in sequentially. a (2 * N) > g N" unfolding g_def using a_2N_sqrt . moreover have "\\<^sub>F N in sequentially. a (2 * N + 1) \ a (2*N)" apply (rule eventuallyI) using \mono a\ by (simp add: incseqD) ultimately show "\\<^sub>F N in sequentially. f N < (a (2 * N + 1))" apply eventually_elim by auto qed have a_nth_prime_gt:"\\<^sub>F n in sequentially. a n / nth_prime n > 1" proof - define f where "f=(\n::nat. (((n-1)/2) powr (3/2) - (n-1)) /2)" have "\\<^sub>F x in sequentially. real (nth_prime x) / (real x * ln (real x)) < 2" using nth_prime_asymptotics[unfolded asymp_equiv_def,THEN order_tendstoD(2),of 2] by simp from this[] eventually_gt_at_top[of 1] have "\\<^sub>F n in sequentially. real (nth_prime n) < 2*(real n * ln n)" apply eventually_elim by (auto simp:field_simps) moreover have *:"\\<^sub>F N in sequentially. f N >0 " unfolding f_def by real_asymp moreover have " \\<^sub>F n in sequentially. f n < a n" using a_n_sqrt unfolding f_def . ultimately have "\\<^sub>F n in sequentially. a n / nth_prime n > f n / (2*(real n * ln n))" apply eventually_elim apply (rule frac_less2) by auto moreover have "\\<^sub>F n in sequentially. (f n)/ (2*(real n * ln n)) > 1" unfolding f_def by real_asymp ultimately show ?thesis by eventually_elim argo qed have a_nth_prime_lt:"\\<^sub>F n in sequentially. a n / nth_prime n < 1" proof - have "liminf (\x. a x / nth_prime x) < 1" using nth_2 by auto from this[unfolded less_Liminf_iff] show ?thesis apply (auto elim!:frequently_elim1) by (meson divide_less_eq_1 ereal_less_eq(7) leD leI nth_prime_nonzero of_nat_eq_0_iff of_nat_less_0_iff order.trans) qed from a_nth_prime_gt a_nth_prime_lt show False by (simp add: eventually_mono frequently_def) qed section\Acknowledgements\ text\A.K.-A. and W.L. were supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council and led by Professor Lawrence Paulson at the University of Cambridge, UK.\ end \ No newline at end of file diff --git a/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy b/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy --- a/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy +++ b/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy @@ -1,1087 +1,1087 @@ (* File: Irrationality_J_Hancl.thy Authors: Angeliki Koutsoukou-Argyraki and Wenda Li, Computer Laboratory, University of Cambridge, UK. Email: ak2110@cam.ac.uk, wl302@cam.ac.uk *) section \Irrational Rapidly Convergent Series\ theory Irrationality_J_Hancl imports "HOL-Analysis.Analysis" "HOL-Decision_Procs.Approximation" begin text \This is the formalisation of a proof by J. Hancl, in particular of the proof of his Theorem 3 in the paper: Irrational Rapidly Convergent Series, Rend. Sem. Mat. Univ. Padova, Vol 107 (2002). The statement asserts the irrationality of the sum of a series consisting of rational numbers defined using sequences that fulfill certain properties. Even though the statement is number-theoretic, the proof uses only arguments from introductory Analysis.\ text \We show the central result (theorem Hancl3) by proof by contradiction, by making use of some of the auxiliary lemmas. To this end, assuming that the sum is a rational number, for a quantity $\textrm{ALPHA}(n)$ we show that $\textrm{ALPHA}(n) \geq 1$ for all $n \in \mathbb{N}$. After that we show that we can find an $n \in \mathbb{N}$ for which $\textrm{ALPHA}(n) < 1$ which yields a contradiction and we thus conclude that the sum of the series is a rational number. We finally give an immediate application of theorem Hancl3 for a specific series (corollary Hancl3corollary, requiring lemma summable\_ln\_plus) which corresponds to Corollary 2 in the original paper by J. Hancl. \ hide_const floatarith.Max subsection \Misc\ lemma filterlim_sequentially_iff: "filterlim f F1 sequentially \ filterlim (\x. f (x+k)) F1 sequentially" unfolding filterlim_iff apply (subst eventually_sequentially_seg) by auto lemma filterlim_realpow_sequentially_at_top: "(x::real) > 1 \ filterlim ((^) x) at_top sequentially" apply (rule LIMSEQ_divide_realpow_zero[THEN filterlim_inverse_at_top,of _ 1,simplified]) by auto lemma filterlim_at_top_powr_real: fixes g::"'b \ real" assumes "filterlim f at_top F" "(g \ g') F" "g'>1" shows "filterlim (\x. g x powr f x) at_top F" proof - have "filterlim (\x. ((g'+1)/2) powr f x) at_top F" proof (subst filterlim_at_top_gt[of _ _ 1],rule+) fix Z assume "Z>(1::real)" have "\\<^sub>F x in F. ln Z \ ln (((g' + 1) / 2) powr f x)" using assms(1)[unfolded filterlim_at_top,rule_format,of "ln Z / ln ((g' + 1) / 2)"] apply (eventually_elim) using \g'>1\ by (auto simp:ln_powr divide_simps) then show "\\<^sub>F x in F. Z \ ((g' + 1) / 2) powr f x" apply (eventually_elim) apply (subst (asm) ln_le_cancel_iff) using \Z>1\ \g'>1\ by auto qed moreover have "\\<^sub>F x in F. ((g'+1)/2) powr f x \ g x powr f x" proof - have "\\<^sub>F x in F. g x > (g'+1)/2" apply (rule order_tendstoD[OF assms(2)]) using \g'>1\ by auto moreover have "\\<^sub>F x in F. f x>0" using assms(1)[unfolded filterlim_at_top_dense,rule_format,of 0] . ultimately show ?thesis apply eventually_elim using \g'>1\ by (auto intro!: powr_mono2) qed ultimately show ?thesis using filterlim_at_top_mono by fast qed lemma powrfinitesum: fixes a::real and s::nat assumes "s \ n" shows " (\j=s..(n::nat).(a powr (2^j))) = a powr (\j=s..(n::nat).(2^j)) " using \s \ n\ proof(induct n) case 0 then show ?case by auto next case (Suc n) have ?case when "s\n" using Suc.hyps by (metis Suc.prems add.commute linorder_not_le powr_add prod.nat_ivl_Suc' sum.cl_ivl_Suc that) moreover have ?case when "s=Suc n" proof- have "(\j = s..Suc n. a powr 2 ^ j) =(a powr 2 ^(Suc n)) " using \s=Suc n\ by simp also have "(a powr 2 ^(Suc n) ) = a powr sum ((^) 2) {s..Suc n} " using that by auto ultimately show " (\j = s..Suc n. a powr 2 ^ j) = a powr sum ((^) 2) {s..Suc n}" using \s\Suc n\ by linarith qed ultimately show ?case using \s\Suc n\ by linarith qed lemma summable_ratio_test_tendsto: fixes f :: "nat \ 'a::banach" assumes "c < 1" and "\n. f n\0" and "(\n. norm (f (Suc n)) / norm (f n)) \ c" shows "summable f" proof - obtain N where N_dist:"\n\N. dist (norm (f (Suc n)) / norm (f n)) c < (1-c)/2" using assms unfolding tendsto_iff eventually_sequentially by (meson diff_gt_0_iff_gt zero_less_divide_iff zero_less_numeral) have "norm (f (Suc n)) / norm (f n) \ (1+c)/2" when "n\N" for n using N_dist[rule_format,OF that] \c<1\ apply (auto simp add:field_simps dist_norm) by argo then have "norm (f (Suc n)) \ (1+c)/2 * norm (f n)" when "n\N" for n using that assms(2)[rule_format, of n] by (auto simp add:divide_simps) moreover have "(1+c)/2 < 1" using \c<1\ by auto ultimately show ?thesis using summable_ratio_test[of _ N f] by blast qed lemma summable_ln_plus: fixes f::"nat \ real" assumes "summable f" "\n. f n>0" shows "summable (\n. ln (1+f n))" proof (rule summable_comparison_test_ev[OF _ assms(1)]) have "ln (1 + f n) > 0" for n by (simp add: assms(2) ln_gt_zero) moreover have "ln (1 + f n) \ f n" for n apply (rule ln_add_one_self_le_self2) using assms(2)[rule_format,of n] by auto ultimately show "\\<^sub>F n in sequentially. norm (ln (1 + f n)) \ f n" by (auto intro!: eventuallyI simp add:less_imp_le) qed lemma suminf_real_offset_le: fixes f :: "nat \ real" assumes f: "\i. 0 \ f i" and "summable f" shows "(\i. f (i + k)) \ suminf f" proof - have "(\n. \i (\i. f (i + k))" using summable_sums[OF \summable f\] by (simp add: assms(2) summable_LIMSEQ summable_ignore_initial_segment) moreover have "(\n. \i (\i. f i)" using summable_sums[OF \summable f\] by (simp add: sums_def atLeast0LessThan f) then have "(\n. \i (\i. f i)" by (rule LIMSEQ_ignore_initial_segment) ultimately show ?thesis proof (rule LIMSEQ_le, safe intro!: exI[of _ k]) fix n assume "k \ n" have "(\ii (\i. i + k)) i)" by simp also have "\ = (\i\(\i. i + k) ` {.. \ sum f {..i sum f {.. n" shows " (\i=s..n. 2^i) < (2^(n+1) :: real) " using assms proof (induct n) case 0 show ?case by auto next case (Suc n) have ?case when "s=n+1" using that by auto moreover have ?case when "s \ n+1" proof - have"(\i=s..(n+1). 2^i ) = (\i=s..n. 2^i ) + (2::real)^(n+1) " using sum.cl_ivl_Suc \s \ Suc n \ by (auto simp add:add.commute) also have "... < (2) ^ (n +1) + (2)^(n+1)" proof - have "s \n" using that \s \ Suc n \ by auto then show ?thesis using Suc.hyps \s \ n\ by linarith qed also have "... = 2^(n+2)" by simp finally show "(\i=s..(Suc n). 2^i )< (2::real)^(Suc n+1)" by auto qed ultimately show ?case by blast qed subsection \Auxiliary lemmas and the main proof\ lemma showpre7: fixes a b ::"nat\int " and q p::int assumes "q>0" and "p>0"and a: "\n. a n>0" and "\n. b n>0" and assumerational:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" shows "q*((\j=1..n. of_int( a j)))*(suminf (\(j::nat). (b (j+n+1)/ a (j+n+1 )))) = ((\j=1..n. of_int( a j)))*(p -q* (\j=1..n. b j / a j)) " proof - define aa where "aa=(\j = 1..n. real_of_int (a j))" define ff where "ff=(\i. real_of_int (b (i+1)) / real_of_int (a (i+1)))" have "(\j. ff (j+n)) = (\n. ff n) - sum ff {..j=1..n. ff (j-1))" proof - have "sum ff {..j=1..n. ff (j-1))" apply (subst sum_bounds_lt_plus1[symmetric]) by simp then show ?thesis unfolding ff_def by auto qed finally have "(\j. ff (j + n)) = p / q - (\j = 1..n. ff (j - 1))" . then have "q*(\j. ff (j + n)) = p - q*(\j = 1..n. ff (j - 1))" using \q>0\ by (auto simp add:field_simps) then have "aa*q*(\j. ff (j + n)) = aa*(p - q*(\j = 1..n. ff (j - 1)))" by auto then show ?thesis unfolding aa_def ff_def by auto qed lemma show7: fixes d::"nat\real" and a b::"nat\int " and q p::int assumes "q \1" and "p \ 1" and a: "\n. a n \ 1" and "\n. b n \ 1" and assumerational:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" shows "q*((\j=1..n. of_int( a j)))*( suminf (\ (j::nat). (b (j+n+1)/ a (j+n+1 )))) \ 1 " (is "?L \ _") proof- define LL where "LL=?L" define aa where "aa=(\j = 1..n. real_of_int (a j))" define ff where "ff=(\i. real_of_int (b (i+1)) / real_of_int (a (i+1)))" have "?L > 0" proof - have "aa > 0" unfolding aa_def using a apply (induct n,auto) by (simp add: int_one_le_iff_zero_less prod_pos) moreover have "(\j. ff (j + n)) > 0" proof (rule suminf_pos) have "summable ff" unfolding ff_def using assumerational using summable_def by blast then show " summable (\j. ff (j + n))" using summable_iff_shift[of ff n] by auto - show " \i. 0 < ff (i + n)" unfolding ff_def using a assms(4) int_one_le_iff_zero_less by auto + show "\i. 0 < ff (i + n)" unfolding ff_def using a assms(4) int_one_le_iff_zero_less by auto qed ultimately show ?thesis unfolding aa_def ff_def using \q\1\ by auto qed moreover have "?L \ \" proof - have "?L = aa *(p -q* ( \j=1..n. b j / a j))" unfolding aa_def apply (rule showpre7) using assms int_one_le_iff_zero_less by auto also have "... = aa * p - q * (\j=1..n. aa * b j / a j)" by (auto simp add:algebra_simps sum_distrib_left) also have "... = prod a {1..n} * p - q * (\j = 1..n. b j * prod a ({1..n} - {j}))" proof - have "(\j=1..n. aa * b j / a j) = (\j=1..n. b j * prod a ({1..n} - {j}))" unfolding of_int_sum proof (rule sum.cong) fix x assume "x \ {1..n}" then have "(\i = 1..n. real_of_int (a i)) = a x * (\i\{1..n} - {x}. real_of_int (a i))" apply (rule_tac prod.remove) by auto then have "aa / real_of_int (a x) = prod a ({1..n} - {x})" unfolding aa_def using a[rule_format,of x] by (auto simp add:field_simps) then show "aa * b x / a x = b x * prod a ({1..n} - {x})" by (metis mult.commute of_int_mult times_divide_eq_right) qed simp moreover have "aa * p = (\j = 1..n. (a j)) * p" unfolding aa_def by auto ultimately show ?thesis by force qed also have "... \ \" using Ints_of_int by blast finally show ?thesis . qed ultimately show ?thesis apply (fold LL_def) by (metis Ints_cases int_one_le_iff_zero_less not_less of_int_0_less_iff of_int_less_1_iff) qed lemma show8: fixes d ::"nat\real " and a :: "nat\int" and s k::nat assumes "A > 1" and d: "\n. d n> 1" and a:"\n. a n>0" and "s>0" and "convergent_prod d" and assu2: "\n \ s. ( A/ (of_int (a n)) powr(1/of_int (2^n)))> prodinf (\j. d(n +j))" shows "\n\s. prodinf (\j. d(j+n)) < A/(Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n}))" proof(rule,rule) fix n assume "s \ n" define sp where "sp = (\n. prodinf (\j. d(j+n)))" define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" have "sp i \ sp n" when "i\n" for i proof - have "(\j. d (j + i)) = (\ia. d (ia + (n - i) + i)) * (\iaconvergent_prod d\ convergent_prod_iff_shift[of d i] by simp subgoal for j using d[rule_format,of "j+i"] by auto done then have "sp i = sp n * (\jn\i\ by (auto simp:algebra_simps) moreover have "sp i>1" "sp n>1" unfolding sp_def using convergent_prod_iff_shift \convergent_prod d\ d by (auto intro!:less_1_prodinf) moreover have "(\j1" apply (rule prod_ge_1) using d less_imp_le by auto ultimately show ?thesis by auto qed moreover have "\j\s. A / ff j > sp j" unfolding ff_def sp_def using assu2 by (auto simp:algebra_simps) ultimately have "\j. s\j \ j\n \ A / ff j > sp n" by force then show "sp n< A / Max (ff ` {s..n})" by (metis (mono_tags, hide_lams) Max_in \n\s\ atLeastAtMost_iff empty_iff finite_atLeastAtMost finite_imageI imageE image_is_empty order_refl) qed lemma auxiliary1_9: fixes d ::"nat\real" and a::"nat\int " and s m::nat assumes d: "\n. d n> 1" and a: "\n. a n>0" and "s>0" and "n \ m" and " m \ s" and auxifalse_assu: "\n\m. (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) < (d (n+1))* (Max ((\ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} ))" shows "(of_int (a (n+1))) powr(1 /of_int (2^(n+1))) < (\j=(m+1)..(n+1). d j) * (Max ((\ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..m}))" proof- define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" have [simp]:"ff j > 0" for j unfolding ff_def by (metis a less_numeral_extra(3) of_int_0_less_iff powr_gt_zero) have ff_asm:"ff (n+1) < d (n+1) * Max (ff ` {s..n})" when "n\m" for n using auxifalse_assu that unfolding ff_def by simp from \n\m\ have Q: "(Max( ff ` {s..n} )) \ (\j=(m+1)..n. d j)* (Max (ff ` {s..m}))" proof(induct n) case 0 then show ?case using \m\s\ by simp next case (Suc n) have ?case when "m=Suc n" using that by auto moreover have ?case when "m\Suc n" proof - have "m \ n" using that Suc(2) by simp then have IH:"Max (ff ` {s..n}) \ prod d {m + 1..n} * Max (ff ` {s..m})" using Suc(1) by linarith have "Max (ff ` {s..Suc n}) = Max (ff ` {s..n} \ {ff (Suc n)})" using Suc.prems assms(5) atLeastAtMostSuc_conv by auto also have "... = max (Max (ff ` {s..n})) (ff (Suc n))" using Suc.prems assms(5) max_def sup_assoc that by auto also have "... \ max (Max (ff ` {s..n})) (d (n+1) * Max (ff ` {s..n}))" apply (rule max.mono) using ff_asm[of n] \ m \ Suc n\ that \s\m\ by auto also have "... \ Max (ff ` {s..n}) * max 1 (d (n+1))" proof - have "Max (ff ` {s..n}) \0" by (metis (mono_tags, hide_lams) Max_in \\j. 0 < ff j\ \m \ n\ assms(5) atLeastAtMost_iff empty_iff finite_atLeastAtMost finite_imageI imageE image_is_empty less_eq_real_def) then show ?thesis using max_mult_distrib_right by (simp add: max_mult_distrib_right mult.commute) qed also have "... = Max (ff ` {s..n}) * d (n+1)" using d[rule_format, of "n+1"] by auto also have "... \ prod d {m + 1..n} * Max (ff ` {s..m}) * d (n+1)" using IH d[rule_format, of "n+1"] by auto also have "... = prod d {m + 1..n+1} * Max (ff ` {s..m})" using \n\m\ by (simp add:prod.nat_ivl_Suc' algebra_simps) finally show ?case by simp qed ultimately show ?case by blast qed then have "d (n+1) * Max (ff ` {s..n} ) \ (\j=(m+1)..(n+1). d j)* (Max (ff ` {s..m}))" using \m\n\ d[rule_format,of "Suc n"] by (simp add:prod.nat_ivl_Suc') then show ?thesis using ff_asm[of n] \s\m\ \m\n\ unfolding ff_def by auto qed lemma show9: fixes d ::"nat\real " and a :: "nat\int " and s ::nat and A::real assumes "A > 1" and d: "\n. d n> 1" and a: "\n. a n>0" and "s>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and "convergent_prod d" and 8: "\n\s. prodinf (\j. d( n+j)) < A/(Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n})) " shows "\m \s. \n\m. ( (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ( ( \ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" proof (rule ccontr) define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" assume assumptioncontra: " \ (\m \s. \n\m. ( (ff (n+1)) \ (d (n+1))* (Max ( ff ` {s..n}))))" then obtain t where "t\s" and ttt: " \n\t. ( (ff (n+1)) < (d (n+1))* (Max ( ff ` {s..n} ) ))" by fastforce define B where "B=prodinf (\j. d(t+1+j))" have "B>0" unfolding B_def apply (rule less_0_prodinf) subgoal using convergent_prod_iff_shift[of d "t+1"] \convergent_prod d\ by (auto simp:algebra_simps) subgoal using d le_less_trans zero_le_one by blast done have "A \ B * Max ( ff ` {s..t})" proof (rule tendsto_le[of sequentially "\n. (\j=(t+1)..(n+1). d j) * Max ( ff ` {s..t})" _ "\n. ff (n+1)"]) show "(\n. ff (n + 1)) \ A" using assu1[folded ff_def] LIMSEQ_ignore_initial_segment by blast have "(\n. prod d {t + 1..n + 1}) \ B" proof - have "(\n. \i\n. d (t + 1 + i)) \ B" apply (rule convergent_prod_LIMSEQ[of "(\j. d(t+1+j))",folded B_def]) using \convergent_prod d\ convergent_prod_iff_shift[of d "t+1"] by (simp add:algebra_simps) then have "(\n. \i\{0..n}. d (i+(t + 1))) \ B" using atLeast0AtMost by (auto simp:algebra_simps) then have "(\n. prod d {(t + 1)..n + (t + 1)}) \ B" apply (subst (asm) prod.shift_bounds_cl_nat_ivl[symmetric]) by simp from seq_offset_neg[OF this,of "t"] show "(\n. prod d {t + 1..n+1}) \ B" apply (elim Lim_transform) apply (rule LIMSEQ_I) apply (rule exI[where x="t+1"]) by auto qed then show "(\n. prod d {t + 1..n + 1} * Max (ff ` {s..t})) \ B * Max (ff ` {s..t})" by (auto intro:tendsto_eq_intros) have "\\<^sub>F n in sequentially. (ff (n+1)) < (\j=(t+1)..(n+1). d j) * (Max ( ff ` {s..t}))" unfolding eventually_sequentially ff_def using auxiliary1_9[OF d a \s>0\ _ \t\s\ ttt[unfolded ff_def]] by blast then show "\\<^sub>F n in sequentially. (ff (n+1)) \ (\j=(t+1)..(n+1). d j) * (Max ( ff ` {s..t}))" by (eventually_elim,simp) qed simp also have "... \ B * Max ( ff ` {s..t+1})" proof - have "Max (ff ` {s..t}) \ Max (ff ` {s..t + 1})" apply (rule Max_mono) using \t\s\ by auto then show ?thesis using \B>0\ by auto qed finally have "A \ B * Max (ff ` {s..t + 1})" unfolding B_def . moreover have "B < A / Max (ff ` {s..t + 1})" using 8[rule_format, of "t+1",folded ff_def B_def] \s\t\ by auto moreover have "Max (ff ` {s..t+1})>0" using \A \ B * Max (ff ` {s..t + 1})\ \B>0\ \A>1\ zero_less_mult_pos [of B "Max (ff ` {s..Suc t})"] by simp ultimately show False by (auto simp add:field_simps) qed lemma show10: fixes d ::"nat\real " and a ::"nat\int " and s::nat assumes d: "\n. d n> 1" and a: "\n. a n>0" and " s>0" and 9: "\m \s. \n\m. ((of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" shows "\m\s. \n\m. (((d (n+1))powr(2^(n+1))) * (\j=1..n. of_int( a j)) * (1/ (\j=1..s-1. (of_int( a j) )))) \ (a (n+1)) " proof (rule,rule) fix m assume "s \ m" from 9[rule_format,OF this] obtain n where "n\m" and asm_9:"( (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ( ( \ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" by auto with \s\m\ have "s\n" by auto have prod:"(\j=1..n. real_of_int( a j)) * ( 1/ (\j=1..s-1. (of_int( a j) ))) = (\j=s..n. of_int( a j))" proof - define f where "f= (\j. real_of_int( a j))" have "{Suc 0..n}= {Suc 0..s - Suc 0} \ {s..n}" using \n\s\ \s >0\ by auto then have "(\j=1..n. f j) = (\j=1..s-1. f j) * (\j=s..n. f j)" apply (subst prod.union_disjoint[symmetric]) by auto moreover have "(\j=1..s-1. f j) > 0 " apply (rule linordered_semidom_class.prod_pos) using a unfolding f_def by auto then have "(\j=1..s-1. f j) \ 0" by argo ultimately show ?thesis unfolding f_def by auto qed then have " (((d (n+1))powr(2^(n+1) )) * (\j=1..n. of_int( a j)) * ( 1/ (\j=1..s-1. (of_int( a j) )))) =(((d (n+1))powr(2^(n+1) )) * (\j=s..n. of_int( a j)))" proof - define f where "f= (\j. real_of_int( a j))" define c where "c = (d (n+1))powr(2^(n+1))" show ?thesis using prod apply (fold f_def c_def) by (metis mult.assoc) qed also have "... \ ((d (n+1))powr(2^(n+1) ) * (\i=s..n. (Max(( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )) powr(2^i)) )" proof (rule mult_left_mono) show "0 \ (d (n + 1)) powr 2 ^ (n + 1)" by auto show "(\j = s..n. real_of_int (a j)) \ (\i = s..n. Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i)" proof (rule prod_mono) fix i assume i: "i \ {s..n}" have "real_of_int (a i) = (real_of_int (a i) powr (1 / real_of_int (2 ^ i))) powr 2 ^ i" unfolding powr_powr by (simp add: a less_eq_real_def) also have "\ \ (Max(( \ (j::nat). ( real_of_int( a j) powr(1 /real_of_int (2^j)))) ` {s..n}))powr(2^i)" proof (rule powr_mono2) show "real_of_int (a i) powr (1 / real_of_int (2 ^ i)) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n})" apply (rule Max_ge) apply auto using i by blast qed simp_all finally have "real_of_int (a i) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i" . then show "0 \ real_of_int (a i) \ real_of_int (a i) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i" using a i by (metis \real_of_int (a i) = (real_of_int (a i) powr (1 / real_of_int (2 ^ i))) powr 2 ^ i\ powr_ge_pzero) qed qed also have "... = ((d (n+1))powr(2^(n+1) )) * (Max(( \ (j::nat). ( of_int( a j) powr (1 /of_int (2^j)) )) ` {s..n } )) powr (\i=s..n. 2^i ) " proof- have "((d (n+1))powr(2^(n+1) ))\1 " by (metis Transcendental.log_one d le_powr_iff zero_le_numeral zero_le_power zero_less_one) moreover have "(\i=s..n. (Max(( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } ) ) powr(2^i)) = (Max(( \ (j::nat). ( of_int( a j) powr(1 /of_int (2^j)) )) ` {s..n } )) powr (\i=s..n. 2^i ) " proof - define ff where "ff = Max (( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )" show ?thesis apply (fold ff_def) using \s\n\ powrfinitesum by auto qed ultimately show ?thesis by auto qed also have "... \ ((d (n+1))powr(2^(n+1) )) * (Max(( \ (j::nat).( of_int( a j) powr(1 /of_int (2^j)) )) ` {s..n })) powr(2^(n+1)) " proof - define ff where "ff = Max (( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )" have " sum ((^) 2) {s..n} < (2::real) ^ (n + 1)" using factt \s\n\ by auto moreover have "1 \ ff" proof - define S where "S=(\(j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n }" have "finite S" unfolding S_def by auto moreover have "S\{}" unfolding S_def using \s\n\ by auto moreover have "\x\S. x\1" proof- have " ( of_int( a s) powr(1 /real_of_int (2^s)) ) \ 1 " apply (rule ge_one_powr_ge_zero) apply auto by (simp add: a int_one_le_iff_zero_less) moreover have " ( of_int( a s) powr(1 /real_of_int (2^s)) ) \ S" unfolding S_def apply (rule rev_image_eqI[where x=s]) using \s\n\ by auto ultimately show ?thesis by auto qed ultimately show ?thesis using Max_ge_iff[of S 1] unfolding S_def ff_def by blast qed moreover have "0 \ (d (n + 1)) powr 2 ^ (n + 1)" by auto ultimately show ?thesis apply(fold ff_def) apply (rule mult_left_mono) apply (rule powr_mono) by auto qed also have "... = ((d (n+1)) * (Max((\j. (of_int( a j) powr(1 /of_int (2^j)))) ` {s..n}))) powr(2^(n+1))" proof - define ss where "ss = (\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}" have "d (n + 1) \0" using d[rule_format,of "n+1"] by argo moreover have "Max ss \0" proof - have "(a s) powr (1 / (2 ^ s)) \ 0" by auto moreover have "(a s) powr (1 / (2 ^ s)) \ ss" unfolding ss_def apply (rule_tac x=s in rev_image_eqI) using \s\n\ by auto moreover have "finite ss" "ss \ {}" unfolding ss_def using \s\n\ by auto ultimately show ?thesis using Max_ge_iff[of ss 0] by blast qed ultimately show ?thesis apply (fold ss_def) using powr_mult by auto qed also have "... \ (real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))) powr 2 ^ (n + 1)" proof - define ss where "ss = (\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}" show ?thesis proof (fold ss_def,rule powr_mono2) have "Max ss \0" \ \NOTE: we are repeating the same proof, so it may be a good idea to put this conclusion in an outer block so that it can be reused (without reproving).\ proof - have "(a s) powr (1 / (2 ^ s)) \ 0" by auto moreover have "(a s) powr (1 / (2 ^ s)) \ ss" unfolding ss_def apply (rule_tac x=s in rev_image_eqI) using \s\n\ by auto moreover have "finite ss" "ss \ {}" unfolding ss_def using \s\n\ by auto ultimately show ?thesis using Max_ge_iff[of ss 0] by blast qed moreover have "d (n + 1) \0" using d[rule_format,of "n+1"] by argo ultimately show "0 \ (d (n + 1)) * Max ss" by auto show "(d (n + 1)) * Max ss \ real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))" using asm_9[folded ss_def] . qed simp qed also have "... = (of_int (a (n+1)))" by (simp add: a less_eq_real_def pos_add_strict powr_powr) finally show "\n\m. d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. real_of_int (a j)) * (1 / (\j = 1..s - 1. real_of_int (a j))) \ real_of_int (a (n + 1))" using \n\m\ \m\s\ apply (rule_tac x=n in exI) by auto qed lemma lasttoshow: fixes d ::"nat\real " and a b ::"nat\int " and q::int and s::nat assumes d: "\n. d n> 1" and a:"\n. a n>0" and "s>0" and "q>0" and "A > 1" and b:"\n. b n>0" and 9: "\m\s. \n\m. ((of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" and assu3: "filterlim( \ n. (d n)^(2^n)/ b n) at_top sequentially " and 5: "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" shows "\n. q*((\j=1..n. real_of_int(a j))) * suminf (\(j::nat). (b (j+n+1)/ a (j+n+1)))<1" proof- define as where "as=(\j = 1..s - 1. real_of_int (a j))" obtain n where "n\s" and n_def1:"real_of_int q * as * 2 * real_of_int (b (n + 1)) / d (n + 1) powr 2 ^ (n + 1) < 1" and n_def2:"d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. real_of_int (a j)) * (1 / as) \ real_of_int (a (n + 1))" and n_def3:"(\j. (b (n+1 + j)) / (a (n+1 + j))) \ 2 * b (n+1) / a (n+1)" proof - have *:"(\n. real_of_int (b n) / d n ^ 2 ^ n) \ 0" using tendsto_inverse_0_at_top[OF assu3] by auto then have "(\n. real_of_int (b n) / d n powr 2 ^ n) \ 0" proof - have "d n ^ 2 ^ n = d n powr (of_nat (2 ^ n))" for n apply (subst powr_realpow) using d[rule_format, of n] by auto then show ?thesis using * by auto qed from tendsto_mult_right_zero[OF this,of "q * as * 2"] have "(\n. q * as * 2 * b n / d n powr 2 ^ n) \ 0" by auto then have "\\<^sub>F n in sequentially. q * as * 2 * b n / d n powr 2 ^ n < 1" by (elim order_tendstoD) simp then have "\\<^sub>F n in sequentially. q * as * 2 * b n / d n powr 2 ^ n < 1 \ (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" using 5 by eventually_elim auto then obtain N where N_def:"\n\N. q * as * 2 * b n / d n powr 2 ^ n < 1 \ (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" unfolding eventually_sequentially by auto obtain n where "n\s" and "n\N" and n_def:"d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. of_int (a j)) * (1 / as) \ real_of_int (a (n + 1))" using show10[OF d a \s>0\ 9,folded as_def,rule_format,of "max s N"] by auto with N_def[rule_format, of "n+1"] that[of n] show ?thesis by auto qed define pa where "pa = (\j = 1..n. real_of_int (a j))" define dn where "dn = d (n + 1) powr 2 ^ (n + 1)" have [simp]:"dn >0" "as > 0" subgoal unfolding dn_def by (metis d not_le numeral_One powr_gt_zero zero_le_numeral) subgoal unfolding as_def by (simp add: a prod_pos) done have [simp]:"pa>0" unfolding pa_def using a by (simp add: prod_pos) have K: "q* (\j=1..n. real_of_int (a j)) * suminf (\ (j::nat). (b (j+n+1)/ a (j+n+1))) \q* (\j=1..n. real_of_int (a j)) *2* (b (n+1))/(a( n+1))" apply (fold pa_def) using mult_left_mono[OF n_def3,of "real_of_int q * pa"] \n\s\ \pa>0\ \q>0\ by (auto simp add:algebra_simps) also have KK:"... \ 2*q* (\j=1..n. real_of_int (a j)) * (b(n+1))* ((\j=1..s-1. real_of_int( a j))/((d (n+1))powr(2^(n+1)) * (\j=1..n. real_of_int ( a j))))" proof - have " dn * pa * (1 / as) \ real_of_int (a (n + 1))" using n_def2 unfolding dn_def pa_def . then show ?thesis apply (fold pa_def dn_def as_def) using \pa>0\ \q>0\ a[rule_format, of "Suc n"] b[rule_format, of "Suc n"] by (auto simp add:divide_simps algebra_simps) qed also have KKK: "... = (q* ((\j=1..(s-1). real_of_int( a j)))*2 * (b (n+1)))/ ((d (n+1))powr(2^(n+1)))" apply (fold as_def pa_def dn_def) apply simp using \0 < pa\ by blast also have KKKK: "... < 1" using n_def1 unfolding as_def by simp finally show ?thesis by auto qed lemma fixes d ::"nat\real " and a b ::"nat\int " and A::real assumes "A > 1" and d: "\n. d n> 1" and a: "\n. a n>0" and b:"\n. b n>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and assu3: "filterlim ( \ n. (d n)^(2^n)/ b n) at_top sequentially" and "convergent_prod d" shows issummable: "summable (\j. b j / a j)" and show5: "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" proof- define c where "c = (\j. b j / a j)" have c_pos:"c j>0" for j using a b unfolding c_def by simp have c_ratio_tendsto:"(\n. c (n+1) / c n ) \ 0" proof - define nn where "nn=(\n. (2::int)^ (Suc n))" define ff where "ff=(\ n. (a n / a (Suc n)) powr(1 /nn n)*(d(Suc n)))" have nn_pos:"nn n>0" and ff_pos:"ff n>0" for n subgoal unfolding nn_def by simp subgoal unfolding ff_def using d[rule_format, of "Suc n"] a[rule_format, of n] a[rule_format, of "Suc n"] by auto done have ff_tendsto:"ff \ 1/ sqrt A" proof - have "(of_int (a n)) powr(1 / (nn n)) = sqrt(of_int (a n) powr(1 /of_int (2^n)))" for n unfolding nn_def using a by (simp add: powr_half_sqrt [symmetric] powr_powr ac_simps) moreover have "(( \ n. sqrt(of_int (a n) powr(1 /of_int (2^n))))\ sqrt A) sequentially " using assu1 tendsto_real_sqrt by blast ultimately have "(( \ n. (of_int (a n)) powr(1 /of_int (nn n)))\ sqrt A) sequentially " by auto from tendsto_divide[OF this assu1[THEN LIMSEQ_ignore_initial_segment[where k=1]]] have "(\n. (a n / a (Suc n)) powr(1 /nn n)) \ 1/sqrt A" using \A>1\ a unfolding nn_def by (auto simp add:powr_divide less_imp_le inverse_eq_divide sqrt_divide_self_eq) moreover have "(\n. d (Suc n))\ 1" apply (rule convergent_prod_imp_LIMSEQ) using convergent_prod_iff_shift[of d 1] \convergent_prod d\ by auto ultimately show ?thesis unfolding ff_def by (auto intro:tendsto_eq_intros) qed have "(\n. (ff n) powr nn n) \ 0" proof - define aa where "aa=(1+1/sqrt A) / 2" have "eventually (\n. ff nA>1\ by (auto simp add:field_simps) moreover have "(\n. aa powr nn n) \ 0" proof - have "(\y. aa ^ (nat \ nn) y) \ 0" apply (rule tendsto_power_zero) subgoal unfolding nn_def comp_def apply (rule filterlim_subseq) by (auto intro:strict_monoI) subgoal unfolding aa_def using \A>1\ by auto done then show ?thesis proof (elim filterlim_mono_eventually) have "aa>0" unfolding aa_def using \A>1\ by (auto simp add:field_simps pos_add_strict) then show " \\<^sub>F x in sequentially. aa ^ (nat \ nn) x = aa powr real_of_int (nn x)" by (auto simp: powr_int order.strict_implies_order[OF nn_pos]) qed auto qed ultimately show ?thesis apply (elim metric_tendsto_imp_tendsto) apply (auto intro!:powr_mono2 elim!:eventually_mono) using nn_pos ff_pos by (meson le_cases not_le)+ qed then have "(\n. (d (Suc n))^(nat (nn n)) * (a n / a (Suc n))) \ 0" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in sequentially. ff x powr (nn x) = d (Suc x) ^ nat (nn x) * (a x / a (Suc x))" apply (rule eventuallyI) subgoal for x unfolding ff_def using a[rule_format,of x] a[rule_format,of "Suc x"] d[rule_format,of "Suc x"] nn_pos[of x] apply (auto simp add:field_simps powr_divide powr_powr powr_mult ) by (simp add: powr_int) done qed auto moreover have "(\n. b (Suc n) / (d (Suc n))^(nat (nn n))) \ 0" using tendsto_inverse_0_at_top[OF assu3,THEN LIMSEQ_ignore_initial_segment[where k=1]] unfolding nn_def by (auto simp add:field_simps nat_mult_distrib nat_power_eq) ultimately have "(\n. b (Suc n) * (a n / a (Suc n))) \ 0" apply - subgoal premises asm using tendsto_mult[OF asm,simplified] apply (elim filterlim_mono_eventually) using d by (auto simp add:algebra_simps,metis (mono_tags, lifting) always_eventually not_one_less_zero) done then have "(\n. (b (Suc n) / b n) * (a n / a (Suc n))) \ 0" apply (elim Lim_transform_bound[rotated]) apply (rule eventuallyI) subgoal for x using a[rule_format, of x] a[rule_format, of "Suc x"] b[rule_format, of x] b[rule_format, of "Suc x"] by (auto simp add:field_simps) done then show ?thesis unfolding c_def by (auto simp add:algebra_simps) qed from c_ratio_tendsto have "(\n. norm (b (Suc n) / a (Suc n)) / norm (b n / a n)) \ 0" unfolding c_def using a b by (force simp add:divide_simps abs_of_pos intro: Lim_transform_eventually) from summable_ratio_test_tendsto[OF _ _ this] a b show "summable c" unfolding c_def apply auto by (metis less_irrefl) have "\\<^sub>F n in sequentially. (\j. c (n + j)) \ 2 * c n" proof - obtain N where N_ratio:"\n\N. c (n + 1) / c n < 1 / 2" proof - have "eventually (\n. c (n+1) / c n < 1/2) sequentially" using c_ratio_tendsto[unfolded tendsto_iff,rule_format, of "1/2",simplified] apply eventually_elim subgoal for n using c_pos[of n] c_pos[of "Suc n"] by auto done then show ?thesis using that unfolding eventually_sequentially by auto qed have "(\j. c (j + n)) \ 2 * c n" when "n\N" for n proof - have "(\j 2*c n * (1 - 1 / 2^(m+1))" for m proof (induct m) case 0 then show ?case using c_pos[of n] by simp next case (Suc m) have "(\ji c n + (\i c (i + n) / 2" for i using N_ratio[rule_format,of "i+n"] \n\N\ c_pos[of "i+n"] by simp then show ?thesis by (auto intro:sum_mono) qed also have "... = c n + (\i c n + c n * (1 - 1 / 2 ^ (m + 1))" using Suc by auto also have "... = 2 * c n * (1 - 1 / 2 ^ (Suc m + 1))" by (auto simp add:field_simps) finally show ?case . qed then have "(\j 2*c n" for m using c_pos[of n] by (smt divide_le_eq_1_pos divide_pos_pos nonzero_mult_div_cancel_left zero_less_power) moreover have "summable (\j. c (j + n))" using \summable c\ by (simp add: summable_iff_shift) ultimately show ?thesis using suminf_le_const[of "\j. c (j+n)" "2*c n"] by auto qed then show ?thesis unfolding eventually_sequentially by (auto simp add:algebra_simps) qed then show "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" unfolding c_def by simp qed theorem Hancl3: fixes d ::"nat\real " and a b ::"nat\int " assumes "A > 1" and d:"\n. d n> 1" and a: "\n. a n>0" and b:"\n. b n>0" and "s>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and assu2: "\n \ s. (A/ (of_int (a n)) powr(1 /of_int (2^n)))> prodinf (\j. d(n +j))" and assu3: "filterlim (\ n. (d n)^(2^n)/ b n) at_top sequentially" and "convergent_prod d" shows "suminf(\ n. b n / a n ) \ Rats" proof(rule ccontr) assume asm:"\ (suminf(\ n. b n / a n ) \ Rats)" have ab_sum:"summable (\j. b j / a j)" using issummable[OF \A>1\ d a b assu1 assu3 \convergent_prod d\] . obtain p q ::int where "q>0" and pq_def:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" proof - from asm have "suminf(\ n. b n / a n ) \ Rats" by auto then have "suminf(\ n. b (n+1) / a (n+1)) \ Rats" apply (subst suminf_minus_initial_segment[OF ab_sum,of 1]) by auto then obtain p' q' ::int where "q'\0" and pq_def:"(\ n. b (n+1) / a (n+1) ) sums (p'/q')" unfolding Rats_eq_int_div_int using summable_ignore_initial_segment[OF ab_sum,of 1,THEN summable_sums] by force define p q where "p=(if q'<0 then - p' else p')" and "q=(if q'<0 then - q' else q')" have "p'/q'=p/q" "q>0" using \q'\0\ unfolding p_def q_def by auto then show ?thesis using that[of q p] pq_def by auto qed define ALPHA where "ALPHA = (\n. (of_int q)*((\j=1..n. of_int(a j)))*(suminf (\ (j::nat). (b (j+n+1)/a (j+n+1)))))" have "ALPHA n \ 1" for n proof - have "suminf(\ n. b (n+1) / a (n+1)) > 0" apply (rule suminf_pos) subgoal using summable_ignore_initial_segment[OF ab_sum,of 1] by auto subgoal using a b by simp done then have "p/q > 0" using sums_unique[OF pq_def,symmetric] by auto then have "q\1" "p\1" using \q>0\ by (auto simp add:divide_simps) moreover have "\n. 1 \ a n" "\n. 1 \ b n" using a b by (auto simp add: int_one_le_iff_zero_less) ultimately show ?thesis unfolding ALPHA_def using show7[OF _ _ _ _ pq_def] by auto qed moreover have "\n. ALPHA n < 1" unfolding ALPHA_def proof (rule lasttoshow[OF d a \s>0\ \q>0\ \A>1\ b _ assu3]) show "\\<^sub>F n in sequentially. (\j. real_of_int (b (n + j)) / real_of_int (a (n + j))) \ real_of_int (2 * b n) / real_of_int (a n)" using show5[OF \A>1\ d a b assu1 assu3 \convergent_prod d\] by simp show "\m\s. \n\m. d (n + 1) * Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) \ real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))" apply (rule show9[OF \A>1\ d a \s>0\ assu1 \convergent_prod d\]) using show8[OF \A>1\ d a \s>0\ \convergent_prod d\ assu2] by (simp add:algebra_simps) qed ultimately show False using not_le by blast qed corollary Hancl3corollary: fixes A::real and a b ::"nat\int " assumes "A > 1" and a: "\n. a n>0" and b:"\n. b n>0" and assu1: "((\ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and asscor2: "\n \6. (of_int (a n)) powr(1 /of_int (2^n ))*(1+ 4*(2/3)^n) \ A \ (b n \2 powr((4/3)^(n-1)) ) " shows "suminf(\ n. b n / a n ) \ Rats" proof- define d::"nat\real" where "d= (\ n. 1+(2/3)^(n+1))" have dgt1:"\n. 1 < d n " unfolding d_def by auto moreover have "convergent_prod d" unfolding d_def apply (rule abs_convergent_prod_imp_convergent_prod) apply (rule summable_imp_abs_convergent_prod) using summable_ignore_initial_segment[OF complete_algebra_summable_geometric [of "2/3::real",simplified],of 1] by simp moreover have "\n\6. (\j. d (n + j)) < A / real_of_int (a n) powr (1 / real_of_int (2 ^ n))" proof rule+ fix n::nat assume "6 \ n" have d_sum:"summable (\j. ln (d j))" unfolding d_def apply (rule summable_ln_plus) apply (rule summable_ignore_initial_segment[OF complete_algebra_summable_geometric [of "2/3::real",simplified],of 1]) by simp have "(\j. ln (d (n + j))) < ln (1+4 * (2 / 3) ^ n)" proof - define c::real where "c=(2 / 3) ^ n" have "0n\6\ apply (subst power_add[symmetric]) by auto also have "... \ (2 / 3)^6" by (auto intro:power_le_one) also have "... < 1/8" by (auto simp add:field_simps) finally show "c<1/8" . qed (simp add:c_def) have "(\j. ln (d (n + j))) \ (\j. (2 / 3) ^ (n + j + 1))" apply (rule suminf_le) subgoal unfolding d_def apply (intro allI ln_add_one_self_le_self2 ) apply (rule order.strict_trans[of _ 0]) by auto subgoal using summable_ignore_initial_segment[OF d_sum,of n] by (auto simp add:algebra_simps) subgoal using summable_geometric[THEN summable_ignore_initial_segment,of "2/3" "n+1"] by (auto simp add:algebra_simps) done also have "... = (\j. (2 / 3)^(n+1)*(2 / 3) ^ j)" by (auto simp add:algebra_simps power_add) also have "... = (2 / 3)^(n+1) * (\j. (2 / 3) ^ j)" apply (rule suminf_mult) by (auto intro:summable_geometric) also have "... = 2 * c" unfolding c_def apply (subst suminf_geometric) by auto also have "... < 4 * c - (4 * c)\<^sup>2" using \0 \c<1/8\ by (sos "((((A<0 * A<1) * R<1) + ((A<=0 * R<1) * (R<1/16 * [1]^2))))") also have "... \ ln (1 + 4 * c)" apply (rule ln_one_plus_pos_lower_bound) using \0 \c<1/8\ by auto finally show ?thesis unfolding c_def by simp qed then have "exp (\j. ln (d (n + j))) < 1 + 4 * (2 / 3) ^ n" by (metis Groups.mult_ac(2) add.right_neutral add_mono_thms_linordered_field(5) divide_inverse divide_less_eq_numeral1(1) divide_pos_pos exp_gt_zero less_eq_real_def ln_exp ln_less_cancel_iff mult_zero_left rel_simps(27) rel_simps(76) zero_less_one zero_less_power) moreover have "exp (\j. ln (d (n + j))) = (\j. d (n + j))" proof (subst exp_suminf_prodinf_real [symmetric]) show "\k. 0 \ ln (d (n + k))" using dgt1 by (simp add: less_imp_le) show "abs_convergent_prod (\na. exp (ln (d (n + na))))" apply (subst exp_ln) subgoal for j using dgt1[rule_format,of "n+j"] by auto subgoal unfolding abs_convergent_prod_def real_norm_def apply (subst abs_of_nonneg) using convergent_prod_iff_shift[of d n] \convergent_prod d\ by (auto simp add: dgt1 less_imp_le algebra_simps) done show "(\na. exp (ln (d (n + na)))) = (\j. d (n + j))" apply (subst exp_ln) subgoal using dgt1 le_less_trans zero_le_one by blast subgoal by simp done qed ultimately have "(\j. d (n + j)) < 1 + 4 * (2 / 3) ^ n" by simp also have "... \ A / (a n) powr (1 / of_int (2 ^ n))" proof - have "a n powr (1 / real_of_int (2 ^ n)) > 0" using a[rule_format,of n] by auto then show ?thesis using asscor2[rule_format,OF \6\n\] by (auto simp add:field_simps) qed finally show "(\j. d (n + j)) < A / real_of_int (a n) powr (1 / of_int (2 ^ n))" . qed moreover have "LIM n sequentially. d n ^ 2 ^ n / real_of_int (b n) :> at_top" proof - have "LIM n sequentially. d n ^ 2 ^ n / 2 powr((4/3)^(n-1)) :> at_top" proof - define n1 where "n1=(\n. (2::real)* (3/2)^(n-1))" define n2 where "n2=(\n. ((4::real)/3)^(n-1))" have "LIM n sequentially. (((1+(8/9)/(n1 n)) powr (n1 n))/2) powr (n2 n) :> at_top" proof (rule filterlim_at_top_powr_real[where g'="exp (8/9) / (2::real)"]) define e1 where "e1=exp (8/9) / (2::real)" show "e1>1" unfolding e1_def by (approximation 4) show "(\n. ((1+(8/9)/(n1 n)) powr (n1 n))/2) \ e1" proof - have "(\n. (1+(8/9)/(n1 n)) powr (n1 n)) \ exp (8/9)" apply (rule filterlim_compose[OF tendsto_exp_limit_at_top]) unfolding n1_def by (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_realpow_sequentially_at_top simp:filterlim_sequentially_iff[of "\x. (3 / 2) ^ (x - Suc 0)" _ 1]) then show ?thesis unfolding e1_def by (intro tendsto_intros,auto) qed show "filterlim n2 at_top sequentially" unfolding n2_def apply (subst filterlim_sequentially_iff[of "\n. (4 / 3) ^ (n - 1)" _ 1]) by (auto intro:filterlim_realpow_sequentially_at_top) qed moreover have "\\<^sub>F n in sequentially. (((1+(8/9)/(n1 n)) powr (n1 n))/2) powr (n2 n) = d n ^ 2 ^ n / 2 powr((4/3)^(n-1))" proof (rule eventually_sequentiallyI[of 1]) fix x assume "x\(1::nat)" have " ((1 + 8 / 9 / n1 x) powr n1 x / 2) powr n2 x = (((1 + 8 / 9 / n1 x) powr n1 x) powr n2 x) / 2 powr (4 / 3) ^ (x - 1)" by (simp add:n1_def n2_def powr_divide) also have "... = (1 + 8 / 9 / n1 x) powr (n1 x * n2 x) / 2 powr (4 / 3) ^ (x - 1)" by (simp add: powr_powr) also have "... = (1 + 8 / 9 / n1 x) powr (2 ^ x) / 2 powr (4 / 3) ^ (x - 1)" proof - have "n1 x * n2 x = 2 ^ x" unfolding n1_def n2_def apply (subst mult.assoc) apply (subst power_mult_distrib[symmetric]) using \x\1\ by (auto simp flip:power_Suc) then show ?thesis by simp qed also have "... = (1 + 8 / 9 / n1 x) ^ (2 ^ x) / 2 powr (4 / 3) ^ (x - 1)" apply (subst (3) powr_realpow[symmetric]) apply (simp_all add: n1_def) by (smt zero_le_divide_iff zero_le_power) also have "... = d x ^ 2 ^ x / 2 powr (4 / 3) ^ (x - 1)" proof - define x1 where "x1=x-1" have *:"x=x1+1" unfolding x1_def using \x\1\ by simp have **: "8 / 9 / n1 x = (2 / 3) ^ (x + 1)" unfolding n1_def using \x\1\ apply (fold x1_def *[symmetric]) by (auto simp add:divide_simps) then show ?thesis unfolding d_def apply (subst **) by auto qed finally show "((1 + 8 / 9 / n1 x) powr n1 x / 2) powr n2 x = d x ^ 2 ^ x / 2 powr (4 / 3) ^ (x - 1) " . qed ultimately show ?thesis using filterlim_cong by fast qed moreover have "\\<^sub>F n in sequentially. d n ^ 2 ^ n / 2 powr((4/3)^(n-1)) \ d n ^ 2 ^ n / real_of_int (b n)" apply (rule eventually_sequentiallyI[of 6]) apply (rule divide_left_mono) subgoal for x using asscor2[rule_format,of x] by auto subgoal for x using \\n. 1 < d n\[rule_format, of x] by auto subgoal for x using b by auto done ultimately show ?thesis by (auto elim: filterlim_at_top_mono) qed ultimately show ?thesis using Hancl3[OF \A>1\ _ a b _ assu1,of d 6] by force qed end diff --git a/thys/Markov_Models/Classifying_Markov_Chain_States.thy b/thys/Markov_Models/Classifying_Markov_Chain_States.thy --- a/thys/Markov_Models/Classifying_Markov_Chain_States.thy +++ b/thys/Markov_Models/Classifying_Markov_Chain_States.thy @@ -1,2341 +1,2340 @@ section \Classifying Markov Chain States\ theory Classifying_Markov_Chain_States imports "HOL-Computational_Algebra.Group_Closure" Discrete_Time_Markov_Chain begin lemma eventually_mult_Gcd: fixes S :: "nat set" assumes S: "\s t. s \ S \ t \ S \ s + t \ S" assumes s: "s \ S" "s > 0" shows "eventually (\m. m * Gcd S \ S) sequentially" proof - define T where "T = insert 0 (int ` S)" with s S have "int s \ T" "0 \ T" and T: "r \ T \ t \ T \ r + t \ T" for r t by (auto simp del: of_nat_add simp add: of_nat_add [symmetric]) have "Gcd T \ group_closure T" by (rule Gcd_in_group_closure) also have "group_closure T = {s - t | s t. s \ T \ t \ T}" proof (auto intro: group_closure.base group_closure.diff) fix x assume "x \ group_closure T" then show "\s t. x = s - t \ s \ T \ t \ T" proof induction case (base x) with \0 \ T\ show ?case apply (rule_tac x=x in exI) apply (rule_tac x=0 in exI) apply auto done next case (diff x y) then obtain a b c d where "a \ T" "b \ T" "x = a - b" "c \ T" "d \ T" "y = c - d" by auto then show ?case apply (rule_tac x="a + d" in exI) apply (rule_tac x="b + c" in exI) apply (auto intro: T) done qed qed finally obtain s' t' :: int where "s' \ T" "t' \ T" "Gcd T = s' - t'" by blast moreover define s and t where "s = nat s'" and "t = nat t'" moreover have "int (Gcd S) = - int t \ S \ {0} \ t = 0" by auto (metis Gcd_dvd_nat dvd_0_right dvd_antisym nat_int nat_zminus_int) ultimately have st: "s = 0 \ s \ S" "t = 0 \ t \ S" and Gcd_S: "Gcd S = s - t" using T_def by safe simp_all with s have "t < s" by (rule_tac ccontr) auto { fix s n have "0 < n \ s \ S \ n * s \ S" proof (induct n) case (Suc n) then show ?case by (cases n) (auto intro: S) qed simp } note cmult_S = this show ?thesis unfolding eventually_sequentially proof cases assume "s = 0 \ t = 0" with st Gcd_S s have *: "Gcd S \ S" by (auto simp: int_eq_iff) then show "\N. \n\N. n * Gcd S \ S" by (auto intro!: exI[of _ 1] cmult_S) next assume "\ (s = 0 \ t = 0)" with st have "s \ S" "t \ S" "t \ 0" by auto then have "Gcd S dvd t" by auto then obtain a where a: "t = Gcd S * a" .. with \t \ 0\ have "0 < a" by auto show "\N. \n\N. n * Gcd S \ S" proof (safe intro!: exI[of _ "a * a"]) fix n define m where "m = (n - a * a) div a" define r where "r = (n - a * a) mod a" with \0 < a\ have "r < a" by simp moreover define am where "am = a + m" ultimately have "r < am" by simp assume "a * a \ n" then have n: "n = a * a + (m * a + r)" unfolding m_def r_def by simp have "n * Gcd S = am * t + r * Gcd S" unfolding n a by (simp add: field_simps am_def) also have "\ = r * s + (am - r) * t" unfolding \Gcd S = s - t\ using \t < s\ \r < am\ by (simp add: field_simps diff_mult_distrib2) also have "\ \ S" using \s \ S\ \t \ S\ \r < am\ by (cases "r = 0") (auto intro!: cmult_S S) finally show "n * Gcd S \ S" . qed qed qed context MC_syntax begin subsection \Expected number of visits\ definition "G s t = (\\<^sup>+\. scount (HLD {t}) (s ## \) \T s)" lemma G_eq: "G s t = (\\<^sup>+\. emeasure (count_space UNIV) {i. (s ## \) !! i = t} \T s)" by (simp add: G_def scount_eq_emeasure HLD_iff) definition "p s t n = \

(\ in T s. (s ## \) !! n = t)" definition "gf_G s t z = (\n. p s t n *\<^sub>R z ^ n)" definition "convergence_G s t z \ summable (\n. p s t n * norm z ^ n)" lemma p_nonneg[simp]: "0 \ p x y n" by (simp add: p_def) lemma p_le_1: "p x y n \ 1" by (simp add: p_def) lemma p_x_x_0[simp]: "p x x 0 = 1" by (simp add: p_def T.prob_space del: space_T) lemma p_0: "p x y 0 = (if x = y then 1 else 0)" by (simp add: p_def T.prob_space del: space_T) lemma p_in_reachable: assumes "(x, y) \ (SIGMA x:UNIV. K x)\<^sup>*" shows "p x y n = 0" unfolding p_def proof (rule T.prob_eq_0_AE) from AE_T_reachable show "AE \ in T x. (x ## \) !! n \ y" proof eventually_elim fix \ assume "alw (HLD ((SIGMA \:UNIV. K \)\<^sup>* `` {x})) \" then have "alw (HLD (- {y})) \" using assms by (auto intro: alw_mono simp: HLD_iff) then show "(x ## \) !! n \ y" using assms by (cases n) (auto simp: alw_HLD_iff_streams streams_iff_snth) qed qed lemma p_Suc: "ennreal (p x y (Suc n)) = (\\<^sup>+ w. p w y n \K x)" unfolding p_def T.emeasure_eq_measure[symmetric] by (subst emeasure_Collect_T) simp_all lemma p_Suc': "p x y (Suc n) = (\x'. p x' y n \K x)" using p_Suc[of x y n] by (subst (asm) nn_integral_eq_integral) (auto simp: p_le_1 intro!: measure_pmf.integrable_const_bound[where B=1]) lemma p_add: "p x y (n + m) = (\\<^sup>+ w. p x w n * p w y m \count_space UNIV)" proof (induction n arbitrary: x) case 0 have [simp]: "\w. (if x = w then 1 else 0) * p w y m = ennreal (p x y m) * indicator {x} w" by auto show ?case by (simp add: p_0 one_ennreal_def[symmetric] max_def) next case (Suc n) define X where "X = (SIGMA x:UNIV. K x)\<^sup>* `` K x" then have X: "countable X" by (blast intro: countable_Image countable_reachable countable_set_pmf) then interpret X: sigma_finite_measure "count_space X" by (rule sigma_finite_measure_count_space_countable) interpret XK: pair_sigma_finite "K x" "count_space X" by unfold_locales have "ennreal (p x y (Suc n + m)) = (\\<^sup>+t. (\\<^sup>+w. p t w n * p w y m \count_space UNIV) \K x)" by (simp add: p_Suc Suc) also have "\ = (\\<^sup>+t. (\\<^sup>+w. ennreal (p t w n * p w y m) * indicator X w \count_space UNIV) \K x)" by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff AE_count_space Image_iff p_in_reachable X_def split: split_indicator) also have "\ = (\\<^sup>+t. (\\<^sup>+w. p t w n * p w y m \count_space X) \K x)" by (subst nn_integral_restrict_space[symmetric]) (simp_all add: restrict_count_space) also have "\ = (\\<^sup>+w. (\\<^sup>+t. p t w n * p w y m \K x) \count_space X)" apply (rule XK.Fubini'[symmetric]) unfolding measurable_split_conv apply (rule measurable_compose_countable'[OF _ measurable_snd X]) apply (rule measurable_compose[OF measurable_fst]) apply simp done also have "\ = (\\<^sup>+w. (\\<^sup>+t. ennreal (p t w n * p w y m) * indicator X w \K x) \count_space UNIV)" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space nn_integral_multc) also have "\ = (\\<^sup>+w. (\\<^sup>+t. ennreal (p t w n * p w y m) \K x) \count_space UNIV)" by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff AE_count_space Image_iff p_in_reachable X_def split: split_indicator) also have "\ = (\\<^sup>+w. (\\<^sup>+t. p t w n \K x) * p w y m \count_space UNIV)" by (simp add: nn_integral_multc[symmetric] ennreal_mult) finally show ?case by (simp add: ennreal_mult p_Suc) qed lemma prob_reachable_le: assumes [simp]: "m \ n" shows "p x y m * p y w (n - m) \ p x w n" proof - have "p x y m * p y w (n - m) = (\\<^sup>+y'. ennreal (p x y m * p y w (n - m)) * indicator {y} y' \count_space UNIV)" by simp also have "\ \ p x w (m + (n - m))" by (subst p_add) (auto intro!: nn_integral_mono split: split_indicator simp del: nn_integral_indicator_singleton) finally show ?thesis by simp qed lemma G_eq_suminf: "G x y = (\i. ennreal (p x y i))" proof - have *: "\i \. indicator {\ \ space S. (x ## \) !! i = y} \ = indicator {i. (x ## \) !! i = y} i" by (auto simp: space_stream_space split: split_indicator) have "G x y = (\\<^sup>+ \. (\i. indicator {\\space (T x). (x ## \) !! i = y} \) \T x)" unfolding G_eq by (simp add: nn_integral_count_space_nat[symmetric] *) also have "\ = (\i. ennreal (p x y i))" by (simp add: T.emeasure_eq_measure[symmetric] p_def nn_integral_suminf) finally show ?thesis . qed lemma G_eq_real_suminf: "convergence_G x y (1::real) \ G x y = ennreal (\i. p x y i)" unfolding G_eq_suminf by (intro suminf_ennreal ennreal_suminf_neq_top p_nonneg) (auto simp: convergence_G_def p_def) lemma convergence_norm_G: "convergence_G x y z \ summable (\n. p x y n * norm z ^ n)" unfolding convergence_G_def . lemma convergence_G: "convergence_G x y (z::'a::{banach, real_normed_div_algebra}) \ summable (\n. p x y n *\<^sub>R z ^ n)" unfolding convergence_G_def by (rule summable_norm_cancel) (simp add: abs_mult norm_power) lemma convergence_G_less_1: fixes z :: "_ :: {banach, real_normed_field}" assumes z: "norm z < 1" shows "convergence_G x y z" unfolding convergence_G_def proof (rule summable_comparison_test) have "\n. p x y n * norm (z ^ n) \ 1 * norm (z ^ n)" by (intro mult_right_mono p_le_1) simp_all then show "\N. \n\N. norm (p x y n * norm z ^ n) \ norm z ^ n" by (simp add: norm_power) qed (simp add: z summable_geometric) lemma lim_gf_G: "((\z. ennreal (gf_G x y z)) \ G x y) (at_left (1::real))" unfolding gf_G_def G_eq_suminf real_scaleR_def by (intro power_series_tendsto_at_left p_nonneg p_le_1 summable_power_series) subsection \Reachability probability\ definition "u x y n = \

(\ in T x. ev_at (HLD {y}) n \)" definition "U s t = \

(\ in T s. ev (HLD {t}) \)" definition "gf_U x y z = (\n. u x y n *\<^sub>R z ^ Suc n)" definition "f x y n = \

(\ in T x. ev_at (HLD {y}) n (x ## \))" definition "F s t = \

(\ in T s. ev (HLD {t}) (s ## \))" definition "gf_F x y z = (\n. f x y n * z ^ n)" lemma f_Suc: "x \ y \ f x y (Suc n) = u x y n" by (simp add: u_def f_def) lemma f_Suc_eq: "f x x (Suc n) = 0" by (simp add: f_def) lemma f_0: "f x y 0 = (if x = y then 1 else 0)" using T.prob_space by (simp add: f_def) lemma shows u_nonneg: "0 \ u x y n" and u_le_1: "u x y n \ 1" by (simp_all add: u_def) lemma shows f_nonneg: "0 \ f x y n" and f_le_1: "f x y n \ 1" by (simp_all add: f_def) lemma U_nonneg[simp]: "0 \ U x y" by (simp add: U_def) lemma U_le_1: "U s t \ 1" by (auto simp add: U_def intro!: antisym) lemma U_cases: "U s s = 1 \ U s s < 1" by (auto simp add: U_def intro!: antisym) lemma u_sums_U: "u x y sums U x y" unfolding u_def[abs_def] U_def ev_iff_ev_at by (intro T.prob_sums) (auto intro: ev_at_unique) lemma gf_U_eq_U: "gf_U x y 1 = U x y" using u_sums_U[THEN sums_unique] by (simp add: gf_U_def U_def) lemma f_sums_F: "f x y sums F x y" unfolding f_def[abs_def] F_def ev_iff_ev_at by (intro T.prob_sums) (auto intro: ev_at_unique) lemma F_nonneg[simp]: "0 \ F x y" by (auto simp: F_def) lemma F_le_1: "F x y \ 1" by (simp add: F_def) lemma gf_F_eq_F: "gf_F x y 1 = F x y" using f_sums_F[THEN sums_unique] by (simp add: gf_F_def F_def) lemma gf_F_le_1: fixes z :: real assumes z: "0 \ z" "z \ 1" shows "gf_F x y z \ 1" proof - have "gf_F x y z \ gf_F x y 1" using z unfolding gf_F_def by (intro suminf_le[OF _ summable_comparison_test[OF _ sums_summable[OF f_sums_F[of x y]]]] mult_left_mono allI f_nonneg) (simp_all add: power_le_one f_nonneg mult_right_le_one_le f_le_1 sums_summable[OF f_sums_F[of x y]]) also have "\ \ 1" by (simp add: gf_F_eq_F F_def) finally show ?thesis . qed lemma u_le_p: "u x y n \ p x y (Suc n)" unfolding u_def p_def by (auto intro!: T.finite_measure_mono dest: ev_at_HLD_imp_snth) lemma f_le_p: "f x y n \ p x y n" unfolding f_def p_def by (auto intro!: T.finite_measure_mono dest: ev_at_HLD_imp_snth) lemma convergence_norm_U: fixes z :: "_ :: real_normed_div_algebra" assumes z: "convergence_G x y z" shows "summable (\n. u x y n * norm z ^ Suc n)" using summable_ignore_initial_segment[OF convergence_norm_G[OF z], of 1] by (rule summable_comparison_test[rotated]) (auto simp add: u_nonneg abs_mult intro!: exI[of _ 0] mult_right_mono u_le_p) lemma convergence_norm_F: fixes z :: "_ :: real_normed_div_algebra" assumes z: "convergence_G x y z" shows "summable (\n. f x y n * norm z ^ n)" using convergence_norm_G[OF z] by (rule summable_comparison_test[rotated]) (auto simp add: f_nonneg abs_mult intro!: exI[of _ 0] mult_right_mono f_le_p) lemma gf_G_nonneg: fixes z :: real shows "0 \ z \ z < 1 \ 0 \ gf_G x y z" unfolding gf_G_def by (intro suminf_nonneg convergence_G convergence_G_less_1) simp_all lemma gf_F_nonneg: fixes z :: real shows "0 \ z \ z < 1 \ 0 \ gf_F x y z" unfolding gf_F_def using convergence_norm_F[OF convergence_G_less_1, of z x y] by (intro suminf_nonneg) (simp_all add: f_nonneg) lemma convergence_U: fixes z :: "_ :: banach" shows "convergence_G x y z \ summable (\n. u x y n * z ^ Suc n)" by (rule summable_norm_cancel) (auto simp add: abs_mult u_nonneg power_abs dest!: convergence_norm_U) lemma p_eq_sum_p_u: "p x y (Suc n) = (\i\n. p y y (n - i) * u x y i)" proof - have "\\. \ !! n = y \ (\i. i \ n \ ev_at (HLD {y}) i \)" proof (induction n) case (Suc n) then obtain i where "i \ n" "ev_at (HLD {y}) i (stl \)" by auto then show ?case by (auto intro!: exI[of _ "if HLD {y} \ then 0 else Suc i"]) qed (simp add: HLD_iff) then have "p x y (Suc n) = (\i\n. \

(\ in T x. ev_at (HLD {y}) i \ \ \ !! n = y))" unfolding p_def by (intro T.prob_sum) (auto intro: ev_at_unique) also have "\ = (\i\n. p y y (n - i) * u x y i)" proof (intro sum.cong refl) fix i assume i: "i \ {.. n}" then have "\\. (Suc i \ n \ \ !! (n - Suc i) = y) \ ((y ## \) !! (n - i) = y)" by (auto simp: Stream_snth diff_Suc split: nat.split) from i have "i \ n" by auto then have "\

(\ in T x. ev_at (HLD {y}) i \ \ \ !! n = y) = (\\'. \

(\ in T y. (y ## \) !! (n - i) = y) * indicator {\'\space (T x). ev_at (HLD {y}) i \' } \' \T x)" by (subst prob_T_split[where n="Suc i"]) (auto simp: ev_at_shift ev_at_HLD_single_imp_snth shift_snth diff_Suc split: split_indicator nat.split intro!: Bochner_Integration.integral_cong arg_cong2[where f=measure] simp del: stake.simps integral_mult_right_zero) then show "\

(\ in T x. ev_at (HLD {y}) i \ \ \ !! n = y) = p y y (n - i) * u x y i" by (simp add: p_def u_def) qed finally show ?thesis . qed lemma p_eq_sum_p_f: "p x y n = (\i\n. p y y (n - i) * f x y i)" by (cases n) (simp_all del: sum.atMost_Suc add: f_0 p_0 p_eq_sum_p_u atMost_Suc_eq_insert_0 zero_notin_Suc_image sum.reindex f_Suc f_Suc_eq) lemma gf_G_eq_gf_F: assumes z: "norm z < 1" shows "gf_G x y z = gf_F x y z * gf_G y y z" proof - have "gf_G x y z = (\n. \i\n. p y y (n - i) * f x y i * z^n)" by (simp add: gf_G_def p_eq_sum_p_f[of x y] sum_distrib_right) also have "\ = (\n. \i\n. (f x y i * z^i) * (p y y (n - i) * z^(n - i)))" by (intro arg_cong[where f=suminf] sum.cong ext atLeast0AtMost[symmetric]) (simp_all add: power_add[symmetric]) also have "\ = (\n. f x y n * z^n) * (\n. p y y n * z^n)" using convergence_norm_F[OF convergence_G_less_1[OF z]] convergence_norm_G[OF convergence_G_less_1[OF z]] by (intro Cauchy_product[symmetric]) (auto simp: f_nonneg abs_mult power_abs) also have "\ = gf_F x y z * gf_G y y z" by (simp add: gf_F_def gf_G_def) finally show ?thesis . qed lemma gf_G_eq_gf_U: fixes z :: "'z :: {banach, real_normed_field}" assumes z: "convergence_G x x z" shows "gf_G x x z = 1 / (1 - gf_U x x z)" "gf_U x x z \ 1" proof - { fix n have "p x x (Suc n) *\<^sub>R z^Suc n = (\i\n. (p x x (n - i) * u x x i) *\<^sub>R z^Suc n)" unfolding scaleR_sum_left[symmetric] by (simp add: p_eq_sum_p_u) also have "\ = (\i\n. (u x x i *\<^sub>R z^Suc i) * (p x x (n - i) *\<^sub>R z^(n - i)))" by (intro sum.cong refl) (simp add: field_simps power_diff cong: disj_cong) finally have "p x x (Suc n) *\<^sub>R z^(Suc n) = (\i\n. (u x x i *\<^sub>R z^Suc i) * (p x x (n - i) *\<^sub>R z^(n - i)))" unfolding atLeast0AtMost . } note gfs_Suc_eq = this have "gf_G x x z = 1 + (\n. p x x (Suc n) *\<^sub>R z^(Suc n))" unfolding gf_G_def by (subst suminf_split_initial_segment[OF convergence_G[OF z], of 1]) simp also have "\ = 1 + (\n. \i\n. (u x x i *\<^sub>R z^Suc i) * (p x x (n - i) *\<^sub>R z^(n - i)))" unfolding gfs_Suc_eq .. also have "\ = 1 + gf_U x x z * gf_G x x z" unfolding gf_U_def gf_G_def by (subst Cauchy_product) (auto simp: u_nonneg norm_power simp del: power_Suc intro!: z convergence_norm_G convergence_norm_U) finally show "gf_G x x z = 1 / (1 - gf_U x x z)" "gf_U x x z \ 1" apply - apply (cases "gf_U x x z = 1") apply (auto simp add: field_simps) done qed lemma gf_U: "(gf_U x y \ U x y) (at_left 1)" proof - have "((\z. ennreal (\n. u x y n * z ^ n)) \ (\n. ennreal (u x y n))) (at_left 1)" using u_le_1 u_nonneg by (intro power_series_tendsto_at_left summable_power_series) also have "(\n. ennreal (u x y n)) = ennreal (suminf (u x y))" by (intro u_nonneg suminf_ennreal ennreal_suminf_neq_top sums_summable[OF u_sums_U]) also have "suminf (u x y) = U x y" using u_sums_U by (rule sums_unique[symmetric]) finally have "((\z. \n. u x y n * z ^ n) \ U x y) (at_left 1)" by (rule tendsto_ennrealD) (auto simp: u_nonneg u_le_1 intro!: suminf_nonneg summable_power_series eventually_at_left_1) then have "((\z. z * (\n. u x y n * z ^ n)) \ 1 * U x y) (at_left 1)" by (intro tendsto_intros) simp then have "((\z. \n. u x y n * z ^ Suc n) \ 1 * U x y) (at_left 1)" apply (rule filterlim_cong[OF refl refl, THEN iffD1, rotated]) apply (rule eventually_at_left_1) apply (subst suminf_mult[symmetric]) apply (auto intro!: summable_power_series u_le_1 u_nonneg) apply (simp add: field_simps) done then show ?thesis by (simp add: gf_U_def[abs_def] U_def) qed lemma gf_U_le_1: assumes z: "0 < z" "z < 1" shows "gf_U x y z \ (1::real)" proof - note u = u_sums_U[of x y, THEN sums_summable] have "gf_U x y z \ gf_U x y 1" using z unfolding gf_U_def real_scaleR_def by (intro suminf_le allI mult_mono power_mono summable_comparison_test_ev[OF _ u] always_eventually) (auto simp: u_nonneg intro!: mult_left_le mult_le_one power_le_one) also have "\ \ 1" unfolding gf_U_eq_U by (rule U_le_1) finally show ?thesis . qed lemma gf_F: "(gf_F x y \ F x y) (at_left 1)" proof - have "((\z. ennreal (\n. f x y n * z ^ n)) \ (\n. ennreal (f x y n))) (at_left 1)" using f_le_1 f_nonneg by (intro power_series_tendsto_at_left summable_power_series) also have "(\n. ennreal (f x y n)) = ennreal (suminf (f x y))" by (intro f_nonneg suminf_ennreal ennreal_suminf_neq_top sums_summable[OF f_sums_F]) also have "suminf (f x y) = F x y" using f_sums_F by (rule sums_unique[symmetric]) finally have "((\z. \n. f x y n * z ^ n) \ F x y) (at_left 1)" by (rule tendsto_ennrealD) (auto simp: f_nonneg f_le_1 intro!: suminf_nonneg summable_power_series eventually_at_left_1) then show ?thesis by (simp add: gf_F_def[abs_def] F_def) qed lemma U_bounded: "0 \ U x y" "U x y \ 1" unfolding U_def by simp_all subsection \Recurrent states\ definition recurrent :: "'s \ bool" where "recurrent s \ (AE \ in T s. ev (HLD {s}) \)" lemma recurrent_iff_U_eq_1: "recurrent s \ U s s = 1" unfolding recurrent_def U_def by (subst T.prob_Collect_eq_1) simp_all definition "H s t = \

(\ in T s. alw (ev (HLD {t})) \)" lemma H_eq: "recurrent s \ H s s = 1" "\ recurrent s \ H s s = 0" "H s t = U s t * H t t" proof - define H' where "H' t n = {\\space S. enat n \ scount (HLD {t::'s}) \}" for t n have [measurable]: "\y n. H' y n \ sets S" by (simp add: H'_def) let ?H' = "\s t n. measure (T s) (H' t n)" { fix x y :: 's and \ have "Suc 0 \ scount (HLD {y}) \ \ ev (HLD {y}) \" using scount_eq_0_iff[of "HLD {y}" \] by (cases "scount (HLD {y}) \" rule: enat_coexhaust) (auto simp: not_ev_iff[symmetric] eSuc_enat[symmetric] enat_0 HLD_iff[abs_def]) } then have H'_1: "\x y. ?H' x y 1 = U x y" unfolding H'_def U_def by simp { fix n and x y :: 's let ?U = "(not (HLD {y}) suntil (HLD {y} aand nxt (\\. enat n \ scount (HLD {y}) \)))" { fix \ have "enat (Suc n) \ scount (HLD {y}) \ \ ?U \" proof assume "enat (Suc n) \ scount (HLD {y}) \" with scount_eq_0_iff[of "HLD {y}" \] have "ev (HLD {y}) \" "enat (Suc n) \ scount (HLD {y}) \" by (auto simp add: not_ev_iff[symmetric] eSuc_enat[symmetric]) then show "?U \" by (induction rule: ev_induct_strong) (auto simp: scount_simps eSuc_enat[symmetric] intro: suntil.intros) next assume "?U \" then show "enat (Suc n) \ scount (HLD {y}) \" by induction (auto simp: scount_simps eSuc_enat[symmetric]) qed } then have "emeasure (T x) (H' y (Suc n)) = emeasure (T x) {\\space (T x). ?U \}" by (simp add: H'_def) also have "\ = U x y * ?H' y y n" by (subst emeasure_suntil_HLD) (simp_all add: T.emeasure_eq_measure U_def H'_def ennreal_mult) finally have "?H' x y (Suc n) = U x y * ?H' y y n" by (simp add: T.emeasure_eq_measure) } note H'_Suc = this { fix m and x :: 's have "?H' x x (Suc m) = U x x^Suc m" using H'_1 H'_Suc by (induct m) auto } note H'_eq = this { fix x y have "?H' x y \ measure (T x) (\i. H' y i)" apply (rule T.finite_Lim_measure_decseq) apply safe apply simp apply (auto simp add: decseq_Suc_iff subset_eq H'_def eSuc_enat[symmetric] intro: ile_eSuc order_trans) done also have "(\i. H' y i) = {\\space (T x). alw (ev (HLD {y})) \}" by (auto simp: H'_def scount_infinite_iff[symmetric]) (metis Suc_ile_eq enat.exhaust neq_iff) finally have "?H' x y \ H x y" unfolding H_def . } note H'_lim = this from H'_lim[of s s, THEN LIMSEQ_Suc] have "(\n. U s s ^ Suc n) \ H s s" by (simp add: H'_eq) then have lim_H: "(\n. U s s ^ n) \ H s s" by (rule LIMSEQ_imp_Suc) have "U s s < 1 \ (\n. U s s ^ n) \ 0" by (rule LIMSEQ_realpow_zero) (simp_all add: U_def) with lim_H have "U s s < 1 \ H s s = 0" by (blast intro: LIMSEQ_unique) moreover have "U s s = 1 \ (\n. U s s ^ n) \ 1" by simp with lim_H have "U s s = 1 \ H s s = 1" by (blast intro: LIMSEQ_unique) moreover note recurrent_iff_U_eq_1 U_cases ultimately show "recurrent s \ H s s = 1" "\ recurrent s \ H s s = 0" by (metis one_neq_zero)+ from H'_lim[of s t, THEN LIMSEQ_Suc] H'_Suc[of s] have "(\n. U s t * ?H' t t n) \ H s t" by simp moreover have "(\n. U s t * ?H' t t n) \ U s t * H t t" by (intro tendsto_intros H'_lim) ultimately show "H s t = U s t * H t t" by (blast intro: LIMSEQ_unique) qed lemma recurrent_iff_G_infinite: "recurrent x \ G x x = \" proof - have "((\z. ennreal (gf_G x x z)) \ G x x) (at_left 1)" by (rule lim_gf_G) then have G: "((\z. ennreal (1 / (1 - gf_U x x z))) \ G x x) (at_left (1::real))" apply (rule filterlim_cong[OF refl refl, THEN iffD1, rotated]) apply (rule eventually_at_left_1) apply (subst gf_G_eq_gf_U) apply (rule convergence_G_less_1) apply simp apply simp done { fix z :: real assume z: "0 < z" "z < 1" have 1: "summable (u x x)" using u_sums_U by (rule sums_summable) have "gf_U x x z \ 1" using gf_G_eq_gf_U[OF convergence_G_less_1[of z]] z by simp moreover have "gf_U x x z \ U x x" unfolding gf_U_def gf_U_eq_U[symmetric] using z by (intro suminf_le) (auto simp add: 1 convergence_U convergence_G_less_1 u_nonneg simp del: power_Suc intro!: mult_right_le_one_le power_le_one) ultimately have "gf_U x x z < 1" using U_bounded[of x x] by simp } note strict = this { assume "U x x = 1" moreover have "((\xa. 1 - gf_U x x xa :: real) \ 1 - U x x) (at_left 1)" by (intro tendsto_intros gf_U) moreover have "eventually (\z. gf_U x x z < 1) (at_left (1::real))" by (auto intro!: eventually_at_left_1 strict simp: \U x x = 1\ gf_U_eq_U) ultimately have "((\z. ennreal (1 / (1 - gf_U x x z))) \ top) (at_left 1)" unfolding ennreal_tendsto_top_eq_at_top by (intro LIM_at_top_divide[where a=1] tendsto_const zero_less_one) (auto simp: field_simps) with G have "G x x = top" by (rule tendsto_unique[rotated]) simp } moreover { assume "U x x < 1" then have "((\xa. ennreal (1 / (1 - gf_U x x xa))) \ 1 / (1 - U x x)) (at_left 1)" by (intro tendsto_intros gf_U tendsto_ennrealI) simp from tendsto_unique[OF _ G this] have "G x x \ \" by simp } ultimately show ?thesis using U_cases recurrent_iff_U_eq_1 by auto qed definition communicating :: "('s \ 's) set" where "communicating = acc \ acc\" definition essential_class :: "'s set \ bool" where "essential_class C \ C \ UNIV // communicating \ acc `` C \ C" lemma accI_U: assumes "0 < U x y" shows "(x, y) \ acc" proof (rule ccontr) assume *: "(x, y) \ acc" { fix \ assume "ev (HLD {y}) \" "alw (HLD (acc `` {x})) \" from this * have False by induction (auto simp: HLD_iff) } with AE_T_reachable[of x] have "U x y = 0" unfolding U_def by (intro T.prob_eq_0_AE) auto with \0 < U x y\ show False by auto qed lemma accD_pos: assumes "(x, y) \ acc" shows "\n. 0 < p x y n" using assms proof induction case base with T.prob_space[of x] show ?case by (auto intro!: exI[of _ 0]) next have [simp]: "\x y. (if x = y then 1 else 0::real) = indicator {y} x" by simp case (step w y) then obtain n where "0 < p x w n" and "0 < pmf (K w) y" by (auto simp: set_pmf_iff less_le) then have "0 < p x w n * pmf (K w) y" by (intro mult_pos_pos) also have "\ \ p x w n * p w y (Suc 0)" by (simp add: p_Suc' p_0 pmf.rep_eq) also have "\ \ p x y (Suc n)" using prob_reachable_le[of n "Suc n" x w y] by simp finally show ?case .. qed lemma accI_pos: "0 < p x y n \ (x, y) \ acc" proof (induct n arbitrary: x) case (Suc n) then have less: "0 < (\x'. p x' y n \K x)" by (simp add: p_Suc') have "\x'\K x. 0 < p x' y n" proof (rule ccontr) assume "\ ?thesis" then have "AE x' in K x. p x' y n = 0" by (simp add: AE_measure_pmf_iff less_le) then have "(\x'. p x' y n \K x) = (\x'. 0 \K x)" by (intro integral_cong_AE) simp_all with less show False by simp qed with Suc show ?case by (auto intro: converse_rtrancl_into_rtrancl) qed (simp add: p_0 split: if_split_asm) lemma recurrent_iffI_communicating: assumes "(x, y) \ communicating" shows "recurrent x \ recurrent y" proof - from assms obtain n m where "0 < p x y n" "0 < p y x m" by (force simp: communicating_def dest: accD_pos) moreover { fix x y n m assume "0 < p x y n" "0 < p y x m" "G y y = \" then have "\ = ennreal (p x y n * p y x m) * G y y" by (auto intro: mult_pos_pos simp: ennreal_mult_top) also have "ennreal (p x y n * p y x m) * G y y = (\i. ennreal (p x y n * p y x m) * p y y i)" unfolding G_eq_suminf by (rule ennreal_suminf_cmult[symmetric]) also have "\ \ (\i. ennreal (p x x (n + i + m)))" proof (intro suminf_le allI) fix i have "(p x y n * p y y ((n + i) - n)) * p y x ((n + i + m) - (n + i)) \ p x y (n + i) * p y x ((n + i + m) - (n + i))" by (intro mult_right_mono prob_reachable_le) simp_all also have "\ \ p x x (n + i + m)" by (intro prob_reachable_le) simp_all finally show "ennreal (p x y n * p y x m) * p y y i \ ennreal (p x x (n + i + m))" by (simp add: ac_simps ennreal_mult'[symmetric]) qed auto also have "\ \ (\i. ennreal (p x x (i + (n + m))))" by (simp add: ac_simps) also have "\ \ (\i. ennreal (p x x i))" by (subst suminf_offset[of "\i. ennreal (p x x i)" "n + m"]) auto also have "\ \ G x x" unfolding G_eq_suminf by (auto intro!: suminf_le_pos) finally have "G x x = \" by (simp add: top_unique) } ultimately show ?thesis using recurrent_iff_G_infinite by blast qed lemma recurrent_acc: assumes "recurrent x" "(x, y) \ acc" shows "U y x = 1" "H y x = 1" "recurrent y" "(x, y) \ communicating" proof - { fix w y assume step: "(x, w) \ acc" "y \ K w" "U w x = 1" "H w x = 1" "recurrent w" "x \ y" have "measure (K w) UNIV = U w x" using step measure_pmf.prob_space[of "K w"] by simp also have "\ = (\v. indicator {x} v + U v x * indicator (- {x}) v \K w)" unfolding U_def by (subst prob_T) (auto intro!: Bochner_Integration.integral_cong arg_cong2[where f=measure] AE_I2 simp: ev_Stream T.prob_eq_1 split: split_indicator) also have "\ = measure (K w) {x} + (\v. U v x * indicator (- {x}) v \K w)" by (subst Bochner_Integration.integral_add) (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: abs_mult mult_le_one U_bounded(2) measure_pmf.emeasure_eq_measure) finally have "measure (K w) UNIV - measure (K w) {x} = (\v. U v x * indicator (- {x}) v \K w)" by simp also have "measure (K w) UNIV - measure (K w) {x} = measure (K w) (UNIV - {x})" by (subst measure_pmf.finite_measure_Diff) auto finally have "0 = (\v. indicator (- {x}) v \K w) - (\v. U v x * indicator (- {x}) v \K w)" by (simp add: measure_pmf.emeasure_eq_measure Compl_eq_Diff_UNIV) also have "\ = (\v. (1 - U v x) * indicator (- {x}) v \K w)" by (subst Bochner_Integration.integral_diff[symmetric]) (auto intro!: measure_pmf.integrable_const_bound[where B=1] Bochner_Integration.integral_cong simp: abs_mult mult_le_one U_bounded(2) split: split_indicator) also have "\ \ (\v. (1 - U y x) * indicator {y} v \K w)" (is "_ \ ?rhs") using \recurrent x\ by (intro integral_mono measure_pmf.integrable_const_bound[where B=1]) (auto simp: abs_mult mult_le_one U_bounded(2) recurrent_iff_U_eq_1 field_simps split: split_indicator) also (xtrans) have "?rhs = (1 - U y x) * pmf (K w) y" by (simp add: measure_pmf.emeasure_eq_measure pmf.rep_eq) finally have "(1 - U y x) * pmf (K w) y = 0" by (auto intro!: antisym simp: U_bounded(2) mult_le_0_iff) with \y \ K w\ have "U y x = 1" by (simp add: set_pmf_iff) then have "U y x = 1" "H y x = 1" using H_eq(3)[of y x] H_eq(1)[of x] by (simp_all add: \recurrent x\) then have "(y, x) \ acc" by (intro accI_U) auto with step have "(x, y) \ communicating" by (auto simp add: communicating_def intro: rtrancl_trans) with \recurrent x\ have "recurrent y" by (simp add: recurrent_iffI_communicating) note this \U y x = 1\ \H y x = 1\ \(x, y) \ communicating\ } note enabled = this from \(x, y) \ acc\ show "U y x = 1" "H y x = 1" "recurrent y" "(x, y) \ communicating" proof induction case base then show "U x x = 1" "H x x = 1" "recurrent x" "(x, x) \ communicating" using \recurrent x\ H_eq(1)[of x] by (auto simp: recurrent_iff_U_eq_1 communicating_def) next case (step w y) with enabled[of w y] \recurrent x\ H_eq(1)[of x] have "U y x = 1 \ H y x = 1 \ recurrent y \ (x, y) \ communicating" by (cases "x = y") (auto simp: recurrent_iff_U_eq_1 communicating_def) then show "U y x = 1" "H y x = 1" "recurrent y" "(x, y) \ communicating" by auto qed qed lemma equiv_communicating: "equiv UNIV communicating" by (auto simp: equiv_def sym_def communicating_def refl_on_def trans_def) lemma recurrent_class: assumes "recurrent x" shows "acc `` {x} = communicating `` {x}" using recurrent_acc(4)[OF \recurrent x\] by (auto simp: communicating_def) lemma irreduccible_recurrent_class: assumes "recurrent x" shows "acc `` {x} \ UNIV // communicating" unfolding recurrent_class[OF \recurrent x\] by (rule quotientI) simp lemma essential_classI: assumes C: "C \ UNIV // communicating" assumes eq: "\x y. x \ C \ (x, y) \ acc \ y \ C" shows "essential_class C" by (auto simp: essential_class_def intro: C) (metis eq) lemma essential_recurrent_class: assumes "recurrent x" shows "essential_class (communicating `` {x})" unfolding recurrent_class[OF \recurrent x\, symmetric] apply (rule essential_classI) apply (rule irreduccible_recurrent_class[OF assms]) apply (auto simp: communicating_def) done lemma essential_classD2: "essential_class C \ x \ C \ (x, y) \ acc \ y \ C" unfolding essential_class_def by auto lemma essential_classD3: "essential_class C \ x \ C \ y \ C \ (x, y) \ communicating" unfolding essential_class_def by (auto elim!: quotientE simp: communicating_def) lemma AE_acc: shows "AE \ in T x. \m. (x, (x ## \) !! m) \ acc" using AE_T_reachable by eventually_elim (auto simp: alw_HLD_iff_streams streams_iff_snth Stream_snth split: nat.splits) lemma finite_essential_class_imp_recurrent: assumes C: "essential_class C" "finite C" and x: "x \ C" shows "recurrent x" proof - have "AE \ in T x. \y\C. alw (ev (HLD {y})) \" using AE_T_reachable proof eventually_elim fix \ assume "alw (HLD (acc `` {x})) \" then have "alw (HLD C) \" by (rule alw_mono) (auto simp: HLD_iff intro: assms essential_classD2) then show "\y\C. alw (ev (HLD {y})) \" by (rule pigeonhole_stream) fact qed then have "1 = \

(\ in T x. \y\C. alw (ev (HLD {y})) \)" by (subst (asm) T.prob_Collect_eq_1[symmetric]) (auto simp: \finite C\) also have "\ = measure (T x) (\y\C. {\\space (T x). alw (ev (HLD {y})) \})" by (intro arg_cong2[where f=measure]) auto also have "\ \ (\y\C. H x y)" unfolding H_def using \finite C\ by (rule T.finite_measure_subadditive_finite) auto also have "\ = (\y\C. U x y * H y y)" by (auto intro!: sum.cong H_eq) finally have "\y\C. recurrent y" by (rule_tac ccontr) (simp add: H_eq(2)) then guess y .. from essential_classD3[OF C(1) x this(1)] recurrent_acc(3)[OF this(2)] show "recurrent x" by (simp add: communicating_def) qed lemma irreducibleD: "C \ UNIV // communicating \ a \ C \ b \ C \ (a, b) \ communicating" by (auto elim!: quotientE simp: communicating_def) lemma irreducibleD2: "C \ UNIV // communicating \ a \ C \ (a, b) \ communicating \ b \ C" by (auto elim!: quotientE simp: communicating_def) lemma essential_class_iff_recurrent: "finite C \ C \ UNIV // communicating \ essential_class C \ (\x\C. recurrent x)" by (metis finite_essential_class_imp_recurrent irreducibleD2 recurrent_acc(4) essential_classI) definition "U' x y = (\\<^sup>+\. eSuc (sfirst (HLD {y}) \) \T x)" lemma U'_neq_zero[simp]: "U' x y \ 0" unfolding U'_def by (simp add: nn_integral_add) definition "gf_U' x y z = (\n. u x y n * Suc n * z ^ n)" definition "pos_recurrent x \ recurrent x \ U' x x \ \" lemma summable_gf_U': assumes z: "norm z < 1" shows "summable (\n. u x y n * Suc n * z ^ n)" proof - have "summable (\n. n * \z\ ^ n)" proof (rule root_test_convergence) have "(\n. root n n * \z\) \ 1 * \z\" by (intro tendsto_intros LIMSEQ_root) then show "(\n. root n (norm (n * \z\ ^ n))) \ \z\" by (rule filterlim_cong[THEN iffD1, rotated 3]) (auto intro!: exI[of _ 1] simp add: abs_mult u_nonneg real_root_mult power_abs eventually_sequentially real_root_power) qed (insert z, simp add: abs_less_iff) note summable_mult[OF this, of "1 / \z\"] from summable_ignore_initial_segment[OF this, of 1] show "summable (\n. u x y n * Suc n * z ^ n)" apply (rule summable_comparison_test[rotated]) using z apply (auto intro!: exI[of _ 1] simp: abs_mult u_nonneg power_abs Suc_le_eq gr0_conv_Suc field_simps le_divide_eq u_le_1 simp del: of_nat_Suc) done qed lemma gf_U'_nonneg[simp]: "0 < z \ z < 1 \ 0 \ gf_U' x y z" unfolding gf_U'_def by (intro suminf_nonneg summable_gf_U') (auto simp: u_nonneg) lemma DERIV_gf_U: fixes z :: real assumes z: "0 < z" "z < 1" shows "DERIV (gf_U x y) z :> gf_U' x y z" unfolding gf_U_def[abs_def] gf_U'_def real_scaleR_def u_def[symmetric] using z by (intro DERIV_power_series'[where R=1] summable_gf_U') auto lemma sfirst_finiteI_recurrent: "recurrent x \ (x, y) \ acc \ AE \ in T x. sfirst (HLD {y}) \ < \" using recurrent_acc(1)[of y x] recurrent_acc[of x y] T.AE_prob_1[of x "{\\space (T x). ev (HLD {y}) \}"] unfolding sfirst_finite U_def by (simp add: space_stream_space communicating_def) lemma U'_eq_suminf: assumes x: "recurrent x" "(x, y) \ acc" shows "U' x y = (\i. ennreal (u x y i * Suc i))" proof - have "(\\<^sup>+\. eSuc (sfirst (HLD {y}) \) \T x) = (\\<^sup>+\. (\i. ennreal (Suc i) * indicator {\\space (T y). ev_at (HLD {y}) i \} \) \T x)" using sfirst_finiteI_recurrent[OF x] proof (intro nn_integral_cong_AE, eventually_elim) fix \ assume "sfirst (HLD {y}) \ < \" then obtain n :: nat where [simp]: "sfirst (HLD {y}) \ = n" by auto show "eSuc (sfirst (HLD {y}) \) = (\i. ennreal (Suc i) * indicator {\\space (T y). ev_at (HLD {y}) i \} \)" by (subst suminf_cmult_indicator[where i=n]) (auto simp: disjoint_family_on_def ev_at_unique space_stream_space sfirst_eq_enat_iff[symmetric] ennreal_of_nat_eq_real_of_nat split: split_indicator) qed also have "\ = (\i. ennreal (Suc i) * emeasure (T x) {\\space (T x). ev_at (HLD {y}) i \})" by (subst nn_integral_suminf) (auto intro!: arg_cong[where f=suminf] nn_integral_cmult_indicator simp: fun_eq_iff) finally show ?thesis by (simp add: U'_def u_def T.emeasure_eq_measure mult_ac ennreal_mult) qed lemma gf_U'_tendsto_U': assumes x: "recurrent x" "(x, y) \ acc" shows "((\z. ennreal (gf_U' x y z)) \ U' x y) (at_left 1)" unfolding U'_eq_suminf[OF x] gf_U'_def by (auto intro!: power_series_tendsto_at_left summable_gf_U' mult_nonneg_nonneg u_nonneg simp del: of_nat_Suc) lemma one_le_integral_t: assumes x: "recurrent x" shows "1 \ U' x x" by (simp add: nn_integral_add T.emeasure_space_1 U'_def del: space_T) lemma gf_U'_pos: fixes z :: real assumes z: "0 < z" "z < 1" and "U x y \ 0" shows "0 < gf_U' x y z" unfolding gf_U'_def proof (subst suminf_pos_iff) show "summable (\n. u x y n * real (Suc n) * z ^ n)" using z by (intro summable_gf_U') simp - show pos: "\n. 0 \ u x y n * real (Suc n) * z ^ n" - using z by (auto intro!: mult_nonneg_nonneg u_nonneg) + show pos: "\n. 0 \ u x y n * real (Suc n) * z ^ n" + using u_nonneg z by auto show "\n. 0 < u x y n * real (Suc n) * z ^ n" proof (rule ccontr) assume "\ (\n. 0 < u x y n * real (Suc n) * z ^ n)" with pos have "\n. u x y n * real (Suc n) * z ^ n = 0" by (intro antisym allI) (simp_all add: not_less) with z have "u x y = (\n. 0)" by (intro ext) simp with u_sums_U[of x y, THEN sums_unique] \U x y \ 0\ show False by simp qed qed lemma inverse_gf_U'_tendsto: assumes "recurrent y" shows "((\x. - 1 / - gf_U' y y x) \ enn2real (1 / U' y y)) (at_left (1::real))" proof cases assume inf: "U' y y = \" with gf_U'_tendsto_U'[of y y] \recurrent y\ have "LIM z (at_left 1). gf_U' y y z :> at_top" by (auto simp: ennreal_tendsto_top_eq_at_top U'_def) then have "LIM z (at_left 1). gf_U' y y z :> at_infinity" by (rule filterlim_mono) (auto simp: at_top_le_at_infinity) with inf show ?thesis by (auto intro!: tendsto_divide_0) next assume fin: "U' y y \ \" then obtain r where r: "U' y y = ennreal r" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: U'_def) then have eq: "enn2real (1 / U' y y) = - 1 / - r" and "1 \ r" using one_le_integral_t[OF \recurrent y\] by (auto simp add: ennreal_1[symmetric] divide_ennreal simp del: ennreal_1) have "((\z. ennreal (gf_U' y y z)) \ ennreal r) (at_left 1)" using gf_U'_tendsto_U'[OF \recurrent y\, of y] r by simp then have gf_U': "(gf_U' y y \ r) (at_left (1::real))" by (rule tendsto_ennrealD) (insert summable_gf_U', auto intro!: eventually_at_left_1 suminf_nonneg simp: gf_U'_def u_nonneg) show ?thesis using \1 \ r\ unfolding eq by (intro tendsto_intros gf_U') simp qed lemma gf_G_pos: fixes z :: real assumes z: "0 < z" "z < 1" and *: "(x, y) \ acc" shows "0 < gf_G x y z" unfolding gf_G_def proof (subst suminf_pos_iff) show "summable (\n. p x y n *\<^sub>R z ^ n)" using z by (intro convergence_G convergence_G_less_1) simp - show pos: "\n. 0 \ p x y n *\<^sub>R z ^ n" + show pos: "\n. 0 \ p x y n *\<^sub>R z ^ n" using z by (auto intro!: mult_nonneg_nonneg p_nonneg) show "\n. 0 < p x y n *\<^sub>R z ^ n" proof (rule ccontr) assume "\ (\n. 0 < p x y n *\<^sub>R z ^ n)" with pos have "\n. p x y n * z ^ n = 0" by (intro antisym allI) (simp_all add: not_less) with z have "\n. p x y n = 0" by simp with *[THEN accD_pos] show False by simp qed qed lemma pos_recurrentI_communicating: assumes y: "pos_recurrent y" and x: "(y, x) \ communicating" shows "pos_recurrent x" proof - from y x have recurrent: "recurrent y" "recurrent x" and fin: "U' y y \ \" by (auto simp: pos_recurrent_def recurrent_iffI_communicating nn_integral_add) have pos: "0 < enn2real (1 / U' y y)" using one_le_integral_t[OF \recurrent y\] fin by (auto simp: U'_def enn2real_positive_iff less_top[symmetric] ennreal_zero_less_divide ennreal_divide_eq_top_iff) from fin obtain r where r: "U' y y = ennreal r" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: U'_def) from x obtain n m where "0 < p x y n" "0 < p y x m" by (auto dest!: accD_pos simp: communicating_def) let ?L = "at_left (1::real)" have le: "eventually (\z. p x y n * p y x m * z^(n + m) \ (1 - gf_U y y z) / (1 - gf_U x x z)) ?L" proof (rule eventually_at_left_1) fix z :: real assume z: "0 < z" "z < 1" then have conv: "\x. convergence_G x x z" by (intro convergence_G_less_1) simp have sums: "(\i. (p x y n * p y x m * z^(n + m)) * (p y y i * z^i)) sums ((p x y n * p y x m * z^(n + m)) * gf_G y y z)" unfolding gf_G_def by (intro sums_mult summable_sums) (auto intro!: conv convergence_G[where 'a=real, simplified]) have "(\i. (p x y n * p y x m * z^(n + m)) * (p y y i * z^i)) \ (\i. p x x (i + (n + m)) * z^(i + (n + m)))" proof (intro allI suminf_le sums_summable[OF sums] summable_ignore_initial_segment convergence_G[where 'a=real, simplified] convergence_G_less_1) show "norm z < 1" using z by simp fix i have "(p x y n * p y y ((n + i) - n)) * p y x ((n + i + m) - (n + i)) \ p x y (n + i) * p y x ((n + i + m) - (n + i))" by (intro mult_right_mono prob_reachable_le) simp_all also have "\ \ p x x (n + i + m)" by (intro prob_reachable_le) simp_all finally show "p x y n * p y x m * z ^ (n + m) * (p y y i * z ^ i) \ p x x (i + (n + m)) * z ^ (i + (n + m))" using z by (auto simp add: ac_simps power_add intro!: mult_left_mono) qed also have "\ \ gf_G x x z" unfolding gf_G_def using z apply (subst (2) suminf_split_initial_segment[where k="n + m"]) apply (intro convergence_G conv) apply (simp add: sum_nonneg) done finally have "(p x y n * p y x m * z^(n + m)) * gf_G y y z \ gf_G x x z" using sums_unique[OF sums] by simp then have "(p x y n * p y x m * z^(n + m)) \ gf_G x x z / gf_G y y z" using z gf_G_pos[of z y y] by (simp add: field_simps) also have "\ = (1 - gf_U y y z) / (1 - gf_U x x z)" unfolding gf_G_eq_gf_U[OF conv] using gf_G_eq_gf_U(2)[OF conv] by (simp add: field_simps ) finally show "p x y n * p y x m * z^(n + m) \ (1 - gf_U y y z) / (1 - gf_U x x z)" . qed have "U' x x \ \" proof assume "U' x x = \" have "((\z. (1 - gf_U y y z) / (1 - gf_U x x z)) \ 0) ?L" proof (rule lhopital_left) show "((\z. 1 - gf_U y y z) \ 0) ?L" using gf_U[of y] recurrent_iff_U_eq_1[of y] \recurrent y\ by (auto intro!: tendsto_eq_intros) show "((\z. 1 - gf_U x x z) \ 0) ?L" using gf_U[of x] recurrent_iff_U_eq_1[of x] \recurrent x\ by (auto intro!: tendsto_eq_intros) show "eventually (\z. 1 - gf_U x x z \ 0) ?L" by (auto intro!: eventually_at_left_1 simp: gf_G_eq_gf_U(2) convergence_G_less_1) show "eventually (\z. - gf_U' x x z \ 0) ?L" using gf_U'_pos[of _ x x] recurrent_iff_U_eq_1[of x] \recurrent x\ by (auto intro!: eventually_at_left_1) (metis less_le) show "eventually (\z. DERIV (\xa. 1 - gf_U x x xa) z :> - gf_U' x x z) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) show "eventually (\z. DERIV (\xa. 1 - gf_U y y xa) z :> - gf_U' y y z) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) have "(gf_U' y y \ U' y y) ?L" using \recurrent y\ by (rule gf_U'_tendsto_U') simp then have *: "(gf_U' y y \ r) ?L" by (auto simp add: r eventually_at_left_1 dest!: tendsto_ennrealD) moreover have "(gf_U' x x \ U' x x) ?L" using \recurrent x\ by (rule gf_U'_tendsto_U') simp then have "LIM z ?L. - gf_U' x x z :> at_bot" by (simp add: ennreal_tendsto_top_eq_at_top \U' x x = \\ filterlim_uminus_at_top del: ennreal_of_enat_eSuc) then have "LIM z ?L. - gf_U' x x z :> at_infinity" by (rule filterlim_mono) (auto simp: at_bot_le_at_infinity) ultimately show "((\z. - gf_U' y y z / - gf_U' x x z) \ 0) ?L" by (intro tendsto_divide_0[where c="- r"] tendsto_intros) qed moreover have "((\z. p x y n * p y x m * z^(n + m)) \ p x y n * p y x m) ?L" by (auto intro!: tendsto_eq_intros) ultimately have "p x y n * p y x m \ 0" using le by (rule tendsto_le[OF trivial_limit_at_left_real]) with \0 < p x y n\ \0 < p y x m\ show False by (auto simp add: mult_le_0_iff) qed with \recurrent x\ show ?thesis by (simp add: pos_recurrent_def nn_integral_add) qed lemma pos_recurrent_iffI_communicating: "(y, x) \ communicating \ pos_recurrent y \ pos_recurrent x" using pos_recurrentI_communicating[of x y] pos_recurrentI_communicating[of y x] by (auto simp add: communicating_def) lemma U_le_F: "U x y \ F x y" by (auto simp: U_def F_def intro!: T.finite_measure_mono) lemma not_empty_irreducible: "C \ UNIV // communicating \ C \ {}" by (auto simp: quotient_def Image_def communicating_def) subsection \Stationary distribution\ definition stat :: "'s set \ 's measure" where "stat C = point_measure UNIV (\x. indicator C x / U' x x)" lemma sets_stat[simp]: "sets (stat C) = sets (count_space UNIV)" by (simp add: stat_def sets_point_measure) lemma space_stat[simp]: "space (stat C) = UNIV" by (simp add: stat_def space_point_measure) lemma stat_subprob: assumes C: "essential_class C" and "countable C" and pos: "\c\C. pos_recurrent c" shows "emeasure (stat C) C \ 1" proof - let ?L = "at_left (1::real)" from finite_sequence_to_countable_set[OF \countable C\] guess A . note A = this then have "(\n. emeasure (stat C) (A n)) \ emeasure (stat C) (\i. A i)" by (intro Lim_emeasure_incseq) (auto simp: incseq_Suc_iff) then have "emeasure (stat C) (\i. A i) \ 1" proof (rule LIMSEQ_le[OF _ tendsto_const], intro exI allI impI) fix n from A(1,3) have A_n: "finite (A n)" by auto from C have "C \ {}" by (simp add: essential_class_def not_empty_irreducible) then obtain x where "x \ C" by auto have "((\z. (\y\A n. gf_F x y z * ((1 - z) / (1 - gf_U y y z)))) \ (\y\A n. F x y * enn2real (1 / U' y y))) ?L" proof (intro tendsto_intros gf_F, rule lhopital_left) fix y assume "y \ A n" with \A n \ C\ have "y \ C" by auto show "((-) 1 \ 0) ?L" by (intro tendsto_eq_intros) simp_all have "recurrent y" using pos[THEN bspec, OF \y\C\] by (simp add: pos_recurrent_def) then have "U y y = 1" by (simp add: recurrent_iff_U_eq_1) show "((\x. 1 - gf_U y y x) \ 0) ?L" using gf_U[of y y] \U y y = 1\ by (intro tendsto_eq_intros) auto show "eventually (\x. 1 - gf_U y y x \ 0) ?L" using gf_G_eq_gf_U(2)[OF convergence_G_less_1, where 'z=real] by (auto intro!: eventually_at_left_1) have "eventually (\x. 0 < gf_U' y y x) ?L" by (intro eventually_at_left_1 gf_U'_pos) (simp_all add: \U y y = 1\) then show "eventually (\x. - gf_U' y y x \ 0) ?L" by eventually_elim simp show "eventually (\x. DERIV (\x. 1 - gf_U y y x) x :> - gf_U' y y x) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) show "eventually (\x. DERIV ((-) 1) x :> - 1) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros) show "((\x. - 1 / - gf_U' y y x) \ enn2real (1 / U' y y)) ?L" using \recurrent y\ by (rule inverse_gf_U'_tendsto) qed also have "(\y\A n. F x y * enn2real (1 / U' y y)) = (\y\A n. enn2real (1 / U' y y))" proof (intro sum.cong refl) fix y assume "y \ A n" with \A n \ C\ have "y \ C" by auto with \x \ C\ have "(x, y) \ communicating" by (rule essential_classD3[OF C]) with \y\C\ have "recurrent y" "(y, x) \ acc" using pos[THEN bspec, of y] by (auto simp add: pos_recurrent_def communicating_def) then have "U x y = 1" by (rule recurrent_acc) with F_le_1[of x y] U_le_F[of x y] have "F x y = 1" by simp then show "F x y * enn2real (1 / U' y y) = enn2real (1 / U' y y)" by simp qed finally have le: "(\y\A n. enn2real (1 / U' y y)) \ 1" proof (rule tendsto_le[OF trivial_limit_at_left_real tendsto_const], intro eventually_at_left_1) fix z :: real assume z: "0 < z" "z < 1" with \x \ C\ have "norm z < 1" by auto then have conv: "\x y. convergence_G x y z" by (simp add: convergence_G_less_1) have "(\y\A n. gf_F x y z / (1 - gf_U y y z)) = (\y\A n. gf_G x y z)" using \norm z < 1\ apply (intro sum.cong refl) apply (subst gf_G_eq_gf_F) apply assumption apply (subst gf_G_eq_gf_U(1)[OF conv]) apply auto done also have "\ = (\y\A n. \n. p x y n * z^n)" by (simp add: gf_G_def) also have "\ = (\i. \y\A n. p x y i *\<^sub>R z^i)" by (subst suminf_sum[OF convergence_G[OF conv]]) simp also have "\ \ (\i. z^i)" proof (intro suminf_le summable_sum convergence_G conv summable_geometric allI) fix l have "(\y\A n. p x y l *\<^sub>R z ^ l) = (\y\A n. p x y l) * z ^ l" by (simp add: sum_distrib_right) also have "\ \ z ^ l" proof (intro mult_left_le_one_le) have "(\y\A n. p x y l) = \

(\ in T x. (x ## \) !! l \ A n)" unfolding p_def using \finite (A n)\ by (subst T.finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def intro!: arg_cong2[where f=measure]) then show "(\y\A n. p x y l) \ 1" by simp qed (insert z, auto simp: sum_nonneg) finally show "(\y\A n. p x y l *\<^sub>R z ^ l) \ z ^ l" . qed fact also have "\ = 1 / (1 - z)" using sums_unique[OF geometric_sums, OF \norm z < 1\] .. finally have "(\y\A n. gf_F x y z / (1 - gf_U y y z)) \ 1 / (1 - z)" . then have "(\y\A n. gf_F x y z / (1 - gf_U y y z)) * (1 - z) \ 1" using z by (simp add: field_simps) then have "(\y\A n. gf_F x y z / (1 - gf_U y y z) * (1 - z)) \ 1" by (simp add: sum_distrib_right) then show "(\y\A n. gf_F x y z * ((1 - z) / (1 - gf_U y y z))) \ 1" by simp qed from A_n have "emeasure (stat C) (A n) = (\y\A n. emeasure (stat C) {y})" by (intro emeasure_eq_sum_singleton) simp_all also have "\ = (\y\A n. inverse (U' y y))" unfolding stat_def U'_def using A(1)[of n] apply (intro sum.cong refl) apply (subst emeasure_point_measure_finite2) apply (auto simp: divide_ennreal_def Collect_conv_if) done also have "\ = ennreal (\y\A n. enn2real (1 / U' y y))" apply (subst sum_ennreal[symmetric], simp) proof (intro sum.cong refl) fix y assume "y \ A n" with \A n \ C\ pos have "pos_recurrent y" by auto with one_le_integral_t[of y] obtain r where "U' y y = ennreal r" "1 \ U' y y" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: pos_recurrent_def nn_integral_add) then show "inverse (U' y y) = ennreal (enn2real (1 / U' y y))" by (simp add: ennreal_1[symmetric] divide_ennreal inverse_ennreal inverse_eq_divide del: ennreal_1) qed also have "\ \ 1" using le by simp finally show "emeasure (stat C) (A n) \ 1" . qed with A show ?thesis by simp qed lemma emeasure_stat_not_C: assumes "y \ C" shows "emeasure (stat C) {y} = 0" unfolding stat_def using \y \ C\ by (subst emeasure_point_measure_finite2) auto definition stationary_distribution :: "'s pmf \ bool" where "stationary_distribution N \ N = bind_pmf N K" lemma stationary_distributionI: assumes le: "\y. (\x. pmf (K x) y \measure_pmf N) \ pmf N y" shows "stationary_distribution N" unfolding stationary_distribution_def proof (rule pmf_eqI antisym)+ fix i show "pmf (bind_pmf N K) i \ pmf N i" by (simp add: pmf_bind le) define \ where "\ = N \ (\i\N. set_pmf (K i))" then have \: "countable \" by (auto intro: countable_set_pmf) then interpret N: sigma_finite_measure "count_space \" by (rule sigma_finite_measure_count_space_countable) interpret pN: pair_sigma_finite N "count_space \" by unfold_locales have measurable_pmf[measurable]: "(\(x, y). pmf (K x) y) \ borel_measurable (N \\<^sub>M count_space \)" unfolding measurable_split_conv apply (rule measurable_compose_countable'[OF _ measurable_snd]) apply (rule measurable_compose[OF measurable_fst]) apply (simp_all add: \) done { assume *: "(\y. pmf (K y) i \N) < pmf N i" have "0 \ (\y. pmf (K y) i \N)" by (intro integral_nonneg_AE) simp with * have i: "i \ set_pmf N" "i \ \" by (auto simp: set_pmf_iff \_def not_le[symmetric]) from * have "0 < pmf N i - (\y. pmf (K y) i \N)" by (simp add: field_simps) also have "\ = (\t. (pmf N i - (\y. pmf (K y) i \N)) * indicator {i} t \count_space \)" by (simp add: i) also have "\ \ (\t. pmf N t - \y. pmf (K y) t \N \count_space \)" using le by (intro integral_mono integrable_diff) (auto simp: i pmf_bind[symmetric] integrable_pmf field_simps split: split_indicator) also have "\ = (\t. pmf N t \count_space \) - (\t. \y. pmf (K y) t \N \count_space \)" by (subst Bochner_Integration.integral_diff) (auto intro!: integrable_pmf simp: pmf_bind[symmetric]) also have "(\t. \y. pmf (K y) t \N \count_space \) = (\y. \t. pmf (K y) t \count_space \ \N)" apply (intro pN.Fubini_integral integrable_iff_bounded[THEN iffD2] conjI) apply (auto simp add: N.nn_integral_fst[symmetric] nn_integral_eq_integral integrable_pmf) unfolding less_top[symmetric] unfolding infinity_ennreal_def[symmetric] apply (intro integrableD) apply (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: AE_measure_pmf_iff integral_nonneg_AE integral_pmf) done also have "(\y. \t. pmf (K y) t \count_space \ \N) = (\y. 1 \N)" by (intro integral_cong_AE) (auto simp: AE_measure_pmf_iff integral_pmf \_def intro!: measure_pmf.prob_eq_1[THEN iffD2]) finally have False using measure_pmf.prob_space[of N] by (simp add: integral_pmf field_simps not_le[symmetric]) } then show "pmf N i \ pmf (bind_pmf N K) i" by (auto simp: pmf_bind not_le[symmetric]) qed lemma stationary_distribution_iterate: assumes N: "stationary_distribution N" shows "ennreal (pmf N y) = (\\<^sup>+x. p x y n \N)" proof (induct n arbitrary: y) have [simp]: "\x y. ennreal (if x = y then 1 else 0) = indicator {y} x" by simp case 0 then show ?case by (simp add: p_0 pmf.rep_eq measure_pmf.emeasure_eq_measure) next case (Suc n) with N show ?case apply (simp add: nn_integral_eq_integral[symmetric] p_le_1 p_Suc' measure_pmf.integrable_const_bound[where B=1]) apply (subst nn_integral_bind[symmetric, where B="count_space UNIV"]) apply (auto simp: stationary_distribution_def measure_pmf_bind[symmetric] simp del: measurable_pmf_measure1) done qed lemma stationary_distribution_iterate': assumes "stationary_distribution N" shows "measure N {y} = (\x. p x y n \N)" using stationary_distribution_iterate[OF assms] by (subst (asm) nn_integral_eq_integral) (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: p_le_1 pmf.rep_eq) lemma stationary_distributionD: assumes C: "essential_class C" "countable C" assumes N: "stationary_distribution N" "N \ C" shows "\x\C. pos_recurrent x" "measure_pmf N = stat C" proof - have integrable_K: "\f x. integrable N (\s. pmf (K s) (f x))" by (rule measure_pmf.integrable_const_bound[where B=1]) (simp_all add: pmf_le_1) have measure_C: "measure N C = 1" and ae_C: "AE x in N. x \ C" using N C measure_pmf.prob_eq_1[of C] by (auto simp: AE_measure_pmf_iff) have integrable_p: "\n y. integrable N (\x. p x y n)" by (rule measure_pmf.integrable_const_bound[where B=1]) (simp_all add: p_le_1) { fix e :: real assume "0 < e" then have [simp]: "0 \ e" by simp have "\A\C. finite A \ 1 - e < measure N A" proof (rule ccontr) assume contr: "\ (\A \ C. finite A \ 1 - e < measure N A)" from finite_sequence_to_countable_set[OF \countable C\] guess F . note F = this then have *: "(\n. measure N (F n)) \ measure N (\i. F i)" by (intro measure_pmf.finite_Lim_measure_incseq) (auto simp: incseq_Suc_iff) with F contr have "measure N (\i. F i) \ 1 - e" by (intro LIMSEQ_le[OF * tendsto_const]) (auto simp: not_less) with F \0 < e\ show False by (simp add: measure_C) qed then obtain A where "A \ C" "finite A" and e: "1 - e < measure N A" by auto { fix y n assume "y \ C" from N(1) have "measure N {y} = (\x. p x y n \N)" by (rule stationary_distribution_iterate') also have "\ \ (\x. p x y n * indicator A x + indicator (C - A) x \N)" using ae_C \A \ C\ by (intro integral_mono_AE) (auto elim!: eventually_mono intro!: integral_add integral_indicator p_le_1 integrable_real_mult_indicator integrable_add split: split_indicator simp: integrable_p less_top[symmetric] top_unique) also have "\ = (\x. p x y n * indicator A x \N) + measure N (C - A)" using ae_C \A \ C\ apply (subst Bochner_Integration.integral_add) apply (auto elim!: eventually_mono intro!: integral_add integral_indicator p_le_1 integrable_real_mult_indicator split: split_indicator simp: integrable_p less_top[symmetric] top_unique) done also have "\ \ (\x. p x y n * indicator A x \N) + e" using e \A \ C\ by (simp add: measure_pmf.finite_measure_Diff measure_C) finally have "measure N {y} \ (\x. p x y n * indicator A x \N) + e" . then have "emeasure N {y} \ ennreal (\x. p x y n * indicator A x \N) + e" by (simp add: measure_pmf.emeasure_eq_measure ennreal_plus[symmetric] del: ennreal_plus) also have "\ = (\\<^sup>+x. ennreal (p x y n) * indicator A x \N) + e" by (subst nn_integral_eq_integral[symmetric]) (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: abs_mult p_le_1 mult_le_one ennreal_indicator ennreal_mult) finally have "emeasure N {y} \ (\\<^sup>+x. ennreal (p x y n) * indicator A x \N) + e" . } note v_le = this { fix y and z :: real assume y: "y \ C" and z: "0 < z" "z < 1" have summable_int_p: "summable (\n. (\ x. p x y n * indicator A x \N) * (1 - z) * z ^ n)" using \y\C\ z \A \ C\ by (auto intro!: summable_comparison_test[OF _ summable_mult[OF summable_geometric[of z], of 1]] exI[of _ 0] mult_le_one measure_pmf.integral_le_const integrable_real_mult_indicator integrable_p AE_I2 p_le_1 simp: abs_mult integral_nonneg_AE) from y z have sums_y: "(\n. measure N {y} * (1 - z) * z ^ n) sums measure N {y}" using sums_mult[OF geometric_sums[of z], of "measure N {y} * (1 - z)"] by simp then have "emeasure N {y} = ennreal (\n. (measure N {y} * (1 - z)) * z ^ n)" by (auto simp add: sums_unique[symmetric] measure_pmf.emeasure_eq_measure) also have "\ = (\n. emeasure N {y} * (1 - z) * z ^ n)" using z summable_mult[OF summable_geometric[of z], of "measure_pmf.prob N {y} * (1 - z)"] by (subst suminf_ennreal[symmetric]) (auto simp: measure_pmf.emeasure_eq_measure ennreal_mult[symmetric] ennreal_suminf_neq_top) also have "\ \ (\n. ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) + e) * (1 - z) * z ^ n)" using \y\C\ z \A \ C\ by (intro suminf_le mult_right_mono v_le allI) (auto simp: measure_pmf.emeasure_eq_measure) also have "\ = (\n. (\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * (1 - z) * z ^ n) + e" using \0 < e\ z sums_mult[OF geometric_sums[of z], of "e * (1 - z)"] \0 \z<1\ by (simp add: distrib_right suminf_add[symmetric] ennreal_suminf_cmult[symmetric] ennreal_mult[symmetric] suminf_ennreal_eq sums_unique[symmetric] del: ennreal_suminf_cmult) also have "\ = (\n. ennreal (1 - z) * ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * z ^ n)) + e" by (simp add: ac_simps) also have "\ = ennreal (1 - z) * (\n. ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * z ^ n)) + e" using z by (subst ennreal_suminf_cmult) simp_all also have "(\n. ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * z ^ n)) = (\n. (\\<^sup>+x. ennreal (p x y n * z ^ n) * indicator A x \N))" using z by (simp add: ac_simps nn_integral_cmult[symmetric] ennreal_mult) also have "\ = (\\<^sup>+x. ennreal (gf_G x y z) * indicator A x \N)" using z apply (subst nn_integral_suminf[symmetric]) apply (auto simp add: gf_G_def simp del: suminf_ennreal intro!: ennreal_mult_right_cong suminf_ennreal2 nn_integral_cong) apply (intro summable_comparison_test[OF _ summable_mult[OF summable_geometric[of z], of 1]] impI) apply (simp_all add: abs_mult p_le_1 mult_le_one power_le_one split: split_indicator) done also have "\ = (\\<^sup>+x. ennreal (gf_F x y z * gf_G y y z) * indicator A x \N)" using z by (intro nn_integral_cong) (simp add: gf_G_eq_gf_F[symmetric]) also have "\ = ennreal (gf_G y y z) * (\\<^sup>+x. ennreal (gf_F x y z) * indicator A x \N)" using z by (subst nn_integral_cmult[symmetric]) (simp_all add: gf_G_nonneg gf_F_nonneg ac_simps ennreal_mult) also have "\ = ennreal (1 / (1 - gf_U y y z)) * (\\<^sup>+x. ennreal (gf_F x y z) * indicator A x \N)" using z \y \ C\ by (subst gf_G_eq_gf_U) (auto intro!: convergence_G_less_1) finally have "emeasure N {y} \ ennreal ((1 - z) / (1 - gf_U y y z)) * (\\<^sup>+x. gf_F x y z * indicator A x \N) + e" using z by (subst (asm) mult.assoc[symmetric]) (simp add: ennreal_indicator[symmetric] ennreal_mult'[symmetric] gf_F_nonneg) then have "measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e" using z by (subst (asm) nn_integral_eq_integral[OF measure_pmf.integrable_const_bound[where B=1]]) (auto simp: gf_F_nonneg gf_U_le_1 gf_F_le_1 measure_pmf.emeasure_eq_measure mult_le_one ennreal_mult''[symmetric] ennreal_plus[symmetric] simp del: ennreal_plus) } then have "\A \ C. finite A \ (\y\C. \z. 0 < z \ z < 1 \ measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e)" using \A \ C\ \finite A\ by auto } note eps = this { fix y A assume "y \ C" "finite A" "A \ C" then have "((\z. \x. gf_F x y z * indicator A x \N) \ \x. F x y * indicator A x \N) (at_left 1)" by (subst (1 2) integral_measure_pmf[of A]) (auto intro!: tendsto_intros gf_F simp: indicator_eq_0_iff) } note int_gf_F = this have all_recurrent: "\y\C. recurrent y" proof (rule ccontr) assume "\ (\y\C. recurrent y)" then obtain x where "x \ C" "\ recurrent x" by auto then have transient: "\x. x \ C \ \ recurrent x" using C by (auto simp: essential_class_def recurrent_iffI_communicating[symmetric] elim!: quotientE) { fix y assume "y \ C" with transient have "U y y < 1" by (metis recurrent_iff_U_eq_1 U_cases) have "measure N {y} \ 0" proof (rule dense_ge) fix e :: real assume "0 < e" from eps[OF this] \y \ C\ obtain A where A: "finite A" "A \ C" and le: "\z. 0 < z \ z < 1 \ measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e" by auto have "((\z. (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e) \ (1 - 1) / (1 - U y y) * (\x. F x y * indicator A x \N) + e) (at_left (1::real))" using A \U y y < 1\ \y \ C\ by (intro tendsto_intros gf_U int_gf_F) auto then have 1: "((\z. (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e) \ e) (at_left (1::real))" by simp with le show "measure N {y} \ e" by (intro tendsto_le[OF trivial_limit_at_left_real _ tendsto_const]) (auto simp: eventually_at_left_1) qed then have "measure N {y} = 0" by (intro antisym measure_nonneg) } then have "emeasure N C = 0" by (subst emeasure_countable_singleton) (auto simp: measure_pmf.emeasure_eq_measure nn_integral_0_iff_AE ae_C C) then show False using \measure N C = 1\ by (simp add: measure_pmf.emeasure_eq_measure) qed then have "\x. x \ C \ U x x = 1" by (metis recurrent_iff_U_eq_1) { fix y assume "y \ C" then have "U y y = 1" "recurrent y" using \y \ C \ U y y = 1\ all_recurrent by auto have "measure N {y} \ enn2real (1 / U' y y)" proof (rule field_le_epsilon) fix e :: real assume "0 < e" from eps[OF \0 < e\] \y \ C\ obtain A where A: "finite A" "A \ C" and le: "\z. 0 < z \ z < 1 \ measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e" by auto let ?L = "at_left (1::real)" have "((\z. (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e) \ enn2real (1 / U' y y) * (\x. F x y * indicator A x \N) + e) ?L" proof (intro tendsto_add tendsto_const tendsto_mult int_gf_F, rule lhopital_left[where f'="\x. - 1" and g'="\z. - gf_U' y y z"]) show "((-) 1 \ 0) ?L" "((\x. 1 - gf_U y y x) \ 0) ?L" using gf_U[of y y] by (auto intro!: tendsto_eq_intros simp: \U y y = 1\) show "y \ C" "finite A" "A \ C" by fact+ show "eventually (\x. 1 - gf_U y y x \ 0) ?L" using gf_G_eq_gf_U(2)[OF convergence_G_less_1, where 'z=real] by (auto intro!: eventually_at_left_1) show "((\x. - 1 / - gf_U' y y x) \ enn2real (1 / U' y y)) ?L" using \recurrent y\ by (rule inverse_gf_U'_tendsto) have "eventually (\x. 0 < gf_U' y y x) ?L" by (intro eventually_at_left_1 gf_U'_pos) (simp_all add: \U y y = 1\) then show "eventually (\x. - gf_U' y y x \ 0) ?L" by eventually_elim simp show "eventually (\x. DERIV (\x. 1 - gf_U y y x) x :> - gf_U' y y x) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) show "eventually (\x. DERIV ((-) 1) x :> - 1) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros) qed then have "measure N {y} \ enn2real (1 / U' y y) * (\x. F x y * indicator A x \N) + e" by (rule tendsto_le[OF trivial_limit_at_left_real _ tendsto_const]) (intro eventually_at_left_1 le) then have "measure N {y} - e \ enn2real (1 / U' y y) * (\x. F x y * indicator A x \N)" by simp also have "\ \ enn2real (1 / U' y y)" using A by (intro mult_left_le measure_pmf.integral_le_const measure_pmf.integrable_const_bound[where B=1]) (auto simp: mult_le_one F_le_1 U'_def) finally show "measure N {y} \ enn2real (1 / U' y y) + e" by simp qed } note measure_y_le = this show pos: "\y\C. pos_recurrent y" proof (rule ccontr) assume "\ (\y\C. pos_recurrent y)" then obtain x where x: "x \ C" "\ pos_recurrent x" by auto { fix y assume "y \ C" with x have "\ pos_recurrent y" using C by (auto simp: essential_class_def pos_recurrent_iffI_communicating[symmetric] elim!: quotientE) with all_recurrent \y \ C\ have "enn2real (1 / U' y y) = 0" by (simp add: pos_recurrent_def nn_integral_add) with measure_y_le[OF \y \ C\] have "measure N {y} = 0" by (auto intro!: antisym simp: pos_recurrent_def) } then have "emeasure N C = 0" by (subst emeasure_countable_singleton) (auto simp: C ae_C measure_pmf.emeasure_eq_measure nn_integral_0_iff_AE) then show False using \measure N C = 1\ by (simp add: measure_pmf.emeasure_eq_measure) qed { fix A :: "'s set" assume [simp]: "countable A" have "emeasure N A = (\\<^sup>+x. emeasure N {x} \count_space A)" by (intro emeasure_countable_singleton) auto also have "\ \ (\\<^sup>+x. emeasure (stat C) {x} \count_space A)" proof (intro nn_integral_mono) fix y assume "y \ space (count_space A)" show "emeasure N {y} \ emeasure (stat C) {y}" proof cases assume "y \ C" with pos have "pos_recurrent y" by auto with one_le_integral_t[of y] obtain r where r: "U' y y = ennreal r" "1 \ U' y y" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: pos_recurrent_def nn_integral_add) from measure_y_le[OF \y \ C\] have "emeasure N {y} \ ennreal (enn2real (1 / U' y y))" by (simp add: measure_pmf.emeasure_eq_measure) also have "\ = emeasure (stat C) {y}" unfolding stat_def using \y \ C\ r by (subst emeasure_point_measure_finite2) (auto simp add: ennreal_1[symmetric] divide_ennreal inverse_ennreal inverse_eq_divide ennreal_mult[symmetric] simp del: ennreal_1) finally show "emeasure N {y} \ emeasure (stat C) {y}" by simp next assume "y \ C" with ae_C have "emeasure N {y} = 0" by (subst AE_iff_measurable[symmetric, where P="\x. x \ y"]) (auto elim!: eventually_mono) moreover have "emeasure (stat C) {y} = 0" using emeasure_stat_not_C[OF \y \ C\] . ultimately show ?thesis by simp qed qed also have "\ = emeasure (stat C) A" by (intro emeasure_countable_singleton[symmetric]) auto finally have "emeasure N A \ emeasure (stat C) A" . } note N_le_C = this from stat_subprob[OF C(1) \countable C\ pos] N_le_C[OF \countable C\] \measure N C = 1\ have stat_C_eq_1: "emeasure (stat C) C = 1" by (auto simp add: measure_pmf.emeasure_eq_measure one_ennreal_def) moreover have "emeasure (stat C) (UNIV - C) = 0" by (subst AE_iff_measurable[symmetric, where P="\x. x \ C"]) (auto simp: stat_def AE_point_measure sets_point_measure space_point_measure split: split_indicator cong del: AE_cong) ultimately have "emeasure (stat C) (space (stat C)) = 1" using plus_emeasure[of C "stat C" "UNIV - C"] by (simp add: Un_absorb1) interpret stat: prob_space "stat C" by standard fact show "measure_pmf N = stat C" proof (rule measure_eqI_countable_AE) show "sets N = UNIV" "sets (stat C) = UNIV" by auto show "countable C" "AE x in N. x \ C" and ae_stat: "AE x in stat C. x \ C" using C ae_C stat_C_eq_1 by (auto intro!: stat.AE_prob_1 simp: stat.emeasure_eq_measure) { assume "\x. emeasure N {x} \ emeasure (stat C) {x}" then obtain x where [simp]: "emeasure N {x} \ emeasure (stat C) {x}" by auto with N_le_C[of "{x}"] have x: "emeasure N {x} < emeasure (stat C) {x}" by (auto simp: less_le) have "1 = emeasure N {x} + emeasure N (C - {x})" using ae_C by (subst plus_emeasure) (auto intro!: measure_pmf.emeasure_eq_1_AE) also have "\ < emeasure (stat C) {x} + emeasure (stat C) (C - {x})" using x N_le_C[of "C - {x}"] C ae_C by (simp add: stat.emeasure_eq_measure measure_pmf.emeasure_eq_measure ennreal_plus[symmetric] ennreal_less_iff del: ennreal_plus) also have "\ = 1" using ae_stat by (subst plus_emeasure) (auto intro!: stat.emeasure_eq_1_AE) finally have False by simp } then show "\x. emeasure N {x} = emeasure (stat C) {x}" by auto qed qed lemma measure_point_measure_singleton: "x \ A \ measure (point_measure A X) {x} = enn2real (X x)" unfolding measure_def by (subst emeasure_point_measure_finite2) auto lemma stationary_distribution_imp_int_t: assumes C: "essential_class C" "countable C" "stationary_distribution N" "N \ C" assumes x: "x \ C" shows "U' x x = 1 / ennreal (pmf N x)" proof - from stationary_distributionD[OF C] have "measure_pmf N = stat C" and *: "\x\C. pos_recurrent x" by auto show ?thesis unfolding \measure_pmf N = stat C\ pmf.rep_eq stat_def using *[THEN bspec, OF x] x apply (simp add: measure_point_measure_singleton) apply (cases "U' x x") subgoal for r by (cases "r = 0") (simp_all add: divide_ennreal_def inverse_ennreal) apply simp done qed definition "period_set x = {i. 0 < i \ 0 < p x x i }" definition "period C = (SOME d. \x\C. d = Gcd (period_set x))" lemma Gcd_period_set_invariant: assumes c: "(x, y) \ communicating" shows "Gcd (period_set x) = Gcd (period_set y)" proof - { fix x y n assume c: "(x, y) \ communicating" "x \ y" and n: "n \ period_set x" from c obtain l k where "0 < p x y l" "0 < p y x k" by (auto simp: communicating_def dest!: accD_pos) moreover with \x \ y\ have "l \ 0 \ k \ 0" by (intro notI conjI) (auto simp: p_0) ultimately have pos: "0 < l" "0 < k" and l: "0 < p x y l" and k: "0 < p y x k" by auto from mult_pos_pos[OF k l] prob_reachable_le[of k "k + l" y x y] c have k_l: "0 < p y y (k + l)" by simp then have "Gcd (period_set y) dvd k + l" using pos by (auto intro!: Gcd_dvd_nat simp: period_set_def) moreover from n have "0 < p x x n" "0 < n" by (auto simp: period_set_def) from mult_pos_pos[OF k this(1)] prob_reachable_le[of k "k + n" y x x] c have "0 < p y x (k + n)" by simp from mult_pos_pos[OF this(1) l] prob_reachable_le[of "k + n" "(k + n) + l" y x y] c have "0 < p y y (k + n + l)" by simp then have "Gcd (period_set y) dvd (k + l) + n" using pos by (auto intro!: Gcd_dvd_nat simp: period_set_def ac_simps) ultimately have "Gcd (period_set y) dvd n" by (metis dvd_add_left_iff add.commute) } note this[of x y] this[of y x] c moreover have "(y, x) \ communicating" using c by (simp add: communicating_def) ultimately show ?thesis by (auto intro: dvd_antisym Gcd_greatest Gcd_dvd) qed lemma period_eq: assumes "C \ UNIV // communicating" "x \ C" shows "period C = Gcd (period_set x)" unfolding period_def using assms by (rule_tac someI2[where a="Gcd (period_set x)"]) (auto intro!: Gcd_period_set_invariant irreducibleD) definition "aperiodic C \ C \ UNIV // communicating \ period C = 1" definition "not_ephemeral C \ C \ UNIV // communicating \ \ (\x. C = {x} \ p x x 1 = 0)" lemma not_ephemeralD: assumes C: "not_ephemeral C" "x \ C" shows "\n>0. 0 < p x x n" proof cases assume "\x. C = {x}" with \x \ C\ have "C = {x}" by auto with C p_nonneg[of x x 1] have "0 < p x x 1" by (auto simp: not_ephemeral_def less_le) with \C = {x}\ show ?thesis by auto next from C have irr: "C \ UNIV // communicating" by (auto simp: not_ephemeral_def) assume "\(\x. C = {x})" then have "\x. C \ {x}" by auto with \x \ C\ obtain y where "y \ C" "x \ y" by blast with irreducibleD[OF irr, of x y] C \x \ C\ have c: "(x, y) \ communicating" by auto with accD_pos[of x y] accD_pos[of y x] obtain k l where pos: "0 < p x y k" "0 < p y x l" by (auto simp: communicating_def) with \x \ y\ have "l \ 0" by (intro notI) (auto simp: p_0) have "0 < p x y k * p y x (k + l - k)" using pos by auto also have "p x y k * p y x (k + l - k) \ p x x (k + l)" using prob_reachable_le[of "k" "k + l" x y x] c by auto finally show ?thesis using \l \ 0\ \x \ C\ by (auto intro!: exI[of _ "k + l"]) qed lemma not_ephemeralD_pos_period: assumes C: "not_ephemeral C" shows "0 < period C" proof - from C not_empty_irreducible[of C] obtain x where "x \ C" by (auto simp: not_ephemeral_def) from not_ephemeralD[OF C this] obtain n where n: "0 < p x x n" "0 < n" by auto have C': "C \ UNIV // communicating" using C by (auto simp: not_ephemeral_def) have "period C \ 0" unfolding period_eq [OF C' \x \ C\] using n by (auto simp: period_set_def) then show ?thesis by auto qed lemma period_posD: assumes C: "C \ UNIV // communicating" and "0 < period C" "x \ C" shows "\n>0. 0 < p x x n" proof - from \0 < period C\ have "period C \ 0" by auto then show ?thesis unfolding period_eq [OF C \x \ C\] unfolding period_set_def by auto qed lemma not_ephemeralD_pos_period': assumes C: "C \ UNIV // communicating" shows "not_ephemeral C \ 0 < period C" proof (auto dest!: not_ephemeralD_pos_period intro: C) from C not_empty_irreducible[of C] obtain x where "x \ C" by (auto simp: not_ephemeral_def) assume "0 < period C" then show "not_ephemeral C" apply (auto simp: not_ephemeral_def C) oops \ \should be easy to finish\ lemma eventually_periodic: assumes C: "C \ UNIV // communicating" "0 < period C" "x \ C" shows "eventually (\m. 0 < p x x (m * period C)) sequentially" proof - from period_posD[OF assms] obtain n where n: "0 < p x x n" "0 < n" by auto have C': "C \ UNIV // communicating" using C by auto have "period C \ 0" unfolding period_eq [OF C' \x \ C\] using n by (auto simp: period_set_def) have "eventually (\m. m * Gcd (period_set x) \ (period_set x)) sequentially" proof (rule eventually_mult_Gcd) show "n > 0" "n \ period_set x" using n by (auto simp add: period_set_def) fix k l assume "k \ period_set x" "l \ period_set x" then have "0 < p x x k * p x x l" "0 < l" "0 < k" by (auto simp: period_set_def) moreover have "p x x k * p x x l \ p x x (k + l)" using prob_reachable_le[of k "k + l" x x x] \x \ C\ by auto ultimately show "k + l \ period_set x" using \0 < l\ by (auto simp: period_set_def) qed with eventually_ge_at_top[of 1] show "eventually (\m. 0 < p x x (m * period C)) sequentially" by eventually_elim (insert \period C \ 0\ period_eq[OF C' \x \ C\, symmetric], auto simp: period_set_def) qed lemma aperiodic_eventually_recurrent: "aperiodic C \ C \ UNIV // communicating \ (\x\C. eventually (\m. 0 < p x x m) sequentially)" proof safe fix x assume "x \ C" "aperiodic C" with eventually_periodic[of C x] show "eventually (\m. 0 < p x x m) sequentially" by (auto simp add: aperiodic_def) next assume "\x\C. eventually (\m. 0 < p x x m) sequentially" and C: "C \ UNIV // communicating" moreover from not_empty_irreducible[OF C] obtain x where "x \ C" by auto ultimately obtain N where "\M. M\N \ 0 < p x x M" by (auto simp: eventually_sequentially) then have "{N <..} \ period_set x" by (auto simp: period_set_def) from C show "aperiodic C" unfolding period_eq [OF C \x \ C\] aperiodic_def proof show "Gcd (period_set x) = 1" proof (rule Gcd_eqI) from one_dvd show "1 dvd q" for q :: nat . fix m assume "\q. q \ period_set x \ m dvd q" moreover from \{N <..} \ period_set x\ have "{Suc N, Suc (Suc N)} \ period_set x" by auto ultimately have "m dvd Suc (Suc N)" and "m dvd Suc N" by auto then have "m dvd Suc (Suc N) - Suc N" by (rule dvd_diff_nat) then show "is_unit m" by simp qed simp qed qed (simp add: aperiodic_def) lemma stationary_distributionD_emeasure: assumes N: "stationary_distribution N" shows "emeasure N A = (\\<^sup>+s. emeasure (K s) A \N)" proof - have "prob_space (measure_pmf N)" by intro_locales then interpret subprob_space "measure_pmf N" by (rule prob_space_imp_subprob_space) show ?thesis unfolding measure_pmf.emeasure_eq_measure apply (subst N[unfolded stationary_distribution_def]) apply (simp add: measure_pmf_bind) apply (subst measure_pmf.measure_bind[where N="count_space UNIV"]) apply (rule measurable_compose[OF _ measurable_measure_pmf]) apply (auto intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1]) done qed lemma communicatingD1: "C \ UNIV // communicating \ (a, b) \ communicating \ a \ C \ b \ C" by (auto elim!: quotientE) (auto simp add: communicating_def) lemma communicatingD2: "C \ UNIV // communicating \ (a, b) \ communicating \ b \ C \ a \ C" by (auto elim!: quotientE) (auto simp add: communicating_def) lemma acc_iff: "(x, y) \ acc \ (\n. 0 < p x y n)" by (blast intro: accD_pos accI_pos) lemma communicating_iff: "(x, y) \ communicating \ (\n. 0 < p x y n) \ (\n. 0 < p y x n)" by (auto simp add: acc_iff communicating_def) end context MC_pair begin lemma p_eq_p1_p2: "p (x1, x2) (y1, y2) n = K1.p x1 y1 n * K2.p x2 y2 n" unfolding p_def K1.p_def K2.p_def by (subst prod_eq_prob_T) (auto intro!: arg_cong2[where f=measure] split: nat.splits simp: Stream_snth) lemma P_accD: assumes "((x1, x2), (y1, y2)) \ acc"shows "(x1, y1) \ K1.acc" "(x2, y2) \ K2.acc" using assms by (auto simp: acc_iff K1.acc_iff K2.acc_iff p_eq_p1_p2 zero_less_mult_iff not_le[of 0, symmetric] cong: conj_cong) lemma aperiodicI_pair: assumes C1: "K1.aperiodic C1" and C2: "K2.aperiodic C2" shows "aperiodic (C1 \ C2)" unfolding aperiodic_eventually_recurrent proof safe from C1[unfolded K1.aperiodic_eventually_recurrent] C2[unfolded K2.aperiodic_eventually_recurrent] have C1: "C1 \ UNIV // K1.communicating" and C2: "C2 \ UNIV // K2.communicating" and ev: "\x. x \ C1 \ eventually (\m. 0 < K1.p x x m) sequentially" "\x. x \ C2 \ eventually (\m. 0 < K2.p x x m) sequentially" by auto { fix x1 x2 assume x: "x1 \ C1" "x2 \ C2" from ev(1)[OF x(1)] ev(2)[OF x(2)] show "eventually (\m. 0 < p (x1, x2) (x1, x2) m) sequentially" by eventually_elim (simp add: p_eq_p1_p2 x) } { fix x1 x2 y1 y2 assume acc: "(x1, y1) \ K1.acc" "(x2, y2) \ K2.acc" "x1 \ C1" "y1 \ C1" "x2 \ C2" "y2 \ C2" then obtain k l where "0 < K1.p x1 y1 l" "0 < K2.p x2 y2 k" by (auto dest!: K1.accD_pos K2.accD_pos) with acc ev(1)[of y1] ev(2)[of y2] have "eventually (\m. 0 < K1.p x1 y1 l * K1.p y1 y1 m \ 0 < K2.p x2 y2 k * K2.p y2 y2 m) sequentially" by (auto elim: eventually_elim2) then have "eventually (\m. 0 < K1.p x1 y1 (m + l) \ 0 < K2.p x2 y2 (m + k)) sequentially" proof eventually_elim fix m assume "0 < K1.p x1 y1 l * K1.p y1 y1 m \ 0 < K2.p x2 y2 k * K2.p y2 y2 m" with acc K1.prob_reachable_le[of l "l + m" x1 y1 y1] K2.prob_reachable_le[of k "k + m" x2 y2 y2] show "0 < K1.p x1 y1 (m + l) \ 0 < K2.p x2 y2 (m + k)" by (auto simp add: ac_simps) qed then have "eventually (\m. 0 < K1.p x1 y1 m \ 0 < K2.p x2 y2 m) sequentially" unfolding eventually_conj_iff by (subst (asm) (1 2) eventually_sequentially_seg) (auto elim: eventually_elim2) then obtain N where "0 < K1.p x1 y1 N" "0 < K2.p x2 y2 N" by (auto simp: eventually_sequentially) with acc have "0 < p (x1, x2) (y1, y2) N" by (auto simp add: p_eq_p1_p2) with acc have "((x1, x2), (y1, y2)) \ acc" by (auto intro!: accI_pos) } note 1 = this { fix x1 x2 y1 y2 assume acc:"((x1, x2), (y1, y2)) \ acc" moreover from acc obtain k where "0 < p (x1, x2) (y1, y2) k" by (auto dest!: accD_pos) ultimately have "(x1, y1) \ K1.acc \ (x2, y2) \ K2.acc" by (subst (asm) p_eq_p1_p2) (auto intro!: K1.accI_pos K2.accI_pos simp: zero_less_mult_iff not_le[of 0, symmetric]) } note 2 = this from K1.not_empty_irreducible[OF C1] K2.not_empty_irreducible[OF C2] obtain x1 x2 where xC: "x1 \ C1" "x2 \ C2" by auto show "C1 \ C2 \ UNIV // communicating" apply (simp add: quotient_def Image_def) apply (safe intro!: exI[of _ x1] exI[of _ x2]) proof - fix y1 y2 assume yC: "y1 \ C1" "y2 \ C2" from K1.irreducibleD[OF C1 \x1 \ C1\ \y1 \ C1\] K2.irreducibleD[OF C2 \x2 \ C2\ \y2 \ C2\] show "((x1, x2), (y1, y2)) \ communicating" using 1[of x1 y1 x2 y2] 1[of y1 x1 y2 x2] xC yC by (auto simp: communicating_def K1.communicating_def K2.communicating_def) next fix y1 y2 assume "((x1, x2), (y1, y2)) \ communicating" with 2[of x1 x2 y1 y2] 2[of y1 y2 x1 x2] have "(x1, y1) \ K1.communicating" "(x2, y2) \ K2.communicating" by (auto simp: communicating_def K1.communicating_def K2.communicating_def) with xC show "y1 \ C1" "y2 \ C2" using K1.communicatingD1[OF C1] K2.communicatingD1[OF C2] by auto qed qed lemma stationary_distributionI_pair: assumes N1: "K1.stationary_distribution N1" assumes N2: "K2.stationary_distribution N2" shows "stationary_distribution (pair_pmf N1 N2)" unfolding stationary_distribution_def unfolding Kp_def pair_pmf_def apply (subst N1[unfolded K1.stationary_distribution_def]) apply (subst N2[unfolded K2.stationary_distribution_def]) apply (simp add: bind_assoc_pmf bind_return_pmf) apply (subst bind_commute_pmf[of N2]) apply simp done end context MC_syntax begin lemma stationary_distribution_imp_limit: assumes C: "aperiodic C" "essential_class C" "countable C" and N: "stationary_distribution N" "N \ C" assumes [simp]: "y \ C" shows "(\n. \x. \p y x n - pmf N x\ \count_space C) \ 0" (is "?L \ 0") proof - from \essential_class C\ have C_comm: "C \ UNIV // communicating" by (simp add: essential_class_def) define K' where "K' = (\Some x \ map_pmf Some (K x) | None \ map_pmf Some N)" interpret K2: MC_syntax K' . interpret KN: MC_pair K K' . from stationary_distributionD[OF C(2,3) N] have pos: "\x. x \ C \ pos_recurrent x" and "measure_pmf N = stat C" by auto have pos: "\x. x \ C \ 0 < emeasure N {x}" using pos unfolding stat_def \measure_pmf N = stat C\ by (subst emeasure_point_measure_finite2) (auto simp: U'_def pos_recurrent_def nn_integral_add ennreal_zero_less_divide less_top) then have rpos: "\x. x \ C \ 0 < pmf N x" by (simp add: measure_pmf.emeasure_eq_measure pmf.rep_eq) have eq: "\x y. (if x = y then 1 else 0) = indicator {y} x" by auto have intK: "\f x. (\x. (f x :: real) \K' (Some x)) = (\x. f (Some x) \K x)" by (simp add: K'_def integral_distr map_pmf_rep_eq) { fix m and x y :: 's have "K2.p (Some x) (Some y) m = p x y m" by (induct m arbitrary: x) (auto intro!: integral_cong simp add: K2.p_Suc' p_Suc' intK K2.p_0 p_0) } note K_p_eq = this { fix n and x :: 's have "K2.p (Some x) None n = 0" by (induct n arbitrary: x) (auto simp: K2.p_Suc' K2.p_0 intK cong: integral_cong) } note K_S_None = this from not_empty_irreducible[OF C_comm] obtain c0 where c0: "c0 \ C" by auto have K2_acc: "\x y. (Some x, y) \ K2.acc \ (\z. y = Some z \ (x, z) \ acc)" apply (auto simp: K2.acc_iff acc_iff K_p_eq) apply (case_tac y) apply (auto simp: K_p_eq K_S_None) done have K2_communicating: "\c x. c \ C \ (Some c, x) \ K2.communicating \ (\c'\C. x = Some c')" proof safe fix x c assume "c \ C" "(Some c, x) \ K2.communicating" then show "\c'\C. x = Some c'" by (cases x) (auto simp: communicating_iff K2.communicating_iff K_p_eq K_S_None intro!: irreducibleD2[OF C_comm \c\C\]) next fix c c' x assume "c \ C" "c' \ C" with irreducibleD[OF C_comm this] show "(Some c, Some c') \ K2.communicating" by (auto simp: K2.communicating_iff communicating_iff K_p_eq) qed have "Some ` C \ UNIV // K2.communicating" by (auto simp add: quotient_def Image_def c0 K2_communicating intro!: exI[of _ "Some c0"]) then have "K2.essential_class (Some ` C)" by (rule K2.essential_classI) (auto simp: K2_acc essential_classD2[OF \essential_class C\]) have "K2.aperiodic (Some ` C)" unfolding K2.aperiodic_eventually_recurrent proof safe fix x assume "x \ C" then show "eventually (\m. 0 < K2.p (Some x) (Some x) m) sequentially" using \aperiodic C\ unfolding aperiodic_eventually_recurrent by (auto elim!: eventually_mono simp: K_p_eq) qed fact then have aperiodic: "KN.aperiodic (C \ Some ` C)" by (rule KN.aperiodicI_pair[OF \aperiodic C\]) have KN_essential: "KN.essential_class (C \ Some ` C)" proof (rule KN.essential_classI) show "C \ Some ` C \ UNIV // KN.communicating" using aperiodic by (simp add: KN.aperiodic_def) next fix x y assume "x \ C \ Some ` C" "(x, y) \ KN.acc" with KN.P_accD[of "fst x" "snd x" "fst y" "snd y"] show "y \ C \ Some ` C" by (cases x y rule: prod.exhaust[case_product prod.exhaust]) (auto simp: K2_acc essential_classD2[OF \essential_class C\]) qed { fix n and x y :: 's have "measure N {y} = \

(\ in K2.T None. (None ## \) !! (Suc n) = Some y)" unfolding stationary_distribution_iterate'[OF N(1), of y n] apply (subst K2.p_def[symmetric]) apply (subst K2.p_Suc') apply (subst K'_def) apply (simp add: map_pmf_rep_eq integral_distr K_p_eq) done then have "measure N {y} = \

(\ in K2.T None. \ !! n = Some y)" by simp } note measure_y_eq = this define D where "D = {x::'s \ 's option. Some (fst x) = snd x}" have [measurable]: "\P::('s \ 's option \ bool). P \ measurable (count_space UNIV) (count_space UNIV)" by simp { fix n and x :: 's have "\

(\ in KN.T (y, None). \i !! n) = Some x \ ev_at (HLD D) i \) = (\i(\ in KN.T (y, None). snd (\ !! n) = Some x \ ev_at (HLD D) i \))" by (subst KN.T.finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def intro!: arg_cong2[where f=measure] dest: ev_at_unique) also have "\ = (\i(\ in KN.T (y, None). fst (\ !! n) = x \ ev_at (HLD D) i \))" proof (intro sum.cong refl) fix i assume i: "i \ {..< n}" show "\

(\ in KN.T (y, None). snd (\ !! n) = Some x \ ev_at (HLD D) i \) = \

(\ in KN.T (y, None). fst (\ !! n) = x \ ev_at (HLD D) i \)" apply (subst (1 2) KN.prob_T_split[where n="Suc i"]) apply (simp_all add: ev_at_shift snth_Stream del: stake.simps KN.space_T) unfolding ev_at_shift snth_Stream proof (intro Bochner_Integration.integral_cong refl) fix \ :: "('s \ 's option) stream" let ?s = "\\'. stake (Suc i) \ @- \'" show "\

(\' in KN.T (\ !! i). snd (?s \' !! n) = Some x \ ev_at (HLD D) i \) = \

(\' in KN.T (\ !! i). fst (?s \' !! n) = x \ ev_at (HLD D) i \)" proof cases assume "ev_at (HLD D) i \" from ev_at_imp_snth[OF this] have eq: "snd (\ !! i) = Some (fst (\ !! i))" by (simp add: D_def HLD_iff) have "\

(\' in KN.T (\ !! i). fst (\' !! (n - Suc i)) = x) = \

(\' in T (fst (\ !! i)). \' !! (n - Suc i) = x) * \

(\' in K2.T (snd (\ !! i)). True)" by (subst KN.prod_eq_prob_T) simp_all also have "\ = p (fst (\ !! i)) x (Suc (n - Suc i))" using K2.T.prob_space by (simp add: p_def) also have "\ = K2.p (snd (\ !! i)) (Some x) (Suc (n - Suc i))" by (simp add: K_p_eq eq) also have "\ = \

(\' in T (fst (\ !! i)). True) * \

(\' in K2.T (snd (\ !! i)). \' !! (n - Suc i) = Some x)" using T.prob_space by (simp add: K2.p_def) also have "\ = \

(\' in KN.T (\ !! i). snd (\' !! (n - Suc i)) = Some x)" by (subst KN.prod_eq_prob_T) simp_all finally show ?thesis using \ev_at (HLD D) i \\ i by (simp del: stake.simps) qed simp qed qed also have "\ = \

(\ in KN.T (y, None). (\i !! n) = x \ ev_at (HLD D) i \))" by (subst KN.T.finite_measure_finite_Union[symmetric]) (auto simp add: disjoint_family_on_def dest: ev_at_unique intro!: arg_cong2[where f=measure]) finally have eq: "\

(\ in KN.T (y, None). (\i !! n) = Some x \ ev_at (HLD D) i \)) = \

(\ in KN.T (y, None). (\i !! n) = x \ ev_at (HLD D) i \))" . have "p y x (Suc n) - measure N {x} = \

(\ in T y. \ !! n = x) - \

(\ in K2.T None. \ !! n = Some x)" unfolding p_def by (subst measure_y_eq) simp_all also have "\

(\ in T y. \ !! n = x) = \

(\ in T y. \ !! n = x) * \

(\ in K2.T None. True)" using K2.T.prob_space by simp also have "\ = \

(\ in KN.T (y, None). fst (\ !! n) = x)" by (subst KN.prod_eq_prob_T) auto also have "\ = \

(\ in KN.T (y, None). (\i !! n) = x \ ev_at (HLD D) i \)) + \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i))" by (subst KN.T.finite_measure_Union[symmetric]) (auto intro!: arg_cong2[where f=measure]) also have "\

(\ in K2.T None. \ !! n = Some x) = \

(\ in T y. True) * \

(\ in K2.T None. \ !! n = Some x)" using T.prob_space by simp also have "\ = \

(\ in KN.T (y, None). snd (\ !! n) = Some x)" by (subst KN.prod_eq_prob_T) auto also have "\ = \

(\ in KN.T (y, None). (\i !! n) = Some x \ ev_at (HLD D) i \)) + \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i))" by (subst KN.T.finite_measure_Union[symmetric]) (auto intro!: arg_cong2[where f=measure]) finally have "\ p y x (Suc n) - measure N {x} \ = \ \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) - \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i)) \" unfolding eq by (simp add: field_simps) also have "\ \ \ \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) \ + \ \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i)) \" by (rule abs_triangle_ineq4) also have "\ \ \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) + \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i))" by simp finally have "\ p y x (Suc n) - measure N {x} \ \ \" . } note mono = this { fix n :: nat have "(\\<^sup>+x. \ p y x (Suc n) - measure N {x} \ \count_space C) \ (\\<^sup>+x. ennreal (\

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i))) + ennreal (\

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i))) \count_space C)" using mono by (intro nn_integral_mono) (simp add: ennreal_plus[symmetric] del: ennreal_plus) also have "\ = (\\<^sup>+x. \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) \count_space C) + (\\<^sup>+x. \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i)) \count_space C)" by (subst nn_integral_add) auto also have "\ = emeasure (KN.T (y, None)) (\x\C. {\\space (KN.T (y, None)). fst (\ !! n) = x \ \ (\i)}) + emeasure (KN.T (y, None)) (\x\C. {\\space (KN.T (y, None)). snd (\ !! n) = Some x \ \ (\i)})" by (subst (1 2) emeasure_UN_countable) (auto simp add: disjoint_family_on_def KN.T.emeasure_eq_measure C) also have "\ \ ennreal (\

(\ in KN.T (y, None). \ (\i))) + ennreal (\

(\ in KN.T (y, None). \ (\i)))" unfolding KN.T.emeasure_eq_measure by (intro add_mono) (auto intro!: KN.T.finite_measure_mono) also have "\ \ 2 * \

(\ in KN.T (y, None). \ (\i))" by (simp add: ennreal_plus[symmetric] del: ennreal_plus) finally have "?L (Suc n) \ 2 * \

(\ in KN.T (y, None). \ (\i))" by (auto intro!: integral_real_bounded simp add: pmf.rep_eq) } note le_2 = this have c0_D: "(c0, Some c0) \ D" by (simp add: D_def c0) let ?N' = "map_pmf Some N" interpret NP: pair_prob_space N ?N' .. have pos_recurrent: "\x\C \ Some ` C. KN.pos_recurrent x" proof (rule KN.stationary_distributionD(1)[OF KN_essential _ KN.stationary_distributionI_pair[OF N(1)]]) show "K2.stationary_distribution ?N'" unfolding K2.stationary_distribution_def by (subst N(1)[unfolded stationary_distribution_def]) (auto intro!: bind_pmf_cong simp: K'_def map_pmf_def bind_assoc_pmf bind_return_pmf) show "countable (C \ Some`C)" using C by auto show "set_pmf (pair_pmf N (map_pmf Some N)) \ C \ Some ` C" using \N \ C\ by auto qed from c0_D have "\

(\ in KN.T (y, None). alw (not (HLD D)) \) \ \

(\ in KN.T (y, None). alw (not (HLD {(c0, Some c0)})) \)" apply (auto intro!: KN.T.finite_measure_mono) apply (rule alw_mono, assumption) apply (auto simp: HLD_iff) done also have "\ = 0" apply (rule KN.T.prob_eq_0_AE) apply (simp add: not_ev_iff[symmetric]) apply (subst KN.AE_T_iff) apply simp proof fix t assume t: "t \ KN.Kp (y, None)" then obtain a b where t_eq: "t = (a, Some b)" "a \ K y" "b \ N" unfolding KN.Kp_def by (auto simp: K'_def) with \y \ C\ have "a \ C" using essential_classD2[OF \essential_class C\ \y \ C\] by auto have "b \ C" using \N \ C\ \b \ N\ by auto from pos_recurrent[THEN bspec, of "(c0, Some c0)"] have recurrent_c0: "KN.recurrent (c0, Some c0)" by (simp add: KN.pos_recurrent_def c0) have "C \ Some ` C \ UNIV // KN.communicating" using aperiodic by (simp add: KN.aperiodic_def) then have "((c0, Some c0), t) \ KN.communicating" by (rule KN.irreducibleD) (simp_all add: t_eq c0 \b \ C\ \a \ C\) then have "((c0, Some c0), t) \ KN.acc" by (simp add: KN.communicating_def) then have "KN.U t (c0, Some c0) = 1" by (rule KN.recurrent_acc(1)[OF recurrent_c0]) then show "AE \ in KN.T t. ev (HLD {(c0, Some c0)}) (t ## \)" unfolding KN.U_def by (subst (asm) KN.T.prob_Collect_eq_1) (auto simp add: ev_Stream) qed finally have "\

(\ in KN.T (y, None). alw (not (HLD D)) \) = 0" by (intro antisym measure_nonneg) have "(\n. \

(\ in KN.T (y, None). \ (\i))) \ measure (KN.T (y, None)) (\n. {\\space (KN.T (y, None)). \ (\i)})" by (rule KN.T.finite_Lim_measure_decseq) (auto simp: decseq_def) also have "(\n. {\\space (KN.T (y, None)). \ (\i)}) = {\\space (KN.T (y, None)). alw (not (HLD D)) \}" by (auto simp: not_ev_iff[symmetric] ev_iff_ev_at) also have "\

(\ in KN.T (y, None). alw (not (HLD D)) \) = 0" by fact finally have *: "(\n. 2 * \

(\ in KN.T (y, None). \ (\i))) \ 0" by (intro tendsto_eq_intros) auto show ?thesis apply (rule LIMSEQ_imp_Suc) apply (rule tendsto_sandwich[OF _ _ tendsto_const *]) using le_2 apply (simp_all add: integral_nonneg_AE) done qed lemma stationary_distribution_imp_p_limit: assumes "aperiodic C" "essential_class C" and [simp]: "countable C" assumes N: "stationary_distribution N" "N \ C" assumes [simp]: "x \ C" "y \ C" shows "p x y \ pmf N y" proof - define D where "D y n = \p x y n - pmf N y\" for y n from stationary_distribution_imp_limit[OF assms(1,2,3,4,5,6)] have INT: "(\n. \y. D y n \count_space C) \ 0" unfolding D_def . { fix n have "D y n \ (\z. D y n * indicator {y} z \count_space C)" by simp also have "\ \ (\y. D y n \count_space C)" by (intro integral_mono) (auto split: split_indicator simp: D_def p_def disjoint_family_on_def intro!: Bochner_Integration.integrable_diff integrable_pmf T.integrable_measure) finally have "D y n \ (\y. D y n \count_space C)" . } note * = this have D_nonneg: "\n. 0 \ D y n" by (simp add: D_def) have "D y \ 0" by (rule tendsto_sandwich[OF _ _ tendsto_const INT]) (auto simp: eventually_sequentially * D_nonneg) then show ?thesis using Lim_null[where l="pmf N y" and net=sequentially and f="p x y"] by (simp add: D_def [abs_def] tendsto_rabs_zero_iff) qed end lemma (in MC_syntax) essential_classI2: assumes "X \ {}" assumes accI: "\x y. x \ X \ y \ X \ (x, y) \ acc" assumes ED: "\x y. x \ X \ y \ set_pmf (K x) \ y \ X" shows "essential_class X" proof (rule essential_classI) { fix x y assume "(x, y) \ acc" "x \ X" then show "y \ X" by induct (auto dest: ED)} note accD = this - from \X \ {}\ obtain x where "x \ X" by auto from \x \ X\ show "X \ UNIV // communicating" by (auto simp add: quotient_def Image_def communicating_def accI dest: accD intro!: exI[of _ x]) qed end diff --git a/thys/Ordinary_Differential_Equations/IVP/Poincare_Map.thy b/thys/Ordinary_Differential_Equations/IVP/Poincare_Map.thy --- a/thys/Ordinary_Differential_Equations/IVP/Poincare_Map.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Poincare_Map.thy @@ -1,2499 +1,2504 @@ theory Poincare_Map imports Flow begin abbreviation "plane n c \ {x. x \ n = c}" lemma eventually_tendsto_compose_within: assumes "eventually P (at l within S)" assumes "P l" assumes "(f \ l) (at x within T)" assumes "eventually (\x. f x \ S) (at x within T)" shows "eventually (\x. P (f x)) (at x within T)" proof - from assms(1) assms(2) obtain U where U: "open U" "l \ U" "\x. x \ U \ x \ S \ P x" by (force simp: eventually_at_topological) from topological_tendstoD[OF assms(3) \open U\ \l \ U\] have "\\<^sub>F x in at x within T. f x \ U" by auto then show ?thesis using assms(4) by eventually_elim (auto intro!: U) qed lemma eventually_eventually_withinI:\ \aha...\ assumes "\\<^sub>F x in at x within A. P x" "P x" shows "\\<^sub>F a in at x within S. \\<^sub>F x in at a within A. P x" using assms unfolding eventually_at_topological by force lemma eventually_not_in_closed: assumes "closed P" assumes "f t \ P" "t \ T" assumes "continuous_on T f" shows "\\<^sub>F t in at t within T. f t \ P" using assms unfolding Compl_iff[symmetric] closed_def continuous_on_topological eventually_at_topological by metis context ll_on_open_it begin lemma existence_ivl_trans': assumes "t + s \ existence_ivl t0 x0" "t \ existence_ivl t0 x0" shows "t + s \ existence_ivl t (flow t0 x0 t)" by (meson assms(1) assms(2) general.existence_ivl_reverse general.flow_solves_ode general.is_interval_existence_ivl general.maximal_existence_flow(1) general.mem_existence_ivl_iv_defined(2) general.mem_existence_ivl_subset local.existence_ivl_subset subsetD) end context auto_ll_on_open\ \TODO: generalize to continuous systems\ begin definition returns_to ::"'a set \ 'a \ bool" where "returns_to P x \ (\\<^sub>F t in at_right 0. flow0 x t \ P) \ (\t>0. t \ existence_ivl0 x \ flow0 x t \ P)" definition return_time :: "'a set \ 'a \ real" where "return_time P x = (if returns_to P x then (SOME t. t > 0 \ t \ existence_ivl0 x \ flow0 x t \ P \ (\s \ {0<.. P)) else 0)" lemma returns_toI: assumes t: "t > 0" "t \ existence_ivl0 x" "flow0 x t \ P" assumes ev: "\\<^sub>F t in at_right 0. flow0 x t \ P" assumes "closed P" shows "returns_to P x" using assms by (auto simp: returns_to_def) lemma returns_to_outsideI: assumes t: "t \ 0" "t \ existence_ivl0 x" "flow0 x t \ P" assumes ev: "x \ P" assumes "closed P" shows "returns_to P x" proof cases assume "t > 0" moreover have "\\<^sub>F s in at 0 within {0 .. t}. flow0 x s \ P" using assms mem_existence_ivl_iv_defined ivl_subset_existence_ivl[OF \t \ _\] \0 < t\ by (auto intro!: eventually_not_in_closed flow_continuous_on continuous_intros simp: eventually_conj_iff) with order_tendstoD(2)[OF tendsto_ident_at \0 < t\, of "{0<..}"] have "\\<^sub>F t in at_right 0. flow0 x t \ P" unfolding eventually_at_filter by eventually_elim (use \t > 0\ in auto) then show ?thesis by (auto intro!: returns_toI assms \0 < t\) qed (use assms in simp) lemma returns_toE: assumes "returns_to P x" obtains t0 t1 where "0 < t0" "t0 \ t1" "t1 \ existence_ivl0 x" "flow0 x t1 \ P" "\t. 0 < t \ t < t0 \ flow0 x t \ P" proof - obtain t0 t1 where t0: "t0 > 0" "\t. 0 < t \ t < t0 \ flow0 x t \ P" and t1: "t1 > 0" "t1 \ existence_ivl0 x" "flow0 x t1 \ P" using assms by (auto simp: returns_to_def eventually_at_right[OF zero_less_one]) moreover have "t0 \ t1" using t0(2)[of t1] t1 t0(1) by force ultimately show ?thesis by (blast intro: that) qed lemma return_time_some: assumes "returns_to P x" shows "return_time P x = (SOME t. t > 0 \ t \ existence_ivl0 x \ flow0 x t \ P \ (\s \ {0<.. P))" using assms by (auto simp: return_time_def) lemma return_time_ex1: assumes "returns_to P x" assumes "closed P" shows "\!t. t > 0 \ t \ existence_ivl0 x \ flow0 x t \ P \ (\s \ {0<.. P)" proof - from returns_toE[OF \returns_to P x\] obtain t0 t1 where t1: "t1 \ t0" "t1 \ existence_ivl0 x" "flow0 x t1 \ P" and t0: "t0 > 0" "\t. 0 < t \ t < t0 \ flow0 x t \ P" by metis from flow_continuous_on have cont: "continuous_on {0 .. t1} (flow0 x)" by (rule continuous_on_subset) (intro ivl_subset_existence_ivl t1) from cont have cont': "continuous_on {t0 .. t1} (flow0 x)" by (rule continuous_on_subset) (use \0 < t0\ in auto) have "compact (flow0 x -` P \ {t0 .. t1})" using \closed P\ cont' by (auto simp: compact_eq_bounded_closed bounded_Int bounded_closed_interval intro!: closed_vimage_Int) have "flow0 x -` P \ {t0..t1} \ {}" using t1 t0 by auto from compact_attains_inf[OF \compact _\ this] t0 t1 obtain rt where rt: "t0 \ rt" "rt \ t1" "flow0 x rt \ P" and least: "\t'. flow0 x t' \ P \ t0 \ t' \ t' \ t1 \ rt \ t'" by auto have "0 < rt" "flow0 x rt \ P" "rt \ existence_ivl0 x" and "0 < t' \ t' < rt \ flow0 x t' \ P" for t' using ivl_subset_existence_ivl[OF \t1 \ existence_ivl0 x\] t0 t1 rt least[of t'] by force+ then show ?thesis by (intro ex_ex1I) force+ qed lemma return_time_pos_returns_to: "return_time P x > 0 \ returns_to P x" by (auto simp: return_time_def split: if_splits) lemma assumes ret: "returns_to P x" assumes "closed P" shows return_time_pos: "return_time P x > 0" using someI_ex[OF return_time_ex1[OF assms, THEN ex1_implies_ex]] unfolding return_time_some[OF ret, symmetric] by auto lemma returns_to_return_time_pos: assumes "closed P" shows "returns_to P x \ return_time P x > 0" by (auto intro!: return_time_pos assms) (auto simp: return_time_def split: if_splits) lemma return_time: assumes ret: "returns_to P x" assumes "closed P" shows "return_time P x > 0" and return_time_exivl: "return_time P x \ existence_ivl0 x" and return_time_returns: "flow0 x (return_time P x) \ P" and return_time_least: "\s. 0 < s \ s < return_time P x \ flow0 x s \ P" using someI_ex[OF return_time_ex1[OF assms, THEN ex1_implies_ex]] unfolding return_time_some[OF ret, symmetric] by auto lemma returns_to_earlierI: assumes ret: "returns_to P (flow0 x t)" "closed P" assumes "t \ 0" "t \ existence_ivl0 x" assumes ev: "\\<^sub>F t in at_right 0. flow0 x t \ P" shows "returns_to P x" proof - from return_time[OF ret] have rt: "0 < return_time P (flow0 x t)" "flow0 (flow0 x t) (return_time P (flow0 x t)) \ P" and "0 < s \ s < return_time P (flow0 x t) \ flow0 (flow0 x t) s \ P" for s by auto let ?t = "t + return_time P (flow0 x t)" show ?thesis proof (rule returns_toI[of ?t]) show "0 < ?t" by (auto intro!: add_nonneg_pos rt \t \ 0\) show "?t \ existence_ivl0 x" by (intro existence_ivl_trans return_time_exivl assms) have "flow0 x (t + return_time P (flow0 x t)) = flow0 (flow0 x t) (return_time P (flow0 x t))" by (intro flow_trans assms return_time_exivl) also have "\ \ P" by (rule return_time_returns[OF ret]) finally show "flow0 x (t + return_time P (flow0 x t)) \ P" . show "closed P" by fact show "\\<^sub>F t in at_right 0. flow0 x t \ P" by fact qed qed lemma return_time_gt: assumes ret: "returns_to P x" "closed P" assumes flow_not: "\s. 0 < s \ s \ t \ flow0 x s \ P" shows "t < return_time P x" using flow_not[of "return_time P x"] return_time_pos[OF ret] return_time_returns[OF ret] by force lemma return_time_le: assumes ret: "returns_to P x" "closed P" assumes flow_not: "flow0 x t \ P" "t > 0" shows "return_time P x \ t" using return_time_least[OF assms(1,2), of t] flow_not by force lemma returns_to_laterI: assumes ret: "returns_to P x" "closed P" assumes t: "t > 0" "t \ existence_ivl0 x" assumes flow_not: "\s. 0 < s \ s \ t \ flow0 x s \ P" shows "returns_to P (flow0 x t)" apply (rule returns_toI[of "return_time P x - t"]) subgoal using flow_not by (auto intro!: return_time_gt ret) subgoal by (auto intro!: existence_ivl_trans' return_time_exivl ret t) subgoal by (subst flow_trans[symmetric]) (auto intro!: existence_ivl_trans' return_time_exivl ret t return_time_returns) subgoal proof - have "\\<^sub>F y in nhds 0. y \ existence_ivl0 (flow0 x t)" apply (rule eventually_nhds_in_open[OF open_existence_ivl[of "flow0 x t"] existence_ivl_zero]) apply (rule flow_in_domain) apply fact done then have "\\<^sub>F s in at_right 0. s \ existence_ivl0 (flow0 x t)" unfolding eventually_at_filter by eventually_elim auto moreover have "\\<^sub>F s in at_right 0. t + s < return_time P x" using return_time_gt[OF ret flow_not, of t] by (auto simp: eventually_at_right[OF zero_less_one] intro!: exI[of _ "return_time P x - t"]) moreover have "\\<^sub>F s in at_right 0. 0 < t + s" by (metis (mono_tags) eventually_at_rightI greaterThanLessThan_iff pos_add_strict t(1)) ultimately show ?thesis apply eventually_elim apply (subst flow_trans[symmetric]) using return_time_least[OF ret] by (auto intro!: existence_ivl_trans' t) qed subgoal by fact done lemma never_returns: assumes "\returns_to P x" assumes "closed P" "t \ 0" "t \ existence_ivl0 x" assumes ev: "\\<^sub>F t in at_right 0. flow0 x t \ P" shows "\returns_to P (flow0 x t)" using returns_to_earlierI[OF _ assms(2-5)] assms(1) by blast lemma return_time_eqI: assumes "closed P" and t_pos: "t > 0" and ex: "t \ existence_ivl0 x" and ret: "flow0 x t \ P" and least: "\s. 0 < s \ s < t \ flow0 x s \ P" shows "return_time P x = t" proof - from least t_pos have "\\<^sub>F t in at_right 0. flow0 x t \ P" by (auto simp: eventually_at_right[OF zero_less_one]) then have "returns_to P x" by (auto intro!: returns_toI[of t] assms) then show ?thesis using least by (auto simp: return_time_def t_pos ex ret intro!: some1_equality[OF return_time_ex1[OF \returns_to _ _\ \closed _\]]) qed lemma return_time_step: assumes "returns_to P (flow0 x t)" assumes "closed P" assumes flow_not: "\s. 0 < s \ s \ t \ flow0 x s \ P" assumes t: "t > 0" "t \ existence_ivl0 x" shows "return_time P (flow0 x t) = return_time P x - t" proof - from flow_not t have "\\<^sub>F t in at_right 0. flow0 x t \ P" by (auto simp: eventually_at_right[OF zero_less_one]) from returns_to_earlierI[OF assms(1,2) less_imp_le, OF t this] have ret: "returns_to P x" . from return_time_gt[OF ret \closed P\ flow_not] have "t < return_time P x" by simp moreover have "0 < s \ s < return_time P x - t \ flow0 (flow0 x t) s = flow0 x (t + s)" for s using ivl_subset_existence_ivl[OF return_time_exivl[OF ret \closed _\]] t by (subst flow_trans) (auto intro!: existence_ivl_trans') ultimately show ?thesis using flow_not assms(1) ret return_time_least t(1) by (auto intro!: return_time_eqI return_time_returns ret simp: flow_trans[symmetric] \closed P\ t(2) existence_ivl_trans' return_time_exivl) qed definition "poincare_map P x = flow0 x (return_time P x)" lemma poincare_map_step_flow: assumes ret: "returns_to P x" "closed P" assumes flow_not: "\s. 0 < s \ s \ t \ flow0 x s \ P" assumes t: "t > 0" "t \ existence_ivl0 x" shows "poincare_map P (flow0 x t) = poincare_map P x" unfolding poincare_map_def apply (subst flow_trans[symmetric]) subgoal by fact subgoal using flow_not by (auto intro!: return_time_exivl returns_to_laterI t ret) subgoal using flow_not by (subst return_time_step) (auto intro!: return_time_exivl returns_to_laterI t ret) done lemma poincare_map_returns: assumes "returns_to P x" "closed P" shows "poincare_map P x \ P" by (auto intro!: return_time_returns assms simp: poincare_map_def) lemma poincare_map_onto: assumes "closed P" assumes "0 < t" "t \ existence_ivl0 x" "\\<^sub>F t in at_right 0. flow0 x t \ P" assumes "flow0 x t \ P" shows "poincare_map P x \ flow0 x ` {0 <.. t} \ P" proof (rule IntI) have "returns_to P x" by (rule returns_toI) (rule assms)+ then have "return_time P x \ {0<..t}" by (auto intro!: return_time_pos assms return_time_le) then show "poincare_map P x \ flow0 x ` {0<..t}" by (auto simp: poincare_map_def) show "poincare_map P x \ P" by (auto intro!: poincare_map_returns \returns_to _ _\ \closed _\) qed end lemma isCont_blinfunD: fixes f'::"'a::metric_space \ 'b::real_normed_vector \\<^sub>L 'c::real_normed_vector" assumes "isCont f' a" "0 < e" shows "\d>0. \x. dist a x < d \ onorm (\v. blinfun_apply (f' x) v - blinfun_apply (f' a) v) < e" proof - have "\\<^sub>F x in at a. dist (f' x) (f' a) < e" using assms isCont_def tendsto_iff by blast then show ?thesis using \e > 0\ norm_eq_zero by (force simp: eventually_at dist_commute dist_norm norm_blinfun.rep_eq simp flip: blinfun.bilinear_simps) qed proposition has_derivative_locally_injective_blinfun: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" and f'::"'n \ 'n \\<^sub>L 'm" and g'::"'m \\<^sub>L 'n" assumes "a \ s" and "open s" and g': "g' o\<^sub>L (f' a) = 1\<^sub>L" and f': "\x. x \ s \ (f has_derivative f' x) (at x)" and c: "isCont f' a" obtains r where "r > 0" "ball a r \ s" "inj_on f (ball a r)" proof - have bl: "bounded_linear (blinfun_apply g')" by (auto simp: blinfun.bounded_linear_right) from g' have g': "blinfun_apply g' \ blinfun_apply (f' a) = id" by transfer (simp add: id_def) from has_derivative_locally_injective[OF \a \ s\ \open s\ bl g' f' isCont_blinfunD[OF c]] obtain r where "0 < r" "ball a r \ s" "inj_on f (ball a r)" by auto then show ?thesis .. qed lift_definition embed1_blinfun::"'a::real_normed_vector \\<^sub>L ('a*'b::real_normed_vector)" is "\x. (x, 0)" by standard (auto intro!: exI[where x=1]) lemma blinfun_apply_embed1_blinfun[simp]: "blinfun_apply embed1_blinfun x = (x, 0)" by transfer simp lift_definition embed2_blinfun::"'a::real_normed_vector \\<^sub>L ('b::real_normed_vector*'a)" is "\x. (0, x)" by standard (auto intro!: exI[where x=1]) lemma blinfun_apply_embed2_blinfun[simp]: "blinfun_apply embed2_blinfun x = (0, x)" by transfer simp lemma blinfun_inverseD: "f o\<^sub>L f' = 1\<^sub>L \ f (f' x) = x" apply transfer unfolding o_def by meson lemmas continuous_on_open_vimageI = continuous_on_open_vimage[THEN iffD1, rule_format] lemmas continuous_on_closed_vimageI = continuous_on_closed_vimage[THEN iffD1, rule_format] lemma ball_times_subset: "ball a (c/2) \ ball b (c/2) \ ball (a, b) c" proof - { fix a' b' have "sqrt ((dist a a')\<^sup>2 + (dist b b')\<^sup>2) \ dist a a' + dist b b'" by (rule real_le_lsqrt) (auto simp: power2_eq_square algebra_simps) also assume "a' \ ball a (c / 2)" then have "dist a a' < c / 2" by (simp add:) also assume "b' \ ball b (c / 2)" then have "dist b b' < c / 2" by (simp add:) finally have "sqrt ((dist a a')\<^sup>2 + (dist b b')\<^sup>2) < c" by simp } thus ?thesis by (auto simp: dist_prod_def mem_cball) qed lemma linear_inverse_blinop_lemma: fixes w::"'a::{banach, perfect_space} blinop" assumes "norm w < 1" shows "summable (\n. (-1)^n *\<^sub>R w^n)" (is ?C) "(\n. (-1)^n *\<^sub>R w^n) * (1 + w) = 1" (is ?I1) "(1 + w) * (\n. (-1)^n *\<^sub>R w^n) = 1" (is ?I2) "norm ((\n. (-1)^n *\<^sub>R w^n) - 1 + w) \ (norm w)\<^sup>2/(1 - norm (w))" (is ?L) proof - have "summable (\n. norm w ^ n)" apply (rule summable_geometric) using assms by auto then have "summable (\n. norm (w ^ n))" by (rule summable_comparison_test'[where N=0]) (auto intro!: norm_power_ineq) then show ?C by (rule summable_comparison_test'[where N=0]) (auto simp: norm_power ) { fix N have 1: "(1 + w) * sum (\n. (-1)^n *\<^sub>R w^n) {..n. (-1)^n *\<^sub>R w^n) {.. = sum (\n. (-1)^n *\<^sub>R w^n - (-1)^Suc n *\<^sub>R w^ Suc n) {.. = 1 - (-1)^N *\<^sub>R w^N" by (subst sum_lessThan_telescope') (auto simp: algebra_simps) finally have "(1 + w) * (\nR w ^ n) = 1 - (- 1) ^ N *\<^sub>R w ^ N" . note 1 this } note nu = this show "?I1" apply (subst suminf_mult2, fact) apply (subst suminf_eq_lim) apply (subst sum_distrib_right[symmetric]) apply (rule limI) apply (subst nu(1)[symmetric]) apply (subst nu(2)) apply (rule tendsto_eq_intros) apply (rule tendsto_intros) apply (rule tendsto_norm_zero_cancel) apply auto apply (rule Lim_transform_bound[where g="\i. norm w ^ i"]) apply (rule eventuallyI) apply simp apply (rule norm_power_ineq) apply (auto intro!: LIMSEQ_power_zero assms) done show "?I2" apply (subst suminf_mult[symmetric], fact) apply (subst suminf_eq_lim) apply (subst sum_distrib_left[symmetric]) apply (rule limI) apply (subst nu(2)) apply (rule tendsto_eq_intros) apply (rule tendsto_intros) apply (rule tendsto_norm_zero_cancel) apply auto apply (rule Lim_transform_bound[where g="\i. norm w ^ i"]) apply (rule eventuallyI) apply simp apply (rule norm_power_ineq) apply (auto intro!: LIMSEQ_power_zero assms) done have *: "(\n. (- 1) ^ n *\<^sub>R w ^ n) - 1 + w = (w\<^sup>2 * (\n. (- 1) ^ n *\<^sub>R w ^ n))" apply (subst suminf_split_initial_segment[where k=2], fact) apply (subst suminf_mult[symmetric], fact) by (auto simp: power2_eq_square algebra_simps eval_nat_numeral) also have "norm \ \ (norm w)\<^sup>2 / (1 - norm w)" - apply (rule order_trans[OF norm_mult_ineq]) - apply (subst divide_inverse) - apply (rule mult_mono) - apply (auto simp: norm_power_ineq inverse_eq_divide assms ) - apply (rule order_trans[OF summable_norm]) - apply auto - apply fact - apply (rule order_trans[OF suminf_le]) - apply (rule allI) apply (rule norm_power_ineq) - apply fact - apply fact - by (auto simp: suminf_geometric assms) + proof - + have \

: "norm (\n. (- 1) ^ n *\<^sub>R w ^ n) \ 1 / (1 - norm w)" + apply (rule order_trans[OF summable_norm]) + apply auto + apply fact + apply (rule order_trans[OF suminf_le]) + apply (rule norm_power_ineq) + apply fact + apply fact + by (auto simp: suminf_geometric assms) + show ?thesis + apply (rule order_trans[OF norm_mult_ineq]) + apply (subst divide_inverse) + apply (rule mult_mono) + apply (auto simp: norm_power_ineq inverse_eq_divide assms \
) + done + qed finally show ?L . qed lemma linear_inverse_blinfun_lemma: fixes w::"'a \\<^sub>L 'a::{banach, perfect_space}" assumes "norm w < 1" obtains I where "I o\<^sub>L (1\<^sub>L + w) = 1\<^sub>L" "(1\<^sub>L + w) o\<^sub>L I = 1\<^sub>L" "norm (I - 1\<^sub>L + w) \ (norm w)\<^sup>2/(1 - norm (w))" proof - define v::"'a blinop" where "v = Blinop w" have "norm v = norm w" unfolding v_def apply transfer by (simp add: bounded_linear_Blinfun_apply norm_blinfun.rep_eq) with assms have "norm v < 1" by simp from linear_inverse_blinop_lemma[OF this] have v: "(\n. (- 1) ^ n *\<^sub>R v ^ n) * (1 + v) = 1" "(1 + v) * (\n. (- 1) ^ n *\<^sub>R v ^ n) = 1" "norm ((\n. (- 1) ^ n *\<^sub>R v ^ n) - 1 + v) \ (norm v)\<^sup>2 / (1 - norm v)" by auto define J::"'a blinop" where "J = (\n. (- 1) ^ n *\<^sub>R v ^ n)" define I::"'a \\<^sub>L 'a" where "I = Blinfun J" have "Blinfun (blinop_apply J) - 1\<^sub>L + w = Rep_blinop (J - 1 + Blinop (blinfun_apply w))" by transfer' (auto simp: blinfun_apply_inverse) then have ne: "norm (Blinfun (blinop_apply J) - 1\<^sub>L + w) = norm (J - 1 + Blinop (blinfun_apply w))" by (auto simp: norm_blinfun_def norm_blinop_def) from v have "I o\<^sub>L (1\<^sub>L + w) = 1\<^sub>L" "(1\<^sub>L + w) o\<^sub>L I = 1\<^sub>L" "norm (I - 1\<^sub>L + w) \ (norm w)\<^sup>2/(1 - norm (w))" apply (auto simp: I_def J_def[symmetric]) unfolding v_def apply (auto simp: blinop.bounded_linear_right bounded_linear_Blinfun_apply intro!: blinfun_eqI) subgoal by transfer (auto simp: blinfun_ext blinfun.bilinear_simps bounded_linear_Blinfun_apply) subgoal by transfer (auto simp: Transfer.Rel_def blinfun_ext blinfun.bilinear_simps bounded_linear_Blinfun_apply) subgoal apply (auto simp: ne) apply transfer by (auto simp: norm_blinfun_def bounded_linear_Blinfun_apply) done then show ?thesis .. qed definition "invertibles_blinfun = {w. \wi. w o\<^sub>L wi = 1\<^sub>L \ wi o\<^sub>L w = 1\<^sub>L}" lemma blinfun_inverse_open:\ \8.3.2 in Dieudonne, TODO: add continuity and derivative\ shows "open (invertibles_blinfun:: ('a::{banach, perfect_space} \\<^sub>L 'b::banach) set)" proof (rule openI) fix u0::"'a \\<^sub>L 'b" assume "u0 \ invertibles_blinfun" then obtain u0i where u0i: "u0 o\<^sub>L u0i = 1\<^sub>L" "u0i o\<^sub>L u0 = 1\<^sub>L" by (auto simp: invertibles_blinfun_def) then have [simp]: "u0i \ 0" apply (auto) by (metis one_blinop.abs_eq zero_blinop.abs_eq zero_neq_one) let ?e = "inverse (norm u0i)" show "\e>0. ball u0 e \ invertibles_blinfun" apply (clarsimp intro!: exI[where x = ?e] simp: invertibles_blinfun_def) subgoal premises prems for u0s proof - define s where "s = u0s - u0" have u0s: "u0s = u0 + s" by (auto simp: s_def) have "norm (u0i o\<^sub>L s) < 1" using prems by (auto simp: dist_norm u0s divide_simps ac_simps intro!: le_less_trans[OF norm_blinfun_compose]) from linear_inverse_blinfun_lemma[OF this] obtain I where I: "I o\<^sub>L 1\<^sub>L + (u0i o\<^sub>L s) = 1\<^sub>L" "1\<^sub>L + (u0i o\<^sub>L s) o\<^sub>L I = 1\<^sub>L" "norm (I - 1\<^sub>L + (u0i o\<^sub>L s)) \ (norm (u0i o\<^sub>L s))\<^sup>2 / (1 - norm (u0i o\<^sub>L s))" by auto have u0s_eq: "u0s = u0 o\<^sub>L (1\<^sub>L + (u0i o\<^sub>L s))" using u0i by (auto simp: s_def blinfun.bilinear_simps blinfun_ext) show ?thesis apply (rule exI[where x="I o\<^sub>L u0i"]) using I u0i apply (auto simp: u0s_eq) by (auto simp: algebra_simps blinfun_ext blinfun.bilinear_simps) qed done qed lemma blinfun_compose_assoc[ac_simps]: "a o\<^sub>L b o\<^sub>L c = a o\<^sub>L (b o\<^sub>L c)" by (auto intro!: blinfun_eqI) text \TODO: move @{thm norm_minus_cancel} to class!\ lemma (in real_normed_vector) norm_minus_cancel [simp]: "norm (- x) = norm x" proof - have scaleR_minus_left: "- a *\<^sub>R x = - (a *\<^sub>R x)" for a x proof - have "\x1 x2. (x2::real) + x1 = x1 + x2" by auto then have f1: "\r ra a. (ra + r) *\<^sub>R (a::'a) = r *\<^sub>R a + ra *\<^sub>R a" using local.scaleR_add_left by presburger have f2: "a + a = 2 * a" by force have f3: "2 * a + - 1 * a = a" by auto have "- a = - 1 * a" by auto then show ?thesis using f3 f2 f1 by (metis local.add_minus_cancel local.add_right_imp_eq) qed have "norm (- x) = norm (scaleR (- 1) x)" by (simp only: scaleR_minus_left scaleR_one) also have "\ = \- 1\ * norm x" by (rule norm_scaleR) finally show ?thesis by simp qed text \TODO: move @{thm norm_minus_commute} to class!\ lemma (in real_normed_vector) norm_minus_commute: "norm (a - b) = norm (b - a)" proof - have "norm (- (b - a)) = norm (b - a)" by (rule norm_minus_cancel) then show ?thesis by simp qed instance euclidean_space \ banach by standard lemma blinfun_apply_Pair_split: "blinfun_apply g (a, b) = blinfun_apply g (a, 0) + blinfun_apply g (0, b)" unfolding blinfun.bilinear_simps[symmetric] by simp lemma blinfun_apply_Pair_add2: "blinfun_apply f (0, a + b) = blinfun_apply f (0, a) + blinfun_apply f (0, b)" unfolding blinfun.bilinear_simps[symmetric] by simp lemma blinfun_apply_Pair_add1: "blinfun_apply f (a + b, 0) = blinfun_apply f (a, 0) + blinfun_apply f (b, 0)" unfolding blinfun.bilinear_simps[symmetric] by simp lemma blinfun_apply_Pair_minus2: "blinfun_apply f (0, a - b) = blinfun_apply f (0, a) - blinfun_apply f (0, b)" unfolding blinfun.bilinear_simps[symmetric] by simp lemma blinfun_apply_Pair_minus1: "blinfun_apply f (a - b, 0) = blinfun_apply f (a, 0) - blinfun_apply f (b, 0)" unfolding blinfun.bilinear_simps[symmetric] by simp lemma implicit_function_theorem: fixes f::"'a::euclidean_space * 'b::euclidean_space \ 'c::euclidean_space"\ \TODO: generalize?!\ assumes [derivative_intros]: "\x. x \ S \ (f has_derivative blinfun_apply (f' x)) (at x)" assumes S: "(x, y) \ S" "open S" assumes "DIM('c) \ DIM('b)" assumes f'C: "isCont f' (x, y)" assumes "f (x, y) = 0" assumes T2: "T o\<^sub>L (f' (x, y) o\<^sub>L embed2_blinfun) = 1\<^sub>L" assumes T1: "(f' (x, y) o\<^sub>L embed2_blinfun) o\<^sub>L T = 1\<^sub>L"\ \TODO: reduce?!\ obtains u e r where "f (x, u x) = 0" "u x = y" "\s. s \ cball x e \ f (s, u s) = 0" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ S" "e > 0" "(u has_derivative - T o\<^sub>L f' (x, y) o\<^sub>L embed1_blinfun) (at x)" "r > 0" "\U v s. v x = y \ (\s. s \ U \ f (s, v s) = 0) \ U \ cball x e \ continuous_on U v \ s \ U \ (s, v s) \ ball (x, y) r \ u s = v s" proof - define H where "H \ \(x, y). (x, f (x, y))" define H' where "H' \ \x. (embed1_blinfun o\<^sub>L fst_blinfun) + (embed2_blinfun o\<^sub>L (f' x))" have f'_inv: "f' (x, y) o\<^sub>L embed2_blinfun \ invertibles_blinfun" using T1 T2 by (auto simp: invertibles_blinfun_def ac_simps intro!: exI[where x=T]) from openE[OF blinfun_inverse_open this] obtain d0 where e0: "0 < d0" "ball (f' (x, y) o\<^sub>L embed2_blinfun) d0 \ invertibles_blinfun" by auto have "isCont (\s. f' s o\<^sub>L embed2_blinfun) (x, y)" by (auto intro!: continuous_intros f'C) from this[unfolded isCont_def, THEN tendstoD, OF \0 < d0\] have "\\<^sub>F s in at (x, y). f' s o\<^sub>L embed2_blinfun \ invertibles_blinfun" apply eventually_elim using e0 by (auto simp: subset_iff dist_commute) then obtain e0 where "e0 > 0" "xa \ (x, y) \ dist xa (x, y) < e0 \ f' xa o\<^sub>L embed2_blinfun \ invertibles_blinfun" for xa unfolding eventually_at by auto then have e0: "e0 > 0" "dist xa (x, y) < e0 \ f' xa o\<^sub>L embed2_blinfun \ invertibles_blinfun" for xa apply - subgoal by simp using f'_inv apply (cases "xa = (x, y)") by auto have H': "x \ S \ (H has_derivative H' x) (at x)" for x unfolding H_def H'_def by (auto intro!: derivative_eq_intros ext simp: blinfun.bilinear_simps) have cH': "isCont H' (x, y)" unfolding H'_def by (auto intro!: continuous_intros assms) have linear_H': "\s. s \ S \ linear (H' s)" using H' assms(2) has_derivative_linear by blast have *: "blinfun_apply T (blinfun_apply (f' (x, y)) (0, b)) = b" for b using blinfun_inverseD[OF T2, of b] by simp have "inj (f' (x, y) o\<^sub>L embed2_blinfun)" by (metis (no_types, lifting) "*" blinfun_apply_blinfun_compose embed2_blinfun.rep_eq injI) then have [simp]: "blinfun_apply (f' (x, y)) (0, b) = 0 \ b = 0" for b apply (subst (asm) linear_injective_0) subgoal apply (rule bounded_linear.linear) apply (rule blinfun.bounded_linear_right) done subgoal by simp done have "inj (H' (x, y))" apply (subst linear_injective_0) apply (rule linear_H') apply fact apply (auto simp: H'_def blinfun.bilinear_simps zero_prod_def) done define Hi where "Hi = (embed1_blinfun o\<^sub>L fst_blinfun) + ((embed2_blinfun o\<^sub>L T o\<^sub>L (snd_blinfun - (f' (x, y) o\<^sub>L embed1_blinfun o\<^sub>L fst_blinfun))))" have Hi': "(\u. snd (blinfun_apply Hi (u, 0))) = - T o\<^sub>L f' (x, y) o\<^sub>L embed1_blinfun" by (auto simp: Hi_def blinfun.bilinear_simps) have Hi: "Hi o\<^sub>L H' (x, y) = 1\<^sub>L" apply (auto simp: H'_def fun_eq_iff blinfun.bilinear_simps Hi_def intro!: ext blinfun_eqI) apply (subst blinfun_apply_Pair_split) by (auto simp: * blinfun.bilinear_simps) from has_derivative_locally_injective_blinfun[OF S this H' cH'] obtain r0 where r0: "0 < r0" "ball (x, y) r0 \ S" and inj: "inj_on H (ball (x, y) r0)" by auto define r where "r = min r0 e0" have r: "0 < r" "ball (x, y) r \ S" and inj: "inj_on H (ball (x, y) r)" and r_inv: "\s. s \ ball (x, y) r \ f' s o\<^sub>L embed2_blinfun \ invertibles_blinfun" subgoal using e0 r0 by (auto simp: r_def) subgoal using e0 r0 by (auto simp: r_def) subgoal using inj apply (rule inj_on_subset) using e0 r0 by (auto simp: r_def) subgoal for s using e0 r0 by (auto simp: r_def dist_commute) done obtain i::'a where "i \ Basis" using nonempty_Basis by blast define undef where "undef \ (x, y) + r *\<^sub>R (i, 0)"\ \really??\ have ud: "\ dist (x, y) undef < r" using \r > 0\ \i \ Basis\ by (auto simp: undef_def dist_norm) define G where "G \ the_inv_into (ball (x, y) r) H" { fix u v assume [simp]: "(u, v) \ H ` ball (x, y) r" note [simp] = inj have "(u, v) = H (G (u, v))" unfolding G_def by (subst f_the_inv_into_f[where f=H]) auto moreover have "\ = H (G (u, v))" by (auto simp: G_def) moreover have "\ = (fst (G (u, v)), f (G (u, v)))" by (auto simp: H_def split_beta') ultimately have "u = fst (G (u, v))" "v = f (G (u, v))" by simp_all then have "f (u, snd (G(u, v))) = v" "u = fst (G (u, v))" by (metis prod.collapse)+ } note uvs = this note uv = uvs(1) moreover have "f (x, snd (G (x, 0))) = 0" apply (rule uv) by (metis (mono_tags, lifting) H_def assms(6) case_prod_beta' centre_in_ball fst_conv image_iff r(1) snd_conv) moreover have cH: "continuous_on S H" apply (rule has_derivative_continuous_on) apply (subst at_within_open) apply (auto intro!: H' assms) done have inj2: "inj_on H (ball (x, y) (r / 2))" apply (rule inj_on_subset, rule inj) using r by auto have oH: "open (H ` ball (x, y) (r/2))" apply (rule invariance_of_domain_gen) apply (auto simp: assms inj) apply (rule continuous_on_subset) apply fact using r apply auto using inj2 apply simp done have "(x, f (x, y)) \ H ` ball (x, y) (r/2)" using \r > 0\ by (auto simp: H_def) from open_contains_cball[THEN iffD1, OF oH, rule_format, OF this] obtain e' where e': "e' > 0" "cball (x, f (x, y)) e' \ H ` ball (x, y) (r/2)" by auto have inv_subset: "the_inv_into (ball (x, y) r) H a = the_inv_into R H a" if "a \ H ` R" "R \ (ball (x, y) r)" for a R apply (rule the_inv_into_f_eq[OF inj]) apply (rule f_the_inv_into_f) apply (rule inj_on_subset[OF inj]) apply fact apply fact apply (rule the_inv_into_into) apply (rule inj_on_subset[OF inj]) apply fact apply fact apply (rule order_trans) apply fact using r apply auto done have GH: "G (H z) = z" if "dist (x, y) z < r" for z by (auto simp: G_def the_inv_into_f_f inj that) define e where "e = min (e' / 2) e0" define r2 where "r2 = r / 2" have r2: "r2 > 0" "r2 < r" using \r > 0\ by (auto simp: r2_def) have "e > 0" using e' e0 by (auto simp: e_def) from cball_times_subset[of "x" e' "f (x, y)"] e' have "cball x e \ cball (f (x, y)) e \ H ` ball (x, y) (r/2)" by (force simp: e_def) then have e_r_subset: "z \ cball x e \ (z, 0) \ H ` ball (x, y) (r/2)" for z using \0 < e\ assms(6) by (auto simp: H_def subset_iff) have u0: "(u, 0) \ H ` ball (x, y) r" if "u \ cball x e" for u apply (rule rev_subsetD) apply (rule e_r_subset) apply fact unfolding r2_def using r2 by auto have G_r: "G (u, 0) \ ball (x, y) r" if "u \ cball x e" for u unfolding G_def apply (rule the_inv_into_into) apply fact apply (auto) apply (rule u0, fact) done note e_r_subset ultimately have G2: "f (x, snd (G (x, 0))) = 0" "snd (G (x, 0)) = y" "\u. u \ cball x e \ f (u, snd (G (u, 0))) = 0" "continuous_on (cball x e) (\u. snd (G (u, 0)))" "(\t. (t, snd (G (t, 0)))) ` cball x e \ S" "e > 0" "((\u. snd (G (u, 0))) has_derivative (\u. snd (Hi (u, 0)))) (at x)" apply (auto simp: G_def split_beta' intro!: continuous_intros continuous_on_compose2[OF cH]) subgoal premises prems proof - have "the_inv_into (ball (x, y) r) H (x, 0) = (x, y)" apply (rule the_inv_into_f_eq) apply fact by (auto simp: H_def assms \r > 0\) then show ?thesis by auto qed using r2(2) r2_def apply fastforce apply (subst continuous_on_cong[OF refl]) apply (rule inv_subset[where R="cball (x, y) r2"]) subgoal using r2 apply auto using r2_def by force subgoal using r2 by (force simp:) subgoal apply (rule continuous_on_compose2[OF continuous_on_inv_into]) using r(2) r2(2) apply (auto simp: r2_def[symmetric] intro!: continuous_on_compose2[OF cH] continuous_intros) apply (rule inj_on_subset) apply (rule inj) using r(2) r2(2) apply force apply force done subgoal premises prems for u proof - from prems have u: "u \ cball x e" by auto note G_r[OF u] also have "ball (x, y) r \ S" using r by simp finally have "(G (u, 0)) \ S" . then show ?thesis unfolding G_def[symmetric] using uvs(2)[OF u0, OF u] by (metis prod.collapse) qed subgoal using \e > 0\ by simp subgoal premises prems proof - have "(x, y) \ cball (x, y) r2" using r2 by auto moreover have "H (x, y) \ interior (H ` cball (x, y) r2)" apply (rule interiorI[OF oH]) using r2 by (auto simp: r2_def) moreover have "cball (x, y) r2 \ S" using r r2 by auto moreover have "\z. z \ cball (x, y) r2 \ G (H z) = z" using r2 by (auto intro!: GH) ultimately have "(G has_derivative Hi) (at (H (x, y)))" proof (rule has_derivative_inverse[where g = G and f = H, OF compact_cball _ _ continuous_on_subset[OF cH] _ H' _ _]) show "blinfun_apply Hi \ blinfun_apply (H' (x, y)) = id" using Hi by transfer auto qed (use S blinfun.bounded_linear_right in auto) then have g': "(G has_derivative Hi) (at (x, 0))" by (auto simp: H_def assms) show ?thesis unfolding G_def[symmetric] H_def[symmetric] apply (auto intro!: derivative_eq_intros) apply (rule has_derivative_compose[where g=G and f="\x. (x, 0)"]) apply (auto intro!: g' derivative_eq_intros) done qed done moreover note \r > 0\ moreover define u where "u \ \x. snd (G (x, 0))" have local_unique: "u s = v s" if solves: "(\s. s \ U \ f (s, v s) = 0)" and i: "v x = y" and v: "continuous_on U v" and s: "s \ U" and s': "(s, v s) \ ball (x, y) r" and U: "U \ cball x e" for U v s proof - have H_eq: "H (s, v s) = H (s, u s)" apply (auto simp: H_def solves[OF s]) unfolding u_def apply (rule G2) apply (rule subsetD; fact) done have "(s, snd (G (s, 0))) = (G (s, 0))" using GH H_def s s' solves by fastforce also have "\ \ ball (x, y) r" unfolding G_def apply (rule the_inv_into_into) apply fact apply (rule u0) apply (rule subsetD; fact) apply (rule order_refl) done finally have "(s, u s) \ ball (x, y) r" unfolding u_def . from inj_onD[OF inj H_eq s' this] show "u s = v s" by auto qed ultimately show ?thesis unfolding u_def Hi' .. qed lemma implicit_function_theorem_unique: fixes f::"'a::euclidean_space * 'b::euclidean_space \ 'c::euclidean_space"\ \TODO: generalize?!\ assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative blinfun_apply (f' x)) (at x)" assumes S: "(x, y) \ S" "open S" assumes D: "DIM('c) \ DIM('b)" assumes f'C: "continuous_on S f'" assumes z: "f (x, y) = 0" assumes T2: "T o\<^sub>L (f' (x, y) o\<^sub>L embed2_blinfun) = 1\<^sub>L" assumes T1: "(f' (x, y) o\<^sub>L embed2_blinfun) o\<^sub>L T = 1\<^sub>L"\ \TODO: reduce?!\ obtains u e where "f (x, u x) = 0" "u x = y" "\s. s \ cball x e \ f (s, u s) = 0" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ S" "e > 0" "(u has_derivative (- T o\<^sub>L f' (x, y) o\<^sub>L embed1_blinfun)) (at x)" "\s. s \ cball x e \ f' (s, u s) o\<^sub>L embed2_blinfun \ invertibles_blinfun" "\U v s. (\s. s \ U \ f (s, v s) = 0) \ u x = v x \ continuous_on U v \ s \ U \ x \ U \ U \ cball x e \ connected U \ open U \ u s = v s" proof - from T1 T2 have f'I: "f' (x, y) o\<^sub>L embed2_blinfun \ invertibles_blinfun" by (auto simp: invertibles_blinfun_def) from assms have f'Cg: "s \ S \ isCont f' s" for s by (auto simp: continuous_on_eq_continuous_at[OF \open S\]) then have f'C: "isCont f' (x, y)" by (auto simp: S) obtain u e1 r where u: "f (x, u x) = 0" "u x = y" "\s. s \ cball x e1 \ f (s, u s) = 0" "continuous_on (cball x e1) u" "(\t. (t, u t)) ` cball x e1 \ S" "e1 > 0" "(u has_derivative (- T o\<^sub>L f' (x, y) o\<^sub>L embed1_blinfun)) (at x)" and unique_u: "r > 0" "(\v s U. v x = y \ (\s. s \ U \ f (s, v s) = 0) \ continuous_on U v \ s \ U \ U \ cball x e1 \ (s, v s) \ ball (x, y) r \ u s = v s)" by (rule implicit_function_theorem[OF f' S D f'C z T2 T1]; blast) from openE[OF blinfun_inverse_open f'I] obtain d where d: "0 < d" "ball (f' (x, y) o\<^sub>L embed2_blinfun) d \ invertibles_blinfun" by auto note [continuous_intros] = continuous_at_compose[OF _ f'Cg, unfolded o_def] from \continuous_on _ u\ have "continuous_on (ball x e1) u" by (rule continuous_on_subset) auto then have "\s. s \ ball x e1 \ isCont u s" unfolding continuous_on_eq_continuous_at[OF open_ball] by auto note [continuous_intros] = continuous_at_compose[OF _ this, unfolded o_def] from assms have f'Ce: "isCont (\s. f' (s, u s) o\<^sub>L embed2_blinfun) x" by (auto simp: u intro!: continuous_intros) from f'Ce[unfolded isCont_def, THEN tendstoD, OF \0 < d\] d obtain e0 where "e0 > 0" "\s. s \ x \ s \ ball x e0 \ (f' (s, u s) o\<^sub>L embed2_blinfun) \ invertibles_blinfun" by (auto simp: eventually_at dist_commute subset_iff u) then have e0: "s \ ball x e0 \ (f' (s, u s) o\<^sub>L embed2_blinfun) \ invertibles_blinfun" for s by (cases "s = x") (auto simp: f'I \0 < d\ u) define e where "e = min (e0/2) (e1/2)" have e: "f (x, u x) = 0" "u x = y" "\s. s \ cball x e \ f (s, u s) = 0" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ S" "e > 0" "(u has_derivative (- T o\<^sub>L f' (x, y) o\<^sub>L embed1_blinfun)) (at x)" "\s. s \ cball x e \ f' (s, u s) o\<^sub>L embed2_blinfun \ invertibles_blinfun" using e0 u \e0 > 0\ by (auto simp: e_def intro: continuous_on_subset) from u(4) have "continuous_on (ball x e1) u" apply (rule continuous_on_subset) using \e1 > 0\ by (auto simp: e_def) then have "\s. s \ cball x e \ isCont u s" using \e0 > 0\ \e1 > 0\ unfolding continuous_on_eq_continuous_at[OF open_ball] by (auto simp: e_def Ball_def dist_commute) note [continuous_intros] = continuous_at_compose[OF _ this, unfolded o_def] have "u s = v s" if solves: "(\s. s \ U \ f (s, v s) = 0)" and i: "u x = v x" and v: "continuous_on U v" and s: "s \ U" and U: "x \ U" "U \ cball x e" "connected U" "open U" for U v s proof - define M where "M = {s \ U. u s = v s}" have "x \ M" using i U by (auto simp: M_def) moreover have "continuous_on U (\s. u s - v s)" by (auto intro!: continuous_intros v continuous_on_subset[OF e(4) U(2)]) from continuous_closedin_preimage[OF this closed_singleton[where a=0]] have "closedin (top_of_set U) M" by (auto simp: M_def vimage_def Collect_conj_eq) moreover have "\s. s \ U \ isCont v s" using v unfolding continuous_on_eq_continuous_at[OF \open U\] by auto note [continuous_intros] = continuous_at_compose[OF _ this, unfolded o_def] { fix a assume "a \ M" then have aU: "a \ U" and u_v: "u a = v a" by (auto simp: M_def) then have a_ball: "a \ cball x e" and a_dist: "dist x a \ e" using U by auto then have a_S: "(a, u a) \ S" using e by auto have fa_z: "f (a, u a) = 0" using \a \ cball x e\ by (auto intro!: e) from e(8)[OF \a \ cball _ _\] obtain Ta where Ta: "Ta o\<^sub>L (f' (a, u a) o\<^sub>L embed2_blinfun) = 1\<^sub>L" "f' (a, u a) o\<^sub>L embed2_blinfun o\<^sub>L Ta = 1\<^sub>L" by (auto simp: invertibles_blinfun_def ac_simps) obtain u' e' r' where "r' > 0" "e' > 0" and u': "\v s U. v a = u a \ (\s. s \ U \ f (s, v s) = 0) \ continuous_on U v \ s \ U \ U \ cball a e' \ (s, v s) \ ball (a, u a) r' \ u' s = v s" by (rule implicit_function_theorem[OF f' a_S \open S\ D f'Cg[OF a_S] fa_z Ta]; blast) from openE[OF \open U\ aU] obtain dU where dU: "dU > 0" "\s. s \ ball a dU \ s \ U" by (auto simp: dist_commute subset_iff) have v_tendsto: "((\s. (s, v s)) \ (a, u a)) (at a)" unfolding u_v by (subst continuous_at[symmetric]) (auto intro!: continuous_intros aU) from tendstoD[OF v_tendsto \0 < r'\, unfolded eventually_at] obtain dv where "dv > 0" "s \ a \ dist s a < dv \ (s, v s) \ ball (a, u a) r'" for s by (auto simp: dist_commute) then have dv: "dist s a < dv \ (s, v s) \ ball (a, u a) r'" for s by (cases "s = a") (auto simp: u_v \0 < r'\) have v_tendsto: "((\s. (s, u s)) \ (a, u a)) (at a)" using a_dist by (subst continuous_at[symmetric]) (auto intro!: continuous_intros) from tendstoD[OF v_tendsto \0 < r'\, unfolded eventually_at] obtain du where "du > 0" "s \ a \ dist s a < du \ (s, u s) \ ball (a, u a) r'" for s by (auto simp: dist_commute) then have du: "dist s a < du \ (s, u s) \ ball (a, u a) r'" for s by (cases "s = a") (auto simp: u_v \0 < r'\) { fix s assume s: "s \ ball a (Min {dU, e', dv, du})" let ?U = "ball a (Min {dU, e', dv, du})" have balls: "ball a (Min {dU, e', dv, du}) \ cball a e'" by auto have dsadv: "dist s a < dv" using s by (auto simp: dist_commute) have dsadu: "dist s a < du" using s by (auto simp: dist_commute) have U_U: "\s. s \ ball a (Min {dU, e', dv, du}) \ s \ U" using dU by auto have U_e: "\s. s \ ball a (Min {dU, e', dv, du}) \ s \ cball x e" using dU U by (auto simp: dist_commute subset_iff) have cv: "continuous_on ?U v" using v apply (rule continuous_on_subset) using dU by auto have cu: "continuous_on ?U u" using e(4) apply (rule continuous_on_subset) using dU U(2) by auto from u'[where v=v, OF u_v[symmetric] solves[OF U_U] cv s balls dv[OF dsadv]] u'[where v=u, OF refl e(3)[OF U_e] cu s balls du[OF dsadu]] have "v s = u s" by auto } then have "\dv>0. \s \ ball a dv. v s = u s" using \0 < dU\ \0 < e'\ \0 < dv\ \0 < du\ by (auto intro!: exI[where x="(Min {dU, e', dv, du})"]) } note ex = this have "openin (top_of_set U) M" unfolding openin_contains_ball apply (rule conjI) subgoal using U by (auto simp: M_def) apply (auto simp:) apply (drule ex) apply auto subgoal for x d by (rule exI[where x=d]) (auto simp: M_def) done ultimately have "M = U" using \connected U\ by (auto simp: connected_clopen) with \s \ U\ show ?thesis by (auto simp: M_def) qed from e this show ?thesis .. qed lemma uniform_limit_compose: assumes ul: "uniform_limit T f l F" assumes uc: "uniformly_continuous_on S s" assumes ev: "\\<^sub>F x in F. f x ` T \ S" assumes subs: "l ` T \ S" shows "uniform_limit T (\i x. s (f i x)) (\x. s (l x)) F" proof (rule uniform_limitI) fix e::real assume "e > 0" from uniformly_continuous_onE[OF uc \e > 0\] obtain d where d: "0 < d" "\t t'. t \ S \ t' \ S \ dist t' t < d \ dist (s t') (s t) < e" by auto from uniform_limitD[OF ul \0 < d\] have "\\<^sub>F n in F. \x\T. dist (f n x) (l x) < d" . then show "\\<^sub>F n in F. \x\T. dist (s (f n x)) (s (l x)) < e" using ev by eventually_elim (use d subs in force) qed lemma uniform_limit_in_open: fixes l::"'a::topological_space\'b::heine_borel" assumes ul: "uniform_limit T f l (at x)" assumes cont: "continuous_on T l" assumes compact: "compact T" and T_ne: "T \ {}" assumes B: "open B" assumes mem: "l ` T \ B" shows "\\<^sub>F y in at x. \t \ T. f y t \ B" proof - have l_ne: "l ` T \ {}" using T_ne by auto have "compact (l ` T)" by (auto intro!: compact_continuous_image cont compact) from compact_in_open_separated[OF l_ne this B mem] obtain e where "e > 0" "{x. infdist x (l ` T) \ e} \ B" by auto from uniform_limitD[OF ul \0 < e\] have "\\<^sub>F n in at x. \x\T. dist (f n x) (l x) < e" . then show ?thesis proof eventually_elim case (elim y) show ?case proof safe fix t assume "t \ T" have "infdist (f y t) (l ` T) \ dist (f y t) (l t)" by (rule infdist_le) (use \t \ T\ in auto) also have "\ < e" using elim \t \ T\ by auto finally have "infdist (f y t) (l ` T) \ e" by simp then have "(f y t) \ {x. infdist x (l ` T) \ e}" by (auto ) also note \\ \ B\ finally show "f y t \ B" . qed qed qed lemma order_uniform_limitD1: fixes l::"'a::topological_space\real"\ \TODO: generalize?!\ assumes ul: "uniform_limit T f l (at x)" assumes cont: "continuous_on T l" assumes compact: "compact T" assumes less: "\t. t \ T \ l t < b" shows "\\<^sub>F y in at x. \t \ T. f y t < b" proof cases assume ne: "T \ {}" from compact_attains_sup[OF compact_continuous_image[OF cont compact], unfolded image_is_empty, OF ne] obtain tmax where tmax: "tmax \ T" "\s. s \ T \ l s \ l tmax" by auto have "b - l tmax > 0" using ne tmax less by auto from uniform_limitD[OF ul this] have "\\<^sub>F n in at x. \x\T. dist (f n x) (l x) < b - l tmax" by auto then show ?thesis apply eventually_elim using tmax by (force simp: dist_real_def abs_real_def split: if_splits) qed auto lemma order_uniform_limitD2: fixes l::"'a::topological_space\real"\ \TODO: generalize?!\ assumes ul: "uniform_limit T f l (at x)" assumes cont: "continuous_on T l" assumes compact: "compact T" assumes less: "\t. t \ T \ l t > b" shows "\\<^sub>F y in at x. \t \ T. f y t > b" proof - have "\\<^sub>F y in at x. \t\T. (- f) y t < - b" by (rule order_uniform_limitD1[of "- f" T "-l" x "- b"]) (auto simp: assms fun_Compl_def intro!: uniform_limit_eq_intros continuous_intros) then show ?thesis by auto qed lemma continuous_on_avoid_cases: fixes l::"'b::topological_space \ 'a::linear_continuum_topology"\ \TODO: generalize!\ assumes cont: "continuous_on T l" and conn: "connected T" assumes avoid: "\t. t \ T \ l t \ b" obtains "\t. t \ T \ l t < b" | "\t. t \ T \ l t > b" apply atomize_elim using connected_continuous_image[OF cont conn] using avoid unfolding connected_iff_interval apply (auto simp: image_iff) using leI by blast lemma order_uniform_limit_ne: fixes l::"'a::topological_space\real"\ \TODO: generalize?!\ assumes ul: "uniform_limit T f l (at x)" assumes cont: "continuous_on T l" assumes compact: "compact T" and conn: "connected T" assumes ne: "\t. t \ T \ l t \ b" shows "\\<^sub>F y in at x. \t \ T. f y t \ b" proof - from continuous_on_avoid_cases[OF cont conn ne] consider "(\t. t \ T \ l t < b)" | "(\t. t \ T \ l t > b)" by blast then show ?thesis proof cases case 1 from order_uniform_limitD1[OF ul cont compact 1] have "\\<^sub>F y in at x. \t\T. f y t < b" by simp then show ?thesis by eventually_elim auto next case 2 from order_uniform_limitD2[OF ul cont compact 2] have "\\<^sub>F y in at x. \t\T. f y t > b" by simp then show ?thesis by eventually_elim auto qed qed lemma open_cballE: assumes "open S" "x\S" obtains e where "e>0" "cball x e \ S" using assms unfolding open_contains_cball by auto lemma pos_half_less: fixes x::real shows "x > 0 \ x / 2 < x" by auto lemma closed_levelset: "closed {x. s x = (c::'a::t1_space)}" if "continuous_on UNIV s" proof - have "{x. s x = c} = s -` {c}" by auto also have "closed \" apply (rule closed_vimage) apply (rule closed_singleton) apply (rule that) done finally show ?thesis . qed lemma closed_levelset_within: "closed {x \ S. s x = (c::'a::t1_space)}" if "continuous_on S s" "closed S" proof - have "{x \ S. s x = c} = s -` {c} \ S" by auto also have "closed \" apply (rule continuous_on_closed_vimageI) apply (rule that) apply (rule that) apply simp done finally show ?thesis . qed context c1_on_open_euclidean begin lemma open_existence_ivlE: assumes "t \ existence_ivl0 x" "t \ 0" obtains e where "e > 0" "cball x e \ {0 .. t + e} \ Sigma X existence_ivl0" proof - from assms have "(x, t) \ Sigma X existence_ivl0" by auto from open_cballE[OF open_state_space this] obtain e0' where e0: "0 < e0'" "cball (x, t) e0' \ Sigma X existence_ivl0" by auto define e0 where "e0 = (e0' / 2)" from cball_times_subset[of x e0' t] pos_half_less[OF \0 < e0'\] half_gt_zero[OF \0 < e0'\] e0 have "cball x e0 \ cball t e0 \ Sigma X existence_ivl0" "0 < e0" "e0 < e0'" unfolding e0_def by auto then have "e0 > 0" "cball x e0 \ {0..t + e0} \ Sigma X existence_ivl0" apply (auto simp: subset_iff dest!: spec[where x=t]) subgoal for a b apply (rule in_existence_between_zeroI) apply (drule spec[where x=a]) apply (drule spec[where x="t + e0"]) apply (auto simp: dist_real_def closed_segment_eq_real_ivl) done done then show ?thesis .. qed lemmas [derivative_intros] = flow0_comp_has_derivative lemma flow_isCont_state_space_comp[continuous_intros]: "t x \ existence_ivl0 (s x) \ isCont s x \ isCont t x \ isCont (\x. flow0 (s x) (t x)) x" using continuous_within_compose3[where g="\(x, t). flow0 x t" and f="\x. (s x, t x)" and x = x and s = UNIV] flow_isCont_state_space by auto lemma closed_plane[simp]: "closed {x. x \ i = c}" using closed_hyperplane[of i c] by (auto simp: inner_commute) lemma flow_tendsto_compose[tendsto_intros]: assumes "(x \ xs) F" "(t \ ts) F" assumes "ts \ existence_ivl0 xs" shows "((\s. flow0 (x s) (t s)) \ flow0 xs ts) F" proof - have ev: "\\<^sub>F s in F. (x s, t s) \ Sigma X existence_ivl0" using tendsto_Pair[OF assms(1,2), THEN topological_tendstoD, OF open_state_space] assms by auto show ?thesis by (rule continuous_on_tendsto_compose[OF flow_continuous_on_state_space tendsto_Pair, unfolded split_beta' fst_conv snd_conv]) (use assms ev in auto) qed lemma returns_to_implicit_function: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P x") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "isCont Ds (poincare_map ?P x)" assumes nz: "Ds (poincare_map ?P x) (f (poincare_map ?P x)) \ 0" obtains u e where "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma X existence_ivl0" "0 < e" "(u has_derivative (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L flowderiv x (return_time ?P x)) o\<^sub>L embed1_blinfun)) (at x)" proof - note [derivative_intros] = has_derivative_compose[OF _ Ds] have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) note cls[simp, intro] = closed_levelset[OF cont_s] let ?t1 = "return_time ?P x" have cls[simp, intro]: "closed {x \ S. s x = 0}" by (rule closed_levelset_within) (auto intro!: cS continuous_on_subset[OF cont_s]) then have xt1: "(x, ?t1) \ Sigma X existence_ivl0" by (auto intro!: return_time_exivl rt) have D: "(\x. x \ Sigma X existence_ivl0 \ ((\(x, t). s (flow0 x t)) has_derivative blinfun_apply (Ds (flow0 (fst x) (snd x)) o\<^sub>L (flowderiv (fst x) (snd x)))) (at x))" by (auto intro!: derivative_eq_intros) have C: "isCont (\x. Ds (flow0 (fst x) (snd x)) o\<^sub>L flowderiv (fst x) (snd x)) (x, ?t1)" using flowderiv_continuous_on[unfolded continuous_on_eq_continuous_within, rule_format, OF xt1] using at_within_open[OF xt1 open_state_space] by (auto intro!: continuous_intros tendsto_eq_intros return_time_exivl rt isCont_tendsto_compose[OF DsC, unfolded poincare_map_def] simp: split_beta' isCont_def) from return_time_returns[OF rt cls] have Z: "(case (x, ?t1) of (x, t) \ s (flow0 x t)) = 0" by (auto simp: ) have I1: "blinfun_scaleR_left (inverse (Ds (flow0 x (?t1))(f (flow0 x (?t1))))) o\<^sub>L ((Ds (flow0 (fst (x, return_time {x \ S. s x = 0} x)) (snd (x, return_time {x \ S. s x = 0} x))) o\<^sub>L flowderiv (fst (x, return_time {x \ S. s x = 0} x)) (snd (x, return_time {x \ S. s x = 0} x))) o\<^sub>L embed2_blinfun) = 1\<^sub>L" using nz by (auto intro!: blinfun_eqI simp: rt flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) have I2: "((Ds (flow0 (fst (x, return_time {x \ S. s x = 0} x)) (snd (x, return_time {x \ S. s x = 0} x))) o\<^sub>L flowderiv (fst (x, return_time {x \ S. s x = 0} x)) (snd (x, return_time {x \ S. s x = 0} x))) o\<^sub>L embed2_blinfun) o\<^sub>L blinfun_scaleR_left (inverse (Ds (flow0 x (?t1))(f (flow0 x (?t1))))) = 1\<^sub>L" using nz by (auto intro!: blinfun_eqI simp: rt flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) show ?thesis apply (rule implicit_function_theorem[where f="\(x, t). s (flow0 x t)" and S="Sigma X existence_ivl0", OF D xt1 open_state_space order_refl C Z I1 I2]) apply blast unfolding split_beta' fst_conv snd_conv poincare_map_def[symmetric] .. qed lemma (in auto_ll_on_open) f_tendsto[tendsto_intros]: assumes g1: "(g1 \ b1) (at s within S)" and "b1 \ X" shows "((\x. f (g1 x)) \ f b1) (at s within S)" apply (rule continuous_on_tendsto_compose[OF continuous tendsto_Pair[OF tendsto_const], unfolded split_beta fst_conv snd_conv, OF g1]) by (auto simp: \b1 \ X\ intro!: topological_tendstoD[OF g1]) lemma flow_avoids_surface_eventually_at_right_pos: assumes "s x > 0 \ s x = 0 \ blinfun_apply (Ds x) (f x) > 0" assumes x: "x \ X" assumes Ds: "\x. (s has_derivative Ds x) (at x)" assumes DsC: "\x. isCont Ds x" shows "\\<^sub>F t in at_right 0. s (flow0 x t) > (0::real)" proof - have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) then have [THEN continuous_on_compose2, continuous_intros]: "continuous_on S s" for S by (rule continuous_on_subset) simp note [derivative_intros] = has_derivative_compose[OF _ Ds] note [tendsto_intros] = continuous_on_tendsto_compose[OF cont_s] isCont_tendsto_compose[OF DsC] from assms(1) consider "s x > 0" | "s x = 0" "blinfun_apply (Ds x) (f x) > 0" by auto then show ?thesis proof cases assume s: "s x > 0" then have "((\t. s (flow0 x t)) \ s x) (at_right 0)" by (auto intro!: tendsto_eq_intros simp: split_beta' x) from order_tendstoD(1)[OF this s] show ?thesis . next assume sz: "s x = 0" and pos: "blinfun_apply (Ds x) (f x) > 0" from x have "0 \ existence_ivl0 x" "open (existence_ivl0 x)" by simp_all then have evex: "\\<^sub>F t in at_right 0. t \ existence_ivl0 x" using eventually_at_topological by blast moreover from evex have "\\<^sub>F xa in at_right 0. flow0 x xa \ X" by (eventually_elim) (auto intro!: ) then have "((\t. (Ds (flow0 x t)) (f (flow0 x t))) \ blinfun_apply (Ds x) (f x)) (at_right 0)" by (auto intro!: tendsto_eq_intros simp: split_beta' x) from order_tendstoD(1)[OF this pos] have "\\<^sub>F z in at_right 0. blinfun_apply (Ds (flow0 x z)) (f (flow0 x z)) > 0" . then obtain t where t: "t > 0" "\z. 0 < z \ z < t \ blinfun_apply (Ds (flow0 x z)) (f (flow0 x z)) > 0" by (auto simp: eventually_at) have "\\<^sub>F z in at_right 0. z < t" using \t > 0\ order_tendstoD(2)[OF tendsto_ident_at \0 < t\] by auto moreover have "\\<^sub>F z in at_right 0. 0 < z" by (simp add: eventually_at_filter) ultimately show ?thesis proof eventually_elim case (elim z) from closed_segment_subset_existence_ivl[OF \z \ existence_ivl0 x\] have csi: "{0..z} \ existence_ivl0 x" by (auto simp add: closed_segment_eq_real_ivl) then have cont: "continuous_on {0..z} (\t. s (flow0 x t))" by (auto intro!: continuous_intros) have "\u. \0 < u; u < z\ \ ((\t. s (flow0 x t)) has_derivative (\t. t * blinfun_apply (Ds (flow0 x u)) (f (flow0 x u)))) (at u)" using csi by (auto intro!: derivative_eq_intros simp: flowderiv_def blinfun.bilinear_simps) from mvt[OF \0 < z\ cont this] obtain w where w: "0 < w" "w < z" and sDs: "s (flow0 x z) = z * blinfun_apply (Ds (flow0 x w)) (f (flow0 x w))" using x sz by auto note sDs also have "\ > 0" using elim t(2)[of w] w by simp finally show ?case . qed qed qed lemma flow_avoids_surface_eventually_at_right_neg: assumes "s x < 0 \ s x = 0 \ blinfun_apply (Ds x) (f x) < 0" assumes x: "x \ X" assumes Ds: "\x. (s has_derivative Ds x) (at x)" assumes DsC: "\x. isCont Ds x" shows "\\<^sub>F t in at_right 0. s (flow0 x t) < (0::real)" apply (rule flow_avoids_surface_eventually_at_right_pos[of "-s" x "-Ds", simplified]) using assms by (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps fun_Compl_def) lemma flow_avoids_surface_eventually_at_right: assumes "x \ S \ s x \ 0 \ blinfun_apply (Ds x) (f x) \ 0" assumes x: "x \ X" and cS: "closed S" assumes Ds: "\x. (s has_derivative Ds x) (at x)" assumes DsC: "\x. isCont Ds x" shows "\\<^sub>F t in at_right 0. (flow0 x t) \ {x \ S. s x = (0::real)}" proof - from assms(1) consider "s x > 0 \ s x = 0 \ blinfun_apply (Ds x) (f x) > 0" | "s x < 0 \ s x = 0 \ blinfun_apply (Ds x) (f x) < 0" | "x \ S" by arith then show ?thesis proof cases case 1 from flow_avoids_surface_eventually_at_right_pos[of s x Ds, OF 1 x Ds DsC] show ?thesis by eventually_elim auto next case 2 from flow_avoids_surface_eventually_at_right_neg[of s x Ds, OF 2 x Ds DsC] show ?thesis by eventually_elim auto next case 3 then have nS: "open (- S)" "x \ - S" using cS by auto have "\\<^sub>F t in at_right 0. (flow0 x t) \ - S" by (rule topological_tendstoD[OF _ nS]) (auto intro!: tendsto_eq_intros simp: x) then show ?thesis by eventually_elim auto qed qed lemma eventually_returns_to: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P x") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "\x. isCont Ds x" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). s x = 0 \ x \ S" assumes nz: "Ds (poincare_map ?P x) (f (poincare_map ?P x)) \ 0" assumes nz0: "x \ S \ s x \ 0 \ Ds x (f x) \ 0" shows "\\<^sub>F x in at x. returns_to ?P x" proof - let ?t1 = "return_time ?P x" have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) have cont_s': "continuous_on S s" for S by (rule continuous_on_subset[OF cont_s subset_UNIV]) note s_tendsto[tendsto_intros] = continuous_on_tendsto_compose[OF cont_s, THEN tendsto_eq_rhs] note cls[simp, intro] = closed_levelset_within[OF cont_s' cS, of 0] note [tendsto_intros] = continuous_on_tendsto_compose[OF cont_s] isCont_tendsto_compose[OF DsC] obtain u e where "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma X existence_ivl0" "0 < e" by (rule returns_to_implicit_function[OF rt cS Ds DsC nz]; blast) then have u: "s (flow0 x (u x)) = 0" "u x = ?t1" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "\z. z \ cball x e \ u z \ existence_ivl0 z" "e > 0" by (force simp: split_beta')+ have "\\<^sub>F y in at x. y \ ball x e" using eventually_at_ball[OF \0 < e\] by eventually_elim auto then have ev_cball: "\\<^sub>F y in at x. y \ cball x e" by eventually_elim (use \e > 0\ in auto) moreover have "continuous_on (ball x e) u" using u by (auto simp: continuous_on_subset) then have [tendsto_intros]: "(u \ u x) (at x)" using \e > 0\ at_within_open[of y "ball x e" for y] by (auto simp: continuous_on_def) then have flow0_u_tendsto: "(\x. flow0 x (u x)) \x\ poincare_map ?P x" by (auto intro!: tendsto_eq_intros u return_time_exivl rt simp: poincare_map_def) have s_imp: "s (poincare_map {x \ S. s x = 0} x) = 0 \ poincare_map {x \ S. s x = 0} x \ S" using poincare_map_returns[OF rt] by auto from eventually_tendsto_compose_within[OF eventually_inside s_imp flow0_u_tendsto] have "\\<^sub>F x in at x. s (flow0 x (u x)) = 0 \ flow0 x (u x) \ S" by auto with ev_cball have "\\<^sub>F x in at x. flow0 x (u x) \ S" by eventually_elim (auto simp: u) moreover { have "x \ X" using u(5) u(6) by force from ev_cball have ev_X: "\\<^sub>F y in at x. y \ X"\ \eigentlich ist das \open X\\ apply eventually_elim apply (rule) by (rule u) moreover { { assume a: "x \ S" then have "open (-S)" "x \ - S" using cS by auto from topological_tendstoD[OF tendsto_ident_at this] have "(\\<^sub>F y in at x. y \ S)" by auto } moreover { assume a: "s x \ 0" have "(\\<^sub>F y in at x. s y \ 0)" by (rule tendsto_imp_eventually_ne[OF _ a]) (auto intro!: tendsto_eq_intros) } moreover { assume a: "(Ds x) (f x) \ 0" have "(\\<^sub>F y in at x. blinfun_apply (Ds y) (f y) \ 0)" by (rule tendsto_imp_eventually_ne[OF _ a]) (auto intro!: tendsto_eq_intros ev_X \x \ X\) } ultimately have "(\\<^sub>F y in at x. y \ S) \ (\\<^sub>F y in at x. s y \ 0) \ (\\<^sub>F y in at x. blinfun_apply (Ds y) (f y) \ 0)" using nz0 by auto then have "\\<^sub>F y in at x. y \ S \ s y \ 0 \ blinfun_apply (Ds y) (f y) \ 0" apply - apply (erule disjE) subgoal by (rule eventually_elim2, assumption, assumption, blast) subgoal apply (erule disjE) subgoal by (rule eventually_elim2, assumption, assumption, blast) subgoal by (rule eventually_elim2, assumption, assumption, blast) done done } ultimately have "\\<^sub>F y in at x. (y \ S \ s y \ 0 \ blinfun_apply (Ds y) (f y) \ 0) \ y \ X" by eventually_elim auto } then have "\\<^sub>F y in at x. \\<^sub>F t in at_right 0. flow0 y t \ {x \ S. s x = 0}" apply eventually_elim by (rule flow_avoids_surface_eventually_at_right[where Ds=Ds]) (auto intro!: Ds DsC cS) moreover have at_eq: "(at x within cball x e) = at x" apply (rule at_within_interior) apply (auto simp: \e > 0\) done have "u x > 0" using u(1) by (auto simp: u rt cont_s' intro!: return_time_pos closed_levelset_within cS) then have "\\<^sub>F y in at x. u y > 0" apply (rule order_tendstoD[rotated]) using u(4) apply (auto simp: continuous_on_def) apply (drule bspec[where x=x]) using \e > 0\ by (auto simp: at_eq) ultimately show "\\<^sub>F y in at x. returns_to ?P y" apply eventually_elim subgoal premises prems for y apply (rule returns_toI[where t="u y"]) subgoal using prems by auto subgoal apply (rule u) apply (rule prems) done subgoal using u(3)[of y] prems by auto subgoal using prems(3) by eventually_elim auto subgoal by simp done done qed lemma return_time_isCont_outside: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P x") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "\x. isCont Ds x" assumes through: "(Ds (poincare_map ?P x)) (f (poincare_map ?P x)) \ 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). s x = 0 \ x \ S" assumes outside: "x \ S \ s x \ 0" shows "isCont (return_time ?P) x" unfolding isCont_def proof (rule tendstoI) fix e_orig::real assume "e_orig > 0" define e where "e = e_orig / 2" have "e > 0" using \e_orig > 0\ by (simp add: e_def) have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) then have s_tendsto: "(s \ s x) (at x)" for x by (auto simp: continuous_on_def) have cont_s': "continuous_on S s" by (rule continuous_on_subset[OF cont_s subset_UNIV]) note cls[simp, intro] = closed_levelset_within[OF cont_s' cS(1)] have "{x. s x = 0} = s -` {0}" by auto have ret_exivl: "return_time ?P x \ existence_ivl0 x" by (rule return_time_exivl; fact) then have [intro, simp]: "x \ X" by auto have isCont_Ds_f: "isCont (\s. Ds s (f s)) (poincare_map ?P x)" apply (auto intro!: continuous_intros DsC) apply (rule has_derivative_continuous) apply (rule derivative_rhs) by (auto simp: poincare_map_def intro!: flow_in_domain return_time_exivl assms) obtain u eu where u: "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\y. y \ cball x eu \ s (flow0 y (u y)) = 0)" "continuous_on (cball x eu) u" "(\t. (t, u t)) ` cball x eu \ Sigma X existence_ivl0" "0 < eu" by (rule returns_to_implicit_function[OF rt cS(1) Ds DsC through]; blast) have u_tendsto: "(u \ u x) (at x)" unfolding isCont_def[symmetric] apply (rule continuous_on_interior[OF u(4)]) using \0 < eu\ by auto have "u x > 0" by (auto simp: u intro!: return_time_pos rt) from order_tendstoD(1)[OF u_tendsto this] have "\\<^sub>F x in at x. 0 < u x" . moreover have "\\<^sub>F y in at x. y \ cball x eu" using eventually_at_ball[OF \0 < eu\, of x] by eventually_elim auto moreover have "x \ S \ s x \ 0 \ blinfun_apply (Ds x) (f x) \ 0" using outside by auto have returns: "\\<^sub>F y in at x. returns_to ?P y" by (rule eventually_returns_to; fact) moreover have "\\<^sub>F y in at x. y \ ball x eu" using eventually_at_ball[OF \0 < eu\] by eventually_elim simp then have ev_cball: "\\<^sub>F y in at x. y \ cball x eu" by eventually_elim (use \e > 0\ in auto) have "continuous_on (ball x eu) u" using u by (auto simp: continuous_on_subset) then have [tendsto_intros]: "(u \ u x) (at x)" using \eu > 0\ at_within_open[of y "ball x eu" for y] by (auto simp: continuous_on_def) then have flow0_u_tendsto: "(\x. flow0 x (u x)) \x\ poincare_map ?P x" by (auto intro!: tendsto_eq_intros u return_time_exivl rt simp: poincare_map_def) have s_imp: "s (poincare_map {x \ S. s x = 0} x) = 0 \ poincare_map {x \ S. s x = 0} x \ S" using poincare_map_returns[OF rt] by auto from eventually_tendsto_compose_within[OF eventually_inside s_imp flow0_u_tendsto] have "\\<^sub>F x in at x. s (flow0 x (u x)) = 0 \ flow0 x (u x) \ S" by auto with ev_cball have "\\<^sub>F x in at x. flow0 x (u x) \ S" by eventually_elim (auto simp: u) ultimately have u_returns_ge: "\\<^sub>F y in at x. returns_to ?P y \ return_time ?P y \ u y" proof eventually_elim case (elim y) then show ?case using u elim by (auto intro!: return_time_le[OF _ cls]) qed moreover have "\\<^sub>F y in at x. u y - return_time ?P x < e" using tendstoD[OF u_tendsto \0 < e\, unfolded u] u_returns_ge by eventually_elim (auto simp: dist_real_def) moreover note 1 = outside define ml where "ml = max (return_time ?P x / 2) (return_time ?P x - e)" have [intro, simp, arith]: "0 < ml" "ml < return_time ?P x" "ml \ return_time ?P x" using return_time_pos[OF rt cls] \0 < e\ by (auto simp: ml_def) have mt_in: "ml \ existence_ivl0 x" using \0 < e\ by (auto intro!: mem_existence_ivl_iv_defined in_existence_between_zeroI[OF ret_exivl] simp: closed_segment_eq_real_ivl ml_def) from open_existence_ivlE[OF mt_in] obtain e0 where e0: "e0 > 0" "cball x e0 \ {0..ml + e0} \ Sigma X existence_ivl0" (is "?D \ _") by auto have uc: "uniformly_continuous_on ((\(x, t). flow0 x t) ` ?D) s" apply (auto intro!: compact_uniformly_continuous continuous_on_subset[OF cont_s]) apply (rule compact_continuous_image) apply (rule continuous_on_subset) apply (rule flow_continuous_on_state_space) apply (rule e0) apply (rule compact_Times) apply (rule compact_cball) apply (rule compact_Icc) done let ?T = "{0..ml}" have ul: "uniform_limit ?T flow0 (flow0 x) (at x)" using \0 < e\ by (intro uniform_limit_flow) (auto intro!: mem_existence_ivl_iv_defined in_existence_between_zeroI[OF ret_exivl] simp: closed_segment_eq_real_ivl ) have "\\<^sub>F y in at x. \t\{0..ml}. flow0 y t \ - {x \ S. s x = 0}" apply (rule uniform_limit_in_open) apply (rule ul) apply (auto intro!: continuous_intros continuous_on_compose2[OF cont_s] simp: split: if_splits) apply (meson atLeastAtMost_iff contra_subsetD local.ivl_subset_existence_ivl mt_in) subgoal for t apply (cases "t = 0") subgoal using 1 by (simp) subgoal using return_time_least[OF rt cls, of t] \ml < return_time {x \ S. s x = 0} x\ by auto done done then have "\\<^sub>F y in at x. return_time ?P y \ return_time ?P x - e" using u_returns_ge proof eventually_elim case (elim y) have "return_time ?P x - e \ ml" by (auto simp: ml_def) also have ry: "returns_to ?P y" "return_time ?P y \ u y" using elim by auto have "ml < return_time ?P y" apply (rule return_time_gt[OF ry(1) cls]) using elim by (auto simp: Ball_def) finally show ?case by simp qed ultimately have "\\<^sub>F y in at x. dist (return_time ?P y) (return_time ?P x) \ e" by eventually_elim (auto simp: dist_real_def abs_real_def algebra_simps) then show "\\<^sub>F y in at x. dist (return_time ?P y) (return_time ?P x) < e_orig" by eventually_elim (use \e_orig > 0\ in \auto simp: e_def\) qed lemma isCont_poincare_map: assumes "isCont (return_time P) x" "returns_to P x" "closed P" shows "isCont (poincare_map P) x" unfolding poincare_map_def by (auto intro!: continuous_intros assms return_time_exivl) lemma poincare_map_tendsto: assumes "(return_time P \ return_time P x) (at x within S)" "returns_to P x" "closed P" shows "(poincare_map P \ poincare_map P x) (at x within S)" unfolding poincare_map_def by (rule tendsto_eq_intros refl assms return_time_exivl)+ lemma return_time_continuous_below: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P x") assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes cS: "closed S" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). s x = 0 \ x \ S" assumes DsC: "\x. isCont Ds x" assumes through: "(Ds (poincare_map ?P x)) (f (poincare_map ?P x)) \ 0" assumes inside: "x \ S" "s x = 0" "Ds x (f x) < 0" shows "continuous (at x within {x. s x \ 0}) (return_time ?P)" unfolding continuous_within proof (rule tendstoI) fix e_orig::real assume "e_orig > 0" define e where "e = e_orig / 2" have "e > 0" using \e_orig > 0\ by (simp add: e_def) note DsC_tendso[tendsto_intros] = isCont_tendsto_compose[OF DsC] have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) then have s_tendsto: "(s \ s x) (at x)" for x by (auto simp: continuous_on_def) note [continuous_intros] = continuous_on_compose2[OF cont_s _ subset_UNIV] note [derivative_intros] = has_derivative_compose[OF _ Ds] have cont_s': "continuous_on S s" by (rule continuous_on_subset[OF cont_s subset_UNIV]) note cls[simp, intro] = closed_levelset_within[OF cont_s' cS(1)] have "{x. s x = 0} = s -` {0}" by auto have ret_exivl: "return_time ?P x \ existence_ivl0 x" by (rule return_time_exivl; fact) then have [intro, simp]: "x \ X" by auto have isCont_Ds_f: "isCont (\s. Ds s (f s)) (poincare_map ?P x)" apply (auto intro!: continuous_intros DsC) apply (rule has_derivative_continuous) apply (rule derivative_rhs) by (auto simp: poincare_map_def intro!: flow_in_domain return_time_exivl assms) have "\\<^sub>F yt in at (x, 0) within UNIV \ {0<..}. (Ds (flow0 (fst yt) (snd yt))) (f (flow0 (fst yt) (snd yt))) < 0" by (rule order_tendstoD) (auto intro!: tendsto_eq_intros inside) moreover have "(x, 0) \ Sigma X existence_ivl0" by auto from topological_tendstoD[OF tendsto_ident_at open_state_space this, of "UNIV \ {0<..}"] have "\\<^sub>F yt in at (x, 0) within UNIV \ {0<..}. snd yt \ existence_ivl0 (fst yt)" by eventually_elim auto moreover from topological_tendstoD[OF tendsto_ident_at open_Times[OF open_dom open_UNIV], of "(x, 0)" "UNIV \ {0<..}"] have "\\<^sub>F yt in at (x, 0) within UNIV \ {0<..}. fst yt \ X" by (auto simp: mem_Times_iff) ultimately have "\\<^sub>F yt in at (x, 0) within UNIV \ {0<..}. (Ds (flow0 (fst yt) (snd yt))) (f (flow0 (fst yt) (snd yt))) < 0 \ snd yt \ existence_ivl0 (fst yt) \ 0 \ existence_ivl0 (fst yt)" by eventually_elim auto then obtain d2 where "0 < d2" and d2_neg: "\y t. (y, t) \ cball (x, 0) d2 \ 0 < t \ (Ds (flow0 y t)) (f (flow0 y t)) < 0" and d2_ex: "\y t. (y, t) \ cball (x, 0) d2 \ 0 < t \ t \ existence_ivl0 y" and d2_ex0: "\y t. (y, t::real) \ cball (x, 0) d2 \ 0 < t \ y \ X" by (auto simp: eventually_at_le dist_commute) define d where "d \ d2 / 2" from \0 < d2\ have "d > 0" by (simp add: d_def) have d_neg: "dist y x< d \ 0 < t \ t \ d \ (Ds (flow0 y t)) (f (flow0 y t)) < 0" for y t using d2_neg[of y t, OF subsetD[OF cball_times_subset[of x d2 0]]] by (auto simp: d_def dist_commute) have d_ex: "t \ existence_ivl0 y" if "dist y x< d" "0 \ t" "t \ d" for y t proof cases assume "t = 0" have "sqrt ((dist x y)\<^sup>2 + (d2 / 2)\<^sup>2) \ dist x y + d2/2" using \0 < d2\ by (intro sqrt_sum_squares_le_sum) auto also have "dist x y \ d2 / 2" using that by (simp add: d_def dist_commute) finally have "sqrt ((dist x y)\<^sup>2 + (d2 / 2)\<^sup>2) \ d2" by simp with \t = 0\ show ?thesis using d2_ex[of y t, OF subsetD[OF cball_times_subset[of x d2 0]]] d2_ex0[of y d] \0 < d2\ by (auto simp: d_def dist_commute dist_prod_def) next assume "t \ 0" then show ?thesis using d2_ex[of y t, OF subsetD[OF cball_times_subset[of x d2 0]]] that by (auto simp: d_def dist_commute) qed have d_mvt: "s (flow0 y t) < s y" if "0 < t" "t \ d" "dist y x < d" for y t proof - have c: "continuous_on {0 .. t} (\t. s (flow0 y t))" using that by (auto intro!: continuous_intros d_ex) have d: "\x. \0 < x; x < t\ \ ((\t. s (flow0 y t)) has_derivative (\t. t * blinfun_apply (Ds (flow0 y x)) (f (flow0 y x)))) (at x)" using that by (auto intro!: derivative_eq_intros d_ex simp: flowderiv_def blinfun.bilinear_simps) from mvt[OF \0 < t\ c d] obtain xi where xi: "0 < xi" "xi < t" and "s (flow0 y t) - s (flow0 y 0) = t * blinfun_apply (Ds (flow0 y xi)) (f (flow0 y xi))" by auto note this(3) also have "\ < 0" using \0 < t\ apply (rule mult_pos_neg) apply (rule d_neg) using that xi by auto also have "flow0 y 0 = y" apply (rule flow_initial_time) apply auto using \0 < d\ d_ex that(3) by fastforce finally show ?thesis by (auto simp: ) qed obtain u eu where u: "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\y. y \ cball x eu \ s (flow0 y (u y)) = 0)" "continuous_on (cball x eu) u" "(\t. (t, u t)) ` cball x eu \ Sigma X existence_ivl0" "0 < eu" by (rule returns_to_implicit_function[OF rt cS(1) Ds DsC through]; blast) have u_tendsto: "(u \ u x) (at x)" unfolding isCont_def[symmetric] apply (rule continuous_on_interior[OF u(4)]) using \0 < eu\ by auto have "u x > 0" by (auto simp: u intro!: return_time_pos rt) from order_tendstoD(1)[OF u_tendsto this] have "\\<^sub>F x in at x. 0 < u x" . moreover have "\\<^sub>F y in at x. y \ cball x eu" using eventually_at_ball[OF \0 < eu\, of x] by eventually_elim auto moreover have "x \ S \ s x \ 0 \ blinfun_apply (Ds x) (f x) \ 0" using inside by auto have returns: "\\<^sub>F y in at x. returns_to ?P y" by (rule eventually_returns_to; fact) moreover have "\\<^sub>F y in at x. y \ ball x eu" using eventually_at_ball[OF \0 < eu\] by eventually_elim simp then have ev_cball: "\\<^sub>F y in at x. y \ cball x eu" by eventually_elim (use \e > 0\ in auto) have "continuous_on (ball x eu) u" using u by (auto simp: continuous_on_subset) then have [tendsto_intros]: "(u \ u x) (at x)" using \eu > 0\ at_within_open[of y "ball x eu" for y] by (auto simp: continuous_on_def) then have flow0_u_tendsto: "(\x. flow0 x (u x)) \x\ poincare_map ?P x" by (auto intro!: tendsto_eq_intros u return_time_exivl rt simp: poincare_map_def) have s_imp: "s (poincare_map {x \ S. s x = 0} x) = 0 \ poincare_map {x \ S. s x = 0} x \ S" using poincare_map_returns[OF rt] by auto from eventually_tendsto_compose_within[OF eventually_inside s_imp flow0_u_tendsto] have "\\<^sub>F x in at x. s (flow0 x (u x)) = 0 \ flow0 x (u x) \ S" by auto with ev_cball have "\\<^sub>F x in at x. flow0 x (u x) \ S" by eventually_elim (auto simp: u) ultimately have u_returns_ge: "\\<^sub>F y in at x. returns_to ?P y \ return_time ?P y \ u y" proof eventually_elim case (elim y) then show ?case using u elim by (auto intro!: return_time_le[OF _ cls]) qed moreover have "\\<^sub>F y in at x. u y - return_time ?P x < e" using tendstoD[OF u_tendsto \0 < e\, unfolded u] u_returns_ge by eventually_elim (auto simp: dist_real_def) moreover have d_less: "d < return_time ?P x" apply (rule return_time_gt) apply fact apply fact subgoal for t using d_mvt[of t x] \s x = 0\ \0 < d\ by auto done note 1 = inside define ml where "ml = Max {return_time ?P x / 2, return_time ?P x - e, d}" have [intro, simp, arith]: "0 < ml" "ml < return_time ?P x" "ml \ return_time ?P x" "d \ ml" using return_time_pos[OF rt cls] \0 < e\ d_less by (auto simp: ml_def) have mt_in: "ml \ existence_ivl0 x" using \0 < e\ \0 < d\ d_less by (auto intro!: mem_existence_ivl_iv_defined in_existence_between_zeroI[OF ret_exivl] simp: closed_segment_eq_real_ivl ml_def) from open_existence_ivlE[OF mt_in] obtain e0 where e0: "e0 > 0" "cball x e0 \ {0..ml + e0} \ Sigma X existence_ivl0" (is "?D \ _") by auto have uc: "uniformly_continuous_on ((\(x, t). flow0 x t) ` ?D) s" apply (auto intro!: compact_uniformly_continuous continuous_on_subset[OF cont_s]) apply (rule compact_continuous_image) apply (rule continuous_on_subset) apply (rule flow_continuous_on_state_space) apply (rule e0) apply (rule compact_Times) apply (rule compact_cball) apply (rule compact_Icc) done let ?T = "{d..ml}" have ul: "uniform_limit ?T flow0 (flow0 x) (at x)" using \0 < e\ \0 < d\ d_less by (intro uniform_limit_flow) (auto intro!: mem_existence_ivl_iv_defined in_existence_between_zeroI[OF ret_exivl] simp: closed_segment_eq_real_ivl ) { have "\\<^sub>F y in at x within {x. s x \ 0}. y \ X" by (rule topological_tendstoD[OF tendsto_ident_at open_dom \x \ X\]) moreover have "\\<^sub>F y in at x within {x. s x \ 0}. s y \ 0" by (auto simp: eventually_at) moreover have "\\<^sub>F y in at x within {x. s x \ 0}. Ds y (f y) < 0" by (rule order_tendstoD) (auto intro!: tendsto_eq_intros inside) moreover from tendstoD[OF tendsto_ident_at \0 < d\] have "\\<^sub>F y in at x within {x. s x \ 0}. dist y x < d" by (auto simp: ) moreover have "d \ existence_ivl0 x" using d_ex[of x d] \0 < d\ by auto have dret: "returns_to {x\S. s x = 0} (flow0 x d)" apply (rule returns_to_laterI) apply fact+ subgoal for u using d_mvt[of u x] \s x = 0\ by auto done have "\\<^sub>F y in at x. \t\{d..ml}. flow0 y t \ - {x \ S. s x = 0}" apply (rule uniform_limit_in_open) apply (rule ul) apply (auto intro!: continuous_intros continuous_on_compose2[OF cont_s] simp: split: if_splits) using \d \ existence_ivl0 x\ mem_is_interval_1_I mt_in apply blast subgoal for t using return_time_least[OF rt cls, of t] \ml < return_time {x \ S. s x = 0} x\ \0 < d\ by auto done then have "\\<^sub>F y in at x within {x. s x \ 0}. \t\{d .. ml}. flow0 y t \ - {x \ S. s x = 0}" by (auto simp add: eventually_at; force) ultimately have "\\<^sub>F y in at x within {x. s x \ 0}. \t\{0<..ml}. flow0 y t \ - {x \ S. s x = 0}" apply eventually_elim apply auto using d_mvt by fastforce moreover have "\\<^sub>F y in at x. returns_to ?P y" by fact then have "\\<^sub>F y in at x within {x. s x \ 0}. returns_to ?P y" by (auto simp: eventually_at) ultimately have "\\<^sub>F y in at x within {x. s x \ 0}. return_time ?P y > ml" apply eventually_elim apply (rule return_time_gt) by auto } then have "\\<^sub>F y in at x within {x. s x \ 0}. return_time ?P y \ return_time ?P x - e" by eventually_elim (auto simp: ml_def) ultimately have "\\<^sub>F y in at x within {x . s x \ 0}. dist (return_time ?P y) (return_time ?P x) \ e" unfolding eventually_at_filter by eventually_elim (auto simp: dist_real_def abs_real_def algebra_simps) then show "\\<^sub>F y in at x within {x. s x \ 0}. dist (return_time ?P y) (return_time ?P x) < e_orig" by eventually_elim (use \e_orig > 0\ in \auto simp: e_def\) qed lemma return_time_continuous_below_plane: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ R. x \ n = c} x" (is "returns_to ?P x") assumes cR: "closed R" assumes through: "f (poincare_map ?P x) \ n \ 0" assumes R: "x \ R" assumes inside: "x \ n = c" "f x \ n < 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). x \ n = c \ x \ R" shows "continuous (at x within {x. x \ n \ c}) (return_time ?P)" apply (rule return_time_continuous_below[of R "\x. x \ n - c", simplified]) using through rt inside cR R eventually_inside by (auto intro!: derivative_eq_intros blinfun_inner_left.rep_eq[symmetric]) lemma poincare_map_in_interior_eventually_return_time_equal: assumes RP: "R \ P" assumes cP: "closed P" assumes cR: "closed R" assumes ret: "returns_to P x" assumes evret: "\\<^sub>F x in at x within S. returns_to P x" assumes evR: "\\<^sub>F x in at x within S. poincare_map P x \ R" shows "\\<^sub>F x in at x within S. returns_to R x \ return_time P x = return_time R x" proof - from evret evR show ?thesis proof eventually_elim case (elim x) from return_time_least[OF elim(1) cP] RP have rtl: "\s. 0 < s \ s < return_time P x \ flow0 x s \ R" by auto from elim(2) have pR: "poincare_map P x \ R" by auto have "\\<^sub>F t in at_right 0. 0 < t" by (simp add: eventually_at_filter) moreover have "\\<^sub>F t in at_right 0. t < return_time P x" using return_time_pos[OF elim(1) cP] by (rule order_tendstoD[OF tendsto_ident_at]) ultimately have evR: "\\<^sub>F t in at_right 0. flow0 x t \ R" proof eventually_elim case et: (elim t) from return_time_least[OF elim(1) cP et] show ?case using RP by auto qed have rtp: "0 < return_time P x" by (intro return_time_pos cP elim) have rtex: "return_time P x \ existence_ivl0 x" by (intro return_time_exivl elim cP) have frR: "flow0 x (return_time P x) \ R" unfolding poincare_map_def[symmetric] by (rule pR) have "returns_to R x" by (rule returns_toI[where t="return_time P x"]; fact) moreover have "return_time R x = return_time P x" by (rule return_time_eqI) fact+ ultimately show ?case by auto qed qed lemma poincare_map_in_planeI: assumes "returns_to (plane n c) x0" shows "poincare_map (plane n c) x0 \ n = c" using poincare_map_returns[OF assms] by fastforce lemma less_return_time_imp_exivl: "h \ existence_ivl0 x'" if "h \ return_time P x'" "returns_to P x'" "closed P" "0 \ h" proof - from return_time_exivl[OF that(2,3)] have "return_time P x' \ existence_ivl0 x'" by auto from ivl_subset_existence_ivl[OF this] that show ?thesis by auto qed lemma eventually_returns_to_continuousI: assumes "returns_to P x" assumes "closed P" assumes "continuous (at x within S) (return_time P)" shows "\\<^sub>F x in at x within S. returns_to P x" proof - have "return_time P x > 0" using assms by (auto simp: return_time_pos) from order_tendstoD(1)[OF assms(3)[unfolded continuous_within] this] have "\\<^sub>F x in at x within S. 0 < return_time P x" . then show ?thesis by eventually_elim (auto simp: return_time_pos_returns_to) qed lemma return_time_implicit_functionE: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P _") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "\x. isCont Ds x" assumes Ds_through: "(Ds (poincare_map ?P x)) (f (poincare_map ?P x)) \ 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). s x = 0 \ x \ S" assumes outside: "x \ S \ s x \ 0" obtains e' where "0 < e'" "\y. y \ ball x e' \ returns_to ?P y" "\y. y \ ball x e' \ s (flow0 y (return_time ?P y)) = 0" "continuous_on (ball x e') (return_time ?P)" "(\y. y \ ball x e' \ Ds (poincare_map ?P y) o\<^sub>L flowderiv y (return_time ?P y) o\<^sub>L embed2_blinfun \ invertibles_blinfun)" "(\U v sa. (\sa. sa \ U \ s (flow0 sa (v sa)) = 0) \ return_time ?P x = v x \ continuous_on U v \ sa \ U \ x \ U \ U \ ball x e' \ connected U \ open U \ return_time ?P sa = v sa)" "(return_time ?P has_derivative - blinfun_scaleR_left (inverse ((Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L Dflow x (return_time ?P x))) (at x)" proof - have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) then have s_tendsto: "(s \ s x) (at x)" for x by (auto simp: continuous_on_def) have cls[simp, intro]: "closed {x \ S. s x = 0}" by (rule closed_levelset_within) (auto intro!: cS continuous_on_subset[OF cont_s]) have cont_Ds: "continuous_on UNIV Ds" using DsC by (auto simp: continuous_on_def isCont_def) note [tendsto_intros] = continuous_on_tendsto_compose[OF cont_Ds _ UNIV_I, simplified] note [continuous_intros] = continuous_on_compose2[OF cont_Ds _ subset_UNIV] have "\\<^sub>F x in at (poincare_map ?P x). s x = 0 \ x \ S" using eventually_inside by auto then obtain U where "open U" "poincare_map ?P x \ U" "\x. x \ U \ s x = 0 \ x \ S" using poincare_map_returns[OF rt cls] by (force simp: eventually_at_topological) have s_imp: "s (poincare_map ?P x) = 0 \ poincare_map ?P x \ S" using poincare_map_returns[OF rt cls] by auto have outside_disj: "x \ S \ s x \ 0 \ blinfun_apply (Ds x) (f x) \ 0" using outside by auto have pm_tendsto: "(poincare_map ?P \ poincare_map ?P x) (at x)" apply (rule poincare_map_tendsto) unfolding isCont_def[symmetric] apply (rule return_time_isCont_outside) using assms by (auto intro!: cls ) have evmemS: "\\<^sub>F x in at x. poincare_map ?P x \ S" using eventually_returns_to[OF rt cS Ds DsC eventually_inside Ds_through outside_disj] apply eventually_elim using poincare_map_returns by auto have "\\<^sub>F x in at x. \\<^sub>F x in at (poincare_map ?P x). s x = 0 \ x \ S" apply (rule eventually_tendsto_compose_within[OF _ _ pm_tendsto]) apply (rule eventually_eventually_withinI) apply (rule eventually_inside) apply (rule s_imp) apply (rule eventually_inside) apply (rule evmemS) done moreover have "eventually (\x. x \ - ?P) (at x)" apply (rule topological_tendstoD) using outside by (auto intro!: ) then have "eventually (\x. x \ S \ s x \ 0) (at x)" by auto moreover have "eventually (\x. (Ds (poincare_map ?P x)) (f (poincare_map ?P x)) \ 0) (at x)" apply (rule tendsto_imp_eventually_ne) apply (rule tendsto_intros) apply (rule tendsto_intros) unfolding poincare_map_def apply (rule tendsto_intros) apply (rule tendsto_intros) apply (subst isCont_def[symmetric]) apply (rule return_time_isCont_outside[OF rt cS Ds DsC Ds_through eventually_inside outside]) apply (rule return_time_exivl[OF rt cls]) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (subst isCont_def[symmetric]) apply (rule return_time_isCont_outside[OF rt cS Ds DsC Ds_through eventually_inside outside]) apply (rule return_time_exivl[OF rt cls]) apply (rule flow_in_domain) apply (rule return_time_exivl[OF rt cls]) unfolding poincare_map_def[symmetric] apply (rule Ds_through) done ultimately have "eventually (\y. returns_to ?P y \ (\\<^sub>F x in at (poincare_map ?P y). s x = 0 \ x \ S) \ (y \ S \ s y \ 0) \ (Ds (poincare_map ?P y)) (f (poincare_map ?P y)) \ 0) (at x)" using eventually_returns_to[OF rt cS Ds DsC eventually_inside Ds_through outside_disj] by eventually_elim auto then obtain Y' where Y': "open Y'" "x \ Y'" "\y. y \ Y' \ returns_to ?P y" "\y. y \ Y' \ (\\<^sub>F x in at (poincare_map ?P y). s x = 0 \ x \ S)" "\y. y \ Y' \ y \ S \ s y \ 0" "\y. y \ Y' \ blinfun_apply (Ds (poincare_map ?P y)) (f (poincare_map ?P y)) \ 0" apply (subst (asm) (3) eventually_at_topological) using rt outside Ds_through eventually_inside by fastforce from openE[OF \open Y'\ \x \ Y'\] obtain eY where eY: "0 < eY" "ball x eY \ Y'" by auto define Y where "Y = ball x eY" then have Y: "open Y" and x: "x \ Y" and Yr: "\y. y \ Y \ returns_to ?P y" and Y_mem: "\y. y \ Y \ (\\<^sub>F x in at (poincare_map ?P y). s x = 0 \ x \ S)" and Y_nz: "\y. y \ Y \ y \ S \ s y \ 0" and Y_fnz: "\y. y \ Y \ Ds (poincare_map ?P y) (f (poincare_map ?P y)) \ 0" and Y_convex: "convex Y" using Y' eY by (auto simp: subset_iff dist_commute) have "isCont (return_time ?P) y" if "y \ Y" for y using return_time_isCont_outside[OF Yr[OF that] cS Ds DsC Y_fnz Y_mem Y_nz, OF that that that] . then have cY: "continuous_on Y (return_time ?P)" by (auto simp: continuous_on_def isCont_def Lim_at_imp_Lim_at_within) note [derivative_intros] = has_derivative_compose[OF _ Ds] let ?t1 = "return_time ?P x" have t1_exivl: "?t1 \ existence_ivl0 x" by (auto intro!: return_time_exivl rt) then have [simp]: "x \ X" by auto have xt1: "(x, ?t1) \ Sigma Y existence_ivl0" by (auto intro!: return_time_exivl rt x) have "Sigma Y existence_ivl0 = Sigma X existence_ivl0 \ fst -` Y" by auto also have "open \" by (rule open_Int[OF open_state_space open_vimage_fst[OF \open Y\]]) finally have "open (Sigma Y existence_ivl0)" . have D: "(\x. x \ Sigma Y existence_ivl0 \ ((\(x, t). s (flow0 x t)) has_derivative blinfun_apply (Ds (flow0 (fst x) (snd x)) o\<^sub>L (flowderiv (fst x) (snd x)))) (at x))" by (auto intro!: derivative_eq_intros) have C: "continuous_on (Sigma Y existence_ivl0) (\x. Ds (flow0 (fst x) (snd x)) o\<^sub>L flowderiv (fst x) (snd x))" by (auto intro!: continuous_intros) from return_time_returns[OF rt cls] have Z: "(case (x, ?t1) of (x, t) \ s (flow0 x t)) = 0" by (auto simp: x) have I1: "blinfun_scaleR_left (inverse (Ds (flow0 x (?t1))(f (flow0 x (?t1))))) o\<^sub>L ((Ds (flow0 (fst (x, return_time ?P x)) (snd (x, return_time ?P x))) o\<^sub>L flowderiv (fst (x, return_time ?P x)) (snd (x, return_time ?P x))) o\<^sub>L embed2_blinfun) = 1\<^sub>L" using Ds_through by (auto intro!: blinfun_eqI simp: rt flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) have I2: "((Ds (flow0 (fst (x, return_time ?P x)) (snd (x, return_time ?P x))) o\<^sub>L flowderiv (fst (x, return_time ?P x)) (snd (x, return_time ?P x))) o\<^sub>L embed2_blinfun) o\<^sub>L blinfun_scaleR_left (inverse (Ds (flow0 x (?t1))(f (flow0 x (?t1))))) = 1\<^sub>L" using Ds_through by (auto intro!: blinfun_eqI simp: rt flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) obtain u e where u: "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\sa. sa \ cball x e \ s (flow0 sa (u sa)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma Y existence_ivl0" "0 < e" "(u has_derivative blinfun_apply (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L flowderiv x (return_time ?P x)) o\<^sub>L embed1_blinfun)) (at x)" "(\s. s \ cball x e \ Ds (flow0 s (u s)) o\<^sub>L flowderiv s (u s) o\<^sub>L embed2_blinfun \ invertibles_blinfun)" and unique: "(\U v sa. (\sa. sa \ U \ s (flow0 sa (v sa)) = 0) \ u x = v x \ continuous_on U v \ sa \ U \ x \ U \ U \ cball x e \ connected U \ open U \ u sa = v sa)" apply (rule implicit_function_theorem_unique[where f="\(x, t). s (flow0 x t)" and S="Sigma Y existence_ivl0", OF D xt1 \open (Sigma Y _)\ order_refl C Z I1 I2]) apply blast unfolding split_beta' fst_conv snd_conv poincare_map_def[symmetric] apply (rule) by (assumption+, blast) have u_rt: "u y = return_time ?P y" if "y \ ball x e \ Y" for y apply (rule unique[of "ball x e \ Y" "return_time ?P"]) subgoal for y unfolding poincare_map_def[symmetric] using poincare_map_returns[OF Yr cls] by auto subgoal by (auto simp: u) subgoal using cY by (rule continuous_on_subset) auto subgoal using that by auto subgoal using x \0 < e\ by auto subgoal by auto subgoal apply (rule convex_connected) apply (rule convex_Int) apply simp apply fact done subgoal by (auto intro!: open_Int \open Y\) done have *: "(- blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L flowderiv x (return_time ?P x)) o\<^sub>L embed1_blinfun) = - blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L Dflow x (return_time ?P x))" by (auto intro!: blinfun_eqI simp: flowderiv_def) define e' where "e' = min e eY" have e'_eq: "ball x e' = ball x e \ Y" by (auto simp: e'_def Y_def) have "0 < e'" "\y. y \ ball x e' \ returns_to ?P y" "\y. y \ ball x e' \ s (flow0 y (return_time ?P y)) = 0" "continuous_on (ball x e') (return_time ?P)" "(\y. y \ ball x e' \ Ds (poincare_map ?P y) o\<^sub>L flowderiv y (return_time ?P y) o\<^sub>L embed2_blinfun \ invertibles_blinfun)" "(\U v sa. (\sa. sa \ U \ s (flow0 sa (v sa)) = 0) \ return_time ?P x = v x \ continuous_on U v \ sa \ U \ x \ U \ U \ ball x e' \ connected U \ open U \ return_time ?P sa = v sa)" "(return_time ?P has_derivative blinfun_apply (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L flowderiv x (return_time ?P x)) o\<^sub>L embed1_blinfun)) (at x)" unfolding e'_eq subgoal by (auto simp: e'_def \0 < e\ \0 < eY\) subgoal by (rule Yr) auto subgoal for y unfolding poincare_map_def[symmetric] using poincare_map_returns[OF Yr cls] by auto subgoal using cY by (rule continuous_on_subset) auto subgoal premises prems for y unfolding poincare_map_def unfolding u_rt[OF prems, symmetric] apply (rule u) using prems by auto subgoal premises prems for U v t apply (subst u_rt[symmetric]) subgoal using prems by force apply (rule unique[of U v]) subgoal by fact subgoal by (auto simp: u prems) subgoal by fact subgoal by fact subgoal by fact subgoal using prems by auto subgoal by fact subgoal by fact done subgoal proof - have "\\<^sub>F x' in at x. x' \ ball x e'" using eventually_at_ball[OF \0 < e'\] by eventually_elim simp then have "\\<^sub>F x' in at x. u x' = return_time ?P x'" unfolding e'_eq by eventually_elim (rule u_rt, auto) from u(7) this show ?thesis by (rule has_derivative_transform_eventually) (auto simp: u) qed done then show ?thesis unfolding * .. qed lemma return_time_has_derivative: fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P _") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "\x. isCont Ds x" assumes Ds_through: "(Ds (poincare_map ?P x)) (f (poincare_map ?P x)) \ 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map {x \ S. s x = 0} x). s x = 0 \ x \ S" assumes outside: "x \ S \ s x \ 0" shows "(return_time ?P has_derivative - blinfun_scaleR_left (inverse ((Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L Dflow x (return_time ?P x))) (at x)" using return_time_implicit_functionE[OF assms] by blast lemma return_time_plane_has_derivative_blinfun: assumes rt: "returns_to {x \ S. x \ i = c} x" (is "returns_to ?P _") assumes cS: "closed S" assumes fnz: "f (poincare_map ?P x) \ i \ 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). x \ i = c \ x \ S" assumes outside: "x \ S \ x \ i \ c" shows "(return_time ?P has_derivative (- blinfun_scaleR_left (inverse ((blinfun_inner_left i) (f (poincare_map ?P x)))) o\<^sub>L (blinfun_inner_left i o\<^sub>L Dflow x (return_time ?P x)))) (at x)" proof - have rt: "returns_to {x \ S. x \ i - c = 0} x" using rt by auto have D: "((\x. x \ i - c) has_derivative blinfun_inner_left i) (at x)" for x by (auto intro!: derivative_eq_intros) have DC: "(\x. isCont (\x. blinfun_inner_left i) x)" by (auto intro!: continuous_intros) have nz: "blinfun_apply (blinfun_inner_left i) (f (poincare_map {x \ S. x \ i - c = 0} x)) \ 0" using fnz by (auto ) from cS have cS: "closed S"by auto have out: "x \ S \ x \ i - c \ 0" using outside by simp from eventually_inside have eventually_inside: "\\<^sub>F x in at (poincare_map {x \ S. x \ i - c = 0} x). x \ i - c = 0 \ x \ S" by auto from return_time_has_derivative[OF rt cS D DC nz eventually_inside out] show ?thesis by auto qed lemma return_time_plane_has_derivative: assumes rt: "returns_to {x \ S. x \ i = c} x" (is "returns_to ?P _") assumes cS: "closed S" assumes fnz: "f (poincare_map ?P x) \ i \ 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). x \ i = c \ x \ S" assumes outside: "x \ S \ x \ i \ c" shows "(return_time ?P has_derivative (\h. - (Dflow x (return_time ?P x)) h \ i / (f (poincare_map ?P x) \ i))) (at x)" by (rule return_time_plane_has_derivative_blinfun[OF assms, THEN has_derivative_eq_rhs]) (auto simp: blinfun.bilinear_simps flowderiv_def inverse_eq_divide intro!: ext) definition "Dpoincare_map i c S x = (\h. (Dflow x (return_time {x \ S. x \ i = c} x)) h - ((Dflow x (return_time {x \ S. x \ i = c} x)) h \ i / (f (poincare_map {x \ S. x \ i = c} x) \ i)) *\<^sub>R f (poincare_map {x \ S. x \ i = c} x))" definition "Dpoincare_map' i c S x = Dflow x (return_time {x \ S. x \ i - c = 0} x) - (blinfun_scaleR_left (f (poincare_map {x \ S. x \ i = c} x)) o\<^sub>L (blinfun_scaleR_left (inverse ((f (poincare_map {x \ S. x \ i = c} x) \ i))) o\<^sub>L (blinfun_inner_left i o\<^sub>L Dflow x (return_time {x \ S. x \ i - c = 0} x))))" theorem poincare_map_plane_has_derivative: assumes rt: "returns_to {x \ S. x \ i = c} x" (is "returns_to ?P _") assumes cS: "closed S" assumes fnz: "f (poincare_map ?P x) \ i \ 0" assumes eventually_inside: "\\<^sub>F x in at (poincare_map ?P x). x \ i = c \ x \ S" assumes outside: "x \ S \ x \ i \ c" notes [derivative_intros] = return_time_plane_has_derivative[OF rt cS fnz eventually_inside outside] shows "(poincare_map ?P has_derivative Dpoincare_map' i c S x) (at x)" unfolding poincare_map_def Dpoincare_map'_def using fnz outside by (auto intro!: derivative_eq_intros return_time_exivl assms ext closed_levelset_within continuous_intros simp: flowderiv_eq poincare_map_def blinfun.bilinear_simps inverse_eq_divide algebra_simps) 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 - from sum2 sum1 have "ln (1 + 1/x) \ T x" - proof (rule sums_le [OF allI, rotated]) + 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] 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