| author | berghofe | 
| Wed, 07 May 2008 10:59:34 +0200 | |
| changeset 26819 | 56036226028b | 
| parent 23441 | ee218296d635 | 
| child 27108 | e447b3107696 | 
| 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  | 
| 10751 | 8  | 
*)  | 
9  | 
||
| 14416 | 10  | 
header{*Finite Summation and Infinite Series*}
 | 
| 10751 | 11  | 
|
| 15131 | 12  | 
theory Series  | 
| 21141 | 13  | 
imports SEQ  | 
| 15131 | 14  | 
begin  | 
| 15561 | 15  | 
|
| 19765 | 16  | 
definition  | 
| 20692 | 17  | 
sums :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21141 
diff
changeset
 | 
18  | 
(infixr "sums" 80) where  | 
| 19765 | 19  | 
   "f sums s = (%n. setsum f {0..<n}) ----> s"
 | 
| 10751 | 20  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21141 
diff
changeset
 | 
21  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21141 
diff
changeset
 | 
22  | 
summable :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool" where  | 
| 19765 | 23  | 
"summable f = (\<exists>s. f sums s)"  | 
| 14416 | 24  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21141 
diff
changeset
 | 
25  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21141 
diff
changeset
 | 
26  | 
suminf :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a" where  | 
| 20688 | 27  | 
"suminf f = (THE s. f sums s)"  | 
| 14416 | 28  | 
|
| 15546 | 29  | 
syntax  | 
| 20692 | 30  | 
  "_suminf" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a" ("\<Sum>_. _" [0, 10] 10)
 | 
| 15546 | 31  | 
translations  | 
| 20770 | 32  | 
"\<Sum>i. b" == "CONST suminf (%i. b)"  | 
| 15546 | 33  | 
|
| 14416 | 34  | 
|
| 15539 | 35  | 
lemma sumr_diff_mult_const:  | 
36  | 
 "setsum f {0..<n} - (real n*r) = setsum (%i. f i - r) {0..<n::nat}"
 | 
|
| 15536 | 37  | 
by (simp add: diff_minus setsum_addf real_of_nat_def)  | 
38  | 
||
| 15542 | 39  | 
lemma real_setsum_nat_ivl_bounded:  | 
40  | 
"(!!p. p < n \<Longrightarrow> f(p) \<le> K)  | 
|
41  | 
      \<Longrightarrow> setsum f {0..<n::nat} \<le> real n * K"
 | 
|
42  | 
using setsum_bounded[where A = "{0..<n}"]
 | 
|
43  | 
by (auto simp:real_of_nat_def)  | 
|
| 14416 | 44  | 
|
| 15539 | 45  | 
(* Generalize from real to some algebraic structure? *)  | 
46  | 
lemma sumr_minus_one_realpow_zero [simp]:  | 
|
| 15543 | 47  | 
"(\<Sum>i=0..<2*n. (-1) ^ Suc i) = (0::real)"  | 
| 15251 | 48  | 
by (induct "n", auto)  | 
| 14416 | 49  | 
|
| 15539 | 50  | 
(* FIXME this is an awful lemma! *)  | 
51  | 
lemma sumr_one_lb_realpow_zero [simp]:  | 
|
52  | 
"(\<Sum>n=Suc 0..<n. f(n) * (0::real) ^ n) = 0"  | 
|
| 20692 | 53  | 
by (rule setsum_0', simp)  | 
| 14416 | 54  | 
|
| 15543 | 55  | 
lemma sumr_group:  | 
| 15539 | 56  | 
     "(\<Sum>m=0..<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {0 ..< n * k}"
 | 
| 15543 | 57  | 
apply (subgoal_tac "k = 0 | 0 < k", auto)  | 
| 15251 | 58  | 
apply (induct "n")  | 
| 15539 | 59  | 
apply (simp_all add: setsum_add_nat_ivl add_commute)  | 
| 14416 | 60  | 
done  | 
| 15539 | 61  | 
|
| 20692 | 62  | 
lemma sumr_offset3:  | 
63  | 
  "setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)) + setsum f {0..<k}"
 | 
|
64  | 
apply (subst setsum_shift_bounds_nat_ivl [symmetric])  | 
|
65  | 
apply (simp add: setsum_add_nat_ivl add_commute)  | 
|
66  | 
done  | 
|
67  | 
||
| 16819 | 68  | 
lemma sumr_offset:  | 
| 20692 | 69  | 
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"  | 
70  | 
  shows "(\<Sum>m=0..<n. f(m+k)) = setsum f {0..<n+k} - setsum f {0..<k}"
 | 
|
71  | 
by (simp add: sumr_offset3)  | 
|
| 16819 | 72  | 
|
73  | 
lemma sumr_offset2:  | 
|
74  | 
 "\<forall>f. (\<Sum>m=0..<n::nat. f(m+k)::real) = setsum f {0..<n+k} - setsum f {0..<k}"
 | 
|
| 20692 | 75  | 
by (simp add: sumr_offset)  | 
| 16819 | 76  | 
|
77  | 
lemma sumr_offset4:  | 
|
| 20692 | 78  | 
  "\<forall>n f. setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)::real) + setsum f {0..<k}"
 | 
79  | 
by (clarify, rule sumr_offset3)  | 
|
| 16819 | 80  | 
|
81  | 
(*  | 
|
82  | 
lemma sumr_from_1_from_0: "0 < n ==>  | 
|
83  | 
(\<Sum>n=Suc 0 ..< Suc n. if even(n) then 0 else  | 
|
84  | 
((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n =  | 
|
85  | 
(\<Sum>n=0..<Suc n. if even(n) then 0 else  | 
|
86  | 
((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n"  | 
|
87  | 
by (rule_tac n1 = 1 in sumr_split_add [THEN subst], auto)  | 
|
88  | 
*)  | 
|
| 14416 | 89  | 
|
90  | 
subsection{* Infinite Sums, by the Properties of Limits*}
 | 
|
91  | 
||
92  | 
(*----------------------  | 
|
93  | 
suminf is the sum  | 
|
94  | 
---------------------*)  | 
|
95  | 
lemma sums_summable: "f sums l ==> summable f"  | 
|
96  | 
by (simp add: sums_def summable_def, blast)  | 
|
97  | 
||
98  | 
lemma summable_sums: "summable f ==> f sums (suminf f)"  | 
|
| 20688 | 99  | 
apply (simp add: summable_def suminf_def sums_def)  | 
100  | 
apply (blast intro: theI LIMSEQ_unique)  | 
|
| 14416 | 101  | 
done  | 
102  | 
||
103  | 
lemma summable_sumr_LIMSEQ_suminf:  | 
|
| 15539 | 104  | 
     "summable f ==> (%n. setsum f {0..<n}) ----> (suminf f)"
 | 
| 20688 | 105  | 
by (rule summable_sums [unfolded sums_def])  | 
| 14416 | 106  | 
|
107  | 
(*-------------------  | 
|
108  | 
sum is unique  | 
|
109  | 
------------------*)  | 
|
110  | 
lemma sums_unique: "f sums s ==> (s = suminf f)"  | 
|
111  | 
apply (frule sums_summable [THEN summable_sums])  | 
|
112  | 
apply (auto intro!: LIMSEQ_unique simp add: sums_def)  | 
|
113  | 
done  | 
|
114  | 
||
| 16819 | 115  | 
lemma sums_split_initial_segment: "f sums s ==>  | 
116  | 
(%n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"  | 
|
117  | 
apply (unfold sums_def);  | 
|
118  | 
apply (simp add: sumr_offset);  | 
|
119  | 
apply (rule LIMSEQ_diff_const)  | 
|
120  | 
apply (rule LIMSEQ_ignore_initial_segment)  | 
|
121  | 
apply assumption  | 
|
122  | 
done  | 
|
123  | 
||
124  | 
lemma summable_ignore_initial_segment: "summable f ==>  | 
|
125  | 
summable (%n. f(n + k))"  | 
|
126  | 
apply (unfold summable_def)  | 
|
127  | 
apply (auto intro: sums_split_initial_segment)  | 
|
128  | 
done  | 
|
129  | 
||
130  | 
lemma suminf_minus_initial_segment: "summable f ==>  | 
|
131  | 
suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)"  | 
|
132  | 
apply (frule summable_ignore_initial_segment)  | 
|
133  | 
apply (rule sums_unique [THEN sym])  | 
|
134  | 
apply (frule summable_sums)  | 
|
135  | 
apply (rule sums_split_initial_segment)  | 
|
136  | 
apply auto  | 
|
137  | 
done  | 
|
138  | 
||
139  | 
lemma suminf_split_initial_segment: "summable f ==>  | 
|
140  | 
suminf f = (SUM i = 0..< k. f i) + suminf (%n. f(n + k))"  | 
|
141  | 
by (auto simp add: suminf_minus_initial_segment)  | 
|
142  | 
||
| 14416 | 143  | 
lemma series_zero:  | 
| 15539 | 144  | 
     "(\<forall>m. n \<le> m --> f(m) = 0) ==> f sums (setsum f {0..<n})"
 | 
| 15537 | 145  | 
apply (simp add: sums_def LIMSEQ_def diff_minus[symmetric], safe)  | 
| 14416 | 146  | 
apply (rule_tac x = n in exI)  | 
| 15542 | 147  | 
apply (clarsimp simp add:setsum_diff[symmetric] cong:setsum_ivl_cong)  | 
| 14416 | 148  | 
done  | 
149  | 
||
| 23121 | 150  | 
lemma sums_zero: "(\<lambda>n. 0) sums 0"  | 
151  | 
unfolding sums_def by (simp add: LIMSEQ_const)  | 
|
| 15539 | 152  | 
|
| 23121 | 153  | 
lemma summable_zero: "summable (\<lambda>n. 0)"  | 
154  | 
by (rule sums_zero [THEN sums_summable])  | 
|
| 16819 | 155  | 
|
| 23121 | 156  | 
lemma suminf_zero: "suminf (\<lambda>n. 0) = 0"  | 
157  | 
by (rule sums_zero [THEN sums_unique, symmetric])  | 
|
| 16819 | 158  | 
|
| 23119 | 159  | 
lemma (in bounded_linear) sums:  | 
160  | 
"(\<lambda>n. X n) sums a \<Longrightarrow> (\<lambda>n. f (X n)) sums (f a)"  | 
|
161  | 
unfolding sums_def by (drule LIMSEQ, simp only: setsum)  | 
|
162  | 
||
163  | 
lemma (in bounded_linear) summable:  | 
|
164  | 
"summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"  | 
|
165  | 
unfolding summable_def by (auto intro: sums)  | 
|
166  | 
||
167  | 
lemma (in bounded_linear) suminf:  | 
|
168  | 
"summable (\<lambda>n. X n) \<Longrightarrow> f (\<Sum>n. X n) = (\<Sum>n. f (X n))"  | 
|
| 23121 | 169  | 
by (intro sums_unique sums summable_sums)  | 
| 23119 | 170  | 
|
| 20692 | 171  | 
lemma sums_mult:  | 
172  | 
fixes c :: "'a::real_normed_algebra"  | 
|
173  | 
shows "f sums a \<Longrightarrow> (\<lambda>n. c * f n) sums (c * a)"  | 
|
| 23127 | 174  | 
by (rule mult_right.sums)  | 
| 14416 | 175  | 
|
| 20692 | 176  | 
lemma summable_mult:  | 
177  | 
fixes c :: "'a::real_normed_algebra"  | 
|
| 23121 | 178  | 
shows "summable f \<Longrightarrow> summable (%n. c * f n)"  | 
| 23127 | 179  | 
by (rule mult_right.summable)  | 
| 16819 | 180  | 
|
| 20692 | 181  | 
lemma suminf_mult:  | 
182  | 
fixes c :: "'a::real_normed_algebra"  | 
|
183  | 
shows "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f";  | 
|
| 23127 | 184  | 
by (rule mult_right.suminf [symmetric])  | 
| 16819 | 185  | 
|
| 20692 | 186  | 
lemma sums_mult2:  | 
187  | 
fixes c :: "'a::real_normed_algebra"  | 
|
188  | 
shows "f sums a \<Longrightarrow> (\<lambda>n. f n * c) sums (a * c)"  | 
|
| 23127 | 189  | 
by (rule mult_left.sums)  | 
| 16819 | 190  | 
|
| 20692 | 191  | 
lemma summable_mult2:  | 
192  | 
fixes c :: "'a::real_normed_algebra"  | 
|
193  | 
shows "summable f \<Longrightarrow> summable (\<lambda>n. f n * c)"  | 
|
| 23127 | 194  | 
by (rule mult_left.summable)  | 
| 16819 | 195  | 
|
| 20692 | 196  | 
lemma suminf_mult2:  | 
197  | 
fixes c :: "'a::real_normed_algebra"  | 
|
198  | 
shows "summable f \<Longrightarrow> suminf f * c = (\<Sum>n. f n * c)"  | 
|
| 23127 | 199  | 
by (rule mult_left.suminf)  | 
| 16819 | 200  | 
|
| 20692 | 201  | 
lemma sums_divide:  | 
202  | 
fixes c :: "'a::real_normed_field"  | 
|
203  | 
shows "f sums a \<Longrightarrow> (\<lambda>n. f n / c) sums (a / c)"  | 
|
| 23127 | 204  | 
by (rule divide.sums)  | 
| 14416 | 205  | 
|
| 20692 | 206  | 
lemma summable_divide:  | 
207  | 
fixes c :: "'a::real_normed_field"  | 
|
208  | 
shows "summable f \<Longrightarrow> summable (\<lambda>n. f n / c)"  | 
|
| 23127 | 209  | 
by (rule divide.summable)  | 
| 16819 | 210  | 
|
| 20692 | 211  | 
lemma suminf_divide:  | 
212  | 
fixes c :: "'a::real_normed_field"  | 
|
213  | 
shows "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"  | 
|
| 23127 | 214  | 
by (rule divide.suminf [symmetric])  | 
| 16819 | 215  | 
|
| 23121 | 216  | 
lemma sums_add: "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) sums (a + b)"  | 
217  | 
unfolding sums_def by (simp add: setsum_addf LIMSEQ_add)  | 
|
| 16819 | 218  | 
|
| 23121 | 219  | 
lemma summable_add: "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n + Y n)"  | 
220  | 
unfolding summable_def by (auto intro: sums_add)  | 
|
| 16819 | 221  | 
|
222  | 
lemma suminf_add:  | 
|
| 23121 | 223  | 
"\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X + suminf Y = (\<Sum>n. X n + Y n)"  | 
224  | 
by (intro sums_unique sums_add summable_sums)  | 
|
| 14416 | 225  | 
|
| 23121 | 226  | 
lemma sums_diff: "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) sums (a - b)"  | 
227  | 
unfolding sums_def by (simp add: setsum_subtractf LIMSEQ_diff)  | 
|
228  | 
||
229  | 
lemma summable_diff: "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n - Y n)"  | 
|
230  | 
unfolding summable_def by (auto intro: sums_diff)  | 
|
| 14416 | 231  | 
|
232  | 
lemma suminf_diff:  | 
|
| 23121 | 233  | 
"\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X - suminf Y = (\<Sum>n. X n - Y n)"  | 
234  | 
by (intro sums_unique sums_diff summable_sums)  | 
|
| 14416 | 235  | 
|
| 23121 | 236  | 
lemma sums_minus: "X sums a ==> (\<lambda>n. - X n) sums (- a)"  | 
237  | 
unfolding sums_def by (simp add: setsum_negf LIMSEQ_minus)  | 
|
| 16819 | 238  | 
|
| 23121 | 239  | 
lemma summable_minus: "summable X \<Longrightarrow> summable (\<lambda>n. - X n)"  | 
240  | 
unfolding summable_def by (auto intro: sums_minus)  | 
|
| 16819 | 241  | 
|
| 23121 | 242  | 
lemma suminf_minus: "summable X \<Longrightarrow> (\<Sum>n. - X n) = - (\<Sum>n. X n)"  | 
243  | 
by (intro sums_unique [symmetric] sums_minus summable_sums)  | 
|
| 14416 | 244  | 
|
245  | 
lemma sums_group:  | 
|
| 15539 | 246  | 
     "[|summable f; 0 < k |] ==> (%n. setsum f {n*k..<n*k+k}) sums (suminf f)"
 | 
| 14416 | 247  | 
apply (drule summable_sums)  | 
| 20692 | 248  | 
apply (simp only: sums_def sumr_group)  | 
249  | 
apply (unfold LIMSEQ_def, safe)  | 
|
250  | 
apply (drule_tac x="r" in spec, safe)  | 
|
251  | 
apply (rule_tac x="no" in exI, safe)  | 
|
252  | 
apply (drule_tac x="n*k" in spec)  | 
|
253  | 
apply (erule mp)  | 
|
254  | 
apply (erule order_trans)  | 
|
255  | 
apply simp  | 
|
| 14416 | 256  | 
done  | 
257  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
258  | 
text{*A summable series of positive terms has limit that is at least as
 | 
| 
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
259  | 
great as any partial sum.*}  | 
| 14416 | 260  | 
|
| 20692 | 261  | 
lemma series_pos_le:  | 
262  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
263  | 
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 \<le> f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} \<le> suminf f"
 | 
|
| 14416 | 264  | 
apply (drule summable_sums)  | 
265  | 
apply (simp add: sums_def)  | 
|
| 15539 | 266  | 
apply (cut_tac k = "setsum f {0..<n}" in LIMSEQ_const)
 | 
267  | 
apply (erule LIMSEQ_le, blast)  | 
|
| 20692 | 268  | 
apply (rule_tac x="n" in exI, clarify)  | 
| 15539 | 269  | 
apply (rule setsum_mono2)  | 
270  | 
apply auto  | 
|
| 14416 | 271  | 
done  | 
272  | 
||
273  | 
lemma series_pos_less:  | 
|
| 20692 | 274  | 
fixes f :: "nat \<Rightarrow> real"  | 
275  | 
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} < suminf f"
 | 
|
276  | 
apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
 | 
|
277  | 
apply simp  | 
|
278  | 
apply (erule series_pos_le)  | 
|
279  | 
apply (simp add: order_less_imp_le)  | 
|
280  | 
done  | 
|
281  | 
||
282  | 
lemma suminf_gt_zero:  | 
|
283  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
284  | 
shows "\<lbrakk>summable f; \<forall>n. 0 < f n\<rbrakk> \<Longrightarrow> 0 < suminf f"  | 
|
285  | 
by (drule_tac n="0" in series_pos_less, simp_all)  | 
|
286  | 
||
287  | 
lemma suminf_ge_zero:  | 
|
288  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
289  | 
shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 \<le> suminf f"  | 
|
290  | 
by (drule_tac n="0" in series_pos_le, simp_all)  | 
|
291  | 
||
292  | 
lemma sumr_pos_lt_pair:  | 
|
293  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
294  | 
shows "\<lbrakk>summable f;  | 
|
295  | 
\<forall>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\<rbrakk>  | 
|
296  | 
      \<Longrightarrow> setsum f {0..<k} < suminf f"
 | 
|
297  | 
apply (subst suminf_split_initial_segment [where k="k"])  | 
|
298  | 
apply assumption  | 
|
299  | 
apply simp  | 
|
300  | 
apply (drule_tac k="k" in summable_ignore_initial_segment)  | 
|
301  | 
apply (drule_tac k="Suc (Suc 0)" in sums_group, simp)  | 
|
302  | 
apply simp  | 
|
303  | 
apply (frule sums_unique)  | 
|
304  | 
apply (drule sums_summable)  | 
|
305  | 
apply simp  | 
|
306  | 
apply (erule suminf_gt_zero)  | 
|
307  | 
apply (simp add: add_ac)  | 
|
| 14416 | 308  | 
done  | 
309  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
310  | 
text{*Sum of a geometric progression.*}
 | 
| 14416 | 311  | 
|
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
16819 
diff
changeset
 | 
312  | 
lemmas sumr_geometric = geometric_sum [where 'a = real]  | 
| 14416 | 313  | 
|
| 20692 | 314  | 
lemma geometric_sums:  | 
| 
22719
 
c51667189bd3
lemma geometric_sum no longer needs class division_by_zero
 
huffman 
parents: 
21404 
diff
changeset
 | 
315  | 
  fixes x :: "'a::{real_normed_field,recpower}"
 | 
| 20692 | 316  | 
shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) sums (1 / (1 - x))"  | 
317  | 
proof -  | 
|
318  | 
assume less_1: "norm x < 1"  | 
|
319  | 
hence neq_1: "x \<noteq> 1" by auto  | 
|
320  | 
hence neq_0: "x - 1 \<noteq> 0" by simp  | 
|
321  | 
from less_1 have lim_0: "(\<lambda>n. x ^ n) ----> 0"  | 
|
322  | 
by (rule LIMSEQ_power_zero)  | 
|
| 
22719
 
c51667189bd3
lemma geometric_sum no longer needs class division_by_zero
 
huffman 
parents: 
21404 
diff
changeset
 | 
323  | 
hence "(\<lambda>n. x ^ n / (x - 1) - 1 / (x - 1)) ----> 0 / (x - 1) - 1 / (x - 1)"  | 
| 20692 | 324  | 
using neq_0 by (intro LIMSEQ_divide LIMSEQ_diff LIMSEQ_const)  | 
325  | 
hence "(\<lambda>n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)"  | 
|
326  | 
by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib)  | 
|
327  | 
thus "(\<lambda>n. x ^ n) sums (1 / (1 - x))"  | 
|
328  | 
by (simp add: sums_def geometric_sum neq_1)  | 
|
329  | 
qed  | 
|
330  | 
||
331  | 
lemma summable_geometric:  | 
|
| 
22719
 
c51667189bd3
lemma geometric_sum no longer needs class division_by_zero
 
huffman 
parents: 
21404 
diff
changeset
 | 
332  | 
  fixes x :: "'a::{real_normed_field,recpower}"
 | 
| 20692 | 333  | 
shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"  | 
334  | 
by (rule geometric_sums [THEN sums_summable])  | 
|
| 14416 | 335  | 
|
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
336  | 
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
 | 
337  | 
|
| 15539 | 338  | 
lemma summable_convergent_sumr_iff:  | 
339  | 
 "summable f = convergent (%n. setsum f {0..<n})"
 | 
|
| 14416 | 340  | 
by (simp add: summable_def sums_def convergent_def)  | 
341  | 
||
| 20689 | 342  | 
lemma summable_LIMSEQ_zero: "summable f \<Longrightarrow> f ----> 0"  | 
343  | 
apply (drule summable_convergent_sumr_iff [THEN iffD1])  | 
|
| 20692 | 344  | 
apply (drule convergent_Cauchy)  | 
| 20689 | 345  | 
apply (simp only: Cauchy_def LIMSEQ_def, safe)  | 
346  | 
apply (drule_tac x="r" in spec, safe)  | 
|
347  | 
apply (rule_tac x="M" in exI, safe)  | 
|
348  | 
apply (drule_tac x="Suc n" in spec, simp)  | 
|
349  | 
apply (drule_tac x="n" in spec, simp)  | 
|
350  | 
done  | 
|
351  | 
||
| 14416 | 352  | 
lemma summable_Cauchy:  | 
| 20848 | 353  | 
"summable (f::nat \<Rightarrow> 'a::banach) =  | 
354  | 
      (\<forall>e > 0. \<exists>N. \<forall>m \<ge> N. \<forall>n. norm (setsum f {m..<n}) < e)"
 | 
|
355  | 
apply (simp only: summable_convergent_sumr_iff Cauchy_convergent_iff [symmetric] Cauchy_def, safe)  | 
|
| 20410 | 356  | 
apply (drule spec, drule (1) mp)  | 
357  | 
apply (erule exE, rule_tac x="M" in exI, clarify)  | 
|
358  | 
apply (rule_tac x="m" and y="n" in linorder_le_cases)  | 
|
359  | 
apply (frule (1) order_trans)  | 
|
360  | 
apply (drule_tac x="n" in spec, drule (1) mp)  | 
|
361  | 
apply (drule_tac x="m" in spec, drule (1) mp)  | 
|
362  | 
apply (simp add: setsum_diff [symmetric])  | 
|
363  | 
apply simp  | 
|
364  | 
apply (drule spec, drule (1) mp)  | 
|
365  | 
apply (erule exE, rule_tac x="N" in exI, clarify)  | 
|
366  | 
apply (rule_tac x="m" and y="n" in linorder_le_cases)  | 
|
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
367  | 
apply (subst norm_minus_commute)  | 
| 20410 | 368  | 
apply (simp add: setsum_diff [symmetric])  | 
369  | 
apply (simp add: setsum_diff [symmetric])  | 
|
| 14416 | 370  | 
done  | 
371  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
372  | 
text{*Comparison test*}
 | 
| 
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
373  | 
|
| 20692 | 374  | 
lemma norm_setsum:  | 
375  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
376  | 
shows "norm (setsum f A) \<le> (\<Sum>i\<in>A. norm (f i))"  | 
|
377  | 
apply (case_tac "finite A")  | 
|
378  | 
apply (erule finite_induct)  | 
|
379  | 
apply simp  | 
|
380  | 
apply simp  | 
|
381  | 
apply (erule order_trans [OF norm_triangle_ineq add_left_mono])  | 
|
382  | 
apply simp  | 
|
383  | 
done  | 
|
384  | 
||
| 14416 | 385  | 
lemma summable_comparison_test:  | 
| 20848 | 386  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
387  | 
shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f"  | 
|
| 20692 | 388  | 
apply (simp add: summable_Cauchy, safe)  | 
389  | 
apply (drule_tac x="e" in spec, safe)  | 
|
390  | 
apply (rule_tac x = "N + Na" in exI, safe)  | 
|
| 14416 | 391  | 
apply (rotate_tac 2)  | 
392  | 
apply (drule_tac x = m in spec)  | 
|
393  | 
apply (auto, rotate_tac 2, drule_tac x = n in spec)  | 
|
| 20848 | 394  | 
apply (rule_tac y = "\<Sum>k=m..<n. norm (f k)" in order_le_less_trans)  | 
395  | 
apply (rule norm_setsum)  | 
|
| 15539 | 396  | 
apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
 | 
| 22998 | 397  | 
apply (auto intro: setsum_mono simp add: abs_less_iff)  | 
| 14416 | 398  | 
done  | 
399  | 
||
| 20848 | 400  | 
lemma summable_norm_comparison_test:  | 
401  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
|
402  | 
shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk>  | 
|
403  | 
\<Longrightarrow> summable (\<lambda>n. norm (f n))"  | 
|
404  | 
apply (rule summable_comparison_test)  | 
|
405  | 
apply (auto)  | 
|
406  | 
done  | 
|
407  | 
||
| 14416 | 408  | 
lemma summable_rabs_comparison_test:  | 
| 20692 | 409  | 
fixes f :: "nat \<Rightarrow> real"  | 
410  | 
shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable (\<lambda>n. \<bar>f n\<bar>)"  | 
|
| 14416 | 411  | 
apply (rule summable_comparison_test)  | 
| 15543 | 412  | 
apply (auto)  | 
| 14416 | 413  | 
done  | 
414  | 
||
| 23084 | 415  | 
text{*Summability of geometric series for real algebras*}
 | 
416  | 
||
417  | 
lemma complete_algebra_summable_geometric:  | 
|
418  | 
  fixes x :: "'a::{real_normed_algebra_1,banach,recpower}"
 | 
|
419  | 
shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"  | 
|
420  | 
proof (rule summable_comparison_test)  | 
|
421  | 
show "\<exists>N. \<forall>n\<ge>N. norm (x ^ n) \<le> norm x ^ n"  | 
|
422  | 
by (simp add: norm_power_ineq)  | 
|
423  | 
show "norm x < 1 \<Longrightarrow> summable (\<lambda>n. norm x ^ n)"  | 
|
424  | 
by (simp add: summable_geometric)  | 
|
425  | 
qed  | 
|
426  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
427  | 
text{*Limit comparison property for series (c.f. jrh)*}
 | 
| 
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
428  | 
|
| 14416 | 429  | 
lemma summable_le:  | 
| 20692 | 430  | 
fixes f g :: "nat \<Rightarrow> real"  | 
431  | 
shows "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"  | 
|
| 14416 | 432  | 
apply (drule summable_sums)+  | 
| 20692 | 433  | 
apply (simp only: sums_def, erule (1) LIMSEQ_le)  | 
| 14416 | 434  | 
apply (rule exI)  | 
| 15539 | 435  | 
apply (auto intro!: setsum_mono)  | 
| 14416 | 436  | 
done  | 
437  | 
||
438  | 
lemma summable_le2:  | 
|
| 20692 | 439  | 
fixes f g :: "nat \<Rightarrow> real"  | 
440  | 
shows "\<lbrakk>\<forall>n. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f \<and> suminf f \<le> suminf g"  | 
|
| 20848 | 441  | 
apply (subgoal_tac "summable f")  | 
442  | 
apply (auto intro!: summable_le)  | 
|
| 22998 | 443  | 
apply (simp add: abs_le_iff)  | 
| 20848 | 444  | 
apply (rule_tac g="g" in summable_comparison_test, simp_all)  | 
| 14416 | 445  | 
done  | 
446  | 
||
| 
19106
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
447  | 
(* specialisation for the common 0 case *)  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
448  | 
lemma suminf_0_le:  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
449  | 
fixes f::"nat\<Rightarrow>real"  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
450  | 
assumes gt0: "\<forall>n. 0 \<le> f n" and sm: "summable f"  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
451  | 
shows "0 \<le> suminf f"  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
452  | 
proof -  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
453  | 
let ?g = "(\<lambda>n. (0::real))"  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
454  | 
from gt0 have "\<forall>n. ?g n \<le> f n" by simp  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
455  | 
moreover have "summable ?g" by (rule summable_zero)  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
456  | 
moreover from sm have "summable f" .  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
457  | 
ultimately have "suminf ?g \<le> suminf f" by (rule summable_le)  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
458  | 
then show "0 \<le> suminf f" by (simp add: suminf_zero)  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
459  | 
qed  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
460  | 
|
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
461  | 
|
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
462  | 
text{*Absolute convergence imples normal convergence*}
 | 
| 20848 | 463  | 
lemma summable_norm_cancel:  | 
464  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
|
465  | 
shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"  | 
|
| 20692 | 466  | 
apply (simp only: summable_Cauchy, safe)  | 
467  | 
apply (drule_tac x="e" in spec, safe)  | 
|
468  | 
apply (rule_tac x="N" in exI, safe)  | 
|
469  | 
apply (drule_tac x="m" in spec, safe)  | 
|
| 20848 | 470  | 
apply (rule order_le_less_trans [OF norm_setsum])  | 
471  | 
apply (rule order_le_less_trans [OF abs_ge_self])  | 
|
| 20692 | 472  | 
apply simp  | 
| 14416 | 473  | 
done  | 
474  | 
||
| 20848 | 475  | 
lemma summable_rabs_cancel:  | 
476  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
477  | 
shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> summable f"  | 
|
478  | 
by (rule summable_norm_cancel, simp)  | 
|
479  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
480  | 
text{*Absolute convergence of series*}
 | 
| 20848 | 481  | 
lemma summable_norm:  | 
482  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
|
483  | 
shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"  | 
|
484  | 
by (auto intro: LIMSEQ_le LIMSEQ_norm summable_norm_cancel  | 
|
485  | 
summable_sumr_LIMSEQ_suminf norm_setsum)  | 
|
486  | 
||
| 14416 | 487  | 
lemma summable_rabs:  | 
| 20692 | 488  | 
fixes f :: "nat \<Rightarrow> real"  | 
489  | 
shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> \<bar>suminf f\<bar> \<le> (\<Sum>n. \<bar>f n\<bar>)"  | 
|
| 20848 | 490  | 
by (fold real_norm_def, rule summable_norm)  | 
| 14416 | 491  | 
|
492  | 
subsection{* The Ratio Test*}
 | 
|
493  | 
||
| 20848 | 494  | 
lemma norm_ratiotest_lemma:  | 
| 22852 | 495  | 
fixes x y :: "'a::real_normed_vector"  | 
| 20848 | 496  | 
shows "\<lbrakk>c \<le> 0; norm x \<le> c * norm y\<rbrakk> \<Longrightarrow> x = 0"  | 
497  | 
apply (subgoal_tac "norm x \<le> 0", simp)  | 
|
498  | 
apply (erule order_trans)  | 
|
499  | 
apply (simp add: mult_le_0_iff)  | 
|
500  | 
done  | 
|
501  | 
||
| 14416 | 502  | 
lemma rabs_ratiotest_lemma: "[| c \<le> 0; abs x \<le> c * abs y |] ==> x = (0::real)"  | 
| 20848 | 503  | 
by (erule norm_ratiotest_lemma, simp)  | 
| 14416 | 504  | 
|
505  | 
lemma le_Suc_ex: "(k::nat) \<le> l ==> (\<exists>n. l = k + n)"  | 
|
506  | 
apply (drule le_imp_less_or_eq)  | 
|
507  | 
apply (auto dest: less_imp_Suc_add)  | 
|
508  | 
done  | 
|
509  | 
||
510  | 
lemma le_Suc_ex_iff: "((k::nat) \<le> l) = (\<exists>n. l = k + n)"  | 
|
511  | 
by (auto simp add: le_Suc_ex)  | 
|
512  | 
||
513  | 
(*All this trouble just to get 0<c *)  | 
|
514  | 
lemma ratio_test_lemma2:  | 
|
| 20848 | 515  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
516  | 
shows "\<lbrakk>\<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> 0 < c \<or> summable f"  | 
|
| 14416 | 517  | 
apply (simp (no_asm) add: linorder_not_le [symmetric])  | 
518  | 
apply (simp add: summable_Cauchy)  | 
|
| 15543 | 519  | 
apply (safe, subgoal_tac "\<forall>n. N < n --> f (n) = 0")  | 
520  | 
prefer 2  | 
|
521  | 
apply clarify  | 
|
522  | 
apply(erule_tac x = "n - 1" in allE)  | 
|
523  | 
apply (simp add:diff_Suc split:nat.splits)  | 
|
| 20848 | 524  | 
apply (blast intro: norm_ratiotest_lemma)  | 
| 14416 | 525  | 
apply (rule_tac x = "Suc N" in exI, clarify)  | 
| 15543 | 526  | 
apply(simp cong:setsum_ivl_cong)  | 
| 14416 | 527  | 
done  | 
528  | 
||
529  | 
lemma ratio_test:  | 
|
| 20848 | 530  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
531  | 
shows "\<lbrakk>c < 1; \<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> summable f"  | 
|
| 14416 | 532  | 
apply (frule ratio_test_lemma2, auto)  | 
| 20848 | 533  | 
apply (rule_tac g = "%n. (norm (f N) / (c ^ N))*c ^ n"  | 
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
534  | 
in summable_comparison_test)  | 
| 14416 | 535  | 
apply (rule_tac x = N in exI, safe)  | 
536  | 
apply (drule le_Suc_ex_iff [THEN iffD1])  | 
|
| 22959 | 537  | 
apply (auto simp add: power_add field_power_not_zero)  | 
| 15539 | 538  | 
apply (induct_tac "na", auto)  | 
| 20848 | 539  | 
apply (rule_tac y = "c * norm (f (N + n))" in order_trans)  | 
| 14416 | 540  | 
apply (auto intro: mult_right_mono simp add: summable_def)  | 
541  | 
apply (simp add: mult_ac)  | 
|
| 20848 | 542  | 
apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)  | 
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
543  | 
apply (rule sums_divide)  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
544  | 
apply (rule sums_mult)  | 
| 
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
545  | 
apply (auto intro!: geometric_sums)  | 
| 14416 | 546  | 
done  | 
547  | 
||
| 23111 | 548  | 
subsection {* Cauchy Product Formula *}
 | 
549  | 
||
550  | 
(* Proof based on Analysis WebNotes: Chapter 07, Class 41  | 
|
551  | 
http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm *)  | 
|
552  | 
||
553  | 
lemma setsum_triangle_reindex:  | 
|
554  | 
fixes n :: nat  | 
|
555  | 
  shows "(\<Sum>(i,j)\<in>{(i,j). i+j < n}. f i j) = (\<Sum>k=0..<n. \<Sum>i=0..k. f i (k - i))"
 | 
|
556  | 
proof -  | 
|
557  | 
  have "(\<Sum>(i, j)\<in>{(i, j). i + j < n}. f i j) =
 | 
|
558  | 
    (\<Sum>(k, i)\<in>(SIGMA k:{0..<n}. {0..k}). f i (k - i))"
 | 
|
559  | 
proof (rule setsum_reindex_cong)  | 
|
560  | 
    show "inj_on (\<lambda>(k,i). (i, k - i)) (SIGMA k:{0..<n}. {0..k})"
 | 
|
561  | 
by (rule inj_on_inverseI [where g="\<lambda>(i,j). (i+j, i)"], auto)  | 
|
562  | 
    show "{(i,j). i + j < n} = (\<lambda>(k,i). (i, k - i)) ` (SIGMA k:{0..<n}. {0..k})"
 | 
|
563  | 
by (safe, rule_tac x="(a+b,a)" in image_eqI, auto)  | 
|
564  | 
show "\<And>a. (\<lambda>(k, i). f i (k - i)) a = split f ((\<lambda>(k, i). (i, k - i)) a)"  | 
|
565  | 
by clarify  | 
|
566  | 
qed  | 
|
567  | 
thus ?thesis by (simp add: setsum_Sigma)  | 
|
568  | 
qed  | 
|
569  | 
||
570  | 
lemma Cauchy_product_sums:  | 
|
571  | 
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
 | 
|
572  | 
assumes a: "summable (\<lambda>k. norm (a k))"  | 
|
573  | 
assumes b: "summable (\<lambda>k. norm (b k))"  | 
|
574  | 
shows "(\<lambda>k. \<Sum>i=0..k. a i * b (k - i)) sums ((\<Sum>k. a k) * (\<Sum>k. b k))"  | 
|
575  | 
proof -  | 
|
576  | 
  let ?S1 = "\<lambda>n::nat. {0..<n} \<times> {0..<n}"
 | 
|
577  | 
  let ?S2 = "\<lambda>n::nat. {(i,j). i + j < n}"
 | 
|
578  | 
have S1_mono: "\<And>m n. m \<le> n \<Longrightarrow> ?S1 m \<subseteq> ?S1 n" by auto  | 
|
579  | 
have S2_le_S1: "\<And>n. ?S2 n \<subseteq> ?S1 n" by auto  | 
|
580  | 
have S1_le_S2: "\<And>n. ?S1 (n div 2) \<subseteq> ?S2 n" by auto  | 
|
581  | 
have finite_S1: "\<And>n. finite (?S1 n)" by simp  | 
|
582  | 
with S2_le_S1 have finite_S2: "\<And>n. finite (?S2 n)" by (rule finite_subset)  | 
|
583  | 
||
584  | 
let ?g = "\<lambda>(i,j). a i * b j"  | 
|
585  | 
let ?f = "\<lambda>(i,j). norm (a i) * norm (b j)"  | 
|
586  | 
have f_nonneg: "\<And>x. 0 \<le> ?f x"  | 
|
587  | 
by (auto simp add: mult_nonneg_nonneg)  | 
|
588  | 
hence norm_setsum_f: "\<And>A. norm (setsum ?f A) = setsum ?f A"  | 
|
589  | 
unfolding real_norm_def  | 
|
590  | 
by (simp only: abs_of_nonneg setsum_nonneg [rule_format])  | 
|
591  | 
||
592  | 
have "(\<lambda>n. (\<Sum>k=0..<n. a k) * (\<Sum>k=0..<n. b k))  | 
|
593  | 
----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
594  | 
by (intro LIMSEQ_mult summable_sumr_LIMSEQ_suminf  | 
|
595  | 
summable_norm_cancel [OF a] summable_norm_cancel [OF b])  | 
|
596  | 
hence 1: "(\<lambda>n. setsum ?g (?S1 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
597  | 
by (simp only: setsum_product setsum_Sigma [rule_format]  | 
|
598  | 
finite_atLeastLessThan)  | 
|
599  | 
||
600  | 
have "(\<lambda>n. (\<Sum>k=0..<n. norm (a k)) * (\<Sum>k=0..<n. norm (b k)))  | 
|
601  | 
----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"  | 
|
602  | 
using a b by (intro LIMSEQ_mult summable_sumr_LIMSEQ_suminf)  | 
|
603  | 
hence "(\<lambda>n. setsum ?f (?S1 n)) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"  | 
|
604  | 
by (simp only: setsum_product setsum_Sigma [rule_format]  | 
|
605  | 
finite_atLeastLessThan)  | 
|
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)  | 
|
610  | 
have "Zseq (\<lambda>n. setsum ?f (?S1 n - ?S2 n))"  | 
|
611  | 
proof (rule ZseqI, simp only: norm_setsum_f)  | 
|
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  | 
|
632  | 
hence "Zseq (\<lambda>n. setsum ?g (?S1 n - ?S2 n))"  | 
|
633  | 
apply (rule Zseq_le [rule_format])  | 
|
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"  | 
|
639  | 
by (simp only: LIMSEQ_Zseq_iff setsum_diff finite_S1 S2_le_S1 diff_0_right)  | 
|
640  | 
||
641  | 
with 1 have "(\<lambda>n. setsum ?g (?S2 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
642  | 
by (rule LIMSEQ_diff_approach_zero2)  | 
|
643  | 
thus ?thesis by (simp only: sums_def setsum_triangle_reindex)  | 
|
644  | 
qed  | 
|
645  | 
||
646  | 
lemma Cauchy_product:  | 
|
647  | 
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
 | 
|
648  | 
assumes a: "summable (\<lambda>k. norm (a k))"  | 
|
649  | 
assumes b: "summable (\<lambda>k. norm (b k))"  | 
|
650  | 
shows "(\<Sum>k. a k) * (\<Sum>k. b k) = (\<Sum>k. \<Sum>i=0..k. a i * b (k - i))"  | 
|
| 23441 | 651  | 
using a b  | 
| 23111 | 652  | 
by (rule Cauchy_product_sums [THEN sums_unique])  | 
653  | 
||
| 14416 | 654  | 
end  |