| author | wenzelm | 
| Mon, 22 Sep 2014 16:15:29 +0200 | |
| changeset 58418 | a04b242a7a01 | 
| parent 57418 | 6ab1c7cb0b8d | 
| child 58729 | e8ecc79aee43 | 
| permissions | -rw-r--r-- | 
| 10751 | 1  | 
(* Title : Series.thy  | 
2  | 
Author : Jacques D. Fleuriot  | 
|
3  | 
Copyright : 1998 University of Cambridge  | 
|
| 14416 | 4  | 
|
5  | 
Converted to Isar and polished by lcp  | 
|
| 15539 | 6  | 
Converted to setsum and polished yet more by TNN  | 
| 16819 | 7  | 
Additional contributions by Jeremy Avigad  | 
| 41970 | 8  | 
*)  | 
| 10751 | 9  | 
|
| 56213 | 10  | 
header {* Infinite Series *}
 | 
| 10751 | 11  | 
|
| 15131 | 12  | 
theory Series  | 
| 
51528
 
66c3a7589de7
Series.thy is based on Limits.thy and not Deriv.thy
 
hoelzl 
parents: 
51526 
diff
changeset
 | 
13  | 
imports Limits  | 
| 15131 | 14  | 
begin  | 
| 15561 | 15  | 
|
| 56213 | 16  | 
subsection {* Definition of infinite summability *}
 | 
17  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
18  | 
definition  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
19  | 
  sums :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
20  | 
(infixr "sums" 80)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
21  | 
where  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
22  | 
"f sums s \<longleftrightarrow> (\<lambda>n. \<Sum>i<n. f i) ----> s"  | 
| 14416 | 23  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
24  | 
definition summable :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> bool" where
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
25  | 
"summable f \<longleftrightarrow> (\<exists>s. f sums s)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
26  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
27  | 
definition  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
28  | 
  suminf :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
29  | 
(binder "\<Sum>" 10)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
30  | 
where  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
31  | 
"suminf f = (THE s. f sums s)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
32  | 
|
| 56213 | 33  | 
subsection {* Infinite summability on topological monoids *}
 | 
34  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
35  | 
lemma sums_subst[trans]: "f = g \<Longrightarrow> g sums z \<Longrightarrow> f sums z"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
36  | 
by simp  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
37  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
38  | 
lemma sums_summable: "f sums l \<Longrightarrow> summable f"  | 
| 41970 | 39  | 
by (simp add: sums_def summable_def, blast)  | 
| 14416 | 40  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
41  | 
lemma summable_iff_convergent: "summable f \<longleftrightarrow> convergent (\<lambda>n. \<Sum>i<n. f i)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
42  | 
by (simp add: summable_def sums_def convergent_def)  | 
| 14416 | 43  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
44  | 
lemma suminf_eq_lim: "suminf f = lim (\<lambda>n. \<Sum>i<n. f i)"  | 
| 41970 | 45  | 
by (simp add: suminf_def sums_def lim_def)  | 
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
46  | 
|
| 56213 | 47  | 
lemma sums_zero[simp, intro]: "(\<lambda>n. 0) sums 0"  | 
48  | 
unfolding sums_def by (simp add: tendsto_const)  | 
|
49  | 
||
50  | 
lemma summable_zero[simp, intro]: "summable (\<lambda>n. 0)"  | 
|
51  | 
by (rule sums_zero [THEN sums_summable])  | 
|
52  | 
||
53  | 
lemma sums_group: "f sums s \<Longrightarrow> 0 < k \<Longrightarrow> (\<lambda>n. setsum f {n * k ..< n * k + k}) sums s"
 | 
|
54  | 
apply (simp only: sums_def setsum_nat_group tendsto_def eventually_sequentially)  | 
|
55  | 
apply safe  | 
|
56  | 
apply (erule_tac x=S in allE)  | 
|
57  | 
apply safe  | 
|
58  | 
apply (rule_tac x="N" in exI, safe)  | 
|
59  | 
apply (drule_tac x="n*k" in spec)  | 
|
60  | 
apply (erule mp)  | 
|
61  | 
apply (erule order_trans)  | 
|
62  | 
apply simp  | 
|
63  | 
done  | 
|
64  | 
||
| 47761 | 65  | 
lemma sums_finite:  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
66  | 
assumes [simp]: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"  | 
| 47761 | 67  | 
shows "f sums (\<Sum>n\<in>N. f n)"  | 
68  | 
proof -  | 
|
69  | 
  { fix n
 | 
|
70  | 
    have "setsum f {..<n + Suc (Max N)} = setsum f N"
 | 
|
71  | 
proof cases  | 
|
72  | 
      assume "N = {}"
 | 
|
73  | 
with f have "f = (\<lambda>x. 0)" by auto  | 
|
74  | 
then show ?thesis by simp  | 
|
75  | 
next  | 
|
76  | 
      assume [simp]: "N \<noteq> {}"
 | 
|
77  | 
show ?thesis  | 
|
| 57418 | 78  | 
proof (safe intro!: setsum.mono_neutral_right f)  | 
| 47761 | 79  | 
fix i assume "i \<in> N"  | 
80  | 
then have "i \<le> Max N" by simp  | 
|
81  | 
then show "i < n + Suc (Max N)" by simp  | 
|
82  | 
qed  | 
|
83  | 
qed }  | 
|
84  | 
note eq = this  | 
|
85  | 
show ?thesis unfolding sums_def  | 
|
86  | 
by (rule LIMSEQ_offset[of _ "Suc (Max N)"])  | 
|
87  | 
(simp add: eq atLeast0LessThan tendsto_const del: add_Suc_right)  | 
|
88  | 
qed  | 
|
89  | 
||
| 56213 | 90  | 
lemma summable_finite: "finite N \<Longrightarrow> (\<And>n. n \<notin> N \<Longrightarrow> f n = 0) \<Longrightarrow> summable f"  | 
91  | 
by (rule sums_summable) (rule sums_finite)  | 
|
92  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
93  | 
lemma sums_If_finite_set: "finite A \<Longrightarrow> (\<lambda>r. if r \<in> A then f r else 0) sums (\<Sum>r\<in>A. f r)"  | 
| 47761 | 94  | 
using sums_finite[of A "(\<lambda>r. if r \<in> A then f r else 0)"] by simp  | 
95  | 
||
| 56213 | 96  | 
lemma summable_If_finite_set[simp, intro]: "finite A \<Longrightarrow> summable (\<lambda>r. if r \<in> A then f r else 0)"  | 
97  | 
by (rule sums_summable) (rule sums_If_finite_set)  | 
|
98  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
99  | 
lemma sums_If_finite: "finite {r. P r} \<Longrightarrow> (\<lambda>r. if P r then f r else 0) sums (\<Sum>r | P r. f r)"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
100  | 
  using sums_If_finite_set[of "{r. P r}"] by simp
 | 
| 16819 | 101  | 
|
| 56213 | 102  | 
lemma summable_If_finite[simp, intro]: "finite {r. P r} \<Longrightarrow> summable (\<lambda>r. if P r then f r else 0)"
 | 
103  | 
by (rule sums_summable) (rule sums_If_finite)  | 
|
104  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
105  | 
lemma sums_single: "(\<lambda>r. if r = i then f r else 0) sums f i"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
106  | 
using sums_If_finite[of "\<lambda>r. r = i"] by simp  | 
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
107  | 
|
| 56213 | 108  | 
lemma summable_single[simp, intro]: "summable (\<lambda>r. if r = i then f r else 0)"  | 
109  | 
by (rule sums_summable) (rule sums_single)  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
110  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
111  | 
context  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
112  | 
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
113  | 
begin  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
114  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
115  | 
lemma summable_sums[intro]: "summable f \<Longrightarrow> f sums (suminf f)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
116  | 
by (simp add: summable_def sums_def suminf_def)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
117  | 
(metis convergent_LIMSEQ_iff convergent_def lim_def)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
118  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
119  | 
lemma summable_LIMSEQ: "summable f \<Longrightarrow> (\<lambda>n. \<Sum>i<n. f i) ----> suminf f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
120  | 
by (rule summable_sums [unfolded sums_def])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
121  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
122  | 
lemma sums_unique: "f sums s \<Longrightarrow> s = suminf f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
123  | 
by (metis limI suminf_eq_lim sums_def)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
124  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
125  | 
lemma sums_iff: "f sums x \<longleftrightarrow> summable f \<and> (suminf f = x)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
126  | 
by (metis summable_sums sums_summable sums_unique)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
127  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
128  | 
lemma suminf_finite:  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
129  | 
assumes N: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
130  | 
shows "suminf f = (\<Sum>n\<in>N. f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
131  | 
using sums_finite[OF assms, THEN sums_unique] by simp  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
132  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
133  | 
end  | 
| 16819 | 134  | 
|
| 41970 | 135  | 
lemma suminf_zero[simp]: "suminf (\<lambda>n. 0::'a::{t2_space, comm_monoid_add}) = 0"
 | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
136  | 
by (rule sums_zero [THEN sums_unique, symmetric])  | 
| 16819 | 137  | 
|
| 56213 | 138  | 
|
139  | 
subsection {* Infinite summability on ordered, topological monoids *}
 | 
|
140  | 
||
141  | 
lemma sums_le:  | 
|
142  | 
  fixes f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
|
143  | 
shows "\<forall>n. f n \<le> g n \<Longrightarrow> f sums s \<Longrightarrow> g sums t \<Longrightarrow> s \<le> t"  | 
|
144  | 
by (rule LIMSEQ_le) (auto intro: setsum_mono simp: sums_def)  | 
|
145  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
146  | 
context  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
147  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
148  | 
begin  | 
| 14416 | 149  | 
|
| 56213 | 150  | 
lemma suminf_le: "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"  | 
151  | 
by (auto dest: sums_summable intro: sums_le)  | 
|
152  | 
||
153  | 
lemma setsum_le_suminf: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> setsum f {..<n} \<le> suminf f"
 | 
|
154  | 
by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto  | 
|
155  | 
||
156  | 
lemma suminf_nonneg: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 \<le> suminf f"  | 
|
157  | 
using setsum_le_suminf[of 0] by simp  | 
|
158  | 
||
159  | 
lemma setsum_less_suminf2: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> n \<le> i \<Longrightarrow> 0 < f i \<Longrightarrow> setsum f {..<n} < suminf f"
 | 
|
160  | 
using  | 
|
161  | 
setsum_le_suminf[of "Suc i"]  | 
|
162  | 
    add_strict_increasing[of "f i" "setsum f {..<n}" "setsum f {..<i}"]
 | 
|
163  | 
    setsum_mono2[of "{..<i}" "{..<n}" f]
 | 
|
164  | 
by (auto simp: less_imp_le ac_simps)  | 
|
165  | 
||
166  | 
lemma setsum_less_suminf: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 < f m \<Longrightarrow> setsum f {..<n} < suminf f"
 | 
|
167  | 
using setsum_less_suminf2[of n n] by (simp add: less_imp_le)  | 
|
168  | 
||
169  | 
lemma suminf_pos2: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 < f i \<Longrightarrow> 0 < suminf f"  | 
|
170  | 
using setsum_less_suminf2[of 0 i] by simp  | 
|
171  | 
||
172  | 
lemma suminf_pos: "summable f \<Longrightarrow> \<forall>n. 0 < f n \<Longrightarrow> 0 < suminf f"  | 
|
173  | 
using suminf_pos2[of 0] by (simp add: less_imp_le)  | 
|
174  | 
||
175  | 
lemma suminf_le_const: "summable f \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
 | 
|
176  | 
by (metis LIMSEQ_le_const2 summable_LIMSEQ)  | 
|
| 14416 | 177  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
178  | 
lemma suminf_eq_zero_iff: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"  | 
| 50999 | 179  | 
proof  | 
180  | 
assume "summable f" "suminf f = 0" and pos: "\<forall>n. 0 \<le> f n"  | 
|
181  | 
then have f: "(\<lambda>n. \<Sum>i<n. f i) ----> 0"  | 
|
| 56213 | 182  | 
using summable_LIMSEQ[of f] by simp  | 
183  | 
  then have "\<And>i. (\<Sum>n\<in>{i}. f n) \<le> 0"
 | 
|
184  | 
proof (rule LIMSEQ_le_const)  | 
|
| 50999 | 185  | 
    fix i show "\<exists>N. \<forall>n\<ge>N. (\<Sum>n\<in>{i}. f n) \<le> setsum f {..<n}"
 | 
186  | 
using pos by (intro exI[of _ "Suc i"] allI impI setsum_mono2) auto  | 
|
187  | 
qed  | 
|
188  | 
with pos show "\<forall>n. f n = 0"  | 
|
189  | 
by (auto intro!: antisym)  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
190  | 
qed (metis suminf_zero fun_eq_iff)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
191  | 
|
| 56213 | 192  | 
lemma suminf_pos_iff: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"  | 
193  | 
using setsum_le_suminf[of 0] suminf_eq_zero_iff by (simp add: less_le)  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
194  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
195  | 
end  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
196  | 
|
| 56213 | 197  | 
lemma summableI_nonneg_bounded:  | 
198  | 
  fixes f:: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology, conditionally_complete_linorder}"
 | 
|
199  | 
assumes pos[simp]: "\<And>n. 0 \<le> f n" and le: "\<And>n. (\<Sum>i<n. f i) \<le> x"  | 
|
200  | 
shows "summable f"  | 
|
201  | 
unfolding summable_def sums_def[abs_def]  | 
|
202  | 
proof (intro exI order_tendstoI)  | 
|
203  | 
have [simp, intro]: "bdd_above (range (\<lambda>n. \<Sum>i<n. f i))"  | 
|
204  | 
using le by (auto simp: bdd_above_def)  | 
|
205  | 
  { fix a assume "a < (SUP n. \<Sum>i<n. f i)"
 | 
|
206  | 
then obtain n where "a < (\<Sum>i<n. f i)"  | 
|
207  | 
by (auto simp add: less_cSUP_iff)  | 
|
208  | 
then have "\<And>m. n \<le> m \<Longrightarrow> a < (\<Sum>i<m. f i)"  | 
|
209  | 
by (rule less_le_trans) (auto intro!: setsum_mono2)  | 
|
210  | 
then show "eventually (\<lambda>n. a < (\<Sum>i<n. f i)) sequentially"  | 
|
211  | 
by (auto simp: eventually_sequentially) }  | 
|
212  | 
  { fix a assume "(SUP n. \<Sum>i<n. f i) < a"
 | 
|
213  | 
moreover have "\<And>n. (\<Sum>i<n. f i) \<le> (SUP n. \<Sum>i<n. f i)"  | 
|
214  | 
by (auto intro: cSUP_upper)  | 
|
215  | 
ultimately show "eventually (\<lambda>n. (\<Sum>i<n. f i) < a) sequentially"  | 
|
216  | 
by (auto intro: le_less_trans simp: eventually_sequentially) }  | 
|
217  | 
qed  | 
|
218  | 
||
219  | 
subsection {* Infinite summability on real normed vector spaces *}
 | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
220  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
221  | 
lemma sums_Suc_iff:  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
222  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
223  | 
shows "(\<lambda>n. f (Suc n)) sums s \<longleftrightarrow> f sums (s + f 0)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
224  | 
proof -  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
225  | 
have "f sums (s + f 0) \<longleftrightarrow> (\<lambda>i. \<Sum>j<Suc i. f j) ----> s + f 0"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
226  | 
by (subst LIMSEQ_Suc_iff) (simp add: sums_def)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
227  | 
also have "\<dots> \<longleftrightarrow> (\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) ----> s + f 0"  | 
| 57418 | 228  | 
by (simp add: ac_simps setsum.reindex image_iff lessThan_Suc_eq_insert_0)  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
229  | 
also have "\<dots> \<longleftrightarrow> (\<lambda>n. f (Suc n)) sums s"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
230  | 
proof  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
231  | 
assume "(\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) ----> s + f 0"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
232  | 
with tendsto_add[OF this tendsto_const, of "- f 0"]  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
233  | 
show "(\<lambda>i. f (Suc i)) sums s"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
234  | 
by (simp add: sums_def)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
235  | 
qed (auto intro: tendsto_add tendsto_const simp: sums_def)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
236  | 
finally show ?thesis ..  | 
| 50999 | 237  | 
qed  | 
238  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
239  | 
context  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
240  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
241  | 
begin  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
242  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
243  | 
lemma sums_add: "f sums a \<Longrightarrow> g sums b \<Longrightarrow> (\<lambda>n. f n + g n) sums (a + b)"  | 
| 57418 | 244  | 
unfolding sums_def by (simp add: setsum.distrib tendsto_add)  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
245  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
246  | 
lemma summable_add: "summable f \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. f n + g n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
247  | 
unfolding summable_def by (auto intro: sums_add)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
248  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
249  | 
lemma suminf_add: "summable f \<Longrightarrow> summable g \<Longrightarrow> suminf f + suminf g = (\<Sum>n. f n + g n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
250  | 
by (intro sums_unique sums_add summable_sums)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
251  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
252  | 
lemma sums_diff: "f sums a \<Longrightarrow> g sums b \<Longrightarrow> (\<lambda>n. f n - g n) sums (a - b)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
253  | 
unfolding sums_def by (simp add: setsum_subtractf tendsto_diff)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
254  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
255  | 
lemma summable_diff: "summable f \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. f n - g n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
256  | 
unfolding summable_def by (auto intro: sums_diff)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
257  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
258  | 
lemma suminf_diff: "summable f \<Longrightarrow> summable g \<Longrightarrow> suminf f - suminf g = (\<Sum>n. f n - g n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
259  | 
by (intro sums_unique sums_diff summable_sums)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
260  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
261  | 
lemma sums_minus: "f sums a \<Longrightarrow> (\<lambda>n. - f n) sums (- a)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
262  | 
unfolding sums_def by (simp add: setsum_negf tendsto_minus)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
263  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
264  | 
lemma summable_minus: "summable f \<Longrightarrow> summable (\<lambda>n. - f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
265  | 
unfolding summable_def by (auto intro: sums_minus)  | 
| 20692 | 266  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
267  | 
lemma suminf_minus: "summable f \<Longrightarrow> (\<Sum>n. - f n) = - (\<Sum>n. f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
268  | 
by (intro sums_unique [symmetric] sums_minus summable_sums)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
269  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
270  | 
lemma sums_Suc: "(\<lambda> n. f (Suc n)) sums l \<Longrightarrow> f sums (l + f 0)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
271  | 
by (simp add: sums_Suc_iff)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
272  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
273  | 
lemma sums_iff_shift: "(\<lambda>i. f (i + n)) sums s \<longleftrightarrow> f sums (s + (\<Sum>i<n. f i))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
274  | 
proof (induct n arbitrary: s)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
275  | 
case (Suc n)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
276  | 
moreover have "(\<lambda>i. f (Suc i + n)) sums s \<longleftrightarrow> (\<lambda>i. f (i + n)) sums (s + f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
277  | 
by (subst sums_Suc_iff) simp  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
278  | 
ultimately show ?case  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
279  | 
by (simp add: ac_simps)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
280  | 
qed simp  | 
| 20692 | 281  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
282  | 
lemma summable_iff_shift: "summable (\<lambda>n. f (n + k)) \<longleftrightarrow> summable f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
283  | 
by (metis diff_add_cancel summable_def sums_iff_shift[abs_def])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
284  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
285  | 
lemma sums_split_initial_segment: "f sums s \<Longrightarrow> (\<lambda>i. f (i + n)) sums (s - (\<Sum>i<n. f i))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
286  | 
by (simp add: sums_iff_shift)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
287  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
288  | 
lemma summable_ignore_initial_segment: "summable f \<Longrightarrow> summable (\<lambda>n. f(n + k))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
289  | 
by (simp add: summable_iff_shift)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
290  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
291  | 
lemma suminf_minus_initial_segment: "summable f \<Longrightarrow> (\<Sum>n. f (n + k)) = (\<Sum>n. f n) - (\<Sum>i<k. f i)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
292  | 
by (rule sums_unique[symmetric]) (auto simp: sums_iff_shift)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
293  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
294  | 
lemma suminf_split_initial_segment: "summable f \<Longrightarrow> suminf f = (\<Sum>n. f(n + k)) + (\<Sum>i<k. f i)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
295  | 
by (auto simp add: suminf_minus_initial_segment)  | 
| 20692 | 296  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
297  | 
lemma suminf_exist_split:  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
298  | 
fixes r :: real assumes "0 < r" and "summable f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
299  | 
shows "\<exists>N. \<forall>n\<ge>N. norm (\<Sum>i. f (i + n)) < r"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
300  | 
proof -  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
301  | 
from LIMSEQ_D[OF summable_LIMSEQ[OF `summable f`] `0 < r`]  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
302  | 
  obtain N :: nat where "\<forall> n \<ge> N. norm (setsum f {..<n} - suminf f) < r" by auto
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
303  | 
thus ?thesis  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
304  | 
by (auto simp: norm_minus_commute suminf_minus_initial_segment[OF `summable f`])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
305  | 
qed  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
306  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
307  | 
lemma summable_LIMSEQ_zero: "summable f \<Longrightarrow> f ----> 0"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
308  | 
apply (drule summable_iff_convergent [THEN iffD1])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
309  | 
apply (drule convergent_Cauchy)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
310  | 
apply (simp only: Cauchy_iff LIMSEQ_iff, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
311  | 
apply (drule_tac x="r" in spec, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
312  | 
apply (rule_tac x="M" in exI, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
313  | 
apply (drule_tac x="Suc n" in spec, simp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
314  | 
apply (drule_tac x="n" in spec, simp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
315  | 
done  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
316  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
317  | 
end  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
318  | 
|
| 57025 | 319  | 
context  | 
320  | 
fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::real_normed_vector" and I :: "'i set"  | 
|
321  | 
begin  | 
|
322  | 
||
323  | 
lemma sums_setsum: "(\<And>i. i \<in> I \<Longrightarrow> (f i) sums (x i)) \<Longrightarrow> (\<lambda>n. \<Sum>i\<in>I. f i n) sums (\<Sum>i\<in>I. x i)"  | 
|
324  | 
by (induct I rule: infinite_finite_induct) (auto intro!: sums_add)  | 
|
325  | 
||
326  | 
lemma suminf_setsum: "(\<And>i. i \<in> I \<Longrightarrow> summable (f i)) \<Longrightarrow> (\<Sum>n. \<Sum>i\<in>I. f i n) = (\<Sum>i\<in>I. \<Sum>n. f i n)"  | 
|
327  | 
using sums_unique[OF sums_setsum, OF summable_sums] by simp  | 
|
328  | 
||
329  | 
lemma summable_setsum: "(\<And>i. i \<in> I \<Longrightarrow> summable (f i)) \<Longrightarrow> summable (\<lambda>n. \<Sum>i\<in>I. f i n)"  | 
|
330  | 
using sums_summable[OF sums_setsum[OF summable_sums]] .  | 
|
331  | 
||
332  | 
end  | 
|
333  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
334  | 
lemma (in bounded_linear) sums: "(\<lambda>n. X n) sums a \<Longrightarrow> (\<lambda>n. f (X n)) sums (f a)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
335  | 
unfolding sums_def by (drule tendsto, simp only: setsum)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
336  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
337  | 
lemma (in bounded_linear) summable: "summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
338  | 
unfolding summable_def by (auto intro: sums)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
339  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
340  | 
lemma (in bounded_linear) suminf: "summable (\<lambda>n. X n) \<Longrightarrow> f (\<Sum>n. X n) = (\<Sum>n. f (X n))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
341  | 
by (intro sums_unique sums summable_sums)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
342  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
343  | 
lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real]  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
344  | 
lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real]  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
345  | 
lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real]  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
346  | 
|
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
347  | 
lemmas sums_scaleR_left = bounded_linear.sums[OF bounded_linear_scaleR_left]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
348  | 
lemmas summable_scaleR_left = bounded_linear.summable[OF bounded_linear_scaleR_left]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
349  | 
lemmas suminf_scaleR_left = bounded_linear.suminf[OF bounded_linear_scaleR_left]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
350  | 
|
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
351  | 
lemmas sums_scaleR_right = bounded_linear.sums[OF bounded_linear_scaleR_right]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
352  | 
lemmas summable_scaleR_right = bounded_linear.summable[OF bounded_linear_scaleR_right]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
353  | 
lemmas suminf_scaleR_right = bounded_linear.suminf[OF bounded_linear_scaleR_right]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57129 
diff
changeset
 | 
354  | 
|
| 56213 | 355  | 
subsection {* Infinite summability on real normed algebras *}
 | 
356  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
357  | 
context  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
358  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_algebra"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
359  | 
begin  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
360  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
361  | 
lemma sums_mult: "f sums a \<Longrightarrow> (\<lambda>n. c * f n) sums (c * a)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
362  | 
by (rule bounded_linear.sums [OF bounded_linear_mult_right])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
363  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
364  | 
lemma summable_mult: "summable f \<Longrightarrow> summable (\<lambda>n. c * f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
365  | 
by (rule bounded_linear.summable [OF bounded_linear_mult_right])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
366  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
367  | 
lemma suminf_mult: "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
368  | 
by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
369  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
370  | 
lemma sums_mult2: "f sums a \<Longrightarrow> (\<lambda>n. f n * c) sums (a * c)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
371  | 
by (rule bounded_linear.sums [OF bounded_linear_mult_left])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
372  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
373  | 
lemma summable_mult2: "summable f \<Longrightarrow> summable (\<lambda>n. f n * c)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
374  | 
by (rule bounded_linear.summable [OF bounded_linear_mult_left])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
375  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
376  | 
lemma suminf_mult2: "summable f \<Longrightarrow> suminf f * c = (\<Sum>n. f n * c)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
377  | 
by (rule bounded_linear.suminf [OF bounded_linear_mult_left])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
378  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
379  | 
end  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
380  | 
|
| 56213 | 381  | 
subsection {* Infinite summability on real normed fields *}
 | 
382  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
383  | 
context  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
384  | 
fixes c :: "'a::real_normed_field"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
385  | 
begin  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
386  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
387  | 
lemma sums_divide: "f sums a \<Longrightarrow> (\<lambda>n. f n / c) sums (a / c)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
388  | 
by (rule bounded_linear.sums [OF bounded_linear_divide])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
389  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
390  | 
lemma summable_divide: "summable f \<Longrightarrow> summable (\<lambda>n. f n / c)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
391  | 
by (rule bounded_linear.summable [OF bounded_linear_divide])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
392  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
393  | 
lemma suminf_divide: "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
394  | 
by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric])  | 
| 14416 | 395  | 
|
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
396  | 
text{*Sum of a geometric progression.*}
 | 
| 14416 | 397  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
398  | 
lemma geometric_sums: "norm c < 1 \<Longrightarrow> (\<lambda>n. c^n) sums (1 / (1 - c))"  | 
| 20692 | 399  | 
proof -  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
400  | 
assume less_1: "norm c < 1"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
401  | 
hence neq_1: "c \<noteq> 1" by auto  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
402  | 
hence neq_0: "c - 1 \<noteq> 0" by simp  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
403  | 
from less_1 have lim_0: "(\<lambda>n. c^n) ----> 0"  | 
| 20692 | 404  | 
by (rule LIMSEQ_power_zero)  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
405  | 
hence "(\<lambda>n. c ^ n / (c - 1) - 1 / (c - 1)) ----> 0 / (c - 1) - 1 / (c - 1)"  | 
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
406  | 
using neq_0 by (intro tendsto_intros)  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
407  | 
hence "(\<lambda>n. (c ^ n - 1) / (c - 1)) ----> 1 / (1 - c)"  | 
| 20692 | 408  | 
by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib)  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
409  | 
thus "(\<lambda>n. c ^ n) sums (1 / (1 - c))"  | 
| 20692 | 410  | 
by (simp add: sums_def geometric_sum neq_1)  | 
411  | 
qed  | 
|
412  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
413  | 
lemma summable_geometric: "norm c < 1 \<Longrightarrow> summable (\<lambda>n. c^n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
414  | 
by (rule geometric_sums [THEN sums_summable])  | 
| 14416 | 415  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
416  | 
lemma suminf_geometric: "norm c < 1 \<Longrightarrow> suminf (\<lambda>n. c^n) = 1 / (1 - c)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
417  | 
by (rule sums_unique[symmetric]) (rule geometric_sums)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
418  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
419  | 
end  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
420  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
421  | 
lemma power_half_series: "(\<lambda>n. (1/2::real)^Suc n) sums 1"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
422  | 
proof -  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
423  | 
have 2: "(\<lambda>n. (1/2::real)^n) sums 2" using geometric_sums [of "1/2::real"]  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
424  | 
by auto  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
425  | 
have "(\<lambda>n. (1/2::real)^Suc n) = (\<lambda>n. (1 / 2) ^ n / 2)"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
426  | 
by simp  | 
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
427  | 
thus ?thesis using sums_divide [OF 2, of 2]  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
428  | 
by simp  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
429  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
430  | 
|
| 56213 | 431  | 
subsection {* Infinite summability on Banach spaces *}
 | 
432  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
433  | 
text{*Cauchy-type criterion for convergence of series (c.f. Harrison)*}
 | 
| 
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
434  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
435  | 
lemma summable_Cauchy:  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
436  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
437  | 
  shows "summable f \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n. norm (setsum f {m..<n}) < e)"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
438  | 
apply (simp only: summable_iff_convergent Cauchy_convergent_iff [symmetric] Cauchy_iff, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
439  | 
apply (drule spec, drule (1) mp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
440  | 
apply (erule exE, rule_tac x="M" in exI, clarify)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
441  | 
apply (rule_tac x="m" and y="n" in linorder_le_cases)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
442  | 
apply (frule (1) order_trans)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
443  | 
apply (drule_tac x="n" in spec, drule (1) mp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
444  | 
apply (drule_tac x="m" in spec, drule (1) mp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
445  | 
apply (simp_all add: setsum_diff [symmetric])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
446  | 
apply (drule spec, drule (1) mp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
447  | 
apply (erule exE, rule_tac x="N" in exI, clarify)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
448  | 
apply (rule_tac x="m" and y="n" in linorder_le_cases)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
449  | 
apply (subst norm_minus_commute)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
450  | 
apply (simp_all add: setsum_diff [symmetric])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
451  | 
done  | 
| 14416 | 452  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
453  | 
context  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
454  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
455  | 
begin  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
456  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
457  | 
text{*Absolute convergence imples normal convergence*}
 | 
| 20689 | 458  | 
|
| 56194 | 459  | 
lemma summable_norm_cancel: "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
460  | 
apply (simp only: summable_Cauchy, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
461  | 
apply (drule_tac x="e" in spec, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
462  | 
apply (rule_tac x="N" in exI, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
463  | 
apply (drule_tac x="m" in spec, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
464  | 
apply (rule order_le_less_trans [OF norm_setsum])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
465  | 
apply (rule order_le_less_trans [OF abs_ge_self])  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
466  | 
apply simp  | 
| 50999 | 467  | 
done  | 
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
468  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
469  | 
lemma summable_norm: "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
470  | 
by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_setsum)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
471  | 
|
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
472  | 
text {* Comparison tests *}
 | 
| 14416 | 473  | 
|
| 56194 | 474  | 
lemma summable_comparison_test: "\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n \<Longrightarrow> summable g \<Longrightarrow> summable f"  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
475  | 
apply (simp add: summable_Cauchy, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
476  | 
apply (drule_tac x="e" in spec, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
477  | 
apply (rule_tac x = "N + Na" in exI, safe)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
478  | 
apply (rotate_tac 2)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
479  | 
apply (drule_tac x = m in spec)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
480  | 
apply (auto, rotate_tac 2, drule_tac x = n in spec)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
481  | 
apply (rule_tac y = "\<Sum>k=m..<n. norm (f k)" in order_le_less_trans)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
482  | 
apply (rule norm_setsum)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
483  | 
  apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
484  | 
apply (auto intro: setsum_mono simp add: abs_less_iff)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
485  | 
done  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
486  | 
|
| 
56217
 
dc429a5b13c4
Some rationalisation of basic lemmas
 
paulson <lp15@cam.ac.uk> 
parents: 
56213 
diff
changeset
 | 
487  | 
(*A better argument order*)  | 
| 
 
dc429a5b13c4
Some rationalisation of basic lemmas
 
paulson <lp15@cam.ac.uk> 
parents: 
56213 
diff
changeset
 | 
488  | 
lemma summable_comparison_test': "summable g \<Longrightarrow> (\<And>n. n \<ge> N \<Longrightarrow> norm(f n) \<le> g n) \<Longrightarrow> summable f"  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
489  | 
by (rule summable_comparison_test) auto  | 
| 
56217
 
dc429a5b13c4
Some rationalisation of basic lemmas
 
paulson <lp15@cam.ac.uk> 
parents: 
56213 
diff
changeset
 | 
490  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
491  | 
subsection {* The Ratio Test*}
 | 
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
492  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
493  | 
lemma summable_ratio_test:  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
494  | 
assumes "c < 1" "\<And>n. n \<ge> N \<Longrightarrow> norm (f (Suc n)) \<le> c * norm (f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
495  | 
shows "summable f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
496  | 
proof cases  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
497  | 
assume "0 < c"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
498  | 
show "summable f"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
499  | 
proof (rule summable_comparison_test)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
500  | 
show "\<exists>N'. \<forall>n\<ge>N'. norm (f n) \<le> (norm (f N) / (c ^ N)) * c ^ n"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
501  | 
proof (intro exI allI impI)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
502  | 
fix n assume "N \<le> n" then show "norm (f n) \<le> (norm (f N) / (c ^ N)) * c ^ n"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
503  | 
proof (induct rule: inc_induct)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
504  | 
case (step m)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
505  | 
moreover have "norm (f (Suc m)) / c ^ Suc m * c ^ n \<le> norm (f m) / c ^ m * c ^ n"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
506  | 
using `0 < c` `c < 1` assms(2)[OF `N \<le> m`] by (simp add: field_simps)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
507  | 
ultimately show ?case by simp  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
508  | 
qed (insert `0 < c`, simp)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
509  | 
qed  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
510  | 
show "summable (\<lambda>n. norm (f N) / c ^ N * c ^ n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
511  | 
using `0 < c` `c < 1` by (intro summable_mult summable_geometric) simp  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
512  | 
qed  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
513  | 
next  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
514  | 
assume c: "\<not> 0 < c"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
515  | 
  { fix n assume "n \<ge> N"
 | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
516  | 
then have "norm (f (Suc n)) \<le> c * norm (f n)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
517  | 
by fact  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
518  | 
also have "\<dots> \<le> 0"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
519  | 
using c by (simp add: not_less mult_nonpos_nonneg)  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
520  | 
finally have "f (Suc n) = 0"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
521  | 
by auto }  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
522  | 
then show "summable f"  | 
| 56194 | 523  | 
    by (intro sums_summable[OF sums_finite, of "{.. Suc N}"]) (auto simp: not_le Suc_less_eq2)
 | 
| 56178 | 524  | 
qed  | 
525  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
526  | 
end  | 
| 14416 | 527  | 
|
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
528  | 
text{*Relations among convergence and absolute convergence for power series.*}
 | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
529  | 
|
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
530  | 
lemma abel_lemma:  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
531  | 
fixes a :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
532  | 
assumes r: "0 \<le> r" and r0: "r < r0" and M: "\<And>n. norm (a n) * r0^n \<le> M"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
533  | 
shows "summable (\<lambda>n. norm (a n) * r^n)"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
534  | 
proof (rule summable_comparison_test')  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
535  | 
show "summable (\<lambda>n. M * (r / r0) ^ n)"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
536  | 
using assms  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
537  | 
by (auto simp add: summable_mult summable_geometric)  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
538  | 
next  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
539  | 
fix n  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
540  | 
show "norm (norm (a n) * r ^ n) \<le> M * (r / r0) ^ n"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
541  | 
using r r0 M [of n]  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
542  | 
apply (auto simp add: abs_mult field_simps power_divide)  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
543  | 
apply (cases "r=0", simp)  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
544  | 
apply (cases n, auto)  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
545  | 
done  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
546  | 
qed  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
547  | 
|
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56217 
diff
changeset
 | 
548  | 
|
| 23084 | 549  | 
text{*Summability of geometric series for real algebras*}
 | 
550  | 
||
551  | 
lemma complete_algebra_summable_geometric:  | 
|
| 31017 | 552  | 
  fixes x :: "'a::{real_normed_algebra_1,banach}"
 | 
| 23084 | 553  | 
shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"  | 
554  | 
proof (rule summable_comparison_test)  | 
|
555  | 
show "\<exists>N. \<forall>n\<ge>N. norm (x ^ n) \<le> norm x ^ n"  | 
|
556  | 
by (simp add: norm_power_ineq)  | 
|
557  | 
show "norm x < 1 \<Longrightarrow> summable (\<lambda>n. norm x ^ n)"  | 
|
558  | 
by (simp add: summable_geometric)  | 
|
559  | 
qed  | 
|
560  | 
||
| 23111 | 561  | 
subsection {* Cauchy Product Formula *}
 | 
562  | 
||
| 54703 | 563  | 
text {*
 | 
564  | 
Proof based on Analysis WebNotes: Chapter 07, Class 41  | 
|
565  | 
  @{url "http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm"}
 | 
|
566  | 
*}  | 
|
| 23111 | 567  | 
|
568  | 
lemma setsum_triangle_reindex:  | 
|
569  | 
fixes n :: nat  | 
|
| 56213 | 570  | 
  shows "(\<Sum>(i,j)\<in>{(i,j). i+j < n}. f i j) = (\<Sum>k<n. \<Sum>i\<le>k. f i (k - i))"
 | 
| 57418 | 571  | 
apply (simp add: setsum.Sigma)  | 
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
572  | 
apply (rule setsum.reindex_bij_witness[where j="\<lambda>(i, j). (i+j, i)" and i="\<lambda>(k, i). (i, k - i)"])  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
573  | 
apply auto  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
574  | 
done  | 
| 23111 | 575  | 
|
576  | 
lemma Cauchy_product_sums:  | 
|
577  | 
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
 | 
|
578  | 
assumes a: "summable (\<lambda>k. norm (a k))"  | 
|
579  | 
assumes b: "summable (\<lambda>k. norm (b k))"  | 
|
| 56213 | 580  | 
shows "(\<lambda>k. \<Sum>i\<le>k. a i * b (k - i)) sums ((\<Sum>k. a k) * (\<Sum>k. b k))"  | 
| 23111 | 581  | 
proof -  | 
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
582  | 
  let ?S1 = "\<lambda>n::nat. {..<n} \<times> {..<n}"
 | 
| 23111 | 583  | 
  let ?S2 = "\<lambda>n::nat. {(i,j). i + j < n}"
 | 
584  | 
have S1_mono: "\<And>m n. m \<le> n \<Longrightarrow> ?S1 m \<subseteq> ?S1 n" by auto  | 
|
585  | 
have S2_le_S1: "\<And>n. ?S2 n \<subseteq> ?S1 n" by auto  | 
|
586  | 
have S1_le_S2: "\<And>n. ?S1 (n div 2) \<subseteq> ?S2 n" by auto  | 
|
587  | 
have finite_S1: "\<And>n. finite (?S1 n)" by simp  | 
|
588  | 
with S2_le_S1 have finite_S2: "\<And>n. finite (?S2 n)" by (rule finite_subset)  | 
|
589  | 
||
590  | 
let ?g = "\<lambda>(i,j). a i * b j"  | 
|
591  | 
let ?f = "\<lambda>(i,j). norm (a i) * norm (b j)"  | 
|
| 56536 | 592  | 
have f_nonneg: "\<And>x. 0 \<le> ?f x" by (auto)  | 
| 23111 | 593  | 
hence norm_setsum_f: "\<And>A. norm (setsum ?f A) = setsum ?f A"  | 
594  | 
unfolding real_norm_def  | 
|
595  | 
by (simp only: abs_of_nonneg setsum_nonneg [rule_format])  | 
|
596  | 
||
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
597  | 
have "(\<lambda>n. (\<Sum>k<n. a k) * (\<Sum>k<n. b k)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
598  | 
by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b])  | 
| 23111 | 599  | 
hence 1: "(\<lambda>n. setsum ?g (?S1 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
| 57418 | 600  | 
by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan)  | 
| 23111 | 601  | 
|
| 
56193
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
602  | 
have "(\<lambda>n. (\<Sum>k<n. norm (a k)) * (\<Sum>k<n. norm (b k))) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"  | 
| 
 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 
hoelzl 
parents: 
56178 
diff
changeset
 | 
603  | 
using a b by (intro tendsto_mult summable_LIMSEQ)  | 
| 23111 | 604  | 
hence "(\<lambda>n. setsum ?f (?S1 n)) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"  | 
| 57418 | 605  | 
by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan)  | 
| 23111 | 606  | 
hence "convergent (\<lambda>n. setsum ?f (?S1 n))"  | 
607  | 
by (rule convergentI)  | 
|
608  | 
hence Cauchy: "Cauchy (\<lambda>n. setsum ?f (?S1 n))"  | 
|
609  | 
by (rule convergent_Cauchy)  | 
|
| 36657 | 610  | 
have "Zfun (\<lambda>n. setsum ?f (?S1 n - ?S2 n)) sequentially"  | 
611  | 
proof (rule ZfunI, simp only: eventually_sequentially norm_setsum_f)  | 
|
| 23111 | 612  | 
fix r :: real  | 
613  | 
assume r: "0 < r"  | 
|
614  | 
from CauchyD [OF Cauchy r] obtain N  | 
|
615  | 
where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" ..  | 
|
616  | 
hence "\<And>m n. \<lbrakk>N \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> norm (setsum ?f (?S1 m - ?S1 n)) < r"  | 
|
617  | 
by (simp only: setsum_diff finite_S1 S1_mono)  | 
|
618  | 
hence N: "\<And>m n. \<lbrakk>N \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> setsum ?f (?S1 m - ?S1 n) < r"  | 
|
619  | 
by (simp only: norm_setsum_f)  | 
|
620  | 
show "\<exists>N. \<forall>n\<ge>N. setsum ?f (?S1 n - ?S2 n) < r"  | 
|
621  | 
proof (intro exI allI impI)  | 
|
622  | 
fix n assume "2 * N \<le> n"  | 
|
623  | 
hence n: "N \<le> n div 2" by simp  | 
|
624  | 
have "setsum ?f (?S1 n - ?S2 n) \<le> setsum ?f (?S1 n - ?S1 (n div 2))"  | 
|
625  | 
by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg  | 
|
626  | 
Diff_mono subset_refl S1_le_S2)  | 
|
627  | 
also have "\<dots> < r"  | 
|
628  | 
using n div_le_dividend by (rule N)  | 
|
629  | 
finally show "setsum ?f (?S1 n - ?S2 n) < r" .  | 
|
630  | 
qed  | 
|
631  | 
qed  | 
|
| 36657 | 632  | 
hence "Zfun (\<lambda>n. setsum ?g (?S1 n - ?S2 n)) sequentially"  | 
633  | 
apply (rule Zfun_le [rule_format])  | 
|
| 23111 | 634  | 
apply (simp only: norm_setsum_f)  | 
635  | 
apply (rule order_trans [OF norm_setsum setsum_mono])  | 
|
636  | 
apply (auto simp add: norm_mult_ineq)  | 
|
637  | 
done  | 
|
638  | 
hence 2: "(\<lambda>n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) ----> 0"  | 
|
| 
36660
 
1cc4ab4b7ff7
make (X ----> L) an abbreviation for (X ---> L) sequentially
 
huffman 
parents: 
36657 
diff
changeset
 | 
639  | 
unfolding tendsto_Zfun_iff diff_0_right  | 
| 36657 | 640  | 
by (simp only: setsum_diff finite_S1 S2_le_S1)  | 
| 23111 | 641  | 
|
642  | 
with 1 have "(\<lambda>n. setsum ?g (?S2 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
643  | 
by (rule LIMSEQ_diff_approach_zero2)  | 
|
644  | 
thus ?thesis by (simp only: sums_def setsum_triangle_reindex)  | 
|
645  | 
qed  | 
|
646  | 
||
647  | 
lemma Cauchy_product:  | 
|
648  | 
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
 | 
|
649  | 
assumes a: "summable (\<lambda>k. norm (a k))"  | 
|
650  | 
assumes b: "summable (\<lambda>k. norm (b k))"  | 
|
| 56213 | 651  | 
shows "(\<Sum>k. a k) * (\<Sum>k. b k) = (\<Sum>k. \<Sum>i\<le>k. a i * b (k - i))"  | 
652  | 
using a b  | 
|
653  | 
by (rule Cauchy_product_sums [THEN sums_unique])  | 
|
654  | 
||
655  | 
subsection {* Series on @{typ real}s *}
 | 
|
656  | 
||
657  | 
lemma summable_norm_comparison_test: "\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. norm (f n))"  | 
|
658  | 
by (rule summable_comparison_test) auto  | 
|
659  | 
||
660  | 
lemma summable_rabs_comparison_test: "\<lbrakk>\<exists>N. \<forall>n\<ge>N. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable (\<lambda>n. \<bar>f n :: real\<bar>)"  | 
|
661  | 
by (rule summable_comparison_test) auto  | 
|
662  | 
||
663  | 
lemma summable_rabs_cancel: "summable (\<lambda>n. \<bar>f n :: real\<bar>) \<Longrightarrow> summable f"  | 
|
664  | 
by (rule summable_norm_cancel) simp  | 
|
665  | 
||
666  | 
lemma summable_rabs: "summable (\<lambda>n. \<bar>f n :: real\<bar>) \<Longrightarrow> \<bar>suminf f\<bar> \<le> (\<Sum>n. \<bar>f n\<bar>)"  | 
|
667  | 
by (fold real_norm_def) (rule summable_norm)  | 
|
| 23111 | 668  | 
|
| 14416 | 669  | 
end  |