src/HOL/Hyperreal/Series.thy
author wenzelm
Fri, 17 Nov 2006 02:20:03 +0100
changeset 21404 eb85850d3eb7
parent 21141 f0b5e6254a1f
child 22719 c51667189bd3
permissions -rw-r--r--
more robust syntax for definition/abbreviation/notation;
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
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     8
*) 
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
21141
f0b5e6254a1f move DERIV_sumr from Series.thy to Lim.thy
huffman
parents: 20848
diff changeset
    13
imports SEQ
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
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    17
   sums  :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21141
diff changeset
    18
     (infixr "sums" 80) where
19765
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
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21141
diff changeset
    22
   summable :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool" where
19765
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
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21141
diff changeset
    26
   suminf   :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a" where
20688
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
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15543
diff changeset
    29
syntax
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
    30
  "_suminf" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a" ("\<Sum>_. _" [0, 10] 10)
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15543
diff changeset
    31
translations
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 20692
diff changeset
    32
  "\<Sum>i. b" == "CONST suminf (%i. b)"
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15543
diff changeset
    33
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
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}"
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15360
diff changeset
    37
by (simp add: diff_minus setsum_addf real_of_nat_def)
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
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    81
(*
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    82
lemma sumr_from_1_from_0: "0 < n ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    83
      (\<Sum>n=Suc 0 ..< Suc n. if even(n) then 0 else
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    84
             ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n =
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    85
      (\<Sum>n=0..<Suc n. if even(n) then 0 else
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    86
             ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    87
by (rule_tac n1 = 1 in sumr_split_add [THEN subst], auto)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
    88
*)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    89
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    90
subsection{* Infinite Sums, by the Properties of Limits*}
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    91
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    92
(*----------------------
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    93
   suminf is the sum   
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    94
 ---------------------*)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    95
lemma sums_summable: "f sums l ==> summable f"
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    96
by (simp add: sums_def summable_def, blast)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    97
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
    98
lemma summable_sums: "summable f ==> f sums (suminf f)"
20688
690d866a165d change definitions from SOME to THE
huffman
parents: 20552
diff changeset
    99
apply (simp add: summable_def suminf_def sums_def)
690d866a165d change definitions from SOME to THE
huffman
parents: 20552
diff changeset
   100
apply (blast intro: theI LIMSEQ_unique)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   101
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   102
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   103
lemma summable_sumr_LIMSEQ_suminf: 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   104
     "summable f ==> (%n. setsum f {0..<n}) ----> (suminf f)"
20688
690d866a165d change definitions from SOME to THE
huffman
parents: 20552
diff changeset
   105
by (rule summable_sums [unfolded sums_def])
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   106
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   107
(*-------------------
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   108
    sum is unique                    
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   109
 ------------------*)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   110
lemma sums_unique: "f sums s ==> (s = suminf f)"
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   111
apply (frule sums_summable [THEN summable_sums])
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   112
apply (auto intro!: LIMSEQ_unique simp add: sums_def)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   113
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   114
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   115
lemma sums_split_initial_segment: "f sums s ==> 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   116
  (%n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   117
  apply (unfold sums_def);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   118
  apply (simp add: sumr_offset); 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   119
  apply (rule LIMSEQ_diff_const)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   120
  apply (rule LIMSEQ_ignore_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   121
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   122
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   123
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   124
lemma summable_ignore_initial_segment: "summable f ==> 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   125
    summable (%n. f(n + k))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   126
  apply (unfold summable_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   127
  apply (auto intro: sums_split_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   128
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   129
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   130
lemma suminf_minus_initial_segment: "summable f ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   131
    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
   132
  apply (frule summable_ignore_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   133
  apply (rule sums_unique [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   134
  apply (frule summable_sums)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   135
  apply (rule sums_split_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   136
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   137
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   138
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   139
lemma suminf_split_initial_segment: "summable f ==> 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   140
    suminf f = (SUM i = 0..< k. f i) + suminf (%n. f(n + k))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   141
by (auto simp add: suminf_minus_initial_segment)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   142
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   143
lemma series_zero: 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   144
     "(\<forall>m. n \<le> m --> f(m) = 0) ==> f sums (setsum f {0..<n})"
15537
5538d3244b4d continued eliminating sumr
nipkow
parents: 15536
diff changeset
   145
apply (simp add: sums_def LIMSEQ_def diff_minus[symmetric], safe)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   146
apply (rule_tac x = n in exI)
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
   147
apply (clarsimp simp add:setsum_diff[symmetric] cong:setsum_ivl_cong)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   148
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   149
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   150
lemma sums_zero: "(%n. 0) sums 0";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   151
  apply (unfold sums_def);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   152
  apply simp;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   153
  apply (rule LIMSEQ_const);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   154
done;
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   155
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   156
lemma summable_zero: "summable (%n. 0)";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   157
  apply (rule sums_summable);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   158
  apply (rule sums_zero);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   159
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   160
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   161
lemma suminf_zero: "suminf (%n. 0) = 0";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   162
  apply (rule sym);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   163
  apply (rule sums_unique);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   164
  apply (rule sums_zero);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   165
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   166
  
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   167
lemma sums_mult:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   168
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   169
  shows "f sums a \<Longrightarrow> (\<lambda>n. c * f n) sums (c * a)"
19279
48b527d0331b Renamed setsum_mult to setsum_right_distrib.
ballarin
parents: 19106
diff changeset
   170
by (auto simp add: sums_def setsum_right_distrib [symmetric]
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   171
         intro!: LIMSEQ_mult intro: LIMSEQ_const)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   172
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   173
lemma summable_mult:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   174
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   175
  shows "summable f \<Longrightarrow> summable (%n. c * f n)";
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   176
  apply (unfold summable_def);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   177
  apply (auto intro: sums_mult);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   178
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   179
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   180
lemma suminf_mult:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   181
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   182
  shows "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f";
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   183
  apply (rule sym);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   184
  apply (rule sums_unique);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   185
  apply (rule sums_mult);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   186
  apply (erule summable_sums);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   187
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   188
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   189
lemma sums_mult2:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   190
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   191
  shows "f sums a \<Longrightarrow> (\<lambda>n. f n * c) sums (a * c)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   192
by (auto simp add: sums_def setsum_left_distrib [symmetric]
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   193
         intro!: LIMSEQ_mult LIMSEQ_const)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   194
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   195
lemma summable_mult2:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   196
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   197
  shows "summable f \<Longrightarrow> summable (\<lambda>n. f n * c)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   198
  apply (unfold summable_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   199
  apply (auto intro: sums_mult2)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   200
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   201
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   202
lemma suminf_mult2:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   203
  fixes c :: "'a::real_normed_algebra"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   204
  shows "summable f \<Longrightarrow> suminf f * c = (\<Sum>n. f n * c)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   205
by (auto intro!: sums_unique sums_mult2 summable_sums)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   206
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   207
lemma sums_divide:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   208
  fixes c :: "'a::real_normed_field"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   209
  shows "f sums a \<Longrightarrow> (\<lambda>n. f n / c) sums (a / c)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   210
by (simp add: divide_inverse sums_mult2)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   211
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   212
lemma summable_divide:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   213
  fixes c :: "'a::real_normed_field"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   214
  shows "summable f \<Longrightarrow> summable (\<lambda>n. f n / c)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   215
  apply (unfold summable_def);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   216
  apply (auto intro: sums_divide);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   217
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   218
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   219
lemma suminf_divide:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   220
  fixes c :: "'a::real_normed_field"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   221
  shows "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   222
  apply (rule sym);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   223
  apply (rule sums_unique);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   224
  apply (rule sums_divide);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   225
  apply (erule summable_sums);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   226
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   227
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   228
lemma sums_add: "[| x sums x0; y sums y0 |] ==> (%n. x n + y n) sums (x0+y0)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   229
by (auto simp add: sums_def setsum_addf intro: LIMSEQ_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   230
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   231
lemma summable_add: "summable f ==> summable g ==> summable (%x. f x + g x)";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   232
  apply (unfold summable_def);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   233
  apply clarify;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   234
  apply (rule exI);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   235
  apply (erule sums_add);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   236
  apply assumption;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   237
done;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   238
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   239
lemma suminf_add:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   240
     "[| summable f; summable g |]   
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   241
      ==> suminf f + suminf g  = (\<Sum>n. f n + g n)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   242
by (auto intro!: sums_add sums_unique summable_sums)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   243
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   244
lemma sums_diff: "[| x sums x0; y sums y0 |] ==> (%n. x n - y n) sums (x0-y0)"
15536
3ce1cb7a24f0 starting to get rid of sumr
nipkow
parents: 15360
diff changeset
   245
by (auto simp add: sums_def setsum_subtractf intro: LIMSEQ_diff)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   246
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   247
lemma summable_diff: "summable f ==> summable g ==> summable (%x. f x - g x)";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   248
  apply (unfold summable_def);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   249
  apply clarify;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   250
  apply (rule exI);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   251
  apply (erule sums_diff);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   252
  apply assumption;
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   253
done;
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   254
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   255
lemma suminf_diff:
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   256
     "[| summable f; summable g |]   
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15543
diff changeset
   257
      ==> suminf f - suminf g  = (\<Sum>n. f n - g n)"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   258
by (auto intro!: sums_diff sums_unique summable_sums)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   259
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   260
lemma sums_minus: "f sums s ==> (%x. - f x) sums (- s)";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   261
  by (simp add: sums_def setsum_negf LIMSEQ_minus);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   262
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   263
lemma summable_minus: "summable f ==> summable (%x. - f x)";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   264
  by (auto simp add: summable_def intro: sums_minus);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   265
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   266
lemma suminf_minus: "summable f ==> suminf (%x. - f x) = - (suminf f)";
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   267
  apply (rule sym);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   268
  apply (rule sums_unique);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   269
  apply (rule sums_minus);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   270
  apply (erule summable_sums);
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16733
diff changeset
   271
done;
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   272
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   273
lemma sums_group:
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   274
     "[|summable f; 0 < k |] ==> (%n. setsum f {n*k..<n*k+k}) sums (suminf f)"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   275
apply (drule summable_sums)
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   276
apply (simp only: sums_def sumr_group)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   277
apply (unfold LIMSEQ_def, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   278
apply (drule_tac x="r" in spec, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   279
apply (rule_tac x="no" in exI, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   280
apply (drule_tac x="n*k" in spec)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   281
apply (erule mp)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   282
apply (erule order_trans)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   283
apply simp
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   284
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   285
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   286
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
   287
great as any partial sum.*}
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   288
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   289
lemma series_pos_le:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   290
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   291
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 \<le> f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} \<le> suminf f"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   292
apply (drule summable_sums)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   293
apply (simp add: sums_def)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   294
apply (cut_tac k = "setsum f {0..<n}" in LIMSEQ_const)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   295
apply (erule LIMSEQ_le, blast)
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   296
apply (rule_tac x="n" in exI, clarify)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   297
apply (rule setsum_mono2)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   298
apply auto
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   299
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   300
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   301
lemma series_pos_less:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   302
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   303
  shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} < suminf f"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   304
apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   305
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   306
apply (erule series_pos_le)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   307
apply (simp add: order_less_imp_le)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   308
done
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   309
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   310
lemma suminf_gt_zero:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   311
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   312
  shows "\<lbrakk>summable f; \<forall>n. 0 < f n\<rbrakk> \<Longrightarrow> 0 < suminf f"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   313
by (drule_tac n="0" in series_pos_less, simp_all)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   314
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   315
lemma suminf_ge_zero:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   316
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   317
  shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 \<le> suminf f"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   318
by (drule_tac n="0" in series_pos_le, simp_all)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   319
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   320
lemma sumr_pos_lt_pair:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   321
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   322
  shows "\<lbrakk>summable f;
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   323
        \<forall>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\<rbrakk>
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   324
      \<Longrightarrow> setsum f {0..<k} < suminf f"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   325
apply (subst suminf_split_initial_segment [where k="k"])
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   326
apply assumption
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   327
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   328
apply (drule_tac k="k" in summable_ignore_initial_segment)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   329
apply (drule_tac k="Suc (Suc 0)" in sums_group, simp)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   330
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   331
apply (frule sums_unique)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   332
apply (drule sums_summable)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   333
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   334
apply (erule suminf_gt_zero)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   335
apply (simp add: add_ac)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   336
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   337
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   338
text{*Sum of a geometric progression.*}
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   339
17149
e2b19c92ef51 Lemmas on dvd, power and finite summation added or strengthened.
ballarin
parents: 16819
diff changeset
   340
lemmas sumr_geometric = geometric_sum [where 'a = real]
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   341
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   342
lemma geometric_sums:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   343
  fixes x :: "'a::{real_normed_field,recpower,division_by_zero}"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   344
  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
   345
proof -
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   346
  assume less_1: "norm x < 1"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   347
  hence neq_1: "x \<noteq> 1" by auto
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   348
  hence neq_0: "x - 1 \<noteq> 0" by simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   349
  from less_1 have lim_0: "(\<lambda>n. x ^ n) ----> 0"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   350
    by (rule LIMSEQ_power_zero)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   351
  hence "(\<lambda>n. x ^ n / (x - 1) - 1 / (x - 1)) ----> 0 / (x - 1) - 1 / (x
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   352
- 1)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   353
    using neq_0 by (intro LIMSEQ_divide LIMSEQ_diff LIMSEQ_const)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   354
  hence "(\<lambda>n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   355
    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
   356
  thus "(\<lambda>n. x ^ n) sums (1 / (1 - x))"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   357
    by (simp add: sums_def geometric_sum neq_1)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   358
qed
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   359
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   360
lemma summable_geometric:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   361
  fixes x :: "'a::{real_normed_field,recpower,division_by_zero}"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   362
  shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   363
by (rule geometric_sums [THEN sums_summable])
14416
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{*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
   366
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   367
lemma summable_convergent_sumr_iff:
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   368
 "summable f = convergent (%n. setsum f {0..<n})"
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   369
by (simp add: summable_def sums_def convergent_def)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   370
20689
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   371
lemma summable_LIMSEQ_zero: "summable f \<Longrightarrow> f ----> 0"
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   372
apply (drule summable_convergent_sumr_iff [THEN iffD1])
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   373
apply (drule convergent_Cauchy)
20689
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   374
apply (simp only: Cauchy_def LIMSEQ_def, safe)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   375
apply (drule_tac x="r" in spec, safe)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   376
apply (rule_tac x="M" in exI, safe)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   377
apply (drule_tac x="Suc n" in spec, simp)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   378
apply (drule_tac x="n" in spec, simp)
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   379
done
4950e45442b8 add proof of summable_LIMSEQ_zero
huffman
parents: 20688
diff changeset
   380
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   381
lemma summable_Cauchy:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   382
     "summable (f::nat \<Rightarrow> 'a::banach) =  
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   383
      (\<forall>e > 0. \<exists>N. \<forall>m \<ge> N. \<forall>n. norm (setsum f {m..<n}) < e)"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   384
apply (simp only: summable_convergent_sumr_iff Cauchy_convergent_iff [symmetric] Cauchy_def, safe)
20410
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   385
apply (drule spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   386
apply (erule exE, rule_tac x="M" in exI, clarify)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   387
apply (rule_tac x="m" and y="n" in linorder_le_cases)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   388
apply (frule (1) order_trans)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   389
apply (drule_tac x="n" in spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   390
apply (drule_tac x="m" in spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   391
apply (simp add: setsum_diff [symmetric])
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   392
apply simp
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   393
apply (drule spec, drule (1) mp)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   394
apply (erule exE, rule_tac x="N" in exI, clarify)
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   395
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
   396
apply (subst norm_minus_commute)
20410
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   397
apply (simp add: setsum_diff [symmetric])
4bd5cd97c547 speed up proof of summable_Cauchy
huffman
parents: 20254
diff changeset
   398
apply (simp add: setsum_diff [symmetric])
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   399
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   400
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   401
text{*Comparison test*}
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   402
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   403
lemma norm_setsum:
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   404
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   405
  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
   406
apply (case_tac "finite A")
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   407
apply (erule finite_induct)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   408
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   409
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   410
apply (erule order_trans [OF norm_triangle_ineq add_left_mono])
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   411
apply simp
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   412
done
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   413
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   414
lemma summable_comparison_test:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   415
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   416
  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
   417
apply (simp add: summable_Cauchy, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   418
apply (drule_tac x="e" in spec, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   419
apply (rule_tac x = "N + Na" in exI, safe)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   420
apply (rotate_tac 2)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   421
apply (drule_tac x = m in spec)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   422
apply (auto, rotate_tac 2, drule_tac x = n in spec)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   423
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
   424
apply (rule norm_setsum)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   425
apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   426
apply (auto intro: setsum_mono simp add: abs_interval_iff)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   427
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   428
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   429
lemma summable_norm_comparison_test:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   430
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   431
  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
   432
         \<Longrightarrow> summable (\<lambda>n. norm (f n))"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   433
apply (rule summable_comparison_test)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   434
apply (auto)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   435
done
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   436
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   437
lemma summable_rabs_comparison_test:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   438
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   439
  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
   440
apply (rule summable_comparison_test)
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   441
apply (auto)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   442
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   443
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   444
text{*Limit comparison property for series (c.f. jrh)*}
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   445
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   446
lemma summable_le:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   447
  fixes f g :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   448
  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
   449
apply (drule summable_sums)+
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   450
apply (simp only: sums_def, erule (1) LIMSEQ_le)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   451
apply (rule exI)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   452
apply (auto intro!: setsum_mono)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   453
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   454
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   455
lemma summable_le2:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   456
  fixes f g :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   457
  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
   458
apply (subgoal_tac "summable f")
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   459
apply (auto intro!: summable_le)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   460
apply (simp add: abs_le_interval_iff)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   461
apply (rule_tac g="g" in summable_comparison_test, simp_all)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   462
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   463
19106
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   464
(* 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
   465
lemma suminf_0_le:
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   466
  fixes f::"nat\<Rightarrow>real"
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   467
  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
   468
  shows "0 \<le> suminf f"
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   469
proof -
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   470
  let ?g = "(\<lambda>n. (0::real))"
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   471
  from gt0 have "\<forall>n. ?g n \<le> f n" by simp
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   472
  moreover have "summable ?g" by (rule summable_zero)
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   473
  moreover from sm have "summable f" .
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   474
  ultimately have "suminf ?g \<le> suminf f" by (rule summable_le)
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   475
  then show "0 \<le> suminf f" by (simp add: suminf_zero)
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   476
qed 
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   477
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 17149
diff changeset
   478
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   479
text{*Absolute convergence imples normal convergence*}
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   480
lemma summable_norm_cancel:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   481
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   482
  shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   483
apply (simp only: summable_Cauchy, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   484
apply (drule_tac x="e" in spec, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   485
apply (rule_tac x="N" in exI, safe)
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   486
apply (drule_tac x="m" in spec, safe)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   487
apply (rule order_le_less_trans [OF norm_setsum])
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   488
apply (rule order_le_less_trans [OF abs_ge_self])
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   489
apply simp
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   490
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   491
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   492
lemma summable_rabs_cancel:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   493
  fixes f :: "nat \<Rightarrow> real"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   494
  shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> summable f"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   495
by (rule summable_norm_cancel, simp)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   496
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15053
diff changeset
   497
text{*Absolute convergence of series*}
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   498
lemma summable_norm:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   499
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   500
  shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   501
by (auto intro: LIMSEQ_le LIMSEQ_norm summable_norm_cancel
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   502
                summable_sumr_LIMSEQ_suminf norm_setsum)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   503
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   504
lemma summable_rabs:
20692
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   505
  fixes f :: "nat \<Rightarrow> real"
6df83a636e67 generalized types of sums, summable, and suminf
huffman
parents: 20689
diff changeset
   506
  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
   507
by (fold real_norm_def, rule summable_norm)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   508
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   509
subsection{* The Ratio Test*}
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   510
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   511
lemma norm_ratiotest_lemma:
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   512
  fixes x y :: "'a::normed"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   513
  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
   514
apply (subgoal_tac "norm x \<le> 0", simp)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   515
apply (erule order_trans)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   516
apply (simp add: mult_le_0_iff)
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   517
done
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   518
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   519
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
   520
by (erule norm_ratiotest_lemma, simp)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   521
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   522
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
   523
apply (drule le_imp_less_or_eq)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   524
apply (auto dest: less_imp_Suc_add)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   525
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   526
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   527
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
   528
by (auto simp add: le_Suc_ex)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   529
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   530
(*All this trouble just to get 0<c *)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   531
lemma ratio_test_lemma2:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   532
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   533
  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
   534
apply (simp (no_asm) add: linorder_not_le [symmetric])
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   535
apply (simp add: summable_Cauchy)
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   536
apply (safe, subgoal_tac "\<forall>n. N < n --> f (n) = 0")
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   537
 prefer 2
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   538
 apply clarify
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   539
 apply(erule_tac x = "n - 1" in allE)
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   540
 apply (simp add:diff_Suc split:nat.splits)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   541
 apply (blast intro: norm_ratiotest_lemma)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   542
apply (rule_tac x = "Suc N" in exI, clarify)
15543
0024472afce7 more setsum tuning
nipkow
parents: 15542
diff changeset
   543
apply(simp cong:setsum_ivl_cong)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   544
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   545
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   546
lemma ratio_test:
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   547
  fixes f :: "nat \<Rightarrow> 'a::banach"
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   548
  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
   549
apply (frule ratio_test_lemma2, auto)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   550
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
   551
       in summable_comparison_test)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   552
apply (rule_tac x = N in exI, safe)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   553
apply (drule le_Suc_ex_iff [THEN iffD1])
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   554
apply (auto simp add: power_add realpow_not_zero)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15537
diff changeset
   555
apply (induct_tac "na", auto)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   556
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
   557
apply (auto intro: mult_right_mono simp add: summable_def)
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   558
apply (simp add: mult_ac)
20848
27a09c3eca1f generalize summability lemmas using class banach
huffman
parents: 20792
diff changeset
   559
apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   560
apply (rule sums_divide) 
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   561
apply (rule sums_mult) 
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   562
apply (auto intro!: geometric_sums)
14416
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   563
done
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   564
1f256287d4f0 converted Hyperreal/Series to Isar script
paulson
parents: 12018
diff changeset
   565
end