paulson@10751: (* Title : Series.thy paulson@10751: Author : Jacques D. Fleuriot paulson@10751: Copyright : 1998 University of Cambridge paulson@14416: paulson@14416: Converted to Isar and polished by lcp nipkow@15539: Converted to setsum and polished yet more by TNN avigad@16819: Additional contributions by Jeremy Avigad hoelzl@41970: *) paulson@10751: hoelzl@56213: header {* Infinite Series *} paulson@10751: nipkow@15131: theory Series hoelzl@51528: imports Limits nipkow@15131: begin nipkow@15561: hoelzl@56213: subsection {* Definition of infinite summability *} hoelzl@56213: hoelzl@56193: definition hoelzl@56193: sums :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a \ bool" hoelzl@56193: (infixr "sums" 80) hoelzl@56193: where hoelzl@56193: "f sums s \ (\n. \i s" paulson@14416: hoelzl@56193: definition summable :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ bool" where hoelzl@56193: "summable f \ (\s. f sums s)" hoelzl@56193: hoelzl@56193: definition hoelzl@56193: suminf :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a" hoelzl@56193: (binder "\" 10) hoelzl@56193: where hoelzl@56193: "suminf f = (THE s. f sums s)" hoelzl@56193: hoelzl@56213: subsection {* Infinite summability on topological monoids *} hoelzl@56213: hoelzl@56193: lemma sums_subst[trans]: "f = g \ g sums z \ f sums z" hoelzl@56193: by simp hoelzl@56193: hoelzl@56193: lemma sums_summable: "f sums l \ summable f" hoelzl@41970: by (simp add: sums_def summable_def, blast) paulson@14416: hoelzl@56193: lemma summable_iff_convergent: "summable f \ convergent (\n. \in. \in. 0) sums 0" hoelzl@56213: unfolding sums_def by (simp add: tendsto_const) hoelzl@56213: hoelzl@56213: lemma summable_zero[simp, intro]: "summable (\n. 0)" hoelzl@56213: by (rule sums_zero [THEN sums_summable]) hoelzl@56213: hoelzl@56213: lemma sums_group: "f sums s \ 0 < k \ (\n. setsum f {n * k ..< n * k + k}) sums s" hoelzl@56213: apply (simp only: sums_def setsum_nat_group tendsto_def eventually_sequentially) hoelzl@56213: apply safe hoelzl@56213: apply (erule_tac x=S in allE) hoelzl@56213: apply safe hoelzl@56213: apply (rule_tac x="N" in exI, safe) hoelzl@56213: apply (drule_tac x="n*k" in spec) hoelzl@56213: apply (erule mp) hoelzl@56213: apply (erule order_trans) hoelzl@56213: apply simp hoelzl@56213: done hoelzl@56213: hoelzl@47761: lemma sums_finite: hoelzl@56193: assumes [simp]: "finite N" and f: "\n. n \ N \ f n = 0" hoelzl@47761: shows "f sums (\n\N. f n)" hoelzl@47761: proof - hoelzl@47761: { fix n hoelzl@47761: have "setsum f {..x. 0)" by auto hoelzl@47761: then show ?thesis by simp hoelzl@47761: next hoelzl@47761: assume [simp]: "N \ {}" hoelzl@47761: show ?thesis haftmann@57418: proof (safe intro!: setsum.mono_neutral_right f) hoelzl@47761: fix i assume "i \ N" hoelzl@47761: then have "i \ Max N" by simp hoelzl@47761: then show "i < n + Suc (Max N)" by simp hoelzl@47761: qed hoelzl@47761: qed } hoelzl@47761: note eq = this hoelzl@47761: show ?thesis unfolding sums_def hoelzl@47761: by (rule LIMSEQ_offset[of _ "Suc (Max N)"]) hoelzl@47761: (simp add: eq atLeast0LessThan tendsto_const del: add_Suc_right) hoelzl@47761: qed hoelzl@47761: hoelzl@56213: lemma summable_finite: "finite N \ (\n. n \ N \ f n = 0) \ summable f" hoelzl@56213: by (rule sums_summable) (rule sums_finite) hoelzl@56213: hoelzl@56193: lemma sums_If_finite_set: "finite A \ (\r. if r \ A then f r else 0) sums (\r\A. f r)" hoelzl@47761: using sums_finite[of A "(\r. if r \ A then f r else 0)"] by simp hoelzl@47761: hoelzl@56213: lemma summable_If_finite_set[simp, intro]: "finite A \ summable (\r. if r \ A then f r else 0)" hoelzl@56213: by (rule sums_summable) (rule sums_If_finite_set) hoelzl@56213: hoelzl@56193: lemma sums_If_finite: "finite {r. P r} \ (\r. if P r then f r else 0) sums (\r | P r. f r)" hoelzl@56193: using sums_If_finite_set[of "{r. P r}"] by simp avigad@16819: hoelzl@56213: lemma summable_If_finite[simp, intro]: "finite {r. P r} \ summable (\r. if P r then f r else 0)" hoelzl@56213: by (rule sums_summable) (rule sums_If_finite) hoelzl@56213: hoelzl@56193: lemma sums_single: "(\r. if r = i then f r else 0) sums f i" hoelzl@56193: using sums_If_finite[of "\r. r = i"] by simp hoelzl@29803: hoelzl@56213: lemma summable_single[simp, intro]: "summable (\r. if r = i then f r else 0)" hoelzl@56213: by (rule sums_summable) (rule sums_single) hoelzl@56193: hoelzl@56193: context hoelzl@56193: fixes f :: "nat \ 'a::{t2_space, comm_monoid_add}" hoelzl@56193: begin hoelzl@56193: hoelzl@56193: lemma summable_sums[intro]: "summable f \ f sums (suminf f)" hoelzl@56193: by (simp add: summable_def sums_def suminf_def) hoelzl@56193: (metis convergent_LIMSEQ_iff convergent_def lim_def) hoelzl@56193: hoelzl@56193: lemma summable_LIMSEQ: "summable f \ (\n. \i suminf f" hoelzl@56193: by (rule summable_sums [unfolded sums_def]) hoelzl@56193: hoelzl@56193: lemma sums_unique: "f sums s \ s = suminf f" hoelzl@56193: by (metis limI suminf_eq_lim sums_def) hoelzl@56193: hoelzl@56193: lemma sums_iff: "f sums x \ summable f \ (suminf f = x)" hoelzl@56193: by (metis summable_sums sums_summable sums_unique) hoelzl@56193: hoelzl@56193: lemma suminf_finite: hoelzl@56193: assumes N: "finite N" and f: "\n. n \ N \ f n = 0" hoelzl@56193: shows "suminf f = (\n\N. f n)" hoelzl@56193: using sums_finite[OF assms, THEN sums_unique] by simp hoelzl@56193: hoelzl@56193: end avigad@16819: hoelzl@41970: lemma suminf_zero[simp]: "suminf (\n. 0::'a::{t2_space, comm_monoid_add}) = 0" hoelzl@56193: by (rule sums_zero [THEN sums_unique, symmetric]) avigad@16819: hoelzl@56213: hoelzl@56213: subsection {* Infinite summability on ordered, topological monoids *} hoelzl@56213: hoelzl@56213: lemma sums_le: hoelzl@56213: fixes f g :: "nat \ 'a::{ordered_comm_monoid_add, linorder_topology}" hoelzl@56213: shows "\n. f n \ g n \ f sums s \ g sums t \ s \ t" hoelzl@56213: by (rule LIMSEQ_le) (auto intro: setsum_mono simp: sums_def) hoelzl@56213: hoelzl@56193: context hoelzl@56193: fixes f :: "nat \ 'a::{ordered_comm_monoid_add, linorder_topology}" hoelzl@56193: begin paulson@14416: hoelzl@56213: lemma suminf_le: "\\n. f n \ g n; summable f; summable g\ \ suminf f \ suminf g" hoelzl@56213: by (auto dest: sums_summable intro: sums_le) hoelzl@56213: hoelzl@56213: lemma setsum_le_suminf: "summable f \ \m\n. 0 \ f m \ setsum f {.. suminf f" hoelzl@56213: by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto hoelzl@56213: hoelzl@56213: lemma suminf_nonneg: "summable f \ \n. 0 \ f n \ 0 \ suminf f" hoelzl@56213: using setsum_le_suminf[of 0] by simp hoelzl@56213: hoelzl@56213: lemma setsum_less_suminf2: "summable f \ \m\n. 0 \ f m \ n \ i \ 0 < f i \ setsum f {.. \m\n. 0 < f m \ setsum f {.. \n. 0 \ f n \ 0 < f i \ 0 < suminf f" hoelzl@56213: using setsum_less_suminf2[of 0 i] by simp hoelzl@56213: hoelzl@56213: lemma suminf_pos: "summable f \ \n. 0 < f n \ 0 < suminf f" hoelzl@56213: using suminf_pos2[of 0] by (simp add: less_imp_le) hoelzl@56213: hoelzl@56213: lemma suminf_le_const: "summable f \ (\n. setsum f {.. x) \ suminf f \ x" hoelzl@56213: by (metis LIMSEQ_le_const2 summable_LIMSEQ) paulson@14416: hoelzl@56193: lemma suminf_eq_zero_iff: "summable f \ \n. 0 \ f n \ suminf f = 0 \ (\n. f n = 0)" hoelzl@50999: proof hoelzl@50999: assume "summable f" "suminf f = 0" and pos: "\n. 0 \ f n" hoelzl@50999: then have f: "(\n. \i 0" hoelzl@56213: using summable_LIMSEQ[of f] by simp hoelzl@56213: then have "\i. (\n\{i}. f n) \ 0" hoelzl@56213: proof (rule LIMSEQ_le_const) hoelzl@50999: fix i show "\N. \n\N. (\n\{i}. f n) \ setsum f {..n. f n = 0" hoelzl@50999: by (auto intro!: antisym) hoelzl@56193: qed (metis suminf_zero fun_eq_iff) hoelzl@56193: hoelzl@56213: lemma suminf_pos_iff: "summable f \ \n. 0 \ f n \ 0 < suminf f \ (\i. 0 < f i)" hoelzl@56213: using setsum_le_suminf[of 0] suminf_eq_zero_iff by (simp add: less_le) hoelzl@56193: hoelzl@56193: end hoelzl@56193: hoelzl@56213: lemma summableI_nonneg_bounded: hoelzl@56213: fixes f:: "nat \ 'a::{ordered_comm_monoid_add, linorder_topology, conditionally_complete_linorder}" hoelzl@56213: assumes pos[simp]: "\n. 0 \ f n" and le: "\n. (\i x" hoelzl@56213: shows "summable f" hoelzl@56213: unfolding summable_def sums_def[abs_def] hoelzl@56213: proof (intro exI order_tendstoI) hoelzl@56213: have [simp, intro]: "bdd_above (range (\n. \iiim. n \ m \ a < (\in. a < (\iin. (\i (SUP n. \in. (\i 'a::real_normed_vector" hoelzl@56193: shows "(\n. f (Suc n)) sums s \ f sums (s + f 0)" hoelzl@56193: proof - hoelzl@56193: have "f sums (s + f 0) \ (\i. \j s + f 0" hoelzl@56193: by (subst LIMSEQ_Suc_iff) (simp add: sums_def) hoelzl@56193: also have "\ \ (\i. (\j s + f 0" haftmann@57418: by (simp add: ac_simps setsum.reindex image_iff lessThan_Suc_eq_insert_0) hoelzl@56193: also have "\ \ (\n. f (Suc n)) sums s" hoelzl@56193: proof hoelzl@56193: assume "(\i. (\j s + f 0" hoelzl@56193: with tendsto_add[OF this tendsto_const, of "- f 0"] hoelzl@56193: show "(\i. f (Suc i)) sums s" hoelzl@56193: by (simp add: sums_def) hoelzl@56193: qed (auto intro: tendsto_add tendsto_const simp: sums_def) hoelzl@56193: finally show ?thesis .. hoelzl@50999: qed hoelzl@50999: hoelzl@56193: context hoelzl@56193: fixes f :: "nat \ 'a::real_normed_vector" hoelzl@56193: begin hoelzl@56193: hoelzl@56193: lemma sums_add: "f sums a \ g sums b \ (\n. f n + g n) sums (a + b)" haftmann@57418: unfolding sums_def by (simp add: setsum.distrib tendsto_add) hoelzl@56193: hoelzl@56193: lemma summable_add: "summable f \ summable g \ summable (\n. f n + g n)" hoelzl@56193: unfolding summable_def by (auto intro: sums_add) hoelzl@56193: hoelzl@56193: lemma suminf_add: "summable f \ summable g \ suminf f + suminf g = (\n. f n + g n)" hoelzl@56193: by (intro sums_unique sums_add summable_sums) hoelzl@56193: hoelzl@56193: lemma sums_diff: "f sums a \ g sums b \ (\n. f n - g n) sums (a - b)" hoelzl@56193: unfolding sums_def by (simp add: setsum_subtractf tendsto_diff) hoelzl@56193: hoelzl@56193: lemma summable_diff: "summable f \ summable g \ summable (\n. f n - g n)" hoelzl@56193: unfolding summable_def by (auto intro: sums_diff) hoelzl@56193: hoelzl@56193: lemma suminf_diff: "summable f \ summable g \ suminf f - suminf g = (\n. f n - g n)" hoelzl@56193: by (intro sums_unique sums_diff summable_sums) hoelzl@56193: hoelzl@56193: lemma sums_minus: "f sums a \ (\n. - f n) sums (- a)" hoelzl@56193: unfolding sums_def by (simp add: setsum_negf tendsto_minus) hoelzl@56193: hoelzl@56193: lemma summable_minus: "summable f \ summable (\n. - f n)" hoelzl@56193: unfolding summable_def by (auto intro: sums_minus) huffman@20692: hoelzl@56193: lemma suminf_minus: "summable f \ (\n. - f n) = - (\n. f n)" hoelzl@56193: by (intro sums_unique [symmetric] sums_minus summable_sums) hoelzl@56193: hoelzl@56193: lemma sums_Suc: "(\ n. f (Suc n)) sums l \ f sums (l + f 0)" hoelzl@56193: by (simp add: sums_Suc_iff) hoelzl@56193: hoelzl@56193: 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)" hoelzl@56193: by (subst sums_Suc_iff) simp hoelzl@56193: ultimately show ?case hoelzl@56193: by (simp add: ac_simps) hoelzl@56193: qed simp huffman@20692: hoelzl@56193: lemma summable_iff_shift: "summable (\n. f (n + k)) \ summable f" hoelzl@56193: by (metis diff_add_cancel summable_def sums_iff_shift[abs_def]) hoelzl@56193: hoelzl@56193: lemma sums_split_initial_segment: "f sums s \ (\i. f (i + n)) sums (s - (\i summable (\n. f(n + k))" hoelzl@56193: by (simp add: summable_iff_shift) hoelzl@56193: hoelzl@56193: lemma suminf_minus_initial_segment: "summable f \ (\n. f (n + k)) = (\n. f n) - (\i suminf f = (\n. f(n + k)) + (\iN. \n\N. norm (\i. f (i + n)) < r" hoelzl@56193: proof - hoelzl@56193: from LIMSEQ_D[OF summable_LIMSEQ[OF `summable f`] `0 < r`] hoelzl@56193: obtain N :: nat where "\ n \ N. norm (setsum f {.. f ----> 0" hoelzl@56193: apply (drule summable_iff_convergent [THEN iffD1]) hoelzl@56193: apply (drule convergent_Cauchy) hoelzl@56193: apply (simp only: Cauchy_iff LIMSEQ_iff, safe) hoelzl@56193: apply (drule_tac x="r" in spec, safe) hoelzl@56193: apply (rule_tac x="M" in exI, safe) hoelzl@56193: apply (drule_tac x="Suc n" in spec, simp) hoelzl@56193: apply (drule_tac x="n" in spec, simp) hoelzl@56193: done hoelzl@56193: hoelzl@56193: end hoelzl@56193: hoelzl@57025: context hoelzl@57025: fixes f :: "'i \ nat \ 'a::real_normed_vector" and I :: "'i set" hoelzl@57025: begin hoelzl@57025: hoelzl@57025: lemma sums_setsum: "(\i. i \ I \ (f i) sums (x i)) \ (\n. \i\I. f i n) sums (\i\I. x i)" hoelzl@57025: by (induct I rule: infinite_finite_induct) (auto intro!: sums_add) hoelzl@57025: hoelzl@57025: lemma suminf_setsum: "(\i. i \ I \ summable (f i)) \ (\n. \i\I. f i n) = (\i\I. \n. f i n)" hoelzl@57025: using sums_unique[OF sums_setsum, OF summable_sums] by simp hoelzl@57025: hoelzl@57025: lemma summable_setsum: "(\i. i \ I \ summable (f i)) \ summable (\n. \i\I. f i n)" hoelzl@57025: using sums_summable[OF sums_setsum[OF summable_sums]] . hoelzl@57025: hoelzl@57025: end hoelzl@57025: hoelzl@56193: lemma (in bounded_linear) sums: "(\n. X n) sums a \ (\n. f (X n)) sums (f a)" hoelzl@56193: unfolding sums_def by (drule tendsto, simp only: setsum) hoelzl@56193: hoelzl@56193: lemma (in bounded_linear) summable: "summable (\n. X n) \ summable (\n. f (X n))" hoelzl@56193: unfolding summable_def by (auto intro: sums) hoelzl@56193: hoelzl@56193: lemma (in bounded_linear) suminf: "summable (\n. X n) \ f (\n. X n) = (\n. f (X n))" hoelzl@56193: by (intro sums_unique sums summable_sums) hoelzl@56193: hoelzl@56193: lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real] hoelzl@56193: lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real] hoelzl@56193: lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real] hoelzl@56193: hoelzl@57275: lemmas sums_scaleR_left = bounded_linear.sums[OF bounded_linear_scaleR_left] hoelzl@57275: lemmas summable_scaleR_left = bounded_linear.summable[OF bounded_linear_scaleR_left] hoelzl@57275: lemmas suminf_scaleR_left = bounded_linear.suminf[OF bounded_linear_scaleR_left] hoelzl@57275: hoelzl@57275: lemmas sums_scaleR_right = bounded_linear.sums[OF bounded_linear_scaleR_right] hoelzl@57275: lemmas summable_scaleR_right = bounded_linear.summable[OF bounded_linear_scaleR_right] hoelzl@57275: lemmas suminf_scaleR_right = bounded_linear.suminf[OF bounded_linear_scaleR_right] hoelzl@57275: hoelzl@56213: subsection {* Infinite summability on real normed algebras *} hoelzl@56213: hoelzl@56193: context hoelzl@56193: fixes f :: "nat \ 'a::real_normed_algebra" hoelzl@56193: begin hoelzl@56193: hoelzl@56193: lemma sums_mult: "f sums a \ (\n. c * f n) sums (c * a)" hoelzl@56193: by (rule bounded_linear.sums [OF bounded_linear_mult_right]) hoelzl@56193: hoelzl@56193: lemma summable_mult: "summable f \ summable (\n. c * f n)" hoelzl@56193: by (rule bounded_linear.summable [OF bounded_linear_mult_right]) hoelzl@56193: hoelzl@56193: lemma suminf_mult: "summable f \ suminf (\n. c * f n) = c * suminf f" hoelzl@56193: by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric]) hoelzl@56193: hoelzl@56193: lemma sums_mult2: "f sums a \ (\n. f n * c) sums (a * c)" hoelzl@56193: by (rule bounded_linear.sums [OF bounded_linear_mult_left]) hoelzl@56193: hoelzl@56193: lemma summable_mult2: "summable f \ summable (\n. f n * c)" hoelzl@56193: by (rule bounded_linear.summable [OF bounded_linear_mult_left]) hoelzl@56193: hoelzl@56193: lemma suminf_mult2: "summable f \ suminf f * c = (\n. f n * c)" hoelzl@56193: by (rule bounded_linear.suminf [OF bounded_linear_mult_left]) hoelzl@56193: hoelzl@56193: end hoelzl@56193: hoelzl@56213: subsection {* Infinite summability on real normed fields *} hoelzl@56213: hoelzl@56193: context hoelzl@56193: fixes c :: "'a::real_normed_field" hoelzl@56193: begin hoelzl@56193: hoelzl@56193: lemma sums_divide: "f sums a \ (\n. f n / c) sums (a / c)" hoelzl@56193: by (rule bounded_linear.sums [OF bounded_linear_divide]) hoelzl@56193: hoelzl@56193: lemma summable_divide: "summable f \ summable (\n. f n / c)" hoelzl@56193: by (rule bounded_linear.summable [OF bounded_linear_divide]) hoelzl@56193: hoelzl@56193: lemma suminf_divide: "summable f \ suminf (\n. f n / c) = suminf f / c" hoelzl@56193: by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric]) paulson@14416: paulson@15085: text{*Sum of a geometric progression.*} paulson@14416: hoelzl@56193: lemma geometric_sums: "norm c < 1 \ (\n. c^n) sums (1 / (1 - c))" huffman@20692: proof - hoelzl@56193: assume less_1: "norm c < 1" hoelzl@56193: hence neq_1: "c \ 1" by auto hoelzl@56193: hence neq_0: "c - 1 \ 0" by simp hoelzl@56193: from less_1 have lim_0: "(\n. c^n) ----> 0" huffman@20692: by (rule LIMSEQ_power_zero) hoelzl@56193: hence "(\n. c ^ n / (c - 1) - 1 / (c - 1)) ----> 0 / (c - 1) - 1 / (c - 1)" huffman@44568: using neq_0 by (intro tendsto_intros) hoelzl@56193: hence "(\n. (c ^ n - 1) / (c - 1)) ----> 1 / (1 - c)" huffman@20692: by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) hoelzl@56193: thus "(\n. c ^ n) sums (1 / (1 - c))" huffman@20692: by (simp add: sums_def geometric_sum neq_1) huffman@20692: qed huffman@20692: hoelzl@56193: lemma summable_geometric: "norm c < 1 \ summable (\n. c^n)" hoelzl@56193: by (rule geometric_sums [THEN sums_summable]) paulson@14416: hoelzl@56193: lemma suminf_geometric: "norm c < 1 \ suminf (\n. c^n) = 1 / (1 - c)" hoelzl@56193: by (rule sums_unique[symmetric]) (rule geometric_sums) hoelzl@56193: hoelzl@56193: end paulson@33271: paulson@33271: lemma power_half_series: "(\n. (1/2::real)^Suc n) sums 1" paulson@33271: proof - paulson@33271: have 2: "(\n. (1/2::real)^n) sums 2" using geometric_sums [of "1/2::real"] paulson@33271: by auto paulson@33271: have "(\n. (1/2::real)^Suc n) = (\n. (1 / 2) ^ n / 2)" paulson@33271: by simp huffman@44282: thus ?thesis using sums_divide [OF 2, of 2] paulson@33271: by simp paulson@33271: qed paulson@33271: hoelzl@56213: subsection {* Infinite summability on Banach spaces *} hoelzl@56213: paulson@15085: text{*Cauchy-type criterion for convergence of series (c.f. Harrison)*} paulson@15085: hoelzl@56193: lemma summable_Cauchy: hoelzl@56193: fixes f :: "nat \ 'a::banach" hoelzl@56193: shows "summable f \ (\e>0. \N. \m\N. \n. norm (setsum f {m.. 'a::banach" hoelzl@56193: begin hoelzl@56193: hoelzl@56193: text{*Absolute convergence imples normal convergence*} huffman@20689: hoelzl@56194: lemma summable_norm_cancel: "summable (\n. norm (f n)) \ summable f" hoelzl@56193: apply (simp only: summable_Cauchy, safe) hoelzl@56193: apply (drule_tac x="e" in spec, safe) hoelzl@56193: apply (rule_tac x="N" in exI, safe) hoelzl@56193: apply (drule_tac x="m" in spec, safe) hoelzl@56193: apply (rule order_le_less_trans [OF norm_setsum]) hoelzl@56193: apply (rule order_le_less_trans [OF abs_ge_self]) hoelzl@56193: apply simp hoelzl@50999: done paulson@32707: hoelzl@56193: lemma summable_norm: "summable (\n. norm (f n)) \ norm (suminf f) \ (\n. norm (f n))" hoelzl@56193: by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_setsum) hoelzl@56193: hoelzl@56193: text {* Comparison tests *} paulson@14416: hoelzl@56194: lemma summable_comparison_test: "\N. \n\N. norm (f n) \ g n \ summable g \ summable f" hoelzl@56193: apply (simp add: summable_Cauchy, safe) hoelzl@56193: apply (drule_tac x="e" in spec, safe) hoelzl@56193: apply (rule_tac x = "N + Na" in exI, safe) hoelzl@56193: apply (rotate_tac 2) hoelzl@56193: apply (drule_tac x = m in spec) hoelzl@56193: apply (auto, rotate_tac 2, drule_tac x = n in spec) hoelzl@56193: apply (rule_tac y = "\k=m.. (\n. n \ N \ norm(f n) \ g n) \ summable f" hoelzl@56369: by (rule summable_comparison_test) auto lp15@56217: hoelzl@56193: subsection {* The Ratio Test*} paulson@15085: hoelzl@56193: lemma summable_ratio_test: hoelzl@56193: assumes "c < 1" "\n. n \ N \ norm (f (Suc n)) \ c * norm (f n)" hoelzl@56193: shows "summable f" hoelzl@56193: proof cases hoelzl@56193: assume "0 < c" hoelzl@56193: show "summable f" hoelzl@56193: proof (rule summable_comparison_test) hoelzl@56193: show "\N'. \n\N'. norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" hoelzl@56193: proof (intro exI allI impI) hoelzl@56193: fix n assume "N \ n" then show "norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" hoelzl@56193: proof (induct rule: inc_induct) hoelzl@56193: case (step m) hoelzl@56193: moreover have "norm (f (Suc m)) / c ^ Suc m * c ^ n \ norm (f m) / c ^ m * c ^ n" hoelzl@56193: using `0 < c` `c < 1` assms(2)[OF `N \ m`] by (simp add: field_simps) hoelzl@56193: ultimately show ?case by simp hoelzl@56193: qed (insert `0 < c`, simp) hoelzl@56193: qed hoelzl@56193: show "summable (\n. norm (f N) / c ^ N * c ^ n)" hoelzl@56193: using `0 < c` `c < 1` by (intro summable_mult summable_geometric) simp hoelzl@56193: qed hoelzl@56193: next hoelzl@56193: assume c: "\ 0 < c" hoelzl@56193: { fix n assume "n \ N" hoelzl@56193: then have "norm (f (Suc n)) \ c * norm (f n)" hoelzl@56193: by fact hoelzl@56193: also have "\ \ 0" hoelzl@56193: using c by (simp add: not_less mult_nonpos_nonneg) hoelzl@56193: finally have "f (Suc n) = 0" hoelzl@56193: by auto } hoelzl@56193: then show "summable f" hoelzl@56194: by (intro sums_summable[OF sums_finite, of "{.. Suc N}"]) (auto simp: not_le Suc_less_eq2) lp15@56178: qed lp15@56178: hoelzl@56193: end paulson@14416: hoelzl@56369: text{*Relations among convergence and absolute convergence for power series.*} hoelzl@56369: hoelzl@56369: lemma abel_lemma: hoelzl@56369: fixes a :: "nat \ 'a::real_normed_vector" hoelzl@56369: assumes r: "0 \ r" and r0: "r < r0" and M: "\n. norm (a n) * r0^n \ M" hoelzl@56369: shows "summable (\n. norm (a n) * r^n)" hoelzl@56369: proof (rule summable_comparison_test') hoelzl@56369: show "summable (\n. M * (r / r0) ^ n)" hoelzl@56369: using assms hoelzl@56369: by (auto simp add: summable_mult summable_geometric) hoelzl@56369: next hoelzl@56369: fix n hoelzl@56369: show "norm (norm (a n) * r ^ n) \ M * (r / r0) ^ n" hoelzl@56369: using r r0 M [of n] hoelzl@56369: apply (auto simp add: abs_mult field_simps power_divide) hoelzl@56369: apply (cases "r=0", simp) hoelzl@56369: apply (cases n, auto) hoelzl@56369: done hoelzl@56369: qed hoelzl@56369: hoelzl@56369: huffman@23084: text{*Summability of geometric series for real algebras*} huffman@23084: huffman@23084: lemma complete_algebra_summable_geometric: haftmann@31017: fixes x :: "'a::{real_normed_algebra_1,banach}" huffman@23084: shows "norm x < 1 \ summable (\n. x ^ n)" huffman@23084: proof (rule summable_comparison_test) huffman@23084: show "\N. \n\N. norm (x ^ n) \ norm x ^ n" huffman@23084: by (simp add: norm_power_ineq) huffman@23084: show "norm x < 1 \ summable (\n. norm x ^ n)" huffman@23084: by (simp add: summable_geometric) huffman@23084: qed huffman@23084: huffman@23111: subsection {* Cauchy Product Formula *} huffman@23111: wenzelm@54703: text {* wenzelm@54703: Proof based on Analysis WebNotes: Chapter 07, Class 41 wenzelm@54703: @{url "http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm"} wenzelm@54703: *} huffman@23111: huffman@23111: lemma setsum_triangle_reindex: huffman@23111: fixes n :: nat hoelzl@56213: shows "(\(i,j)\{(i,j). i+j < n}. f i j) = (\ki\k. f i (k - i))" haftmann@57418: apply (simp add: setsum.Sigma) hoelzl@57129: apply (rule setsum.reindex_bij_witness[where j="\(i, j). (i+j, i)" and i="\(k, i). (i, k - i)"]) hoelzl@57129: apply auto hoelzl@57129: done huffman@23111: huffman@23111: lemma Cauchy_product_sums: huffman@23111: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" huffman@23111: assumes a: "summable (\k. norm (a k))" huffman@23111: assumes b: "summable (\k. norm (b k))" hoelzl@56213: shows "(\k. \i\k. a i * b (k - i)) sums ((\k. a k) * (\k. b k))" huffman@23111: proof - hoelzl@56193: let ?S1 = "\n::nat. {.. {..m n. m \ n \ ?S1 m \ ?S1 n" by auto huffman@23111: have S2_le_S1: "\n. ?S2 n \ ?S1 n" by auto huffman@23111: have S1_le_S2: "\n. ?S1 (n div 2) \ ?S2 n" by auto huffman@23111: have finite_S1: "\n. finite (?S1 n)" by simp huffman@23111: with S2_le_S1 have finite_S2: "\n. finite (?S2 n)" by (rule finite_subset) huffman@23111: huffman@23111: let ?g = "\(i,j). a i * b j" huffman@23111: let ?f = "\(i,j). norm (a i) * norm (b j)" nipkow@56536: have f_nonneg: "\x. 0 \ ?f x" by (auto) huffman@23111: hence norm_setsum_f: "\A. norm (setsum ?f A) = setsum ?f A" huffman@23111: unfolding real_norm_def huffman@23111: by (simp only: abs_of_nonneg setsum_nonneg [rule_format]) huffman@23111: hoelzl@56193: have "(\n. (\kk (\k. a k) * (\k. b k)" hoelzl@56193: by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b]) huffman@23111: hence 1: "(\n. setsum ?g (?S1 n)) ----> (\k. a k) * (\k. b k)" haftmann@57418: by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan) huffman@23111: hoelzl@56193: have "(\n. (\kk (\k. norm (a k)) * (\k. norm (b k))" hoelzl@56193: using a b by (intro tendsto_mult summable_LIMSEQ) huffman@23111: hence "(\n. setsum ?f (?S1 n)) ----> (\k. norm (a k)) * (\k. norm (b k))" haftmann@57418: by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan) huffman@23111: hence "convergent (\n. setsum ?f (?S1 n))" huffman@23111: by (rule convergentI) huffman@23111: hence Cauchy: "Cauchy (\n. setsum ?f (?S1 n))" huffman@23111: by (rule convergent_Cauchy) huffman@36657: have "Zfun (\n. setsum ?f (?S1 n - ?S2 n)) sequentially" huffman@36657: proof (rule ZfunI, simp only: eventually_sequentially norm_setsum_f) huffman@23111: fix r :: real huffman@23111: assume r: "0 < r" huffman@23111: from CauchyD [OF Cauchy r] obtain N huffman@23111: where "\m\N. \n\N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" .. huffman@23111: hence "\m n. \N \ n; n \ m\ \ norm (setsum ?f (?S1 m - ?S1 n)) < r" huffman@23111: by (simp only: setsum_diff finite_S1 S1_mono) huffman@23111: hence N: "\m n. \N \ n; n \ m\ \ setsum ?f (?S1 m - ?S1 n) < r" huffman@23111: by (simp only: norm_setsum_f) huffman@23111: show "\N. \n\N. setsum ?f (?S1 n - ?S2 n) < r" huffman@23111: proof (intro exI allI impI) huffman@23111: fix n assume "2 * N \ n" huffman@23111: hence n: "N \ n div 2" by simp huffman@23111: have "setsum ?f (?S1 n - ?S2 n) \ setsum ?f (?S1 n - ?S1 (n div 2))" huffman@23111: by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg huffman@23111: Diff_mono subset_refl S1_le_S2) huffman@23111: also have "\ < r" huffman@23111: using n div_le_dividend by (rule N) huffman@23111: finally show "setsum ?f (?S1 n - ?S2 n) < r" . huffman@23111: qed huffman@23111: qed huffman@36657: hence "Zfun (\n. setsum ?g (?S1 n - ?S2 n)) sequentially" huffman@36657: apply (rule Zfun_le [rule_format]) huffman@23111: apply (simp only: norm_setsum_f) huffman@23111: apply (rule order_trans [OF norm_setsum setsum_mono]) huffman@23111: apply (auto simp add: norm_mult_ineq) huffman@23111: done huffman@23111: hence 2: "(\n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) ----> 0" huffman@36660: unfolding tendsto_Zfun_iff diff_0_right huffman@36657: by (simp only: setsum_diff finite_S1 S2_le_S1) huffman@23111: huffman@23111: with 1 have "(\n. setsum ?g (?S2 n)) ----> (\k. a k) * (\k. b k)" huffman@23111: by (rule LIMSEQ_diff_approach_zero2) huffman@23111: thus ?thesis by (simp only: sums_def setsum_triangle_reindex) huffman@23111: qed huffman@23111: huffman@23111: lemma Cauchy_product: huffman@23111: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" huffman@23111: assumes a: "summable (\k. norm (a k))" huffman@23111: assumes b: "summable (\k. norm (b k))" hoelzl@56213: shows "(\k. a k) * (\k. b k) = (\k. \i\k. a i * b (k - i))" hoelzl@56213: using a b hoelzl@56213: by (rule Cauchy_product_sums [THEN sums_unique]) hoelzl@56213: hoelzl@56213: subsection {* Series on @{typ real}s *} hoelzl@56213: hoelzl@56213: lemma summable_norm_comparison_test: "\N. \n\N. norm (f n) \ g n \ summable g \ summable (\n. norm (f n))" hoelzl@56213: by (rule summable_comparison_test) auto hoelzl@56213: hoelzl@56213: lemma summable_rabs_comparison_test: "\\N. \n\N. \f n\ \ g n; summable g\ \ summable (\n. \f n :: real\)" hoelzl@56213: by (rule summable_comparison_test) auto hoelzl@56213: hoelzl@56213: lemma summable_rabs_cancel: "summable (\n. \f n :: real\) \ summable f" hoelzl@56213: by (rule summable_norm_cancel) simp hoelzl@56213: hoelzl@56213: lemma summable_rabs: "summable (\n. \f n :: real\) \ \suminf f\ \ (\n. \f n\)" hoelzl@56213: by (fold real_norm_def) (rule summable_norm) huffman@23111: paulson@14416: end