src/HOL/Series.thy
author blanchet
Thu, 16 Jan 2014 16:20:17 +0100
changeset 55017 2df6ad1dbd66
parent 54703 499f92dc6e45
child 56178 2a6f58938573
permissions -rw-r--r--
adapted to move of Wfrec
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : Series.thy
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
     4
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
     5
Converted to Isar and polished by lcp
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
     6
Converted to setsum and polished yet more by TNN
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
     7
Additional contributions by Jeremy Avigad
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
     8
*)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     9
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    10
header{*Finite Summation and Infinite Series*}
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    11
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    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
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    14
begin
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15546
diff changeset
    15
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19279
diff changeset
    16
definition
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    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
dfe940911617 misc cleanup;
wenzelm
parents: 19279
diff changeset
    19
   "f sums s = (%n. setsum f {0..<n}) ----> s"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    20
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21141
diff changeset
    21
definition
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    22
   summable :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> bool" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19279
diff changeset
    23
   "summable f = (\<exists>s. f sums s)"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    24
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21141
diff changeset
    25
definition
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    26
   suminf   :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a" where
20688
690d866a165d change definitions from SOME to THE
huffman
parents: 20552
diff changeset
    27
   "suminf f = (THE s. f sums s)"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    28
44289
d81d09cdab9c optimize some proofs
huffman
parents: 44282
diff changeset
    29
notation suminf (binder "\<Sum>" 10)
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15543
diff changeset
    30
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    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
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    35
lemma sumr_diff_mult_const:
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    36
 "setsum f {0..<n} - (real n*r) = setsum (%i. f i - r) {0..<n::nat}"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53602
diff changeset
    37
  by (simp add: setsum_subtractf real_of_nat_def)
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15360
diff changeset
    38
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
    39
lemma real_setsum_nat_ivl_bounded:
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
    40
     "(!!p. p < n \<Longrightarrow> f(p) \<le> K)
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
    41
      \<Longrightarrow> setsum f {0..<n::nat} \<le> real n * K"
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
    42
using setsum_bounded[where A = "{0..<n}"]
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
    43
by (auto simp:real_of_nat_def)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    44
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    45
(* Generalize from real to some algebraic structure? *)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    46
lemma sumr_minus_one_realpow_zero [simp]:
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
    47
  "(\<Sum>i=0..<2*n. (-1) ^ Suc i) = (0::real)"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
    48
by (induct "n", auto)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    49
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    50
(* FIXME this is an awful lemma! *)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    51
lemma sumr_one_lb_realpow_zero [simp]:
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    52
  "(\<Sum>n=Suc 0..<n. f(n) * (0::real) ^ n) = 0"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    53
by (rule setsum_0', simp)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    54
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
    55
lemma sumr_group:
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    56
     "(\<Sum>m=0..<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {0 ..< n * k}"
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
    57
apply (subgoal_tac "k = 0 | 0 < k", auto)
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
    58
apply (induct "n")
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    59
apply (simp_all add: setsum_add_nat_ivl add_commute)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    60
done
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
    61
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    62
lemma sumr_offset3:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    63
  "setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)) + setsum f {0..<k}"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    64
apply (subst setsum_shift_bounds_nat_ivl [symmetric])
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    65
apply (simp add: setsum_add_nat_ivl add_commute)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    66
done
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    67
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    68
lemma sumr_offset:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    69
  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    70
  shows "(\<Sum>m=0..<n. f(m+k)) = setsum f {0..<n+k} - setsum f {0..<k}"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    71
by (simp add: sumr_offset3)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    72
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    73
lemma sumr_offset2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    74
 "\<forall>f. (\<Sum>m=0..<n::nat. f(m+k)::real) = setsum f {0..<n+k} - setsum f {0..<k}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    75
by (simp add: sumr_offset)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    76
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    77
lemma sumr_offset4:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    78
  "\<forall>n f. setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)::real) + setsum f {0..<k}"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    79
by (clarify, rule sumr_offset3)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    80
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    81
subsection{* Infinite Sums, by the Properties of Limits*}
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    82
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    83
(*----------------------
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    84
   suminf is the sum
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    85
 ---------------------*)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    86
lemma sums_summable: "f sums l ==> summable f"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    87
  by (simp add: sums_def summable_def, blast)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    88
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    89
lemma summable_sums:
46904
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 44727
diff changeset
    90
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 44727
diff changeset
    91
  assumes "summable f"
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 44727
diff changeset
    92
  shows "f sums (suminf f)"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    93
proof -
46904
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 44727
diff changeset
    94
  from assms obtain s where s: "(\<lambda>n. setsum f {0..<n}) ----> s"
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 44727
diff changeset
    95
    unfolding summable_def sums_def [abs_def] ..
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 44727
diff changeset
    96
  then show ?thesis unfolding sums_def [abs_def] suminf_def
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    97
    by (rule theI, auto intro!: tendsto_unique[OF trivial_limit_sequentially])
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
    98
qed
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    99
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   100
lemma summable_sumr_LIMSEQ_suminf:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   101
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   102
  shows "summable f \<Longrightarrow> (\<lambda>n. setsum f {0..<n}) ----> suminf f"
20688
690d866a165d change definitions from SOME to THE
huffman
parents: 20552
diff changeset
   103
by (rule summable_sums [unfolded sums_def])
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   108
(*-------------------
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   109
    sum is unique
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   110
 ------------------*)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   111
lemma sums_unique:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   112
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   113
  shows "f sums s \<Longrightarrow> (s = suminf f)"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   114
apply (frule sums_summable[THEN summable_sums])
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   115
apply (auto intro!: tendsto_unique[OF trivial_limit_sequentially] simp add: sums_def)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   116
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   117
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   118
lemma sums_iff:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   119
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   123
lemma sums_finite:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   124
  assumes [simp]: "finite N"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   125
  assumes f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   126
  shows "f sums (\<Sum>n\<in>N. f n)"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   127
proof -
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   128
  { fix n
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   129
    have "setsum f {..<n + Suc (Max N)} = setsum f N"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   130
    proof cases
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   131
      assume "N = {}"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   132
      with f have "f = (\<lambda>x. 0)" by auto
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   133
      then show ?thesis by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   134
    next
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   135
      assume [simp]: "N \<noteq> {}"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   136
      show ?thesis
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   137
      proof (safe intro!: setsum_mono_zero_right f)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   138
        fix i assume "i \<in> N"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   139
        then have "i \<le> Max N" by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   140
        then show "i < n + Suc (Max N)" by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   141
      qed
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   142
    qed }
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   143
  note eq = this
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   144
  show ?thesis unfolding sums_def
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   145
    by (rule LIMSEQ_offset[of _ "Suc (Max N)"])
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   146
       (simp add: eq atLeast0LessThan tendsto_const del: add_Suc_right)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   147
qed
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   148
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   149
lemma suminf_finite:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   150
  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add,t2_space}"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   151
  assumes N: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   152
  shows "suminf f = (\<Sum>n\<in>N. f n)"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   153
  using sums_finite[OF assms, THEN sums_unique] by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   154
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   155
lemma sums_If_finite_set:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   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)"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   157
  using sums_finite[of A "(\<lambda>r. if r \<in> A then f r else 0)"] by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   158
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   159
lemma sums_If_finite:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   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)"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   161
  using sums_If_finite_set[of "{r. P r}" f] by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   162
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   163
lemma sums_single:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   164
  "(\<lambda>r. if r = i then f r else 0::'a::{comm_monoid_add,t2_space}) sums f i"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   165
  using sums_If_finite[of "\<lambda>r. r = i" f] by simp
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47108
diff changeset
   166
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   167
lemma sums_split_initial_segment:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   168
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   169
  shows "f sums s ==> (\<lambda>n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   170
  apply (unfold sums_def)
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   171
  apply (simp add: sumr_offset)
44710
9caf6883f1f4 remove redundant lemmas about LIMSEQ
huffman
parents: 44568
diff changeset
   172
  apply (rule tendsto_diff [OF _ tendsto_const])
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   173
  apply (rule LIMSEQ_ignore_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   174
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   175
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   176
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   177
lemma summable_ignore_initial_segment:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   178
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   179
  shows "summable f ==> summable (%n. f(n + k))"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   180
  apply (unfold summable_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   181
  apply (auto intro: sums_split_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   182
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   183
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   184
lemma suminf_minus_initial_segment:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   185
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   186
  shows "summable f ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   187
    suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   188
  apply (frule summable_ignore_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   189
  apply (rule sums_unique [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   190
  apply (frule summable_sums)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   191
  apply (rule sums_split_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   192
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   193
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   194
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   195
lemma suminf_split_initial_segment:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   196
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   197
  shows "summable f ==>
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   198
    suminf f = (SUM i = 0..< k. f i) + (\<Sum>n. f(n + k))"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   199
by (auto simp add: suminf_minus_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   210
lemma sums_Suc:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   211
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
9caf6883f1f4 remove redundant lemmas about LIMSEQ
huffman
parents: 44568
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   221
lemma series_zero:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   222
  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   223
  assumes "\<forall>m. n \<le> m \<longrightarrow> f m = 0"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   224
  shows "f sums (setsum f {0..<n})"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   225
proof -
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   226
  { fix k :: nat have "setsum f {0..<k + n} = setsum f {0..<n}"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   227
      using assms by (induct k) auto }
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   228
  note setsum_const = this
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   229
  show ?thesis
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   230
    unfolding sums_def
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   231
    apply (rule LIMSEQ_offset[of _ n])
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   234
    done
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   235
qed
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   236
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   239
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   240
lemma summable_zero[simp, intro]: "summable (\<lambda>n. 0)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   241
by (rule sums_zero [THEN sums_summable])
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   242
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   243
lemma suminf_zero[simp]: "suminf (\<lambda>n. 0::'a::{t2_space, comm_monoid_add}) = 0"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   244
by (rule sums_zero [THEN sums_unique, symmetric])
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   245
23119
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   246
lemma (in bounded_linear) sums:
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   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
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   249
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   250
lemma (in bounded_linear) summable:
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   251
  "summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   252
unfolding summable_def by (auto intro: sums)
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   253
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   254
lemma (in bounded_linear) suminf:
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   255
  "summable (\<lambda>n. X n) \<Longrightarrow> f (\<Sum>n. X n) = (\<Sum>n. f (X n))"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   256
by (intro sums_unique sums summable_sums)
23119
0082459a255b add bounded_linear lemmas
huffman
parents: 23111
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   262
lemma sums_mult:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   263
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   266
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   267
lemma summable_mult:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   268
  fixes c :: "'a::real_normed_algebra"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   271
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   272
lemma suminf_mult:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   273
  fixes c :: "'a::real_normed_algebra"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   276
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   277
lemma sums_mult2:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   278
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   281
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   282
lemma summable_mult2:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   283
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   286
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   287
lemma suminf_mult2:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   288
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   291
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   292
lemma sums_divide:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   293
  fixes c :: "'a::real_normed_field"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   296
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   297
lemma summable_divide:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   298
  fixes c :: "'a::real_normed_field"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   301
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   302
lemma suminf_divide:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   303
  fixes c :: "'a::real_normed_field"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   306
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   307
lemma sums_add:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   308
  fixes a b :: "'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   311
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   312
lemma summable_add:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   313
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   314
  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n + Y n)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   315
unfolding summable_def by (auto intro: sums_add)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   316
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   317
lemma suminf_add:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   318
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   319
  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X + suminf Y = (\<Sum>n. X n + Y n)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   320
by (intro sums_unique sums_add summable_sums)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   321
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   322
lemma sums_diff:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   323
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   326
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   327
lemma summable_diff:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   328
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   329
  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n - Y n)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   330
unfolding summable_def by (auto intro: sums_diff)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   331
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   332
lemma suminf_diff:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   333
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   334
  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X - suminf Y = (\<Sum>n. X n - Y n)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   335
by (intro sums_unique sums_diff summable_sums)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   336
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   337
lemma sums_minus:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   338
  fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   341
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   342
lemma summable_minus:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   343
  fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   344
  shows "summable X \<Longrightarrow> summable (\<lambda>n. - X n)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   345
unfolding summable_def by (auto intro: sums_minus)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   346
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   347
lemma suminf_minus:
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   348
  fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   349
  shows "summable X \<Longrightarrow> (\<Sum>n. - X n) = - (\<Sum>n. X n)"
23121
5feeb93b3ba8 cleaned up some proofs
huffman
parents: 23119
diff changeset
   350
by (intro sums_unique [symmetric] sums_minus summable_sums)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   351
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   352
lemma sums_group:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   357
apply (drule_tac x="r" in spec, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   358
apply (rule_tac x="no" in exI, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   359
apply (drule_tac x="n*k" in spec)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   360
apply (erule mp)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   361
apply (erule order_trans)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   362
apply simp
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   363
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   385
lemma series_pos_le:
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   386
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   387
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 \<le> f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} \<le> suminf f"
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   388
  apply (drule summable_sums)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   389
  apply (simp add: sums_def)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   390
  apply (rule LIMSEQ_le_const)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   391
  apply assumption
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   392
  apply (intro exI[of _ n])
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   393
  apply (auto intro!: setsum_mono2)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   394
  done
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   395
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   396
lemma series_pos_less:
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   397
  fixes f :: "nat \<Rightarrow> 'a::{ordered_ab_semigroup_add_imp_le, ordered_comm_monoid_add, linorder_topology}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   398
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} < suminf f"
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   399
  apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   400
  using add_less_cancel_left [of "setsum f {0..<n}" 0 "f n"]
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   401
  apply simp
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   402
  apply (erule series_pos_le)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   403
  apply (simp add: order_less_imp_le)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   404
  done
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   405
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   406
lemma suminf_eq_zero_iff:
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   407
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   408
  shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   409
proof
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   410
  assume "summable f" "suminf f = 0" and pos: "\<forall>n. 0 \<le> f n"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   411
  then have "f sums 0"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   412
    by (simp add: sums_iff)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   413
  then have f: "(\<lambda>n. \<Sum>i<n. f i) ----> 0"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   414
    by (simp add: sums_def atLeast0LessThan)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   415
  have "\<And>i. (\<Sum>n\<in>{i}. f n) \<le> 0"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   416
  proof (rule LIMSEQ_le_const[OF f])
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   417
    fix i show "\<exists>N. \<forall>n\<ge>N. (\<Sum>n\<in>{i}. f n) \<le> setsum f {..<n}"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   418
      using pos by (intro exI[of _ "Suc i"] allI impI setsum_mono2) auto
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   419
  qed
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   420
  with pos show "\<forall>n. f n = 0"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   421
    by (auto intro!: antisym)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   422
next
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   423
  assume "\<forall>n. f n = 0"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   424
  then have "f = (\<lambda>n. 0)"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   425
    by auto
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   426
  then show "suminf f = 0"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   427
    by simp
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   428
qed
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   429
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   430
lemma suminf_gt_zero_iff:
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   431
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   432
  shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   433
  using series_pos_le[of f 0] suminf_eq_zero_iff[of f]
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   434
  by (simp add: less_le)
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   435
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   436
lemma suminf_gt_zero:
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   437
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   438
  shows "\<lbrakk>summable f; \<forall>n. 0 < f n\<rbrakk> \<Longrightarrow> 0 < suminf f"
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   439
  using suminf_gt_zero_iff[of f] by (simp add: less_imp_le)
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   440
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   441
lemma suminf_ge_zero:
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   442
  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   443
  shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 \<le> suminf f"
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   444
  by (drule_tac n="0" in series_pos_le, simp_all)
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   445
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   446
lemma sumr_pos_lt_pair:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   447
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   448
  shows "\<lbrakk>summable f;
53602
0ae3db699a3e tuned proofs
haftmann
parents: 51528
diff changeset
   449
        \<And>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\<rbrakk>
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   452
apply (subst suminf_split_initial_segment [where k="k"])
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   453
apply assumption
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   454
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   457
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   458
apply (frule sums_unique)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   459
apply (drule sums_summable)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   460
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   461
apply (erule suminf_gt_zero)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   462
apply (simp add: add_ac)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   463
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   468
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   469
lemma geometric_sums:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30649
diff changeset
   470
  fixes x :: "'a::{real_normed_field}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   471
  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) sums (1 / (1 - x))"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   472
proof -
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   473
  assume less_1: "norm x < 1"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   474
  hence neq_1: "x \<noteq> 1" by auto
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   475
  hence neq_0: "x - 1 \<noteq> 0" by simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   476
  from less_1 have lim_0: "(\<lambda>n. x ^ n) ----> 0"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   480
  hence "(\<lambda>n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   481
    by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   482
  thus "(\<lambda>n. x ^ n) sums (1 / (1 - x))"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   483
    by (simp add: sums_def geometric_sum neq_1)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   484
qed
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   485
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   486
lemma summable_geometric:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30649
diff changeset
   487
  fixes x :: "'a::{real_normed_field}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   488
  shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   489
by (rule geometric_sums [THEN sums_summable])
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   506
lemma summable_convergent_sumr_iff:
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   507
 "summable f = convergent (%n. setsum f {0..<n})"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   508
by (simp add: summable_def sums_def convergent_def)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   509
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   510
lemma summable_LIMSEQ_zero:
44726
8478eab380e9 generalize some lemmas
huffman
parents: 44710
diff changeset
   511
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   512
  shows "summable f \<Longrightarrow> f ----> 0"
20689
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   513
apply (drule summable_convergent_sumr_iff [THEN iffD1])
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   516
apply (drule_tac x="r" in spec, safe)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   517
apply (rule_tac x="M" in exI, safe)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   518
apply (drule_tac x="Suc n" in spec, simp)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   519
apply (drule_tac x="n" in spec, simp)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   520
done
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   521
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 31336
diff changeset
   522
lemma suminf_le:
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   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
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   525
  apply (drule summable_sums)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   526
  apply (simp add: sums_def)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   527
  apply (rule LIMSEQ_le_const2)
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   528
  apply assumption
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   529
  apply auto
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   530
  done
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 31336
diff changeset
   531
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   532
lemma summable_Cauchy:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   533
     "summable (f::nat \<Rightarrow> 'a::banach) =
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   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
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   536
apply (drule spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   537
apply (erule exE, rule_tac x="M" in exI, clarify)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   538
apply (rule_tac x="m" and y="n" in linorder_le_cases)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   539
apply (frule (1) order_trans)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   540
apply (drule_tac x="n" in spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   541
apply (drule_tac x="m" in spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   542
apply (simp add: setsum_diff [symmetric])
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   543
apply simp
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   544
apply (drule spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   545
apply (erule exE, rule_tac x="N" in exI, clarify)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   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
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   548
apply (simp add: setsum_diff [symmetric])
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   549
apply (simp add: setsum_diff [symmetric])
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   550
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   554
lemma norm_setsum:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   555
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   556
  shows "norm (setsum f A) \<le> (\<Sum>i\<in>A. norm (f i))"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   557
apply (case_tac "finite A")
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   558
apply (erule finite_induct)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   559
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   560
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   561
apply (erule order_trans [OF norm_triangle_ineq add_left_mono])
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   562
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   563
done
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   564
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   565
lemma summable_comparison_test:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   566
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   567
  shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   568
apply (simp add: summable_Cauchy, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   569
apply (drule_tac x="e" in spec, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   570
apply (rule_tac x = "N + Na" in exI, safe)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   571
apply (rotate_tac 2)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   572
apply (drule_tac x = m in spec)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   573
apply (auto, rotate_tac 2, drule_tac x = n in spec)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   574
apply (rule_tac y = "\<Sum>k=m..<n. norm (f k)" in order_le_less_trans)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   575
apply (rule norm_setsum)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   576
apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22959
diff changeset
   577
apply (auto intro: setsum_mono simp add: abs_less_iff)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   578
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   579
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   580
lemma summable_norm_comparison_test:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   581
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   582
  shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk>
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   583
         \<Longrightarrow> summable (\<lambda>n. norm (f n))"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   584
apply (rule summable_comparison_test)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   585
apply (auto)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   586
done
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   587
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   588
lemma summable_rabs_comparison_test:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   589
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   591
apply (rule summable_comparison_test)
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   592
apply (auto)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   593
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   594
23084
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   595
text{*Summability of geometric series for real algebras*}
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   596
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   597
lemma complete_algebra_summable_geometric:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30649
diff changeset
   598
  fixes x :: "'a::{real_normed_algebra_1,banach}"
23084
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   599
  shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   600
proof (rule summable_comparison_test)
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   601
  show "\<exists>N. \<forall>n\<ge>N. norm (x ^ n) \<le> norm x ^ n"
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   602
    by (simp add: norm_power_ineq)
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   603
  show "norm x < 1 \<Longrightarrow> summable (\<lambda>n. norm x ^ n)"
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   604
    by (simp add: summable_geometric)
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   605
qed
bc000fc64fce add lemma complete_algebra_summable_geometric
huffman
parents: 22998
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   609
lemma summable_le:
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   610
  fixes f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   611
  shows "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   612
apply (drule summable_sums)+
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   613
apply (simp only: sums_def, erule (1) LIMSEQ_le)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   614
apply (rule exI)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   615
apply (auto intro!: setsum_mono)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   616
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   617
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   618
lemma summable_le2:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   619
  fixes f g :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   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
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   621
apply (subgoal_tac "summable f")
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   622
apply (auto intro!: summable_le)
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22959
diff changeset
   623
apply (simp add: abs_le_iff)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   624
apply (rule_tac g="g" in summable_comparison_test, simp_all)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   625
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   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
3de230ed0547 introduce order topology
hoelzl
parents: 50331
diff changeset
   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
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   635
lemma summable_norm_cancel:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   636
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   637
  shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   638
apply (simp only: summable_Cauchy, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   639
apply (drule_tac x="e" in spec, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   640
apply (rule_tac x="N" in exI, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   641
apply (drule_tac x="m" in spec, safe)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   642
apply (rule order_le_less_trans [OF norm_setsum])
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   643
apply (rule order_le_less_trans [OF abs_ge_self])
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   644
apply simp
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   645
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   646
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   647
lemma summable_rabs_cancel:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   648
  fixes f :: "nat \<Rightarrow> real"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   649
  shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> summable f"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   650
by (rule summable_norm_cancel, simp)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   651
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   652
text{*Absolute convergence of series*}
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   653
lemma summable_norm:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   654
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   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
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   657
                summable_sumr_LIMSEQ_suminf norm_setsum)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   658
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   659
lemma summable_rabs:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   660
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   661
  shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> \<bar>suminf f\<bar> \<le> (\<Sum>n. \<bar>f n\<bar>)"
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   662
by (fold real_norm_def, rule summable_norm)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   663
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   664
subsection{* The Ratio Test*}
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   665
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   666
lemma norm_ratiotest_lemma:
22852
2490d4b4671a clean up RealVector classes
huffman
parents: 22719
diff changeset
   667
  fixes x y :: "'a::real_normed_vector"
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   668
  shows "\<lbrakk>c \<le> 0; norm x \<le> c * norm y\<rbrakk> \<Longrightarrow> x = 0"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   669
apply (subgoal_tac "norm x \<le> 0", simp)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   670
apply (erule order_trans)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   671
apply (simp add: mult_le_0_iff)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   672
done
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   673
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   674
lemma rabs_ratiotest_lemma: "[| c \<le> 0; abs x \<le> c * abs y |] ==> x = (0::real)"
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   675
by (erule norm_ratiotest_lemma, simp)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   676
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 47761
diff changeset
   677
(* TODO: MOVE *)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   678
lemma le_Suc_ex: "(k::nat) \<le> l ==> (\<exists>n. l = k + n)"
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   679
apply (drule le_imp_less_or_eq)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   680
apply (auto dest: less_imp_Suc_add)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   681
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   682
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   683
lemma le_Suc_ex_iff: "((k::nat) \<le> l) = (\<exists>n. l = k + n)"
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   684
by (auto simp add: le_Suc_ex)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   685
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   686
(*All this trouble just to get 0<c *)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   687
lemma ratio_test_lemma2:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   688
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   689
  shows "\<lbrakk>\<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> 0 < c \<or> summable f"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   690
apply (simp (no_asm) add: linorder_not_le [symmetric])
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   691
apply (simp add: summable_Cauchy)
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   692
apply (safe, subgoal_tac "\<forall>n. N < n --> f (n) = 0")
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   693
 prefer 2
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   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
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   696
 apply (simp add:diff_Suc split:nat.splits)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   697
 apply (blast intro: norm_ratiotest_lemma)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   698
apply (rule_tac x = "Suc N" in exI, clarify)
44710
9caf6883f1f4 remove redundant lemmas about LIMSEQ
huffman
parents: 44568
diff changeset
   699
apply(simp cong del: setsum_cong cong: setsum_ivl_cong)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   700
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   701
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   702
lemma ratio_test:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   703
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   704
  shows "\<lbrakk>c < 1; \<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> summable f"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   705
apply (frule ratio_test_lemma2, auto)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   708
apply (rule_tac x = N in exI, safe)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   709
apply (drule le_Suc_ex_iff [THEN iffD1])
22959
07a7c2900877 remove redundant lemmas
huffman
parents: 22852
diff changeset
   710
apply (auto simp add: power_add field_power_not_zero)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   711
apply (induct_tac "na", auto)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   712
apply (rule_tac y = "c * norm (f (N + n))" in order_trans)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   713
apply (auto intro: mult_right_mono simp add: summable_def)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   714
apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 36660
diff changeset
   715
apply (rule sums_divide)
27108
e447b3107696 whitespace tuning
haftmann
parents: 23441
diff changeset
   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
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   718
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   719
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   720
subsection {* Cauchy Product Formula *}
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   721
54703
499f92dc6e45 more antiquotations;
wenzelm
parents: 54230
diff changeset
   722
text {*
499f92dc6e45 more antiquotations;
wenzelm
parents: 54230
diff changeset
   723
  Proof based on Analysis WebNotes: Chapter 07, Class 41
499f92dc6e45 more antiquotations;
wenzelm
parents: 54230
diff changeset
   724
  @{url "http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm"}
499f92dc6e45 more antiquotations;
wenzelm
parents: 54230
diff changeset
   725
*}
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   726
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   727
lemma setsum_triangle_reindex:
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   728
  fixes n :: nat
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   729
  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))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   730
proof -
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   731
  have "(\<Sum>(i, j)\<in>{(i, j). i + j < n}. f i j) =
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   732
    (\<Sum>(k, i)\<in>(SIGMA k:{0..<n}. {0..k}). f i (k - i))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   733
  proof (rule setsum_reindex_cong)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   734
    show "inj_on (\<lambda>(k,i). (i, k - i)) (SIGMA k:{0..<n}. {0..k})"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   735
      by (rule inj_on_inverseI [where g="\<lambda>(i,j). (i+j, i)"], auto)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   736
    show "{(i,j). i + j < n} = (\<lambda>(k,i). (i, k - i)) ` (SIGMA k:{0..<n}. {0..k})"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   737
      by (safe, rule_tac x="(a+b,a)" in image_eqI, auto)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   738
    show "\<And>a. (\<lambda>(k, i). f i (k - i)) a = split f ((\<lambda>(k, i). (i, k - i)) a)"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   739
      by clarify
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   740
  qed
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   741
  thus ?thesis by (simp add: setsum_Sigma)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   742
qed
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   743
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   744
lemma Cauchy_product_sums:
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   745
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   746
  assumes a: "summable (\<lambda>k. norm (a k))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   747
  assumes b: "summable (\<lambda>k. norm (b k))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   748
  shows "(\<lambda>k. \<Sum>i=0..k. a i * b (k - i)) sums ((\<Sum>k. a k) * (\<Sum>k. b k))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   749
proof -
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   750
  let ?S1 = "\<lambda>n::nat. {0..<n} \<times> {0..<n}"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   751
  let ?S2 = "\<lambda>n::nat. {(i,j). i + j < n}"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   752
  have S1_mono: "\<And>m n. m \<le> n \<Longrightarrow> ?S1 m \<subseteq> ?S1 n" by auto
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   753
  have S2_le_S1: "\<And>n. ?S2 n \<subseteq> ?S1 n" by auto
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   754
  have S1_le_S2: "\<And>n. ?S1 (n div 2) \<subseteq> ?S2 n" by auto
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   755
  have finite_S1: "\<And>n. finite (?S1 n)" by simp
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   756
  with S2_le_S1 have finite_S2: "\<And>n. finite (?S2 n)" by (rule finite_subset)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   757
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   758
  let ?g = "\<lambda>(i,j). a i * b j"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   759
  let ?f = "\<lambda>(i,j). norm (a i) * norm (b j)"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   760
  have f_nonneg: "\<And>x. 0 \<le> ?f x"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   761
    by (auto simp add: mult_nonneg_nonneg)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   762
  hence norm_setsum_f: "\<And>A. norm (setsum ?f A) = setsum ?f A"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   763
    unfolding real_norm_def
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   764
    by (simp only: abs_of_nonneg setsum_nonneg [rule_format])
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   765
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   766
  have "(\<lambda>n. (\<Sum>k=0..<n. a k) * (\<Sum>k=0..<n. b k))
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   767
           ----> (\<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
   768
    by (intro tendsto_mult summable_sumr_LIMSEQ_suminf
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   769
        summable_norm_cancel [OF a] summable_norm_cancel [OF b])
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   770
  hence 1: "(\<lambda>n. setsum ?g (?S1 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   771
    by (simp only: setsum_product setsum_Sigma [rule_format]
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   772
                   finite_atLeastLessThan)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   773
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   774
  have "(\<lambda>n. (\<Sum>k=0..<n. norm (a k)) * (\<Sum>k=0..<n. norm (b k)))
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   775
       ----> (\<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
   776
    using a b by (intro tendsto_mult summable_sumr_LIMSEQ_suminf)
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   777
  hence "(\<lambda>n. setsum ?f (?S1 n)) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   778
    by (simp only: setsum_product setsum_Sigma [rule_format]
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   779
                   finite_atLeastLessThan)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   780
  hence "convergent (\<lambda>n. setsum ?f (?S1 n))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   781
    by (rule convergentI)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   782
  hence Cauchy: "Cauchy (\<lambda>n. setsum ?f (?S1 n))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   783
    by (rule convergent_Cauchy)
36657
f376af79f6b7 remove unneeded constant Zseq
huffman
parents: 36409
diff changeset
   784
  have "Zfun (\<lambda>n. setsum ?f (?S1 n - ?S2 n)) sequentially"
f376af79f6b7 remove unneeded constant Zseq
huffman
parents: 36409
diff changeset
   785
  proof (rule ZfunI, simp only: eventually_sequentially norm_setsum_f)
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   786
    fix r :: real
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   787
    assume r: "0 < r"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   788
    from CauchyD [OF Cauchy r] obtain N
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   789
    where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" ..
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   790
    hence "\<And>m n. \<lbrakk>N \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> norm (setsum ?f (?S1 m - ?S1 n)) < r"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   791
      by (simp only: setsum_diff finite_S1 S1_mono)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   792
    hence N: "\<And>m n. \<lbrakk>N \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> setsum ?f (?S1 m - ?S1 n) < r"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   793
      by (simp only: norm_setsum_f)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   794
    show "\<exists>N. \<forall>n\<ge>N. setsum ?f (?S1 n - ?S2 n) < r"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   795
    proof (intro exI allI impI)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   796
      fix n assume "2 * N \<le> n"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   797
      hence n: "N \<le> n div 2" by simp
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   798
      have "setsum ?f (?S1 n - ?S2 n) \<le> setsum ?f (?S1 n - ?S1 (n div 2))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   799
        by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   800
                  Diff_mono subset_refl S1_le_S2)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   801
      also have "\<dots> < r"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   802
        using n div_le_dividend by (rule N)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   803
      finally show "setsum ?f (?S1 n - ?S2 n) < r" .
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   804
    qed
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   805
  qed
36657
f376af79f6b7 remove unneeded constant Zseq
huffman
parents: 36409
diff changeset
   806
  hence "Zfun (\<lambda>n. setsum ?g (?S1 n - ?S2 n)) sequentially"
f376af79f6b7 remove unneeded constant Zseq
huffman
parents: 36409
diff changeset
   807
    apply (rule Zfun_le [rule_format])
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   808
    apply (simp only: norm_setsum_f)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   809
    apply (rule order_trans [OF norm_setsum setsum_mono])
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   810
    apply (auto simp add: norm_mult_ineq)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   811
    done
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   812
  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
   813
    unfolding tendsto_Zfun_iff diff_0_right
36657
f376af79f6b7 remove unneeded constant Zseq
huffman
parents: 36409
diff changeset
   814
    by (simp only: setsum_diff finite_S1 S2_le_S1)
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   815
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   816
  with 1 have "(\<lambda>n. setsum ?g (?S2 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   817
    by (rule LIMSEQ_diff_approach_zero2)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   818
  thus ?thesis by (simp only: sums_def setsum_triangle_reindex)
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   819
qed
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   820
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   821
lemma Cauchy_product:
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   822
  fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   823
  assumes a: "summable (\<lambda>k. norm (a k))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   824
  assumes b: "summable (\<lambda>k. norm (b k))"
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   825
  shows "(\<Sum>k. a k) * (\<Sum>k. b k) = (\<Sum>k. \<Sum>i=0..k. a i * b (k - i))"
23441
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23127
diff changeset
   826
using a b
23111
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   827
by (rule Cauchy_product_sums [THEN sums_unique])
f8583c2a491a new proof of Cauchy product formula for series
huffman
parents: 23084
diff changeset
   828
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   829
end