| author | blanchet | 
| Thu, 06 Jun 2013 15:49:08 +0200 | |
| changeset 52319 | 37a3b00759dc | 
| parent 51528 | 66c3a7589de7 | 
| child 53602 | 0ae3db699a3e | 
| 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  | 
|
| 14416 | 10  | 
header{*Finite Summation and 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  | 
|
| 19765 | 16  | 
definition  | 
| 41970 | 17  | 
   sums  :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<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  | 
| 41970 | 22  | 
   summable :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<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  | 
| 41970 | 26  | 
   suminf   :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a" where
 | 
| 20688 | 27  | 
"suminf f = (THE s. f sums s)"  | 
| 14416 | 28  | 
|
| 44289 | 29  | 
notation suminf (binder "\<Sum>" 10)  | 
| 15546 | 30  | 
|
| 14416 | 31  | 
|
| 
32877
 
6f09346c7c08
New lemmas connected with the reals and infinite series
 
paulson 
parents: 
32707 
diff
changeset
 | 
32  | 
lemma [trans]: "f=g ==> g sums z ==> f sums z"  | 
| 
 
6f09346c7c08
New lemmas connected with the reals and infinite series
 
paulson 
parents: 
32707 
diff
changeset
 | 
33  | 
by simp  | 
| 
 
6f09346c7c08
New lemmas connected with the reals and infinite series
 
paulson 
parents: 
32707 
diff
changeset
 | 
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  | 
|
| 14416 | 81  | 
subsection{* Infinite Sums, by the Properties of Limits*}
 | 
82  | 
||
83  | 
(*----------------------  | 
|
| 41970 | 84  | 
suminf is the sum  | 
| 14416 | 85  | 
---------------------*)  | 
86  | 
lemma sums_summable: "f sums l ==> summable f"  | 
|
| 41970 | 87  | 
by (simp add: sums_def summable_def, blast)  | 
| 14416 | 88  | 
|
| 41970 | 89  | 
lemma summable_sums:  | 
| 46904 | 90  | 
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
 | 
91  | 
assumes "summable f"  | 
|
92  | 
shows "f sums (suminf f)"  | 
|
| 41970 | 93  | 
proof -  | 
| 46904 | 94  | 
  from assms obtain s where s: "(\<lambda>n. setsum f {0..<n}) ----> s"
 | 
95  | 
unfolding summable_def sums_def [abs_def] ..  | 
|
96  | 
then show ?thesis unfolding sums_def [abs_def] suminf_def  | 
|
| 41970 | 97  | 
by (rule theI, auto intro!: tendsto_unique[OF trivial_limit_sequentially])  | 
98  | 
qed  | 
|
| 14416 | 99  | 
|
| 41970 | 100  | 
lemma summable_sumr_LIMSEQ_suminf:  | 
101  | 
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
 | 
|
102  | 
  shows "summable f \<Longrightarrow> (\<lambda>n. setsum f {0..<n}) ----> suminf f"
 | 
|
| 20688 | 103  | 
by (rule summable_sums [unfolded sums_def])  | 
| 14416 | 104  | 
|
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
105  | 
lemma suminf_eq_lim: "suminf f = lim (%n. setsum f {0..<n})"
 | 
| 41970 | 106  | 
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
 | 
107  | 
|
| 14416 | 108  | 
(*-------------------  | 
| 41970 | 109  | 
sum is unique  | 
| 14416 | 110  | 
------------------*)  | 
| 41970 | 111  | 
lemma sums_unique:  | 
112  | 
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
 | 
|
113  | 
shows "f sums s \<Longrightarrow> (s = suminf f)"  | 
|
114  | 
apply (frule sums_summable[THEN summable_sums])  | 
|
115  | 
apply (auto intro!: tendsto_unique[OF trivial_limit_sequentially] simp add: sums_def)  | 
|
| 14416 | 116  | 
done  | 
117  | 
||
| 41970 | 118  | 
lemma sums_iff:  | 
119  | 
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
 | 
|
120  | 
shows "f sums x \<longleftrightarrow> summable f \<and> (suminf f = x)"  | 
|
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
121  | 
by (metis summable_sums sums_summable sums_unique)  | 
| 
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
122  | 
|
| 47761 | 123  | 
lemma sums_finite:  | 
124  | 
assumes [simp]: "finite N"  | 
|
125  | 
assumes f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"  | 
|
126  | 
shows "f sums (\<Sum>n\<in>N. f n)"  | 
|
127  | 
proof -  | 
|
128  | 
  { fix n
 | 
|
129  | 
    have "setsum f {..<n + Suc (Max N)} = setsum f N"
 | 
|
130  | 
proof cases  | 
|
131  | 
      assume "N = {}"
 | 
|
132  | 
with f have "f = (\<lambda>x. 0)" by auto  | 
|
133  | 
then show ?thesis by simp  | 
|
134  | 
next  | 
|
135  | 
      assume [simp]: "N \<noteq> {}"
 | 
|
136  | 
show ?thesis  | 
|
137  | 
proof (safe intro!: setsum_mono_zero_right f)  | 
|
138  | 
fix i assume "i \<in> N"  | 
|
139  | 
then have "i \<le> Max N" by simp  | 
|
140  | 
then show "i < n + Suc (Max N)" by simp  | 
|
141  | 
qed  | 
|
142  | 
qed }  | 
|
143  | 
note eq = this  | 
|
144  | 
show ?thesis unfolding sums_def  | 
|
145  | 
by (rule LIMSEQ_offset[of _ "Suc (Max N)"])  | 
|
146  | 
(simp add: eq atLeast0LessThan tendsto_const del: add_Suc_right)  | 
|
147  | 
qed  | 
|
148  | 
||
149  | 
lemma suminf_finite:  | 
|
150  | 
  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add,t2_space}"
 | 
|
151  | 
assumes N: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"  | 
|
152  | 
shows "suminf f = (\<Sum>n\<in>N. f n)"  | 
|
153  | 
using sums_finite[OF assms, THEN sums_unique] by simp  | 
|
154  | 
||
155  | 
lemma sums_If_finite_set:  | 
|
156  | 
  "finite A \<Longrightarrow> (\<lambda>r. if r \<in> A then f r else 0 :: 'a::{comm_monoid_add,t2_space}) sums (\<Sum>r\<in>A. f r)"
 | 
|
157  | 
using sums_finite[of A "(\<lambda>r. if r \<in> A then f r else 0)"] by simp  | 
|
158  | 
||
159  | 
lemma sums_If_finite:  | 
|
160  | 
  "finite {r. P r} \<Longrightarrow> (\<lambda>r. if P r then f r else 0 :: 'a::{comm_monoid_add,t2_space}) sums (\<Sum>r | P r. f r)"
 | 
|
161  | 
  using sums_If_finite_set[of "{r. P r}" f] by simp
 | 
|
162  | 
||
163  | 
lemma sums_single:  | 
|
164  | 
  "(\<lambda>r. if r = i then f r else 0::'a::{comm_monoid_add,t2_space}) sums f i"
 | 
|
165  | 
using sums_If_finite[of "\<lambda>r. r = i" f] by simp  | 
|
166  | 
||
| 41970 | 167  | 
lemma sums_split_initial_segment:  | 
168  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
|
169  | 
shows "f sums s ==> (\<lambda>n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"  | 
|
170  | 
apply (unfold sums_def)  | 
|
171  | 
apply (simp add: sumr_offset)  | 
|
| 44710 | 172  | 
apply (rule tendsto_diff [OF _ tendsto_const])  | 
| 16819 | 173  | 
apply (rule LIMSEQ_ignore_initial_segment)  | 
174  | 
apply assumption  | 
|
175  | 
done  | 
|
176  | 
||
| 41970 | 177  | 
lemma summable_ignore_initial_segment:  | 
178  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
|
179  | 
shows "summable f ==> summable (%n. f(n + k))"  | 
|
| 16819 | 180  | 
apply (unfold summable_def)  | 
181  | 
apply (auto intro: sums_split_initial_segment)  | 
|
182  | 
done  | 
|
183  | 
||
| 41970 | 184  | 
lemma suminf_minus_initial_segment:  | 
185  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
|
186  | 
shows "summable f ==>  | 
|
| 16819 | 187  | 
suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)"  | 
188  | 
apply (frule summable_ignore_initial_segment)  | 
|
189  | 
apply (rule sums_unique [THEN sym])  | 
|
190  | 
apply (frule summable_sums)  | 
|
191  | 
apply (rule sums_split_initial_segment)  | 
|
192  | 
apply auto  | 
|
193  | 
done  | 
|
194  | 
||
| 41970 | 195  | 
lemma suminf_split_initial_segment:  | 
196  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
|
197  | 
shows "summable f ==>  | 
|
198  | 
suminf f = (SUM i = 0..< k. f i) + (\<Sum>n. f(n + k))"  | 
|
| 16819 | 199  | 
by (auto simp add: suminf_minus_initial_segment)  | 
200  | 
||
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
201  | 
lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable a"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
202  | 
shows "\<exists> N. \<forall> n \<ge> N. \<bar> \<Sum> i. a (i + n) \<bar> < r"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
203  | 
proof -  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
204  | 
from LIMSEQ_D[OF summable_sumr_LIMSEQ_suminf[OF `summable a`] `0 < r`]  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
205  | 
  obtain N :: nat where "\<forall> n \<ge> N. norm (setsum a {0..<n} - suminf a) < r" by auto
 | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
206  | 
thus ?thesis unfolding suminf_minus_initial_segment[OF `summable a` refl] abs_minus_commute real_norm_def  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
207  | 
by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
208  | 
qed  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
209  | 
|
| 41970 | 210  | 
lemma sums_Suc:  | 
211  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
|
212  | 
assumes sumSuc: "(\<lambda> n. f (Suc n)) sums l" shows "f sums (l + f 0)"  | 
|
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
213  | 
proof -  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
214  | 
from sumSuc[unfolded sums_def]  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
215  | 
have "(\<lambda>i. \<Sum>n = Suc 0..<Suc i. f n) ----> l" unfolding setsum_reindex[OF inj_Suc] image_Suc_atLeastLessThan[symmetric] comp_def .  | 
| 44710 | 216  | 
from tendsto_add[OF this tendsto_const, where b="f 0"]  | 
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
217  | 
have "(\<lambda>i. \<Sum>n = 0..<Suc i. f n) ----> l + f 0" unfolding add_commute setsum_head_upt_Suc[OF zero_less_Suc] .  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
218  | 
thus ?thesis unfolding sums_def by (rule LIMSEQ_imp_Suc)  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
219  | 
qed  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29197 
diff
changeset
 | 
220  | 
|
| 41970 | 221  | 
lemma series_zero:  | 
222  | 
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
 | 
|
223  | 
assumes "\<forall>m. n \<le> m \<longrightarrow> f m = 0"  | 
|
224  | 
  shows "f sums (setsum f {0..<n})"
 | 
|
225  | 
proof -  | 
|
226  | 
  { fix k :: nat have "setsum f {0..<k + n} = setsum f {0..<n}"
 | 
|
227  | 
using assms by (induct k) auto }  | 
|
228  | 
note setsum_const = this  | 
|
229  | 
show ?thesis  | 
|
230  | 
unfolding sums_def  | 
|
231  | 
apply (rule LIMSEQ_offset[of _ n])  | 
|
232  | 
unfolding setsum_const  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
233  | 
apply (rule tendsto_const)  | 
| 41970 | 234  | 
done  | 
235  | 
qed  | 
|
| 14416 | 236  | 
|
| 41970 | 237  | 
lemma sums_zero[simp, intro]: "(\<lambda>n. 0) sums 0"  | 
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
238  | 
unfolding sums_def by (simp add: tendsto_const)  | 
| 15539 | 239  | 
|
| 41970 | 240  | 
lemma summable_zero[simp, intro]: "summable (\<lambda>n. 0)"  | 
| 23121 | 241  | 
by (rule sums_zero [THEN sums_summable])  | 
| 16819 | 242  | 
|
| 41970 | 243  | 
lemma suminf_zero[simp]: "suminf (\<lambda>n. 0::'a::{t2_space, comm_monoid_add}) = 0"
 | 
| 23121 | 244  | 
by (rule sums_zero [THEN sums_unique, symmetric])  | 
| 41970 | 245  | 
|
| 23119 | 246  | 
lemma (in bounded_linear) sums:  | 
247  | 
"(\<lambda>n. X n) sums a \<Longrightarrow> (\<lambda>n. f (X n)) sums (f a)"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
248  | 
unfolding sums_def by (drule tendsto, simp only: setsum)  | 
| 23119 | 249  | 
|
250  | 
lemma (in bounded_linear) summable:  | 
|
251  | 
"summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"  | 
|
252  | 
unfolding summable_def by (auto intro: sums)  | 
|
253  | 
||
254  | 
lemma (in bounded_linear) suminf:  | 
|
255  | 
"summable (\<lambda>n. X n) \<Longrightarrow> f (\<Sum>n. X n) = (\<Sum>n. f (X n))"  | 
|
| 23121 | 256  | 
by (intro sums_unique sums summable_sums)  | 
| 23119 | 257  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
258  | 
lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real]  | 
| 
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
259  | 
lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real]  | 
| 
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
260  | 
lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real]  | 
| 
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
261  | 
|
| 20692 | 262  | 
lemma sums_mult:  | 
263  | 
fixes c :: "'a::real_normed_algebra"  | 
|
264  | 
shows "f sums a \<Longrightarrow> (\<lambda>n. c * f n) sums (c * a)"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
265  | 
by (rule bounded_linear.sums [OF bounded_linear_mult_right])  | 
| 14416 | 266  | 
|
| 20692 | 267  | 
lemma summable_mult:  | 
268  | 
fixes c :: "'a::real_normed_algebra"  | 
|
| 23121 | 269  | 
shows "summable f \<Longrightarrow> summable (%n. c * f n)"  | 
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
270  | 
by (rule bounded_linear.summable [OF bounded_linear_mult_right])  | 
| 16819 | 271  | 
|
| 20692 | 272  | 
lemma suminf_mult:  | 
273  | 
fixes c :: "'a::real_normed_algebra"  | 
|
| 41970 | 274  | 
shows "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f"  | 
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
275  | 
by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric])  | 
| 16819 | 276  | 
|
| 20692 | 277  | 
lemma sums_mult2:  | 
278  | 
fixes c :: "'a::real_normed_algebra"  | 
|
279  | 
shows "f sums a \<Longrightarrow> (\<lambda>n. f n * c) sums (a * c)"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
280  | 
by (rule bounded_linear.sums [OF bounded_linear_mult_left])  | 
| 16819 | 281  | 
|
| 20692 | 282  | 
lemma summable_mult2:  | 
283  | 
fixes c :: "'a::real_normed_algebra"  | 
|
284  | 
shows "summable f \<Longrightarrow> summable (\<lambda>n. f n * c)"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
285  | 
by (rule bounded_linear.summable [OF bounded_linear_mult_left])  | 
| 16819 | 286  | 
|
| 20692 | 287  | 
lemma suminf_mult2:  | 
288  | 
fixes c :: "'a::real_normed_algebra"  | 
|
289  | 
shows "summable f \<Longrightarrow> suminf f * c = (\<Sum>n. f n * c)"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
290  | 
by (rule bounded_linear.suminf [OF bounded_linear_mult_left])  | 
| 16819 | 291  | 
|
| 20692 | 292  | 
lemma sums_divide:  | 
293  | 
fixes c :: "'a::real_normed_field"  | 
|
294  | 
shows "f sums a \<Longrightarrow> (\<lambda>n. f n / c) sums (a / c)"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
295  | 
by (rule bounded_linear.sums [OF bounded_linear_divide])  | 
| 14416 | 296  | 
|
| 20692 | 297  | 
lemma summable_divide:  | 
298  | 
fixes c :: "'a::real_normed_field"  | 
|
299  | 
shows "summable f \<Longrightarrow> summable (\<lambda>n. f n / c)"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
300  | 
by (rule bounded_linear.summable [OF bounded_linear_divide])  | 
| 16819 | 301  | 
|
| 20692 | 302  | 
lemma suminf_divide:  | 
303  | 
fixes c :: "'a::real_normed_field"  | 
|
304  | 
shows "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
305  | 
by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric])  | 
| 16819 | 306  | 
|
| 41970 | 307  | 
lemma sums_add:  | 
308  | 
fixes a b :: "'a::real_normed_field"  | 
|
309  | 
shows "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) sums (a + b)"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
310  | 
unfolding sums_def by (simp add: setsum_addf tendsto_add)  | 
| 16819 | 311  | 
|
| 41970 | 312  | 
lemma summable_add:  | 
313  | 
fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"  | 
|
314  | 
shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n + Y n)"  | 
|
| 23121 | 315  | 
unfolding summable_def by (auto intro: sums_add)  | 
| 16819 | 316  | 
|
317  | 
lemma suminf_add:  | 
|
| 41970 | 318  | 
fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"  | 
319  | 
shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X + suminf Y = (\<Sum>n. X n + Y n)"  | 
|
| 23121 | 320  | 
by (intro sums_unique sums_add summable_sums)  | 
| 14416 | 321  | 
|
| 41970 | 322  | 
lemma sums_diff:  | 
323  | 
fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"  | 
|
324  | 
shows "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) sums (a - b)"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
325  | 
unfolding sums_def by (simp add: setsum_subtractf tendsto_diff)  | 
| 23121 | 326  | 
|
| 41970 | 327  | 
lemma summable_diff:  | 
328  | 
fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"  | 
|
329  | 
shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n - Y n)"  | 
|
| 23121 | 330  | 
unfolding summable_def by (auto intro: sums_diff)  | 
| 14416 | 331  | 
|
332  | 
lemma suminf_diff:  | 
|
| 41970 | 333  | 
fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"  | 
334  | 
shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X - suminf Y = (\<Sum>n. X n - Y n)"  | 
|
| 23121 | 335  | 
by (intro sums_unique sums_diff summable_sums)  | 
| 14416 | 336  | 
|
| 41970 | 337  | 
lemma sums_minus:  | 
338  | 
fixes X :: "nat \<Rightarrow> 'a::real_normed_field"  | 
|
339  | 
shows "X sums a ==> (\<lambda>n. - X n) sums (- a)"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
340  | 
unfolding sums_def by (simp add: setsum_negf tendsto_minus)  | 
| 16819 | 341  | 
|
| 41970 | 342  | 
lemma summable_minus:  | 
343  | 
fixes X :: "nat \<Rightarrow> 'a::real_normed_field"  | 
|
344  | 
shows "summable X \<Longrightarrow> summable (\<lambda>n. - X n)"  | 
|
| 23121 | 345  | 
unfolding summable_def by (auto intro: sums_minus)  | 
| 16819 | 346  | 
|
| 41970 | 347  | 
lemma suminf_minus:  | 
348  | 
fixes X :: "nat \<Rightarrow> 'a::real_normed_field"  | 
|
349  | 
shows "summable X \<Longrightarrow> (\<Sum>n. - X n) = - (\<Sum>n. X n)"  | 
|
| 23121 | 350  | 
by (intro sums_unique [symmetric] sums_minus summable_sums)  | 
| 14416 | 351  | 
|
352  | 
lemma sums_group:  | 
|
| 41970 | 353  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_field"  | 
| 
44727
 
d45acd50a894
modify lemma sums_group, and shorten proofs that use it
 
huffman 
parents: 
44726 
diff
changeset
 | 
354  | 
  shows "\<lbrakk>f sums s; 0 < k\<rbrakk> \<Longrightarrow> (\<lambda>n. setsum f {n*k..<n*k+k}) sums s"
 | 
| 20692 | 355  | 
apply (simp only: sums_def sumr_group)  | 
| 
31336
 
e17f13cd1280
generalize constants in SEQ.thy to class metric_space
 
huffman 
parents: 
31017 
diff
changeset
 | 
356  | 
apply (unfold LIMSEQ_iff, safe)  | 
| 20692 | 357  | 
apply (drule_tac x="r" in spec, safe)  | 
358  | 
apply (rule_tac x="no" in exI, safe)  | 
|
359  | 
apply (drule_tac x="n*k" in spec)  | 
|
360  | 
apply (erule mp)  | 
|
361  | 
apply (erule order_trans)  | 
|
362  | 
apply simp  | 
|
| 14416 | 363  | 
done  | 
364  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
365  | 
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
 | 
366  | 
great as any partial sum.*}  | 
| 14416 | 367  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
368  | 
lemma pos_summable:  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
369  | 
fixes f:: "nat \<Rightarrow> real"  | 
| 50999 | 370  | 
  assumes pos: "\<And>n. 0 \<le> f n" and le: "\<And>n. setsum f {0..<n} \<le> x"
 | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
371  | 
shows "summable f"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
372  | 
proof -  | 
| 41970 | 373  | 
  have "convergent (\<lambda>n. setsum f {0..<n})"
 | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
374  | 
proof (rule Bseq_mono_convergent)  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
375  | 
      show "Bseq (\<lambda>n. setsum f {0..<n})"
 | 
| 
51477
 
2990382dc066
clean up lemma_nest_unique and renamed to nested_sequence_unique
 
hoelzl 
parents: 
50999 
diff
changeset
 | 
376  | 
by (intro BseqI'[of _ x]) (auto simp add: setsum_nonneg pos intro: le)  | 
| 41970 | 377  | 
next  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
378  | 
      show "\<forall>m n. m \<le> n \<longrightarrow> setsum f {0..<m} \<le> setsum f {0..<n}"
 | 
| 41970 | 379  | 
by (auto intro: setsum_mono2 pos)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
380  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
381  | 
thus ?thesis  | 
| 
51477
 
2990382dc066
clean up lemma_nest_unique and renamed to nested_sequence_unique
 
hoelzl 
parents: 
50999 
diff
changeset
 | 
382  | 
by (force simp add: summable_def sums_def convergent_def)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
383  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
384  | 
|
| 20692 | 385  | 
lemma series_pos_le:  | 
| 50999 | 386  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
| 20692 | 387  | 
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 \<le> f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} \<le> suminf f"
 | 
| 50999 | 388  | 
apply (drule summable_sums)  | 
389  | 
apply (simp add: sums_def)  | 
|
390  | 
apply (rule LIMSEQ_le_const)  | 
|
391  | 
apply assumption  | 
|
392  | 
apply (intro exI[of _ n])  | 
|
393  | 
apply (auto intro!: setsum_mono2)  | 
|
394  | 
done  | 
|
| 14416 | 395  | 
|
396  | 
lemma series_pos_less:  | 
|
| 50999 | 397  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_ab_semigroup_add_imp_le, ordered_comm_monoid_add, linorder_topology}"
 | 
| 20692 | 398  | 
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} < suminf f"
 | 
| 50999 | 399  | 
  apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
 | 
400  | 
  using add_less_cancel_left [of "setsum f {0..<n}" 0 "f n"]
 | 
|
401  | 
apply simp  | 
|
402  | 
apply (erule series_pos_le)  | 
|
403  | 
apply (simp add: order_less_imp_le)  | 
|
404  | 
done  | 
|
405  | 
||
406  | 
lemma suminf_eq_zero_iff:  | 
|
407  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
|
408  | 
shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"  | 
|
409  | 
proof  | 
|
410  | 
assume "summable f" "suminf f = 0" and pos: "\<forall>n. 0 \<le> f n"  | 
|
411  | 
then have "f sums 0"  | 
|
412  | 
by (simp add: sums_iff)  | 
|
413  | 
then have f: "(\<lambda>n. \<Sum>i<n. f i) ----> 0"  | 
|
414  | 
by (simp add: sums_def atLeast0LessThan)  | 
|
415  | 
  have "\<And>i. (\<Sum>n\<in>{i}. f n) \<le> 0"
 | 
|
416  | 
proof (rule LIMSEQ_le_const[OF f])  | 
|
417  | 
    fix i show "\<exists>N. \<forall>n\<ge>N. (\<Sum>n\<in>{i}. f n) \<le> setsum f {..<n}"
 | 
|
418  | 
using pos by (intro exI[of _ "Suc i"] allI impI setsum_mono2) auto  | 
|
419  | 
qed  | 
|
420  | 
with pos show "\<forall>n. f n = 0"  | 
|
421  | 
by (auto intro!: antisym)  | 
|
422  | 
next  | 
|
423  | 
assume "\<forall>n. f n = 0"  | 
|
424  | 
then have "f = (\<lambda>n. 0)"  | 
|
425  | 
by auto  | 
|
426  | 
then show "suminf f = 0"  | 
|
427  | 
by simp  | 
|
428  | 
qed  | 
|
429  | 
||
430  | 
lemma suminf_gt_zero_iff:  | 
|
431  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
|
432  | 
shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"  | 
|
433  | 
using series_pos_le[of f 0] suminf_eq_zero_iff[of f]  | 
|
434  | 
by (simp add: less_le)  | 
|
| 20692 | 435  | 
|
436  | 
lemma suminf_gt_zero:  | 
|
| 50999 | 437  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
| 20692 | 438  | 
shows "\<lbrakk>summable f; \<forall>n. 0 < f n\<rbrakk> \<Longrightarrow> 0 < suminf f"  | 
| 50999 | 439  | 
using suminf_gt_zero_iff[of f] by (simp add: less_imp_le)  | 
| 20692 | 440  | 
|
441  | 
lemma suminf_ge_zero:  | 
|
| 50999 | 442  | 
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
| 20692 | 443  | 
shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 \<le> suminf f"  | 
| 50999 | 444  | 
by (drule_tac n="0" in series_pos_le, simp_all)  | 
| 20692 | 445  | 
|
446  | 
lemma sumr_pos_lt_pair:  | 
|
447  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
448  | 
shows "\<lbrakk>summable f;  | 
|
449  | 
\<forall>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\<rbrakk>  | 
|
450  | 
      \<Longrightarrow> setsum f {0..<k} < suminf f"
 | 
|
| 
30082
 
43c5b7bfc791
make more proofs work whether or not One_nat_def is a simp rule
 
huffman 
parents: 
29803 
diff
changeset
 | 
451  | 
unfolding One_nat_def  | 
| 20692 | 452  | 
apply (subst suminf_split_initial_segment [where k="k"])  | 
453  | 
apply assumption  | 
|
454  | 
apply simp  | 
|
455  | 
apply (drule_tac k="k" in summable_ignore_initial_segment)  | 
|
| 
44727
 
d45acd50a894
modify lemma sums_group, and shorten proofs that use it
 
huffman 
parents: 
44726 
diff
changeset
 | 
456  | 
apply (drule_tac k="Suc (Suc 0)" in sums_group [OF summable_sums], simp)  | 
| 20692 | 457  | 
apply simp  | 
458  | 
apply (frule sums_unique)  | 
|
459  | 
apply (drule sums_summable)  | 
|
460  | 
apply simp  | 
|
461  | 
apply (erule suminf_gt_zero)  | 
|
462  | 
apply (simp add: add_ac)  | 
|
| 14416 | 463  | 
done  | 
464  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
465  | 
text{*Sum of a geometric progression.*}
 | 
| 14416 | 466  | 
|
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
16819 
diff
changeset
 | 
467  | 
lemmas sumr_geometric = geometric_sum [where 'a = real]  | 
| 14416 | 468  | 
|
| 20692 | 469  | 
lemma geometric_sums:  | 
| 31017 | 470  | 
  fixes x :: "'a::{real_normed_field}"
 | 
| 20692 | 471  | 
shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) sums (1 / (1 - x))"  | 
472  | 
proof -  | 
|
473  | 
assume less_1: "norm x < 1"  | 
|
474  | 
hence neq_1: "x \<noteq> 1" by auto  | 
|
475  | 
hence neq_0: "x - 1 \<noteq> 0" by simp  | 
|
476  | 
from less_1 have lim_0: "(\<lambda>n. x ^ n) ----> 0"  | 
|
477  | 
by (rule LIMSEQ_power_zero)  | 
|
| 
22719
 
c51667189bd3
lemma geometric_sum no longer needs class division_by_zero
 
huffman 
parents: 
21404 
diff
changeset
 | 
478  | 
hence "(\<lambda>n. x ^ n / (x - 1) - 1 / (x - 1)) ----> 0 / (x - 1) - 1 / (x - 1)"  | 
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
479  | 
using neq_0 by (intro tendsto_intros)  | 
| 20692 | 480  | 
hence "(\<lambda>n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)"  | 
481  | 
by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib)  | 
|
482  | 
thus "(\<lambda>n. x ^ n) sums (1 / (1 - x))"  | 
|
483  | 
by (simp add: sums_def geometric_sum neq_1)  | 
|
484  | 
qed  | 
|
485  | 
||
486  | 
lemma summable_geometric:  | 
|
| 31017 | 487  | 
  fixes x :: "'a::{real_normed_field}"
 | 
| 20692 | 488  | 
shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"  | 
489  | 
by (rule geometric_sums [THEN sums_summable])  | 
|
| 14416 | 490  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46904 
diff
changeset
 | 
491  | 
lemma half: "0 < 1 / (2::'a::linordered_field)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46904 
diff
changeset
 | 
492  | 
by simp  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
493  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
494  | 
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
 | 
495  | 
proof -  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
496  | 
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
 | 
497  | 
by auto  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
498  | 
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
 | 
499  | 
by simp  | 
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
41970 
diff
changeset
 | 
500  | 
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
 | 
501  | 
by simp  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
502  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents: 
32877 
diff
changeset
 | 
503  | 
|
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
504  | 
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
 | 
505  | 
|
| 15539 | 506  | 
lemma summable_convergent_sumr_iff:  | 
507  | 
 "summable f = convergent (%n. setsum f {0..<n})"
 | 
|
| 14416 | 508  | 
by (simp add: summable_def sums_def convergent_def)  | 
509  | 
||
| 41970 | 510  | 
lemma summable_LIMSEQ_zero:  | 
| 44726 | 511  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
| 41970 | 512  | 
shows "summable f \<Longrightarrow> f ----> 0"  | 
| 20689 | 513  | 
apply (drule summable_convergent_sumr_iff [THEN iffD1])  | 
| 20692 | 514  | 
apply (drule convergent_Cauchy)  | 
| 
31336
 
e17f13cd1280
generalize constants in SEQ.thy to class metric_space
 
huffman 
parents: 
31017 
diff
changeset
 | 
515  | 
apply (simp only: Cauchy_iff LIMSEQ_iff, safe)  | 
| 20689 | 516  | 
apply (drule_tac x="r" in spec, safe)  | 
517  | 
apply (rule_tac x="M" in exI, safe)  | 
|
518  | 
apply (drule_tac x="Suc n" in spec, simp)  | 
|
519  | 
apply (drule_tac x="n" in spec, simp)  | 
|
520  | 
done  | 
|
521  | 
||
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
522  | 
lemma suminf_le:  | 
| 50999 | 523  | 
  fixes x :: "'a :: {ordered_comm_monoid_add, linorder_topology}"
 | 
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
524  | 
  shows "summable f \<Longrightarrow> (!!n. setsum f {0..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
 | 
| 50999 | 525  | 
apply (drule summable_sums)  | 
526  | 
apply (simp add: sums_def)  | 
|
527  | 
apply (rule LIMSEQ_le_const2)  | 
|
528  | 
apply assumption  | 
|
529  | 
apply auto  | 
|
530  | 
done  | 
|
| 
32707
 
836ec9d0a0c8
New lemmas involving the real numbers, especially limits and series
 
paulson 
parents: 
31336 
diff
changeset
 | 
531  | 
|
| 14416 | 532  | 
lemma summable_Cauchy:  | 
| 41970 | 533  | 
"summable (f::nat \<Rightarrow> 'a::banach) =  | 
| 20848 | 534  | 
      (\<forall>e > 0. \<exists>N. \<forall>m \<ge> N. \<forall>n. norm (setsum f {m..<n}) < e)"
 | 
| 
31336
 
e17f13cd1280
generalize constants in SEQ.thy to class metric_space
 
huffman 
parents: 
31017 
diff
changeset
 | 
535  | 
apply (simp only: summable_convergent_sumr_iff Cauchy_convergent_iff [symmetric] Cauchy_iff, safe)  | 
| 20410 | 536  | 
apply (drule spec, drule (1) mp)  | 
537  | 
apply (erule exE, rule_tac x="M" in exI, clarify)  | 
|
538  | 
apply (rule_tac x="m" and y="n" in linorder_le_cases)  | 
|
539  | 
apply (frule (1) order_trans)  | 
|
540  | 
apply (drule_tac x="n" in spec, drule (1) mp)  | 
|
541  | 
apply (drule_tac x="m" in spec, drule (1) mp)  | 
|
542  | 
apply (simp add: setsum_diff [symmetric])  | 
|
543  | 
apply simp  | 
|
544  | 
apply (drule spec, drule (1) mp)  | 
|
545  | 
apply (erule exE, rule_tac x="N" in exI, clarify)  | 
|
546  | 
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
 | 
547  | 
apply (subst norm_minus_commute)  | 
| 20410 | 548  | 
apply (simp add: setsum_diff [symmetric])  | 
549  | 
apply (simp add: setsum_diff [symmetric])  | 
|
| 14416 | 550  | 
done  | 
551  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
552  | 
text{*Comparison test*}
 | 
| 
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
553  | 
|
| 20692 | 554  | 
lemma norm_setsum:  | 
555  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
556  | 
shows "norm (setsum f A) \<le> (\<Sum>i\<in>A. norm (f i))"  | 
|
557  | 
apply (case_tac "finite A")  | 
|
558  | 
apply (erule finite_induct)  | 
|
559  | 
apply simp  | 
|
560  | 
apply simp  | 
|
561  | 
apply (erule order_trans [OF norm_triangle_ineq add_left_mono])  | 
|
562  | 
apply simp  | 
|
563  | 
done  | 
|
564  | 
||
| 14416 | 565  | 
lemma summable_comparison_test:  | 
| 20848 | 566  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
567  | 
shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f"  | 
|
| 20692 | 568  | 
apply (simp add: summable_Cauchy, safe)  | 
569  | 
apply (drule_tac x="e" in spec, safe)  | 
|
570  | 
apply (rule_tac x = "N + Na" in exI, safe)  | 
|
| 14416 | 571  | 
apply (rotate_tac 2)  | 
572  | 
apply (drule_tac x = m in spec)  | 
|
573  | 
apply (auto, rotate_tac 2, drule_tac x = n in spec)  | 
|
| 20848 | 574  | 
apply (rule_tac y = "\<Sum>k=m..<n. norm (f k)" in order_le_less_trans)  | 
575  | 
apply (rule norm_setsum)  | 
|
| 15539 | 576  | 
apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
 | 
| 22998 | 577  | 
apply (auto intro: setsum_mono simp add: abs_less_iff)  | 
| 14416 | 578  | 
done  | 
579  | 
||
| 20848 | 580  | 
lemma summable_norm_comparison_test:  | 
581  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
|
582  | 
shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk>  | 
|
583  | 
\<Longrightarrow> summable (\<lambda>n. norm (f n))"  | 
|
584  | 
apply (rule summable_comparison_test)  | 
|
585  | 
apply (auto)  | 
|
586  | 
done  | 
|
587  | 
||
| 14416 | 588  | 
lemma summable_rabs_comparison_test:  | 
| 20692 | 589  | 
fixes f :: "nat \<Rightarrow> real"  | 
590  | 
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 | 591  | 
apply (rule summable_comparison_test)  | 
| 15543 | 592  | 
apply (auto)  | 
| 14416 | 593  | 
done  | 
594  | 
||
| 23084 | 595  | 
text{*Summability of geometric series for real algebras*}
 | 
596  | 
||
597  | 
lemma complete_algebra_summable_geometric:  | 
|
| 31017 | 598  | 
  fixes x :: "'a::{real_normed_algebra_1,banach}"
 | 
| 23084 | 599  | 
shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"  | 
600  | 
proof (rule summable_comparison_test)  | 
|
601  | 
show "\<exists>N. \<forall>n\<ge>N. norm (x ^ n) \<le> norm x ^ n"  | 
|
602  | 
by (simp add: norm_power_ineq)  | 
|
603  | 
show "norm x < 1 \<Longrightarrow> summable (\<lambda>n. norm x ^ n)"  | 
|
604  | 
by (simp add: summable_geometric)  | 
|
605  | 
qed  | 
|
606  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
607  | 
text{*Limit comparison property for series (c.f. jrh)*}
 | 
| 
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
608  | 
|
| 14416 | 609  | 
lemma summable_le:  | 
| 50999 | 610  | 
  fixes f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
 | 
| 20692 | 611  | 
shows "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"  | 
| 14416 | 612  | 
apply (drule summable_sums)+  | 
| 20692 | 613  | 
apply (simp only: sums_def, erule (1) LIMSEQ_le)  | 
| 14416 | 614  | 
apply (rule exI)  | 
| 15539 | 615  | 
apply (auto intro!: setsum_mono)  | 
| 14416 | 616  | 
done  | 
617  | 
||
618  | 
lemma summable_le2:  | 
|
| 20692 | 619  | 
fixes f g :: "nat \<Rightarrow> real"  | 
620  | 
shows "\<lbrakk>\<forall>n. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f \<and> suminf f \<le> suminf g"  | 
|
| 20848 | 621  | 
apply (subgoal_tac "summable f")  | 
622  | 
apply (auto intro!: summable_le)  | 
|
| 22998 | 623  | 
apply (simp add: abs_le_iff)  | 
| 20848 | 624  | 
apply (rule_tac g="g" in summable_comparison_test, simp_all)  | 
| 14416 | 625  | 
done  | 
626  | 
||
| 
19106
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
627  | 
(* 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
 | 
628  | 
lemma suminf_0_le:  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
629  | 
fixes f::"nat\<Rightarrow>real"  | 
| 
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
630  | 
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
 | 
631  | 
shows "0 \<le> suminf f"  | 
| 50999 | 632  | 
using suminf_ge_zero[OF sm gt0] by simp  | 
| 
19106
 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 
kleing 
parents: 
17149 
diff
changeset
 | 
633  | 
|
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
634  | 
text{*Absolute convergence imples normal convergence*}
 | 
| 20848 | 635  | 
lemma summable_norm_cancel:  | 
636  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
|
637  | 
shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"  | 
|
| 20692 | 638  | 
apply (simp only: summable_Cauchy, safe)  | 
639  | 
apply (drule_tac x="e" in spec, safe)  | 
|
640  | 
apply (rule_tac x="N" in exI, safe)  | 
|
641  | 
apply (drule_tac x="m" in spec, safe)  | 
|
| 20848 | 642  | 
apply (rule order_le_less_trans [OF norm_setsum])  | 
643  | 
apply (rule order_le_less_trans [OF abs_ge_self])  | 
|
| 20692 | 644  | 
apply simp  | 
| 14416 | 645  | 
done  | 
646  | 
||
| 20848 | 647  | 
lemma summable_rabs_cancel:  | 
648  | 
fixes f :: "nat \<Rightarrow> real"  | 
|
649  | 
shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> summable f"  | 
|
650  | 
by (rule summable_norm_cancel, simp)  | 
|
651  | 
||
| 
15085
 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 
paulson 
parents: 
15053 
diff
changeset
 | 
652  | 
text{*Absolute convergence of series*}
 | 
| 20848 | 653  | 
lemma summable_norm:  | 
654  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
|
655  | 
shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
656  | 
by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel  | 
| 20848 | 657  | 
summable_sumr_LIMSEQ_suminf norm_setsum)  | 
658  | 
||
| 14416 | 659  | 
lemma summable_rabs:  | 
| 20692 | 660  | 
fixes f :: "nat \<Rightarrow> real"  | 
661  | 
shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> \<bar>suminf f\<bar> \<le> (\<Sum>n. \<bar>f n\<bar>)"  | 
|
| 20848 | 662  | 
by (fold real_norm_def, rule summable_norm)  | 
| 14416 | 663  | 
|
664  | 
subsection{* The Ratio Test*}
 | 
|
665  | 
||
| 20848 | 666  | 
lemma norm_ratiotest_lemma:  | 
| 22852 | 667  | 
fixes x y :: "'a::real_normed_vector"  | 
| 20848 | 668  | 
shows "\<lbrakk>c \<le> 0; norm x \<le> c * norm y\<rbrakk> \<Longrightarrow> x = 0"  | 
669  | 
apply (subgoal_tac "norm x \<le> 0", simp)  | 
|
670  | 
apply (erule order_trans)  | 
|
671  | 
apply (simp add: mult_le_0_iff)  | 
|
672  | 
done  | 
|
673  | 
||
| 14416 | 674  | 
lemma rabs_ratiotest_lemma: "[| c \<le> 0; abs x \<le> c * abs y |] ==> x = (0::real)"  | 
| 20848 | 675  | 
by (erule norm_ratiotest_lemma, simp)  | 
| 14416 | 676  | 
|
| 50331 | 677  | 
(* TODO: MOVE *)  | 
| 14416 | 678  | 
lemma le_Suc_ex: "(k::nat) \<le> l ==> (\<exists>n. l = k + n)"  | 
679  | 
apply (drule le_imp_less_or_eq)  | 
|
680  | 
apply (auto dest: less_imp_Suc_add)  | 
|
681  | 
done  | 
|
682  | 
||
683  | 
lemma le_Suc_ex_iff: "((k::nat) \<le> l) = (\<exists>n. l = k + n)"  | 
|
684  | 
by (auto simp add: le_Suc_ex)  | 
|
685  | 
||
686  | 
(*All this trouble just to get 0<c *)  | 
|
687  | 
lemma ratio_test_lemma2:  | 
|
| 20848 | 688  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
689  | 
shows "\<lbrakk>\<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> 0 < c \<or> summable f"  | 
|
| 14416 | 690  | 
apply (simp (no_asm) add: linorder_not_le [symmetric])  | 
691  | 
apply (simp add: summable_Cauchy)  | 
|
| 15543 | 692  | 
apply (safe, subgoal_tac "\<forall>n. N < n --> f (n) = 0")  | 
693  | 
prefer 2  | 
|
694  | 
apply clarify  | 
|
| 
30082
 
43c5b7bfc791
make more proofs work whether or not One_nat_def is a simp rule
 
huffman 
parents: 
29803 
diff
changeset
 | 
695  | 
apply(erule_tac x = "n - Suc 0" in allE)  | 
| 15543 | 696  | 
apply (simp add:diff_Suc split:nat.splits)  | 
| 20848 | 697  | 
apply (blast intro: norm_ratiotest_lemma)  | 
| 14416 | 698  | 
apply (rule_tac x = "Suc N" in exI, clarify)  | 
| 44710 | 699  | 
apply(simp cong del: setsum_cong cong: setsum_ivl_cong)  | 
| 14416 | 700  | 
done  | 
701  | 
||
702  | 
lemma ratio_test:  | 
|
| 20848 | 703  | 
fixes f :: "nat \<Rightarrow> 'a::banach"  | 
704  | 
shows "\<lbrakk>c < 1; \<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> summable f"  | 
|
| 14416 | 705  | 
apply (frule ratio_test_lemma2, auto)  | 
| 41970 | 706  | 
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
 | 
707  | 
in summable_comparison_test)  | 
| 14416 | 708  | 
apply (rule_tac x = N in exI, safe)  | 
709  | 
apply (drule le_Suc_ex_iff [THEN iffD1])  | 
|
| 22959 | 710  | 
apply (auto simp add: power_add field_power_not_zero)  | 
| 15539 | 711  | 
apply (induct_tac "na", auto)  | 
| 20848 | 712  | 
apply (rule_tac y = "c * norm (f (N + n))" in order_trans)  | 
| 14416 | 713  | 
apply (auto intro: mult_right_mono simp add: summable_def)  | 
| 20848 | 714  | 
apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)  | 
| 41970 | 715  | 
apply (rule sums_divide)  | 
| 27108 | 716  | 
apply (rule sums_mult)  | 
| 
15234
 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 
paulson 
parents: 
15229 
diff
changeset
 | 
717  | 
apply (auto intro!: geometric_sums)  | 
| 14416 | 718  | 
done  | 
719  | 
||
| 23111 | 720  | 
subsection {* Cauchy Product Formula *}
 | 
721  | 
||
722  | 
(* Proof based on Analysis WebNotes: Chapter 07, Class 41  | 
|
723  | 
http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm *)  | 
|
724  | 
||
725  | 
lemma setsum_triangle_reindex:  | 
|
726  | 
fixes n :: nat  | 
|
727  | 
  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))"
 | 
|
728  | 
proof -  | 
|
729  | 
  have "(\<Sum>(i, j)\<in>{(i, j). i + j < n}. f i j) =
 | 
|
730  | 
    (\<Sum>(k, i)\<in>(SIGMA k:{0..<n}. {0..k}). f i (k - i))"
 | 
|
731  | 
proof (rule setsum_reindex_cong)  | 
|
732  | 
    show "inj_on (\<lambda>(k,i). (i, k - i)) (SIGMA k:{0..<n}. {0..k})"
 | 
|
733  | 
by (rule inj_on_inverseI [where g="\<lambda>(i,j). (i+j, i)"], auto)  | 
|
734  | 
    show "{(i,j). i + j < n} = (\<lambda>(k,i). (i, k - i)) ` (SIGMA k:{0..<n}. {0..k})"
 | 
|
735  | 
by (safe, rule_tac x="(a+b,a)" in image_eqI, auto)  | 
|
736  | 
show "\<And>a. (\<lambda>(k, i). f i (k - i)) a = split f ((\<lambda>(k, i). (i, k - i)) a)"  | 
|
737  | 
by clarify  | 
|
738  | 
qed  | 
|
739  | 
thus ?thesis by (simp add: setsum_Sigma)  | 
|
740  | 
qed  | 
|
741  | 
||
742  | 
lemma Cauchy_product_sums:  | 
|
743  | 
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
 | 
|
744  | 
assumes a: "summable (\<lambda>k. norm (a k))"  | 
|
745  | 
assumes b: "summable (\<lambda>k. norm (b k))"  | 
|
746  | 
shows "(\<lambda>k. \<Sum>i=0..k. a i * b (k - i)) sums ((\<Sum>k. a k) * (\<Sum>k. b k))"  | 
|
747  | 
proof -  | 
|
748  | 
  let ?S1 = "\<lambda>n::nat. {0..<n} \<times> {0..<n}"
 | 
|
749  | 
  let ?S2 = "\<lambda>n::nat. {(i,j). i + j < n}"
 | 
|
750  | 
have S1_mono: "\<And>m n. m \<le> n \<Longrightarrow> ?S1 m \<subseteq> ?S1 n" by auto  | 
|
751  | 
have S2_le_S1: "\<And>n. ?S2 n \<subseteq> ?S1 n" by auto  | 
|
752  | 
have S1_le_S2: "\<And>n. ?S1 (n div 2) \<subseteq> ?S2 n" by auto  | 
|
753  | 
have finite_S1: "\<And>n. finite (?S1 n)" by simp  | 
|
754  | 
with S2_le_S1 have finite_S2: "\<And>n. finite (?S2 n)" by (rule finite_subset)  | 
|
755  | 
||
756  | 
let ?g = "\<lambda>(i,j). a i * b j"  | 
|
757  | 
let ?f = "\<lambda>(i,j). norm (a i) * norm (b j)"  | 
|
758  | 
have f_nonneg: "\<And>x. 0 \<le> ?f x"  | 
|
759  | 
by (auto simp add: mult_nonneg_nonneg)  | 
|
760  | 
hence norm_setsum_f: "\<And>A. norm (setsum ?f A) = setsum ?f A"  | 
|
761  | 
unfolding real_norm_def  | 
|
762  | 
by (simp only: abs_of_nonneg setsum_nonneg [rule_format])  | 
|
763  | 
||
764  | 
have "(\<lambda>n. (\<Sum>k=0..<n. a k) * (\<Sum>k=0..<n. b k))  | 
|
765  | 
----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
766  | 
by (intro tendsto_mult summable_sumr_LIMSEQ_suminf  | 
| 23111 | 767  | 
summable_norm_cancel [OF a] summable_norm_cancel [OF b])  | 
768  | 
hence 1: "(\<lambda>n. setsum ?g (?S1 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
769  | 
by (simp only: setsum_product setsum_Sigma [rule_format]  | 
|
770  | 
finite_atLeastLessThan)  | 
|
771  | 
||
772  | 
have "(\<lambda>n. (\<Sum>k=0..<n. norm (a k)) * (\<Sum>k=0..<n. norm (b k)))  | 
|
773  | 
----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"  | 
|
| 
44568
 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 
huffman 
parents: 
44289 
diff
changeset
 | 
774  | 
using a b by (intro tendsto_mult summable_sumr_LIMSEQ_suminf)  | 
| 23111 | 775  | 
hence "(\<lambda>n. setsum ?f (?S1 n)) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"  | 
776  | 
by (simp only: setsum_product setsum_Sigma [rule_format]  | 
|
777  | 
finite_atLeastLessThan)  | 
|
778  | 
hence "convergent (\<lambda>n. setsum ?f (?S1 n))"  | 
|
779  | 
by (rule convergentI)  | 
|
780  | 
hence Cauchy: "Cauchy (\<lambda>n. setsum ?f (?S1 n))"  | 
|
781  | 
by (rule convergent_Cauchy)  | 
|
| 36657 | 782  | 
have "Zfun (\<lambda>n. setsum ?f (?S1 n - ?S2 n)) sequentially"  | 
783  | 
proof (rule ZfunI, simp only: eventually_sequentially norm_setsum_f)  | 
|
| 23111 | 784  | 
fix r :: real  | 
785  | 
assume r: "0 < r"  | 
|
786  | 
from CauchyD [OF Cauchy r] obtain N  | 
|
787  | 
where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" ..  | 
|
788  | 
hence "\<And>m n. \<lbrakk>N \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> norm (setsum ?f (?S1 m - ?S1 n)) < r"  | 
|
789  | 
by (simp only: setsum_diff finite_S1 S1_mono)  | 
|
790  | 
hence N: "\<And>m n. \<lbrakk>N \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> setsum ?f (?S1 m - ?S1 n) < r"  | 
|
791  | 
by (simp only: norm_setsum_f)  | 
|
792  | 
show "\<exists>N. \<forall>n\<ge>N. setsum ?f (?S1 n - ?S2 n) < r"  | 
|
793  | 
proof (intro exI allI impI)  | 
|
794  | 
fix n assume "2 * N \<le> n"  | 
|
795  | 
hence n: "N \<le> n div 2" by simp  | 
|
796  | 
have "setsum ?f (?S1 n - ?S2 n) \<le> setsum ?f (?S1 n - ?S1 (n div 2))"  | 
|
797  | 
by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg  | 
|
798  | 
Diff_mono subset_refl S1_le_S2)  | 
|
799  | 
also have "\<dots> < r"  | 
|
800  | 
using n div_le_dividend by (rule N)  | 
|
801  | 
finally show "setsum ?f (?S1 n - ?S2 n) < r" .  | 
|
802  | 
qed  | 
|
803  | 
qed  | 
|
| 36657 | 804  | 
hence "Zfun (\<lambda>n. setsum ?g (?S1 n - ?S2 n)) sequentially"  | 
805  | 
apply (rule Zfun_le [rule_format])  | 
|
| 23111 | 806  | 
apply (simp only: norm_setsum_f)  | 
807  | 
apply (rule order_trans [OF norm_setsum setsum_mono])  | 
|
808  | 
apply (auto simp add: norm_mult_ineq)  | 
|
809  | 
done  | 
|
810  | 
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
 | 
811  | 
unfolding tendsto_Zfun_iff diff_0_right  | 
| 36657 | 812  | 
by (simp only: setsum_diff finite_S1 S2_le_S1)  | 
| 23111 | 813  | 
|
814  | 
with 1 have "(\<lambda>n. setsum ?g (?S2 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"  | 
|
815  | 
by (rule LIMSEQ_diff_approach_zero2)  | 
|
816  | 
thus ?thesis by (simp only: sums_def setsum_triangle_reindex)  | 
|
817  | 
qed  | 
|
818  | 
||
819  | 
lemma Cauchy_product:  | 
|
820  | 
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
 | 
|
821  | 
assumes a: "summable (\<lambda>k. norm (a k))"  | 
|
822  | 
assumes b: "summable (\<lambda>k. norm (b k))"  | 
|
823  | 
shows "(\<Sum>k. a k) * (\<Sum>k. b k) = (\<Sum>k. \<Sum>i=0..k. a i * b (k - i))"  | 
|
| 23441 | 824  | 
using a b  | 
| 23111 | 825  | 
by (rule Cauchy_product_sums [THEN sums_unique])  | 
826  | 
||
| 14416 | 827  | 
end  |