# HG changeset patch # User wenzelm # Date 1469522019 -7200 # Node ID 3a0f40a6fa42362ea4a7e90e44cd4becea25fe37 # Parent b0d31c7def86ff0fb5978b375978dad1b45c5e22 misc tuning and modernization; diff -r b0d31c7def86 -r 3a0f40a6fa42 src/HOL/Series.thy --- a/src/HOL/Series.thy Mon Jul 25 21:50:04 2016 +0200 +++ b/src/HOL/Series.thy Tue Jul 26 10:33:39 2016 +0200 @@ -15,20 +15,16 @@ 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 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 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)" +definition suminf :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a" + (binder "\" 10) + where "suminf f = (THE s. f sums s)" lemma sums_def': "f sums s \ (\n. \i = 0..n. f i) \ s" apply (simp add: sums_def) @@ -36,6 +32,7 @@ apply (simp only: lessThan_Suc_atMost atLeast0AtMost) done + subsection \Infinite summability on topological monoids\ lemma sums_subst[trans]: "f = g \ g sums z \ f sums z" @@ -50,8 +47,7 @@ lemma summable_iff_convergent: "summable f \ convergent (\n. \i convergent (\n. setsum f {..n})" +lemma summable_iff_convergent': "summable f \ convergent (\n. setsum f {..n})" by (simp_all only: summable_iff_convergent convergent_def lessThan_Suc_atMost [symmetric] LIMSEQ_Suc_iff[of "\n. setsum f {..x. f x = (g x :: 'a :: real_normed_vector)) sequentially" - shows "summable f = summable g" + 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) + from assms obtain N where N: "\n\N. f n = g n" + by (auto simp: eventually_at_top_linorder) define C where "C = (\kn. setsum f {..n. setsum f {.. N" - from n have "{.. {N.. {N.. {N.. {N.. {N..n. n \ N \ f n = 0" + assumes [simp]: "finite N" + and f: "\n. n \ N \ f n = 0" shows "f sums (\n\N. f n)" proof - - { fix n - have "setsum f {..x. 0)" by auto - then show ?thesis by simp - next - assume [simp]: "N \ {}" - show ?thesis - proof (safe intro!: setsum.mono_neutral_right f) - fix i assume "i \ N" - then have "i \ Max N" by simp - then show "i < n + Suc (Max N)" by simp - qed - qed } - note eq = this - show ?thesis unfolding sums_def + have eq: "setsum f {..x. 0)" by auto + then show ?thesis by simp + next + case [simp]: False + show ?thesis + proof (safe intro!: setsum.mono_neutral_right f) + fix i + assume "i \ N" + then have "i \ Max N" by simp + then show "i < n + Suc (Max N)" by simp + qed + qed + show ?thesis + unfolding sums_def by (rule LIMSEQ_offset[of _ "Suc (Max N)"]) (simp add: eq atLeast0LessThan del: add_Suc_right) qed -corollary sums_0: - "(\n. f n = 0) \ (f sums 0)" +corollary sums_0: "(\n. f n = 0) \ (f sums 0)" by (metis (no_types) finite.emptyI setsum.empty sums_finite) lemma summable_finite: "finite N \ (\n. n \ N \ f n = 0) \ summable f" @@ -155,7 +156,7 @@ by (rule sums_summable) (rule sums_single) context - fixes f :: "nat \ 'a::{t2_space, comm_monoid_add}" + fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" begin lemma summable_sums[intro]: "summable f \ f sums (suminf f)" @@ -168,20 +169,19 @@ 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)" +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 :: nat \ 'a :: {comm_monoid_add,t2_space}) \ f sums suminf f" +lemma summable_sums_iff: "summable f \ f sums suminf f" by (auto simp: sums_iff summable_sums) -lemma sums_unique2: - fixes a b :: "'a::{comm_monoid_add,t2_space}" - shows "f sums a \ f sums b \ a = b" -by (simp add: sums_iff) +lemma sums_unique2: "f sums a \ f sums b \ a = b" + for a b :: 'a + by (simp add: sums_iff) lemma suminf_finite: - assumes N: "finite N" and f: "\n. n \ N \ f n = 0" + 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 @@ -193,16 +193,15 @@ subsection \Infinite summability on ordered, topological monoids\ -lemma sums_le: - fixes f g :: "nat \ 'a::{ordered_comm_monoid_add, linorder_topology}" - shows "\n. f n \ g n \ f sums s \ g sums t \ s \ t" +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: setsum_mono simp: sums_def) context - fixes f :: "nat \ 'a::{ordered_comm_monoid_add, linorder_topology}" + 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" +lemma suminf_le: "\n. f n \ g n \ summable f \ summable g \ suminf f \ suminf g" by (auto dest: sums_summable intro: sums_le) lemma setsum_le_suminf: "summable f \ \m\n. 0 \ f m \ setsum f {.. suminf f" @@ -221,15 +220,14 @@ using summable_LIMSEQ[of f] by simp then have "\i. (\n\{i}. f n) \ 0" proof (rule LIMSEQ_le_const) - fix i show "\N. \n\N. (\n\{i}. f n) \ setsum f {..N. \n\N. (\n\{i}. f n) \ setsum f {..n. f n = 0" by (auto intro!: antisym) qed (metis suminf_zero fun_eq_iff) -lemma suminf_pos_iff: - "summable f \ \n. 0 \ f n \ 0 < suminf f \ (\i. 0 < f i)" +lemma suminf_pos_iff: "summable f \ \n. 0 \ f n \ 0 < suminf f \ (\i. 0 < f i)" using setsum_le_suminf[of 0] suminf_eq_zero_iff by (simp add: less_le) lemma suminf_pos2: @@ -249,14 +247,14 @@ end context - fixes f :: "nat \ 'a::{ordered_cancel_comm_monoid_add, linorder_topology}" + fixes f :: "nat \ 'a::{ordered_cancel_comm_monoid_add,linorder_topology}" begin -lemma setsum_less_suminf2: "summable f \ \m\n. 0 \ f m \ n \ i \ 0 < f i \ setsum f {.. \m\n. 0 \ f m \ n \ i \ 0 < f i \ setsum f {.. \m\n. 0 < f m \ setsum f {.. 'a::{ordered_comm_monoid_add, linorder_topology, conditionally_complete_linorder}" - assumes pos[simp]: "\n. 0 \ f n" and le: "\n. (\i x" + fixes f :: "nat \ '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] + unfolding summable_def sums_def [abs_def] proof (rule exI LIMSEQ_incseq_SUP)+ show "bdd_above (range (\n. setsum f {.. 'a::{canonically_ordered_monoid_add, linorder_topology, complete_linorder}" - shows "summable f" +lemma summableI[intro, simp]: "summable f" + for f :: "nat \ 'a::{canonically_ordered_monoid_add,linorder_topology,complete_linorder}" by (intro summableI_nonneg_bounded[where x=top] zero_le top_greatest) + subsection \Infinite summability on topological monoids\ context - fixes f g :: "nat \ 'a :: {t2_space, topological_comm_monoid_add}" + 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)" + 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)" @@ -311,7 +311,8 @@ end context - fixes f :: "'i \ nat \ 'a::{t2_space, topological_comm_monoid_add}" and I :: "'i set" + fixes f :: "'i \ nat \ 'a::{t2_space,topological_comm_monoid_add}" + and I :: "'i set" begin lemma sums_setsum: "(\i. i \ I \ (f i) sums (x i)) \ (\n. \i\I. f i n) sums (\i\I. x i)" @@ -340,8 +341,7 @@ 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" + 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 .. @@ -350,9 +350,12 @@ lemma summable_Suc_iff: "summable (\n. f (Suc n)) = summable f" proof assume "summable f" - hence "f sums suminf f" by (rule summable_sums) - hence "(\n. f (Suc n)) sums (suminf f - f 0)" by (simp add: sums_Suc_iff) - thus "summable (\n. f (Suc n))" unfolding summable_def by blast + 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" @@ -360,7 +363,7 @@ end -context \\Separate contexts are necessary to allow general use of the results above, here.\ +context (* Separate contexts are necessary to allow general use of the results above, here. *) fixes f :: "nat \ 'a::real_normed_vector" begin @@ -384,12 +387,15 @@ 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)" + then have "(\i. f (Suc i + n)) sums s \ (\i. f (i + n)) sums (s + f n)" by (subst sums_Suc_iff) simp - ultimately show ?case + with Suc show ?case by (simp add: ac_simps) -qed simp +qed corollary sums_iff_shift': "(\i. f (i + n)) sums (s - (\i f sums s" by (simp add: sums_iff_shift) @@ -397,10 +403,10 @@ 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) + by (simp add: assms sums_iff_shift) lemma summable_iff_shift: "summable (\n. f (n + k)) \ summable f" - by (metis diff_add_cancel summable_def sums_iff_shift[abs_def]) + 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 - (\iN. \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 (setsum f {.. n \ N. norm (setsum f {..summable f\]) qed lemma summable_LIMSEQ_zero: "summable f \ f \ 0" apply (drule summable_iff_convergent [THEN iffD1]) apply (drule convergent_Cauchy) - apply (simp only: Cauchy_iff LIMSEQ_iff, safe) - apply (drule_tac x="r" in spec, safe) - apply (rule_tac x="M" in exI, safe) - apply (drule_tac x="Suc n" in spec, simp) - apply (drule_tac x="n" in spec, simp) + apply (simp only: Cauchy_iff LIMSEQ_iff) + apply safe + apply (drule_tac x="r" in spec) + apply safe + apply (rule_tac x="M" in exI) + apply safe + apply (drule_tac x="Suc n" in spec) + apply simp + apply (drule_tac x="n" in spec) + apply simp done lemma summable_imp_convergent: "summable f \ convergent f" @@ -445,13 +458,12 @@ end -lemma summable_minus_iff: - fixes f :: "nat \ 'a::real_normed_vector" - shows "summable (\n. - f n) \ summable f" - by (auto dest: summable_minus) \\used two ways, hence must be outside the context above\ +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: setsum) + unfolding sums_def by (drule tendsto) (simp only: setsum) lemma (in bounded_linear) summable: "summable (\n. X n) \ summable (\n. f (X n))" unfolding summable_def by (auto intro: sums) @@ -471,19 +483,21 @@ 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 :: 'a :: real_normed_vector) = 0" +lemma summable_const_iff: "summable (\_. c) \ c = 0" + for c :: "'a::real_normed_vector" proof - - { - assume "c \ 0" - hence "filterlim (\n. of_nat n * norm c) at_top sequentially" + 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) - hence "\convergent (\n. norm (\k convergent (\n. norm (\ksummable (\_. c)" unfolding summable_iff_convergent using convergent_norm by blast - } - thus ?thesis by auto + (simp_all add: setsum_constant_scaleR) + then show ?thesis + unfolding summable_iff_convergent using convergent_norm by blast + qed + then show ?thesis by auto qed @@ -514,18 +528,20 @@ end lemma sums_mult_iff: + fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" - shows "(\n. c * f n :: 'a :: {real_normed_algebra,field}) sums (c * d) \ f sums d" + 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: - assumes "c \ (0 :: 'a :: {real_normed_algebra, field})" + 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" + "(\n. of_real (f n) :: 'a::real_normed_div_algebra) sums of_real c \ f sums c" by (simp add: sums_def of_real_setsum[symmetric] tendsto_of_real_iff del: of_real_setsum) @@ -544,26 +560,28 @@ lemma suminf_divide: "summable f \ suminf (\n. f n / c) = suminf f / c" by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric]) -lemma sums_mult_D: "\(\n. c * f n) sums a; c \ 0\ \ f sums (a/c)" +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" +lemma summable_mult_D: "summable (\n. c * f n) \ c \ 0 \ summable f" by (auto dest: summable_divide) -text\Sum of a geometric progression.\ + +text \Sum of a geometric progression.\ -lemma geometric_sums: "norm c < 1 \ (\n. c^n) sums (1 / (1 - c))" +lemma geometric_sums: + assumes less_1: "norm c < 1" + shows "(\n. c^n) sums (1 / (1 - c))" proof - - assume less_1: "norm c < 1" - hence neq_1: "c \ 1" by auto - hence neq_0: "c - 1 \ 0" by simp + from less_1 have neq_1: "c \ 1" by auto + then have neq_0: "c - 1 \ 0" by simp from less_1 have lim_0: "(\n. c^n) \ 0" by (rule LIMSEQ_power_zero) - hence "(\n. c ^ n / (c - 1) - 1 / (c - 1)) \ 0 / (c - 1) - 1 / (c - 1)" + then have "(\n. c ^ n / (c - 1) - 1 / (c - 1)) \ 0 / (c - 1) - 1 / (c - 1)" using neq_0 by (intro tendsto_intros) - hence "(\n. (c ^ n - 1) / (c - 1)) \ 1 / (1 - c)" + then have "(\n. (c ^ n - 1) / (c - 1)) \ 1 / (1 - c)" by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) - thus "(\n. c ^ n) sums (1 / (1 - c))" + then show "(\n. c ^ n) sums (1 / (1 - c))" by (simp add: sums_def geometric_sum neq_1) qed @@ -576,88 +594,106 @@ lemma summable_geometric_iff: "summable (\n. c ^ n) \ norm c < 1" proof assume "summable (\n. c ^ n :: 'a :: real_normed_field)" - hence "(\n. norm c ^ n) \ 0" + 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) - thus "norm c < 1" using one_le_power[of "norm c" n] by (cases "norm c \ 1") (linarith, simp) + 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 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 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) - thus ?thesis using sums_divide [OF 2, of 2] - by simp + then show ?thesis + using sums_divide [OF 2, of 2] by simp qed subsection \Telescoping\ lemma telescope_sums: - assumes "f \ (c :: 'a :: real_normed_vector)" - shows "(\n. f (Suc n) - f n) sums (c - f 0)" + 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 LIMSEQ_Suc_iff [symmetric]) have "(\n. \kn. f (Suc n) - f 0)" by (simp add: lessThan_Suc_atMost atLeast0AtMost [symmetric] setsum_Suc_diff) - also have "\ \ c - f 0" by (intro tendsto_diff LIMSEQ_Suc[OF assms] tendsto_const) + 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': - assumes "f \ (c :: 'a :: real_normed_vector)" - shows "(\n. f n - f (Suc n)) sums (f 0 - c)" + 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: - assumes "f \ (c :: 'a :: real_normed_vector)" - shows "summable (\n. f (Suc n) - f n)" + 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': - assumes "f \ (c :: 'a :: real_normed_vector)" - shows "summable (\n. f n - f (Suc n))" + 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)\ +text \Cauchy-type criterion for convergence of series (c.f. Harrison).\ -lemma summable_Cauchy: - fixes f :: "nat \ 'a::banach" - shows "summable f \ (\e>0. \N. \m\N. \n. norm (setsum f {m.. (\e>0. \N. \m\N. \n. norm (setsum f {m.. 'a::banach" + apply (simp only: summable_iff_convergent Cauchy_convergent_iff [symmetric] Cauchy_iff) + apply safe + apply (drule spec) + apply (drule (1) mp) + apply (erule exE) + apply (rule_tac x="M" in exI) + apply clarify + apply (rule_tac x="m" and y="n" in linorder_le_cases) + apply (frule (1) order_trans) + apply (drule_tac x="n" in spec) + apply (drule (1) mp) + apply (drule_tac x="m" in spec) + apply (drule (1) mp) + apply (simp_all add: setsum_diff [symmetric]) + apply (drule spec) + apply (drule (1) mp) + apply (erule exE) + apply (rule_tac x="N" in exI) + apply clarify apply (rule_tac x="m" and y="n" in linorder_le_cases) - apply (frule (1) order_trans) - apply (drule_tac x="n" in spec, drule (1) mp) - apply (drule_tac x="m" in spec, drule (1) mp) - apply (simp_all add: setsum_diff [symmetric]) - apply (drule spec, drule (1) mp) - apply (erule exE, rule_tac x="N" in exI, clarify) - apply (rule_tac x="m" and y="n" in linorder_le_cases) - apply (subst norm_minus_commute) - apply (simp_all add: setsum_diff [symmetric]) + apply (subst norm_minus_commute) + apply (simp_all add: setsum_diff [symmetric]) done context fixes f :: "nat \ 'a::banach" begin -text\Absolute convergence imples normal convergence\ +text \Absolute convergence imples normal convergence.\ lemma summable_norm_cancel: "summable (\n. norm (f n)) \ summable f" - apply (simp only: summable_Cauchy, safe) - apply (drule_tac x="e" in spec, safe) - apply (rule_tac x="N" in exI, safe) - apply (drule_tac x="m" in spec, safe) + apply (simp only: summable_Cauchy) + apply safe + apply (drule_tac x="e" in spec) + apply safe + apply (rule_tac x="N" in exI) + apply safe + apply (drule_tac x="m" in spec) + apply safe apply (rule order_le_less_trans [OF norm_setsum]) apply (rule order_le_less_trans [OF abs_ge_self]) apply simp @@ -666,99 +702,117 @@ 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_setsum) -text \Comparison tests\ +text \Comparison tests.\ lemma summable_comparison_test: "\N. \n\N. norm (f n) \ g n \ summable g \ summable f" - apply (simp add: summable_Cauchy, safe) - apply (drule_tac x="e" in spec, safe) - apply (rule_tac x = "N + Na" in exI, safe) + apply (simp add: summable_Cauchy) + apply safe + apply (drule_tac x="e" in spec) + apply safe + apply (rule_tac x = "N + Na" in exI) + apply safe apply (rotate_tac 2) apply (drule_tac x = m in spec) - apply (auto, rotate_tac 2, drule_tac x = n in spec) + apply auto + apply (rotate_tac 2) + apply (drule_tac x = n in spec) apply (rule_tac y = "\k=m..n. norm (f n) \ g n) sequentially \ summable g \ summable f" + "eventually (\n. norm (f n) \ g n) sequentially \ summable g \ summable f" by (rule summable_comparison_test) (auto simp: eventually_at_top_linorder) -(*A better argument order*) -lemma summable_comparison_test': "summable g \ (\n. n \ N \ norm(f n) \ g n) \ summable f" +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 - assume "0 < c" +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" + 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) - moreover have "norm (f (Suc m)) / c ^ Suc m * c ^ n \ norm (f m) / c ^ m * c ^ n" + 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) - ultimately show ?case by simp - qed (insert \0 < c\, simp) + 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 - assume c: "\ 0 < c" - { fix n assume "n \ N" - then have "norm (f (Suc n)) \ c * norm (f n)" - by fact + 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 c by (simp add: not_less mult_nonpos_nonneg) - finally have "f (Suc n) = 0" - by auto } + 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.\ + +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)" + 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) -next - fix n - show "norm (norm (a n) * r ^ n) \ M * (r / r0) ^ n" + show "norm (norm (a n) * r ^ n) \ M * (r / r0) ^ n" for n using r r0 M [of n] apply (auto simp add: abs_mult field_simps) - apply (cases "r=0", simp) - apply (cases n, auto) + apply (cases "r = 0") + apply simp + apply (cases n) + apply auto done qed -text\Summability of geometric series for real algebras\ +text \Summability of geometric series for real algebras.\ lemma complete_algebra_summable_geometric: fixes x :: "'a::{real_normed_algebra_1,banach}" - shows "norm x < 1 \ summable (\n. x ^ n)" + 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) - show "norm x < 1 \ summable (\n. norm x ^ n)" + from assms show "summable (\n. norm x ^ n)" by (simp add: summable_geometric) qed + subsection \Cauchy Product Formula\ text \ @@ -769,7 +823,7 @@ lemma Cauchy_product_sums: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes a: "summable (\k. norm (a k))" - assumes b: "summable (\k. norm (b 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. {.. {..x. 0 \ ?f x" by (auto) - hence norm_setsum_f: "\A. norm (setsum ?f A) = setsum ?f A" + have f_nonneg: "\x. 0 \ ?f x" by auto + then have norm_setsum_f: "\A. norm (setsum ?f A) = setsum ?f A" unfolding real_norm_def by (simp only: abs_of_nonneg setsum_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]) - hence 1: "(\n. setsum ?g (?S1 n)) \ (\k. a k) * (\k. b k)" + then have 1: "(\n. setsum ?g (?S1 n)) \ (\k. a k) * (\k. b k)" by (simp only: setsum_product setsum.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) - hence "(\n. setsum ?f (?S1 n)) \ (\k. norm (a k)) * (\k. norm (b k))" + then have "(\n. setsum ?f (?S1 n)) \ (\k. norm (a k)) * (\k. norm (b k))" by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan) - hence "convergent (\n. setsum ?f (?S1 n))" + then have "convergent (\n. setsum ?f (?S1 n))" by (rule convergentI) - hence Cauchy: "Cauchy (\n. setsum ?f (?S1 n))" + then have Cauchy: "Cauchy (\n. setsum ?f (?S1 n))" by (rule convergent_Cauchy) have "Zfun (\n. setsum ?f (?S1 n - ?S2 n)) sequentially" proof (rule ZfunI, simp only: eventually_sequentially norm_setsum_f) fix r :: real assume r: "0 < r" from CauchyD [OF Cauchy r] obtain N - where "\m\N. \n\N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" .. - hence "\m n. \N \ n; n \ m\ \ norm (setsum ?f (?S1 m - ?S1 n)) < r" + where "\m\N. \n\N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" .. + then have "\m n. N \ n \ n \ m \ norm (setsum ?f (?S1 m - ?S1 n)) < r" by (simp only: setsum_diff finite_S1 S1_mono) - hence N: "\m n. \N \ n; n \ m\ \ setsum ?f (?S1 m - ?S1 n) < r" + then have N: "\m n. N \ n \ n \ m \ setsum ?f (?S1 m - ?S1 n) < r" by (simp only: norm_setsum_f) show "\N. \n\N. setsum ?f (?S1 n - ?S2 n) < r" proof (intro exI allI impI) - fix n assume "2 * N \ n" - hence n: "N \ n div 2" by simp + fix n + assume "2 * N \ n" + then have n: "N \ n div 2" by simp have "setsum ?f (?S1 n - ?S2 n) \ setsum ?f (?S1 n - ?S1 (n div 2))" - by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg - Diff_mono subset_refl S1_le_S2) + by (intro setsum_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 "setsum ?f (?S1 n - ?S2 n) < r" . qed qed - hence "Zfun (\n. setsum ?g (?S1 n - ?S2 n)) sequentially" + then have "Zfun (\n. setsum ?g (?S1 n - ?S2 n)) sequentially" apply (rule Zfun_le [rule_format]) apply (simp only: norm_setsum_f) apply (rule order_trans [OF norm_setsum setsum_mono]) apply (auto simp add: norm_mult_ineq) done - hence 2: "(\n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) \ 0" + then have 2: "(\n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) \ 0" unfolding tendsto_Zfun_iff diff_0_right by (simp only: setsum_diff finite_S1 S2_le_S1) - with 1 have "(\n. setsum ?g (?S2 n)) \ (\k. a k) * (\k. b k)" by (rule Lim_transform2) - thus ?thesis by (simp only: sums_def setsum_triangle_reindex) + then show ?thesis + by (simp only: sums_def setsum_triangle_reindex) qed lemma Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" - assumes a: "summable (\k. norm (a k))" - assumes b: "summable (\k. norm (b k))" + 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 a b - by (rule Cauchy_product_sums [THEN sums_unique]) + using assms by (rule Cauchy_product_sums [THEN sums_unique]) lemma summable_Cauchy_product: - assumes "summable (\k. norm (a k :: 'a :: {real_normed_algebra,banach}))" - "summable (\k. norm (b k))" - shows "summable (\k. \i\k. a i * b (k - i))" + 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))" +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 :: real\)" +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 :: real\) \ summable f" +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 :: real\) \ \suminf f\ \ (\n. \f n\)" +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})" +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) + 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})" +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) @@ -882,33 +942,37 @@ 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" + 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 "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) + 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)" +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 - thus ?thesis by (subst A) simp_all + 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)" + "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 - assume "summable (\n. f (Suc n) * z ^ n)" - from summable_mult2[OF this, of z] show "summable (\n. f (Suc n) * z ^ Suc n)" + show ?rhs if ?lhs + using summable_mult2[OF that, of z] by (simp add: power_commutes algebra_simps) - next - assume "summable (\n. f (Suc n) * z ^ Suc n)" - from summable_mult2[OF this, of "inverse z"] show "summable (\n. f (Suc n) * z ^ n)" + 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) @@ -916,120 +980,133 @@ qed lemma powser_split_head: - assumes "summable (\n. f n * z ^ n :: 'a :: {real_normed_div_algebra,banach})" - shows "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" - "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" - "summable (\n. f (Suc n) * z ^ n)" + 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 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 - thus "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" by simp + 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" - assumes summable: "summable f" and e: "e > (0::real)" + 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 ((\kk N" - have "norm (\k=m..n. f k) < e" - proof (cases "n \ m") - assume n: "n \ m" - with m have "norm ((\kkkkk=m..n. f k)" - by (subst setsum_diff [symmetric]) (simp_all add: setsum_last_plus) - finally show ?thesis . - qed (insert e, simp_all) - } - thus ?thesis by (rule that) + from CauchyD [OF this e] obtain N + where N: "\m 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 setsum_diff [symmetric]) (simp_all add: setsum_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" + "(\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 - thus ?thesis by (simp add: sums_single) + 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 o g)" - and suminf_reindex_mono: "suminf (f o g) \ suminf f" - and suminf_reindex: "(\x. x \ range g \ f x = 0) \ suminf (f \ g) = suminf f" + 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 + 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 + by (metis Max_ge finite_imageI finite_lessThan not_le not_less_eq) + then obtain m where n: "\n'. n' < n \ g n' < m" + by blast have "(\i \ (\i \ suminf f" - using \summable f\ - by (rule setsum_le_suminf) (simp add: pos) - finally show "(\i g) i) \ suminf f" by simp + using \summable f\ by (rule setsum_le_suminf) (simp 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) - hence "(f \ g) sums L" by (simp add: sums_def) - thus "summable (f o g)" by (auto simp add: sums_iff) + then have "(f \ g) sums L" + by (simp add: sums_def) + then show "summable (f \ g)" + by (auto simp add: sums_iff) - hence "(\n. \i g) i) \ suminf (f \ g)" - by(rule summable_LIMSEQ) - thus le: "suminf (f \ g) \ suminf f" + 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) + proof (rule suminf_le_const) fix n have "\ n' \ (g -` {..n'. g n' < n \ n' < m" by blast - + then obtain m where n: "\n'. g n' < n \ n' < m" + by blast have "(\ii\{.. range g. f i)" using f by(auto intro: setsum.mono_neutral_cong_right) also have "\ = (\i\g -` {.. g) i)" - by(rule setsum.reindex_cong[where l=g])(auto) + by (rule setsum.reindex_cong[where l=g])(auto) also have "\ \ (\i g) i)" - by(rule setsum_mono3)(auto simp add: pos n) + by (rule setsum_mono3)(auto simp add: pos n) also have "\ \ suminf (f \ g)" - using \summable (f o g)\ - by(rule setsum_le_suminf)(simp add: pos) + using \summable (f \ g)\ by (rule setsum_le_suminf) (simp add: pos) finally show "setsum f {.. suminf (f \ g)" . qed - with le show "suminf (f \ g) = suminf f" by(rule antisym) + with le show "suminf (f \ g) = suminf f" + by (rule antisym) qed lemma sums_mono_reindex: - assumes subseq: "subseq g" and zero: "\n. n \ range g \ f n = 0" - shows "(\n. f (g n)) sums c \ f sums c" -unfolding sums_def + assumes subseq: "subseq 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. \k = (\kkk \ c" unfolding o_def . + also from LIMSEQ_subseq_LIMSEQ[OF lim subseq] have "\ \ 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) - hence 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 + 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]) (insert k l, simp_all) - with subseq have "l < g_inv n" by (simp add: subseq_strict_mono strict_mono_less) - with k l show False by simp + 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: subseq_strict_mono strict_mono_less) + with k l show False + by simp qed - hence "f k = 0" by (rule zero) + then have "f k = 0" + by (rule zero) } with g_inv_least' g_inv have "(\kk\g`{.. = (\k = (\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 subseq_strict_mono) + fix K n :: nat + assume "g K \ 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 subseq_strict_mono) } - hence "filterlim g_inv at_top sequentially" + then have "filterlim g_inv at_top sequentially" by (auto simp: filterlim_at_top eventually_at_top_linorder) - from lim and this have "(\n. \k c" by (rule filterlim_compose) + with lim have "(\n. \k c" + by (rule filterlim_compose) finally show "(\n. \k c" . qed lemma summable_mono_reindex: - assumes subseq: "subseq g" and zero: "\n. n \ range g \ f n = 0" - shows "summable (\n. f (g n)) \ summable f" + assumes subseq: "subseq 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: - assumes "subseq g" "\n. n \ range g \ f n = (0 :: 'a :: {t2_space,comm_monoid_add})" + fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" + assumes "subseq 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 - hence "\(\c. f sums c)" unfolding summable_def by blast - hence "suminf f = The (\_. False)" by (simp add: suminf_def) - moreover from False have "\summable (\n. f (g n))" + 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 - hence "\(\c. (\n. f (g n)) sums c)" unfolding summable_def by blast - hence "suminf (\n. f (g n)) = The (\_. False)" by (simp add: suminf_def) + 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 (insert sums_mono_reindex[of g f, OF assms] summable_mono_reindex[of g f, OF assms], - simp_all add: sums_iff) +qed end