diff --git a/src/HOL/Series.thy b/src/HOL/Series.thy --- a/src/HOL/Series.thy +++ b/src/HOL/Series.thy @@ -1,1333 +1,1329 @@ (* Title : Series.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Converted to Isar and polished by lcp Converted to sum and polished yet more by TNN Additional contributions by Jeremy Avigad *) section \Infinite Series\ theory Series imports Limits Inequalities begin subsection \Definition of infinite summability\ definition sums :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a \ bool" (infixr "sums" 80) where "f sums s \ (\n. \i s" definition summable :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ bool" where "summable f \ (\s. f sums s)" definition suminf :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a" (binder "\" 10) where "suminf f = (THE s. f sums s)" text\Variants of the definition\ lemma sums_def': "f sums s \ (\n. \i = 0..n. f i) \ s" unfolding sums_def apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp only: lessThan_Suc_atMost atLeast0AtMost) done lemma sums_def_le: "f sums s \ (\n. \i\n. f i) \ s" by (simp add: sums_def' atMost_atLeast0) lemma bounded_imp_summable: fixes a :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology,linordered_comm_semiring_strict}" assumes 0: "\n. a n \ 0" and bounded: "\n. (\k\n. a k) \ B" shows "summable a" proof - have "bdd_above (range(\n. \k\n. a k))" by (meson bdd_aboveI2 bounded) moreover have "incseq (\n. \k\n. a k)" by (simp add: mono_def "0" sum_mono2) ultimately obtain s where "(\n. \k\n. a k) \ s" using LIMSEQ_incseq_SUP by blast then show ?thesis by (auto simp: sums_def_le summable_def) qed subsection \Infinite summability on topological monoids\ lemma sums_subst[trans]: "f = g \ g sums z \ f sums z" by simp lemma sums_cong: "(\n. f n = g n) \ f sums c \ g sums c" - by (drule ext) simp + by presburger lemma sums_summable: "f sums l \ summable f" by (simp add: sums_def summable_def, blast) lemma summable_iff_convergent: "summable f \ convergent (\n. \i convergent (\n. sum f {..n})" - by (simp_all only: summable_iff_convergent convergent_def - lessThan_Suc_atMost [symmetric] filterlim_sequentially_Suc[of "\n. sum f {..n. \in. 0) sums 0" unfolding sums_def by simp lemma summable_zero[simp, intro]: "summable (\n. 0)" by (rule sums_zero [THEN sums_summable]) lemma sums_group: "f sums s \ 0 < k \ (\n. sum f {n * k ..< n * k + k}) sums s" apply (simp only: sums_def sum.nat_group tendsto_def eventually_sequentially) apply (erule all_forward imp_forward exE| assumption)+ - apply (rule_tac x="N" in exI) by (metis le_square mult.commute mult.left_neutral mult_le_cancel2 mult_le_mono) lemma suminf_cong: "(\n. f n = g n) \ suminf f = suminf g" - by (rule arg_cong[of f g], rule ext) simp + by presburger lemma summable_cong: fixes f g :: "nat \ 'a::real_normed_vector" assumes "eventually (\x. f x = g x) sequentially" shows "summable f = summable g" proof - from assms obtain N where N: "\n\N. f n = g n" by (auto simp: eventually_at_top_linorder) define C where "C = (\kn. sum f {.. {N.. {N.. {N..n. n \ N \ f n = 0" shows "f sums (\n\N. f n)" proof - have eq: "sum f {..n. f n = 0) \ (f sums 0)" by (metis (no_types) finite.emptyI sum.empty sums_finite) lemma summable_finite: "finite N \ (\n. n \ N \ f n = 0) \ summable f" by (rule sums_summable) (rule sums_finite) lemma sums_If_finite_set: "finite A \ (\r. if r \ A then f r else 0) sums (\r\A. f r)" using sums_finite[of A "(\r. if r \ A then f r else 0)"] by simp lemma summable_If_finite_set[simp, intro]: "finite A \ summable (\r. if r \ A then f r else 0)" by (rule sums_summable) (rule sums_If_finite_set) lemma sums_If_finite: "finite {r. P r} \ (\r. if P r then f r else 0) sums (\r | P r. f r)" using sums_If_finite_set[of "{r. P r}"] by simp lemma summable_If_finite[simp, intro]: "finite {r. P r} \ summable (\r. if P r then f r else 0)" by (rule sums_summable) (rule sums_If_finite) lemma sums_single: "(\r. if r = i then f r else 0) sums f i" using sums_If_finite[of "\r. r = i"] by simp lemma summable_single[simp, intro]: "summable (\r. if r = i then f r else 0)" by (rule sums_summable) (rule sums_single) context fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" begin lemma summable_sums[intro]: "summable f \ f sums (suminf f)" by (simp add: summable_def sums_def suminf_def) (metis convergent_LIMSEQ_iff convergent_def lim_def) lemma summable_LIMSEQ: "summable f \ (\n. \i suminf f" by (rule summable_sums [unfolded sums_def]) lemma summable_LIMSEQ': "summable f \ (\n. \i\n. f i) \ suminf f" using sums_def_le by blast lemma sums_unique: "f sums s \ s = suminf f" by (metis limI suminf_eq_lim sums_def) lemma sums_iff: "f sums x \ summable f \ suminf f = x" by (metis summable_sums sums_summable sums_unique) lemma summable_sums_iff: "summable f \ f sums suminf f" by (auto simp: sums_iff summable_sums) lemma sums_unique2: "f sums a \ f sums b \ a = b" for a b :: 'a by (simp add: sums_iff) lemma sums_Uniq: "\\<^sub>\\<^sub>1a. f sums a" for a b :: 'a by (simp add: sums_unique2 Uniq_def) lemma suminf_finite: assumes N: "finite N" and f: "\n. n \ N \ f n = 0" shows "suminf f = (\n\N. f n)" using sums_finite[OF assms, THEN sums_unique] by simp end lemma suminf_zero[simp]: "suminf (\n. 0::'a::{t2_space, comm_monoid_add}) = 0" by (rule sums_zero [THEN sums_unique, symmetric]) subsection \Infinite summability on ordered, topological monoids\ lemma sums_le: "(\n. f n \ g n) \ f sums s \ g sums t \ s \ t" for f g :: "nat \ 'a::{ordered_comm_monoid_add,linorder_topology}" by (rule LIMSEQ_le) (auto intro: sum_mono simp: sums_def) context fixes f :: "nat \ 'a::{ordered_comm_monoid_add,linorder_topology}" begin lemma suminf_le: "(\n. f n \ g n) \ summable f \ summable g \ suminf f \ suminf g" using sums_le by blast lemma sum_le_suminf: shows "summable f \ finite I \ (\n. n \- I \ 0 \ f n) \ sum f I \ suminf f" by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto lemma suminf_nonneg: "summable f \ (\n. 0 \ f n) \ 0 \ suminf f" using sum_le_suminf by force lemma suminf_le_const: "summable f \ (\n. sum f {.. x) \ suminf f \ x" by (metis LIMSEQ_le_const2 summable_LIMSEQ) lemma suminf_eq_zero_iff: assumes "summable f" and pos: "\n. 0 \ f n" shows "suminf f = 0 \ (\n. f n = 0)" proof - assume "suminf f = 0" + assume L: "suminf f = 0" then have f: "(\n. \i 0" using summable_LIMSEQ[of f] assms by simp then have "\i. (\n\{i}. f n) \ 0" - proof (rule LIMSEQ_le_const) - show "\N. \n\N. (\n\{i}. f n) \ sum f {..summable f\ order_refl pos sum.infinite sum_le_suminf) with pos show "\n. f n = 0" - by (auto intro!: antisym) + by (simp add: order.antisym) qed (metis suminf_zero fun_eq_iff) lemma suminf_pos_iff: "summable f \ (\n. 0 \ f n) \ 0 < suminf f \ (\i. 0 < f i)" using sum_le_suminf[of "{}"] suminf_eq_zero_iff by (simp add: less_le) lemma suminf_pos2: assumes "summable f" "\n. 0 \ f n" "0 < f i" shows "0 < suminf f" proof - have "0 < (\n \ suminf f" using assms by (intro sum_le_suminf) auto finally show ?thesis . qed lemma suminf_pos: "summable f \ (\n. 0 < f n) \ 0 < suminf f" by (intro suminf_pos2[where i=0]) (auto intro: less_imp_le) end context fixes f :: "nat \ 'a::{ordered_cancel_comm_monoid_add,linorder_topology}" begin lemma sum_less_suminf2: "summable f \ (\m. m\n \ 0 \ f m) \ n \ i \ 0 < f i \ sum f {.. (\m. m\n \ 0 < f m) \ sum f {.. 'a::{ordered_comm_monoid_add,linorder_topology,conditionally_complete_linorder}" assumes pos[simp]: "\n. 0 \ f n" and le: "\n. (\i x" shows "summable f" unfolding summable_def sums_def [abs_def] proof (rule exI LIMSEQ_incseq_SUP)+ show "bdd_above (range (\n. sum f {..n. sum f {.. 'a::{canonically_ordered_monoid_add,linorder_topology,complete_linorder}" by (intro summableI_nonneg_bounded[where x=top] zero_le top_greatest) lemma suminf_eq_SUP_real: assumes X: "summable X" "\i. 0 \ X i" shows "suminf X = (SUP i. \nInfinite summability on topological monoids\ context fixes f g :: "nat \ 'a::{t2_space,topological_comm_monoid_add}" begin lemma sums_Suc: assumes "(\n. f (Suc n)) sums l" shows "f sums (l + f 0)" proof - have "(\n. (\i l + f 0" using assms by (auto intro!: tendsto_add simp: sums_def) moreover have "(\ii g sums b \ (\n. f n + g n) sums (a + b)" unfolding sums_def by (simp add: sum.distrib tendsto_add) lemma summable_add: "summable f \ summable g \ summable (\n. f n + g n)" unfolding summable_def by (auto intro: sums_add) lemma suminf_add: "summable f \ summable g \ suminf f + suminf g = (\n. f n + g n)" by (intro sums_unique sums_add summable_sums) end context fixes f :: "'i \ nat \ 'a::{t2_space,topological_comm_monoid_add}" and I :: "'i set" begin lemma sums_sum: "(\i. i \ I \ (f i) sums (x i)) \ (\n. \i\I. f i n) sums (\i\I. x i)" by (induct I rule: infinite_finite_induct) (auto intro!: sums_add) lemma suminf_sum: "(\i. i \ I \ summable (f i)) \ (\n. \i\I. f i n) = (\i\I. \n. f i n)" using sums_unique[OF sums_sum, OF summable_sums] by simp lemma summable_sum: "(\i. i \ I \ summable (f i)) \ summable (\n. \i\I. f i n)" using sums_summable[OF sums_sum[OF summable_sums]] . end lemma sums_If_finite_set': fixes f g :: "nat \ 'a::{t2_space,topological_ab_group_add}" assumes "g sums S" and "finite A" and "S' = S + (\n\A. f n - g n)" shows "(\n. if n \ A then f n else g n) sums S'" proof - have "(\n. g n + (if n \ A then f n - g n else 0)) sums (S + (\n\A. f n - g n))" by (intro sums_add assms sums_If_finite_set) also have "(\n. g n + (if n \ A then f n - g n else 0)) = (\n. if n \ A then f n else g n)" by (simp add: fun_eq_iff) finally show ?thesis using assms by simp qed subsection \Infinite summability on real normed vector spaces\ context fixes f :: "nat \ 'a::real_normed_vector" begin lemma sums_Suc_iff: "(\n. f (Suc n)) sums s \ f sums (s + f 0)" proof - have "f sums (s + f 0) \ (\i. \j s + f 0" by (subst filterlim_sequentially_Suc) (simp add: sums_def) also have "\ \ (\i. (\j s + f 0" by (simp add: ac_simps lessThan_Suc_eq_insert_0 image_Suc_lessThan sum.atLeast1_atMost_eq) also have "\ \ (\n. f (Suc n)) sums s" proof assume "(\i. (\j s + f 0" with tendsto_add[OF this tendsto_const, of "- f 0"] show "(\i. f (Suc i)) sums s" by (simp add: sums_def) qed (auto intro: tendsto_add simp: sums_def) finally show ?thesis .. qed lemma summable_Suc_iff: "summable (\n. f (Suc n)) = summable f" proof assume "summable f" then have "f sums suminf f" by (rule summable_sums) then have "(\n. f (Suc n)) sums (suminf f - f 0)" by (simp add: sums_Suc_iff) then show "summable (\n. f (Suc n))" unfolding summable_def by blast qed (auto simp: sums_Suc_iff summable_def) lemma sums_Suc_imp: "f 0 = 0 \ (\n. f (Suc n)) sums s \ (\n. f n) sums s" using sums_Suc_iff by simp end context (* Separate contexts are necessary to allow general use of the results above, here. *) fixes f :: "nat \ 'a::real_normed_vector" begin lemma sums_diff: "f sums a \ g sums b \ (\n. f n - g n) sums (a - b)" unfolding sums_def by (simp add: sum_subtractf tendsto_diff) lemma summable_diff: "summable f \ summable g \ summable (\n. f n - g n)" unfolding summable_def by (auto intro: sums_diff) lemma suminf_diff: "summable f \ summable g \ suminf f - suminf g = (\n. f n - g n)" by (intro sums_unique sums_diff summable_sums) lemma sums_minus: "f sums a \ (\n. - f n) sums (- a)" unfolding sums_def by (simp add: sum_negf tendsto_minus) lemma summable_minus: "summable f \ summable (\n. - f n)" unfolding summable_def by (auto intro: sums_minus) lemma suminf_minus: "summable f \ (\n. - f n) = - (\n. f n)" by (intro sums_unique [symmetric] sums_minus summable_sums) lemma sums_iff_shift: "(\i. f (i + n)) sums s \ f sums (s + (\ii. f (Suc i + n)) sums s \ (\i. f (i + n)) sums (s + f n)" by (subst sums_Suc_iff) simp with Suc show ?case by (simp add: ac_simps) qed corollary sums_iff_shift': "(\i. f (i + n)) sums (s - (\i f sums s" by (simp add: sums_iff_shift) lemma sums_zero_iff_shift: assumes "\i. i < n \ f i = 0" shows "(\i. f (i+n)) sums s \ (\i. f i) sums s" by (simp add: assms sums_iff_shift) lemma summable_iff_shift [simp]: "summable (\n. f (n + k)) \ summable f" by (metis diff_add_cancel summable_def sums_iff_shift [abs_def]) lemma sums_split_initial_segment: "f sums s \ (\i. f (i + n)) sums (s - (\i summable (\n. f(n + k))" by (simp add: summable_iff_shift) lemma suminf_minus_initial_segment: "summable f \ (\n. f (n + k)) = (\n. f n) - (\i suminf f = (\n. f(n + k)) + (\i (\n. f (Suc n)) = suminf f - f 0" using suminf_split_initial_segment[of 1] by simp lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable f" shows "\N. \n\N. norm (\i. f (i + n)) < r" proof - from LIMSEQ_D[OF summable_LIMSEQ[OF \summable f\] \0 < r\] obtain N :: nat where "\ n \ N. norm (sum f {..summable f\]) qed lemma summable_LIMSEQ_zero: assumes "summable f" shows "f \ 0" proof - have "Cauchy (\n. sum f {.. convergent f" by (force dest!: summable_LIMSEQ_zero simp: convergent_def) lemma summable_imp_Bseq: "summable f \ Bseq f" by (simp add: convergent_imp_Bseq summable_imp_convergent) end lemma summable_minus_iff: "summable (\n. - f n) \ summable f" for f :: "nat \ 'a::real_normed_vector" by (auto dest: summable_minus) (* used two ways, hence must be outside the context above *) lemma (in bounded_linear) sums: "(\n. X n) sums a \ (\n. f (X n)) sums (f a)" unfolding sums_def by (drule tendsto) (simp only: sum) lemma (in bounded_linear) summable: "summable (\n. X n) \ summable (\n. f (X n))" unfolding summable_def by (auto intro: sums) lemma (in bounded_linear) suminf: "summable (\n. X n) \ f (\n. X n) = (\n. f (X n))" by (intro sums_unique sums summable_sums) lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real] lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real] lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real] lemmas sums_scaleR_left = bounded_linear.sums[OF bounded_linear_scaleR_left] lemmas summable_scaleR_left = bounded_linear.summable[OF bounded_linear_scaleR_left] lemmas suminf_scaleR_left = bounded_linear.suminf[OF bounded_linear_scaleR_left] lemmas sums_scaleR_right = bounded_linear.sums[OF bounded_linear_scaleR_right] lemmas summable_scaleR_right = bounded_linear.summable[OF bounded_linear_scaleR_right] lemmas suminf_scaleR_right = bounded_linear.suminf[OF bounded_linear_scaleR_right] lemma summable_const_iff: "summable (\_. c) \ c = 0" for c :: "'a::real_normed_vector" proof - have "\ summable (\_. c)" if "c \ 0" proof - from that have "filterlim (\n. of_nat n * norm c) at_top sequentially" by (subst mult.commute) (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_real_sequentially) then have "\ convergent (\n. norm (\kInfinite summability on real normed algebras\ context fixes f :: "nat \ 'a::real_normed_algebra" begin lemma sums_mult: "f sums a \ (\n. c * f n) sums (c * a)" by (rule bounded_linear.sums [OF bounded_linear_mult_right]) lemma summable_mult: "summable f \ summable (\n. c * f n)" by (rule bounded_linear.summable [OF bounded_linear_mult_right]) lemma suminf_mult: "summable f \ suminf (\n. c * f n) = c * suminf f" by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric]) lemma sums_mult2: "f sums a \ (\n. f n * c) sums (a * c)" by (rule bounded_linear.sums [OF bounded_linear_mult_left]) lemma summable_mult2: "summable f \ summable (\n. f n * c)" by (rule bounded_linear.summable [OF bounded_linear_mult_left]) lemma suminf_mult2: "summable f \ suminf f * c = (\n. f n * c)" by (rule bounded_linear.suminf [OF bounded_linear_mult_left]) end lemma sums_mult_iff: fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" shows "(\n. c * f n) sums (c * d) \ f sums d" using sums_mult[of f d c] sums_mult[of "\n. c * f n" "c * d" "inverse c"] by (force simp: field_simps assms) lemma sums_mult2_iff: fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" shows "(\n. f n * c) sums (d * c) \ f sums d" using sums_mult_iff[OF assms, of f d] by (simp add: mult.commute) lemma sums_of_real_iff: "(\n. of_real (f n) :: 'a::real_normed_div_algebra) sums of_real c \ f sums c" by (simp add: sums_def of_real_sum[symmetric] tendsto_of_real_iff del: of_real_sum) subsection \Infinite summability on real normed fields\ context fixes c :: "'a::real_normed_field" begin lemma sums_divide: "f sums a \ (\n. f n / c) sums (a / c)" by (rule bounded_linear.sums [OF bounded_linear_divide]) lemma summable_divide: "summable f \ summable (\n. f n / c)" by (rule bounded_linear.summable [OF bounded_linear_divide]) lemma suminf_divide: "summable f \ suminf (\n. f n / c) = suminf f / c" by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric]) lemma summable_inverse_divide: "summable (inverse \ f) \ summable (\n. c / f n)" by (auto dest: summable_mult [of _ c] simp: field_simps) lemma sums_mult_D: "(\n. c * f n) sums a \ c \ 0 \ f sums (a/c)" using sums_mult_iff by fastforce lemma summable_mult_D: "summable (\n. c * f n) \ c \ 0 \ summable f" by (auto dest: summable_divide) text \Sum of a geometric progression.\ lemma geometric_sums: assumes "norm c < 1" shows "(\n. c^n) sums (1 / (1 - c))" proof - have neq_0: "c - 1 \ 0" using assms by auto then have "(\n. c ^ n / (c - 1) - 1 / (c - 1)) \ 0 / (c - 1) - 1 / (c - 1)" by (intro tendsto_intros assms) then have "(\n. (c ^ n - 1) / (c - 1)) \ 1 / (1 - c)" by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) with neq_0 show "(\n. c ^ n) sums (1 / (1 - c))" by (simp add: sums_def geometric_sum) qed lemma summable_geometric: "norm c < 1 \ summable (\n. c^n)" by (rule geometric_sums [THEN sums_summable]) lemma suminf_geometric: "norm c < 1 \ suminf (\n. c^n) = 1 / (1 - c)" by (rule sums_unique[symmetric]) (rule geometric_sums) lemma summable_geometric_iff [simp]: "summable (\n. c ^ n) \ norm c < 1" proof assume "summable (\n. c ^ n :: 'a :: real_normed_field)" then have "(\n. norm c ^ n) \ 0" by (simp add: norm_power [symmetric] tendsto_norm_zero_iff summable_LIMSEQ_zero) from order_tendstoD(2)[OF this zero_less_one] obtain n where "norm c ^ n < 1" by (auto simp: eventually_at_top_linorder) then show "norm c < 1" using one_le_power[of "norm c" n] by (cases "norm c \ 1") (linarith, simp) qed (rule summable_geometric) end text \Biconditional versions for constant factors\ context fixes c :: "'a::real_normed_field" begin lemma summable_cmult_iff [simp]: "summable (\n. c * f n) \ c=0 \ summable f" proof - have "\summable (\n. c * f n); c \ 0\ \ summable f" using summable_mult_D by blast then show ?thesis by (auto simp: summable_mult) qed lemma summable_divide_iff [simp]: "summable (\n. f n / c) \ c=0 \ summable f" proof - have "\summable (\n. f n / c); c \ 0\ \ summable f" by (auto dest: summable_divide [where c = "1/c"]) then show ?thesis by (auto simp: summable_divide) qed end lemma power_half_series: "(\n. (1/2::real)^Suc n) sums 1" proof - have 2: "(\n. (1/2::real)^n) sums 2" using geometric_sums [of "1/2::real"] by auto have "(\n. (1/2::real)^Suc n) = (\n. (1 / 2) ^ n / 2)" by (simp add: mult.commute) then show ?thesis using sums_divide [OF 2, of 2] by simp qed subsection \Telescoping\ lemma telescope_sums: fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "(\n. f (Suc n) - f n) sums (c - f 0)" unfolding sums_def proof (subst filterlim_sequentially_Suc [symmetric]) have "(\n. \kn. f (Suc n) - f 0)" by (simp add: lessThan_Suc_atMost atLeast0AtMost [symmetric] sum_Suc_diff) also have "\ \ c - f 0" by (intro tendsto_diff LIMSEQ_Suc[OF assms] tendsto_const) finally show "(\n. \n c - f 0" . qed lemma telescope_sums': fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "(\n. f n - f (Suc n)) sums (f 0 - c)" using sums_minus[OF telescope_sums[OF assms]] by (simp add: algebra_simps) lemma telescope_summable: fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "summable (\n. f (Suc n) - f n)" using telescope_sums[OF assms] by (simp add: sums_iff) lemma telescope_summable': fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "summable (\n. f n - f (Suc n))" using summable_minus[OF telescope_summable[OF assms]] by (simp add: algebra_simps) subsection \Infinite summability on Banach spaces\ text \Cauchy-type criterion for convergence of series (c.f. Harrison).\ lemma summable_Cauchy: "summable f \ (\e>0. \N. \m\N. \n. norm (sum f {m.. 'a::banach" proof assume f: "summable f" show ?rhs proof clarify fix e :: real assume "0 < e" then obtain M where M: "\m n. \m\M; n\M\ \ norm (sum f {.. M" for m n proof (cases m n rule: linorder_class.le_cases) assume "m \ n" then show ?thesis by (metis (mono_tags, hide_lams) M atLeast0LessThan order_trans sum_diff_nat_ivl that zero_le) next assume "n \ m" then show ?thesis by (simp add: \0 < e\) qed then show "\N. \m\N. \n. norm (sum f {m..m n. m \ N \ norm (sum f {m..N" "n\N" for m n proof (cases m n rule: linorder_class.le_cases) assume "m \ n" then show ?thesis - by (metis Groups_Big.sum_diff N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff norm_minus_commute \m\N\) + by (metis N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff norm_minus_commute sum_diff \m\N\) next assume "n \ m" then show ?thesis - by (metis Groups_Big.sum_diff N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff \n\N\) + by (metis N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff sum_diff \n\N\) qed then show "\M. \m\M. \n\M. norm (sum f {.. 'a :: banach" - assumes "eventually (\m. \n\m. norm (sum f {m.. g m) sequentially" - assumes "filterlim g (nhds 0) sequentially" + assumes ev: "eventually (\m. \n\m. norm (sum f {m.. g m) sequentially" + assumes g0: "g \ 0" shows "summable f" proof (subst summable_Cauchy, intro allI impI, goal_cases) case (1 e) - from order_tendstoD(2)[OF assms(2) this] and assms(1) - have "eventually (\m. \n. norm (sum f {m..\<^sub>F x in sequentially. g x < e" + using g0 order_tendstoD(2) by blast + with ev have "eventually (\m. \n. norm (sum f {m.. m") auto qed qed thus ?case by (auto simp: eventually_at_top_linorder) qed context fixes f :: "nat \ 'a::banach" begin text \Absolute convergence imples normal convergence.\ lemma summable_norm_cancel: "summable (\n. norm (f n)) \ summable f" unfolding summable_Cauchy apply (erule all_forward imp_forward ex_forward | assumption)+ apply (fastforce simp add: order_le_less_trans [OF norm_sum] order_le_less_trans [OF abs_ge_self]) done lemma summable_norm: "summable (\n. norm (f n)) \ norm (suminf f) \ (\n. norm (f n))" by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_sum) text \Comparison tests.\ lemma summable_comparison_test: assumes fg: "\N. \n\N. norm (f n) \ g n" and g: "summable g" shows "summable f" proof - obtain N where N: "\n. n\N \ norm (f n) \ g n" using assms by blast show ?thesis proof (clarsimp simp add: summable_Cauchy) fix e :: real assume "0 < e" then obtain Ng where Ng: "\m n. m \ Ng \ norm (sum g {m..max N Ng" for m n proof - have "norm (sum f {m.. sum g {m.. norm (sum g {m..N. \m\N. \n. norm (sum f {m..n. norm (f n) \ g n) sequentially \ summable g \ summable f" by (rule summable_comparison_test) (auto simp: eventually_at_top_linorder) text \A better argument order.\ lemma summable_comparison_test': "summable g \ (\n. n \ N \ norm (f n) \ g n) \ summable f" by (rule summable_comparison_test) auto subsection \The Ratio Test\ lemma summable_ratio_test: assumes "c < 1" "\n. n \ N \ norm (f (Suc n)) \ c * norm (f n)" shows "summable f" proof (cases "0 < c") case True show "summable f" proof (rule summable_comparison_test) show "\N'. \n\N'. norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" proof (intro exI allI impI) fix n assume "N \ n" then show "norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" proof (induct rule: inc_induct) case base with True show ?case by simp next case (step m) have "norm (f (Suc m)) / c ^ Suc m * c ^ n \ norm (f m) / c ^ m * c ^ n" using \0 < c\ \c < 1\ assms(2)[OF \N \ m\] by (simp add: field_simps) with step show ?case by simp qed qed show "summable (\n. norm (f N) / c ^ N * c ^ n)" using \0 < c\ \c < 1\ by (intro summable_mult summable_geometric) simp qed next case False have "f (Suc n) = 0" if "n \ N" for n proof - from that have "norm (f (Suc n)) \ c * norm (f n)" by (rule assms(2)) also have "\ \ 0" using False by (simp add: not_less mult_nonpos_nonneg) finally show ?thesis by auto qed then show "summable f" by (intro sums_summable[OF sums_finite, of "{.. Suc N}"]) (auto simp: not_le Suc_less_eq2) qed end text \Relations among convergence and absolute convergence for power series.\ lemma Abel_lemma: fixes a :: "nat \ 'a::real_normed_vector" assumes r: "0 \ r" and r0: "r < r0" and M: "\n. norm (a n) * r0^n \ M" shows "summable (\n. norm (a n) * r^n)" proof (rule summable_comparison_test') show "summable (\n. M * (r / r0) ^ n)" using assms by (auto simp add: summable_mult summable_geometric) show "norm (norm (a n) * r ^ n) \ M * (r / r0) ^ n" for n using r r0 M [of n] dual_order.order_iff_strict by (fastforce simp add: abs_mult field_simps) qed text \Summability of geometric series for real algebras.\ lemma complete_algebra_summable_geometric: fixes x :: "'a::{real_normed_algebra_1,banach}" assumes "norm x < 1" shows "summable (\n. x ^ n)" proof (rule summable_comparison_test) show "\N. \n\N. norm (x ^ n) \ norm x ^ n" by (simp add: norm_power_ineq) from assms show "summable (\n. norm x ^ n)" by (simp add: summable_geometric) qed subsection \Cauchy Product Formula\ text \ Proof based on Analysis WebNotes: Chapter 07, Class 41 \<^url>\http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm\ \ lemma Cauchy_product_sums: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes a: "summable (\k. norm (a k))" and b: "summable (\k. norm (b k))" shows "(\k. \i\k. a i * b (k - i)) sums ((\k. a k) * (\k. b k))" proof - let ?S1 = "\n::nat. {.. {..m n. m \ n \ ?S1 m \ ?S1 n" by auto have S2_le_S1: "\n. ?S2 n \ ?S1 n" by auto have S1_le_S2: "\n. ?S1 (n div 2) \ ?S2 n" by auto have finite_S1: "\n. finite (?S1 n)" by simp with S2_le_S1 have finite_S2: "\n. finite (?S2 n)" by (rule finite_subset) let ?g = "\(i,j). a i * b j" let ?f = "\(i,j). norm (a i) * norm (b j)" have f_nonneg: "\x. 0 \ ?f x" by auto then have norm_sum_f: "\A. norm (sum ?f A) = sum ?f A" unfolding real_norm_def by (simp only: abs_of_nonneg sum_nonneg [rule_format]) have "(\n. (\kk (\k. a k) * (\k. b k)" by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b]) then have 1: "(\n. sum ?g (?S1 n)) \ (\k. a k) * (\k. b k)" by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan) have "(\n. (\kk (\k. norm (a k)) * (\k. norm (b k))" using a b by (intro tendsto_mult summable_LIMSEQ) then have "(\n. sum ?f (?S1 n)) \ (\k. norm (a k)) * (\k. norm (b k))" by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan) then have "convergent (\n. sum ?f (?S1 n))" by (rule convergentI) then have Cauchy: "Cauchy (\n. sum ?f (?S1 n))" by (rule convergent_Cauchy) have "Zfun (\n. sum ?f (?S1 n - ?S2 n)) sequentially" proof (rule ZfunI, simp only: eventually_sequentially norm_sum_f) fix r :: real assume r: "0 < r" from CauchyD [OF Cauchy r] obtain N where "\m\N. \n\N. norm (sum ?f (?S1 m) - sum ?f (?S1 n)) < r" .. then have "\m n. N \ n \ n \ m \ norm (sum ?f (?S1 m - ?S1 n)) < r" by (simp only: sum_diff finite_S1 S1_mono) then have N: "\m n. N \ n \ n \ m \ sum ?f (?S1 m - ?S1 n) < r" by (simp only: norm_sum_f) show "\N. \n\N. sum ?f (?S1 n - ?S2 n) < r" proof (intro exI allI impI) fix n assume "2 * N \ n" then have n: "N \ n div 2" by simp have "sum ?f (?S1 n - ?S2 n) \ sum ?f (?S1 n - ?S1 (n div 2))" by (intro sum_mono2 finite_Diff finite_S1 f_nonneg Diff_mono subset_refl S1_le_S2) also have "\ < r" using n div_le_dividend by (rule N) finally show "sum ?f (?S1 n - ?S2 n) < r" . qed qed then have "Zfun (\n. sum ?g (?S1 n - ?S2 n)) sequentially" apply (rule Zfun_le [rule_format]) apply (simp only: norm_sum_f) apply (rule order_trans [OF norm_sum sum_mono]) apply (auto simp add: norm_mult_ineq) done then have 2: "(\n. sum ?g (?S1 n) - sum ?g (?S2 n)) \ 0" unfolding tendsto_Zfun_iff diff_0_right by (simp only: sum_diff finite_S1 S2_le_S1) with 1 have "(\n. sum ?g (?S2 n)) \ (\k. a k) * (\k. b k)" by (rule Lim_transform2) then show ?thesis by (simp only: sums_def sum.triangle_reindex) qed lemma Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes "summable (\k. norm (a k))" and "summable (\k. norm (b k))" shows "(\k. a k) * (\k. b k) = (\k. \i\k. a i * b (k - i))" using assms by (rule Cauchy_product_sums [THEN sums_unique]) lemma summable_Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes "summable (\k. norm (a k))" and "summable (\k. norm (b k))" shows "summable (\k. \i\k. a i * b (k - i))" using Cauchy_product_sums[OF assms] by (simp add: sums_iff) subsection \Series on \<^typ>\real\s\ lemma summable_norm_comparison_test: "\N. \n\N. norm (f n) \ g n \ summable g \ summable (\n. norm (f n))" by (rule summable_comparison_test) auto lemma summable_rabs_comparison_test: "\N. \n\N. \f n\ \ g n \ summable g \ summable (\n. \f n\)" for f :: "nat \ real" by (rule summable_comparison_test) auto lemma summable_rabs_cancel: "summable (\n. \f n\) \ summable f" for f :: "nat \ real" by (rule summable_norm_cancel) simp lemma summable_rabs: "summable (\n. \f n\) \ \suminf f\ \ (\n. \f n\)" for f :: "nat \ real" by (fold real_norm_def) (rule summable_norm) lemma summable_zero_power [simp]: "summable (\n. 0 ^ n :: 'a::{comm_ring_1,topological_space})" proof - have "(\n. 0 ^ n :: 'a) = (\n. if n = 0 then 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp qed lemma summable_zero_power' [simp]: "summable (\n. f n * 0 ^ n :: 'a::{ring_1,topological_space})" proof - have "(\n. f n * 0 ^ n :: 'a) = (\n. if n = 0 then f 0 * 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp qed lemma summable_power_series: fixes z :: real assumes le_1: "\i. f i \ 1" and nonneg: "\i. 0 \ f i" and z: "0 \ z" "z < 1" shows "summable (\i. f i * z^i)" proof (rule summable_comparison_test[OF _ summable_geometric]) show "norm z < 1" using z by (auto simp: less_imp_le) show "\n. \N. \na\N. norm (f na * z ^ na) \ z ^ na" using z by (auto intro!: exI[of _ 0] mult_left_le_one_le simp: abs_mult nonneg power_abs less_imp_le le_1) qed lemma summable_0_powser: "summable (\n. f n * 0 ^ n :: 'a::real_normed_div_algebra)" proof - have A: "(\n. f n * 0 ^ n) = (\n. if n = 0 then f n else 0)" by (intro ext) auto then show ?thesis by (subst A) simp_all qed lemma summable_powser_split_head: "summable (\n. f (Suc n) * z ^ n :: 'a::real_normed_div_algebra) = summable (\n. f n * z ^ n)" proof - have "summable (\n. f (Suc n) * z ^ n) \ summable (\n. f (Suc n) * z ^ Suc n)" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs using summable_mult2[OF that, of z] by (simp add: power_commutes algebra_simps) show ?lhs if ?rhs using summable_mult2[OF that, of "inverse z"] by (cases "z \ 0", subst (asm) power_Suc2) (simp_all add: algebra_simps) qed also have "\ \ summable (\n. f n * z ^ n)" by (rule summable_Suc_iff) finally show ?thesis . qed lemma summable_powser_ignore_initial_segment: fixes f :: "nat \ 'a :: real_normed_div_algebra" shows "summable (\n. f (n + m) * z ^ n) \ summable (\n. f n * z ^ n)" proof (induction m) case (Suc m) have "summable (\n. f (n + Suc m) * z ^ n) = summable (\n. f (Suc n + m) * z ^ n)" by simp also have "\ = summable (\n. f (n + m) * z ^ n)" by (rule summable_powser_split_head) also have "\ = summable (\n. f n * z ^ n)" by (rule Suc.IH) finally show ?case . qed simp_all lemma powser_split_head: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" assumes "summable (\n. f n * z ^ n)" shows "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" and "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" and "summable (\n. f (Suc n) * z ^ n)" proof - from assms show "summable (\n. f (Suc n) * z ^ n)" by (subst summable_powser_split_head) from suminf_mult2[OF this, of z] have "(\n. f (Suc n) * z ^ n) * z = (\n. f (Suc n) * z ^ Suc n)" by (simp add: power_commutes algebra_simps) also from assms have "\ = suminf (\n. f n * z ^ n) - f 0" by (subst suminf_split_head) simp_all finally show "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" by simp then show "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" by simp qed lemma summable_partial_sum_bound: fixes f :: "nat \ 'a :: banach" and e :: real assumes summable: "summable f" and e: "e > 0" obtains N where "\m n. m \ N \ norm (\k=m..n. f k) < e" proof - from summable have "Cauchy (\n. \km n. m \ N \ n \ N \ norm ((\kkk=m..n. f k) < e" if m: "m \ N" for m n proof (cases "n \ m") case True with m have "norm ((\kkkkk=m..n. f k)" by (subst sum_diff [symmetric]) (simp_all add: sum.last_plus) finally show ?thesis . next case False with e show ?thesis by simp_all qed then show ?thesis by (rule that) qed lemma powser_sums_if: "(\n. (if n = m then (1 :: 'a::{ring_1,topological_space}) else 0) * z^n) sums z^m" proof - have "(\n. (if n = m then 1 else 0) * z^n) = (\n. if n = m then z^n else 0)" by (intro ext) auto then show ?thesis by (simp add: sums_single) qed lemma fixes f :: "nat \ real" assumes "summable f" and "inj g" and pos: "\x. 0 \ f x" shows summable_reindex: "summable (f \ g)" and suminf_reindex_mono: "suminf (f \ g) \ suminf f" and suminf_reindex: "(\x. x \ range g \ f x = 0) \ suminf (f \ g) = suminf f" proof - from \inj g\ have [simp]: "\A. inj_on g A" by (rule subset_inj_on) simp have smaller: "\n. (\i g) i) \ suminf f" proof fix n have "\ n' \ (g ` {..n'. n' < n \ g n' < m" by blast have "(\i \ (\i \ suminf f" using \summable f\ by (rule sum_le_suminf) (simp_all add: pos) finally show "(\i g) i) \ suminf f" by simp qed have "incseq (\n. \i g) i)" by (rule incseq_SucI) (auto simp add: pos) then obtain L where L: "(\ n. \i g) i) \ L" using smaller by(rule incseq_convergent) then have "(f \ g) sums L" by (simp add: sums_def) then show "summable (f \ g)" by (auto simp add: sums_iff) then have "(\n. \i g) i) \ suminf (f \ g)" by (rule summable_LIMSEQ) then show le: "suminf (f \ g) \ suminf f" by(rule LIMSEQ_le_const2)(blast intro: smaller[rule_format]) assume f: "\x. x \ range g \ f x = 0" from \summable f\ have "suminf f \ suminf (f \ g)" proof (rule suminf_le_const) fix n have "\ n' \ (g -` {..n'. g n' < n \ n' < m" by blast have "(\ii\{.. range g. f i)" using f by(auto intro: sum.mono_neutral_cong_right) also have "\ = (\i\g -` {.. g) i)" by (rule sum.reindex_cong[where l=g])(auto) also have "\ \ (\i g) i)" by (rule sum_mono2)(auto simp add: pos n) also have "\ \ suminf (f \ g)" using \summable (f \ g)\ by (rule sum_le_suminf) (simp_all add: pos) finally show "sum f {.. suminf (f \ g)" . qed with le show "suminf (f \ g) = suminf f" by (rule antisym) qed lemma sums_mono_reindex: assumes subseq: "strict_mono g" and zero: "\n. n \ range g \ f n = 0" shows "(\n. f (g n)) sums c \ f sums c" unfolding sums_def proof assume lim: "(\n. \k c" have "(\n. \kn. \kkk\g`{.. = (\kkk \ c" by (simp only: o_def) finally show "(\n. \k c" . next assume lim: "(\n. \k c" define g_inv where "g_inv n = (LEAST m. g m \ n)" for n from filterlim_subseq[OF subseq] have g_inv_ex: "\m. g m \ n" for n by (auto simp: filterlim_at_top eventually_at_top_linorder) then have g_inv: "g (g_inv n) \ n" for n unfolding g_inv_def by (rule LeastI_ex) have g_inv_least: "m \ g_inv n" if "g m \ n" for m n using that unfolding g_inv_def by (rule Least_le) have g_inv_least': "g m < n" if "m < g_inv n" for m n using that g_inv_least[of n m] by linarith have "(\n. \kn. \k {.. range g" proof (rule notI, elim imageE) fix l assume l: "k = g l" have "g l < g (g_inv n)" by (rule less_le_trans[OF _ g_inv]) (use k l in simp_all) with subseq have "l < g_inv n" by (simp add: strict_mono_less) with k l show False by simp qed then have "f k = 0" by (rule zero) } with g_inv_least' g_inv have "(\kk\g`{.. = (\kkk n" also have "n \ g (g_inv n)" by (rule g_inv) finally have "K \ g_inv n" using subseq by (simp add: strict_mono_less_eq) } then have "filterlim g_inv at_top sequentially" by (auto simp: filterlim_at_top eventually_at_top_linorder) with lim have "(\n. \k c" by (rule filterlim_compose) finally show "(\n. \k c" . qed lemma summable_mono_reindex: assumes subseq: "strict_mono g" and zero: "\n. n \ range g \ f n = 0" shows "summable (\n. f (g n)) \ summable f" using sums_mono_reindex[of g f, OF assms] by (simp add: summable_def) lemma suminf_mono_reindex: fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" assumes "strict_mono g" "\n. n \ range g \ f n = 0" shows "suminf (\n. f (g n)) = suminf f" proof (cases "summable f") case True with sums_mono_reindex [of g f, OF assms] and summable_mono_reindex [of g f, OF assms] show ?thesis by (simp add: sums_iff) next case False then have "\(\c. f sums c)" unfolding summable_def by blast then have "suminf f = The (\_. False)" by (simp add: suminf_def) moreover from False have "\ summable (\n. f (g n))" using summable_mono_reindex[of g f, OF assms] by simp then have "\(\c. (\n. f (g n)) sums c)" unfolding summable_def by blast then have "suminf (\n. f (g n)) = The (\_. False)" by (simp add: suminf_def) ultimately show ?thesis by simp qed lemma summable_bounded_partials: fixes f :: "nat \ 'a :: {real_normed_vector,complete_space}" assumes bound: "eventually (\x0. \a\x0. \b>a. norm (sum f {a<..b}) \ g a) sequentially" assumes g: "g \ 0" shows "summable f" unfolding summable_iff_convergent' proof (intro Cauchy_convergent CauchyI', goal_cases) case (1 \) with g have "eventually (\x. \g x\ < \) sequentially" by (auto simp: tendsto_iff) from eventually_conj[OF this bound] obtain x0 where x0: "\x. x \ x0 \ \g x\ < \" "\a b. x0 \ a \ a < b \ norm (sum f {a<..b}) \ g a" unfolding eventually_at_top_linorder by auto show ?case proof (intro exI[of _ x0] allI impI) fix m n assume mn: "x0 \ m" "m < n" have "dist (sum f {..m}) (sum f {..n}) = norm (sum f {..n} - sum f {..m})" by (simp add: dist_norm norm_minus_commute) also have "sum f {..n} - sum f {..m} = sum f ({..n} - {..m})" using mn by (intro Groups_Big.sum_diff [symmetric]) auto also have "{..n} - {..m} = {m<..n}" using mn by auto also have "norm (sum f {m<..n}) \ g m" using mn by (intro x0) auto also have "\ \ \g m\" by simp also have "\ < \" using mn by (intro x0) auto finally show "dist (sum f {..m}) (sum f {..n}) < \" . qed qed end