src/HOL/SEQ.thy
author hoelzl
Fri, 12 Mar 2010 15:35:41 +0100
changeset 35748 5f35613d9a65
parent 35292 e4a431b6d9b7
child 36625 2ba6525f9905
permissions -rw-r--r--
Equality of integral and infinite sum.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     1
(*  Title:      HOL/SEQ.thy
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     2
    Author:     Jacques D. Fleuriot, University of Cambridge
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     3
    Author:     Lawrence C Paulson
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     4
    Author:     Jeremy Avigad
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     5
    Author:     Brian Huffman
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     6
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
     7
Convergence of sequences and series.
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
     8
*)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     9
22631
7ae5a6ab7bd6 moved nonstandard stuff from SEQ.thy into new file HSEQ.thy
huffman
parents: 22629
diff changeset
    10
header {* Sequences and Convergence *}
17439
a358da1a0218 add header
huffman
parents: 17429
diff changeset
    11
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15102
diff changeset
    12
theory SEQ
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
    13
imports Limits
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15102
diff changeset
    14
begin
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    15
19765
dfe940911617 misc cleanup;
wenzelm
parents: 18585
diff changeset
    16
definition
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    17
  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    18
    --{*Standard definition of sequence converging to zero*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    19
  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    20
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    21
definition
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
    22
  LIMSEQ :: "[nat \<Rightarrow> 'a::metric_space, 'a] \<Rightarrow> bool"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    23
    ("((_)/ ----> (_))" [60, 60] 60) where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    24
    --{*Standard definition of convergence of sequence*}
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
    25
  [code del]: "X ----> L = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    26
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    27
definition
31404
05d2eddc5d41 generalize type of constant lim
huffman
parents: 31403
diff changeset
    28
  lim :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> 'a" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    29
    --{*Standard definition of limit using choice operator*}
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
    30
  "lim X = (THE L. X ----> L)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    31
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    32
definition
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
    33
  convergent :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    34
    --{*Standard definition of convergence*}
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
    35
  "convergent X = (\<exists>L. X ----> L)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    36
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    37
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    38
  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    39
    --{*Standard definition for bounded sequence*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    40
  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    41
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    42
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    43
  monoseq :: "(nat=>real)=>bool" where
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    44
    --{*Definition of monotonicity. 
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    45
        The use of disjunction here complicates proofs considerably. 
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    46
        One alternative is to add a Boolean argument to indicate the direction. 
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    47
        Another is to develop the notions of increasing and decreasing first.*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    48
  [code del]: "monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) | (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    49
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    50
definition
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    51
  incseq :: "(nat=>real)=>bool" where
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    52
    --{*Increasing sequence*}
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    53
  [code del]: "incseq X = (\<forall>m. \<forall>n\<ge>m. X m \<le> X n)"
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    54
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    55
definition
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    56
  decseq :: "(nat=>real)=>bool" where
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    57
    --{*Increasing sequence*}
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    58
  [code del]: "decseq X = (\<forall>m. \<forall>n\<ge>m. X n \<le> X m)"
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    59
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    60
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    61
  subseq :: "(nat => nat) => bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    62
    --{*Definition of subsequence*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    63
  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    64
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    65
definition
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
    66
  Cauchy :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    67
    --{*Standard definition of the Cauchy condition*}
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
    68
  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < e)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    69
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    70
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    71
subsection {* Bounded Sequences *}
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    72
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
    73
lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    74
unfolding Bseq_def
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    75
proof (intro exI conjI allI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    76
  show "0 < max K 1" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    77
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    78
  fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    79
  have "norm (X n) \<le> K" by (rule K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    80
  thus "norm (X n) \<le> max K 1" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    81
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    82
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    83
lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    84
unfolding Bseq_def by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    85
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
    86
lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
    87
proof (rule BseqI')
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    88
  let ?A = "norm ` X ` {..N}"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    89
  have 1: "finite ?A" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    90
  fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    91
  show "norm (X n) \<le> max K (Max ?A)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    92
  proof (cases rule: linorder_le_cases)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    93
    assume "n \<ge> N"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    94
    hence "norm (X n) \<le> K" using K by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    95
    thus "norm (X n) \<le> max K (Max ?A)" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    96
  next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    97
    assume "n \<le> N"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    98
    hence "norm (X n) \<in> ?A" by simp
26757
e775accff967 thms Max_ge, Min_le: dropped superfluous premise
haftmann
parents: 26312
diff changeset
    99
    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   100
    thus "norm (X n) \<le> max K (Max ?A)" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   101
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   102
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   103
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   104
lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   105
unfolding Bseq_def by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   106
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   107
lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   108
apply (erule BseqE)
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
   109
apply (rule_tac N="k" and K="K" in BseqI2')
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   110
apply clarify
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   111
apply (drule_tac x="n - k" in spec, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   112
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   113
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   114
lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   115
unfolding Bfun_def eventually_sequentially
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   116
apply (rule iffI)
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   117
apply (simp add: Bseq_def)
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   118
apply (auto intro: BseqI2')
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   119
done
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   120
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   121
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   122
subsection {* Sequences That Converge to Zero *}
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   123
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   124
lemma ZseqI:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   125
  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   126
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   127
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   128
lemma ZseqD:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   129
  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   130
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   131
31392
69570155ddf8 replace filters with filter bases
huffman
parents: 31355
diff changeset
   132
lemma Zseq_conv_Zfun: "Zseq X \<longleftrightarrow> Zfun X sequentially"
69570155ddf8 replace filters with filter bases
huffman
parents: 31355
diff changeset
   133
unfolding Zseq_def Zfun_def eventually_sequentially ..
69570155ddf8 replace filters with filter bases
huffman
parents: 31355
diff changeset
   134
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   135
lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   136
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   137
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   138
lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   139
unfolding Zseq_def by force
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   140
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   141
lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   142
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   143
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   144
lemma Zseq_imp_Zseq:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   145
  assumes X: "Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   146
  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   147
  shows "Zseq (\<lambda>n. Y n)"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   148
using X Y Zfun_imp_Zfun [of X sequentially Y K]
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   149
unfolding Zseq_conv_Zfun by simp
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   150
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   151
lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   152
by (erule_tac K="1" in Zseq_imp_Zseq, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   153
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   154
lemma Zseq_add:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   155
  "Zseq X \<Longrightarrow> Zseq Y \<Longrightarrow> Zseq (\<lambda>n. X n + Y n)"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   156
unfolding Zseq_conv_Zfun by (rule Zfun_add)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   157
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   158
lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   159
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   160
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   161
lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   162
by (simp only: diff_minus Zseq_add Zseq_minus)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   163
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   164
lemma (in bounded_linear) Zseq:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   165
  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. f (X n))"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   166
unfolding Zseq_conv_Zfun by (rule Zfun)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   167
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   168
lemma (in bounded_bilinear) Zseq:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   169
  "Zseq X \<Longrightarrow> Zseq Y \<Longrightarrow> Zseq (\<lambda>n. X n ** Y n)"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   170
unfolding Zseq_conv_Zfun by (rule Zfun)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   171
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   172
lemma (in bounded_bilinear) Zseq_prod_Bseq:
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   173
  "Zseq X \<Longrightarrow> Bseq Y \<Longrightarrow> Zseq (\<lambda>n. X n ** Y n)"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   174
unfolding Zseq_conv_Zfun Bseq_conv_Bfun
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   175
by (rule Zfun_prod_Bfun)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   176
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   177
lemma (in bounded_bilinear) Bseq_prod_Zseq:
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   178
  "Bseq X \<Longrightarrow> Zseq Y \<Longrightarrow> Zseq (\<lambda>n. X n ** Y n)"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   179
unfolding Zseq_conv_Zfun Bseq_conv_Bfun
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   180
by (rule Bfun_prod_Zfun)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   181
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   182
lemma (in bounded_bilinear) Zseq_left:
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   183
  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   184
by (rule bounded_linear_left [THEN bounded_linear.Zseq])
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   185
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   186
lemma (in bounded_bilinear) Zseq_right:
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   187
  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   188
by (rule bounded_linear_right [THEN bounded_linear.Zseq])
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   189
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   190
lemmas Zseq_mult = mult.Zseq
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   191
lemmas Zseq_mult_right = mult.Zseq_right
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   192
lemmas Zseq_mult_left = mult.Zseq_left
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   193
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   194
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   195
subsection {* Limits of Sequences *}
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   196
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   197
lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   198
  by simp
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   199
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31404
diff changeset
   200
lemma LIMSEQ_conv_tendsto: "(X ----> L) \<longleftrightarrow> (X ---> L) sequentially"
31488
5691ccb8d6b5 generalize tendsto to class topological_space
huffman
parents: 31487
diff changeset
   201
unfolding LIMSEQ_def tendsto_iff eventually_sequentially ..
31392
69570155ddf8 replace filters with filter bases
huffman
parents: 31355
diff changeset
   202
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   203
lemma LIMSEQ_iff:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   204
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   205
  shows "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   206
unfolding LIMSEQ_def dist_norm ..
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   207
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33042
diff changeset
   208
lemma LIMSEQ_iff_nz: "X ----> L = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33042
diff changeset
   209
  by (auto simp add: LIMSEQ_def) (metis Suc_leD zero_less_Suc)  
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33042
diff changeset
   210
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   211
lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   212
by (simp only: LIMSEQ_iff Zseq_def)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   213
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   214
lemma metric_LIMSEQ_I:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   215
  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X ----> L"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   216
by (simp add: LIMSEQ_def)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   217
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   218
lemma metric_LIMSEQ_D:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   219
  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   220
by (simp add: LIMSEQ_def)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   221
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   222
lemma LIMSEQ_I:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   223
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   224
  shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   225
by (simp add: LIMSEQ_iff)
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   226
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   227
lemma LIMSEQ_D:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   228
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   229
  shows "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   230
by (simp add: LIMSEQ_iff)
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   231
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   232
lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   233
by (simp add: LIMSEQ_def)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   234
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   235
lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l \<longleftrightarrow> k = l"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   236
apply (safe intro!: LIMSEQ_const)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   237
apply (rule ccontr)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   238
apply (drule_tac r="dist k l" in metric_LIMSEQ_D)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   239
apply (simp add: zero_less_dist_iff)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   240
apply auto
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   241
done
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   242
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   243
lemma LIMSEQ_norm:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   244
  fixes a :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   245
  shows "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   246
unfolding LIMSEQ_conv_tendsto by (rule tendsto_norm)
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   247
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   248
lemma LIMSEQ_ignore_initial_segment:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   249
  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   250
apply (rule metric_LIMSEQ_I)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   251
apply (drule (1) metric_LIMSEQ_D)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   252
apply (erule exE, rename_tac N)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   253
apply (rule_tac x=N in exI)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   254
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   255
done
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   256
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   257
lemma LIMSEQ_offset:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   258
  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   259
apply (rule metric_LIMSEQ_I)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   260
apply (drule (1) metric_LIMSEQ_D)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   261
apply (erule exE, rename_tac N)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   262
apply (rule_tac x="N + k" in exI)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   263
apply clarify
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   264
apply (drule_tac x="n - k" in spec)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   265
apply (simp add: le_diff_conv2)
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   266
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   267
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   268
lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   269
by (drule_tac k="Suc 0" in LIMSEQ_ignore_initial_segment, simp)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   270
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   271
lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   272
by (rule_tac k="Suc 0" in LIMSEQ_offset, simp)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   273
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   274
lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   275
by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   276
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   277
lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   278
  unfolding LIMSEQ_def
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   279
  by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   280
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   281
lemma LIMSEQ_add:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   282
  fixes a b :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   283
  shows "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   284
unfolding LIMSEQ_conv_tendsto by (rule tendsto_add)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   285
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   286
lemma LIMSEQ_minus:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   287
  fixes a :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   288
  shows "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   289
unfolding LIMSEQ_conv_tendsto by (rule tendsto_minus)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   290
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   291
lemma LIMSEQ_minus_cancel:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   292
  fixes a :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   293
  shows "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   294
by (drule LIMSEQ_minus, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   295
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   296
lemma LIMSEQ_diff:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   297
  fixes a b :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   298
  shows "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   299
unfolding LIMSEQ_conv_tendsto by (rule tendsto_diff)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   300
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   301
lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   302
apply (rule ccontr)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   303
apply (drule_tac r="dist a b / 2" in metric_LIMSEQ_D, simp add: zero_less_dist_iff)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   304
apply (drule_tac r="dist a b / 2" in metric_LIMSEQ_D, simp add: zero_less_dist_iff)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   305
apply (clarify, rename_tac M N)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   306
apply (subgoal_tac "dist a b < dist a b / 2 + dist a b / 2", simp)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   307
apply (subgoal_tac "dist a b \<le> dist (X (max M N)) a + dist (X (max M N)) b")
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   308
apply (erule le_less_trans, rule add_strict_mono, simp, simp)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   309
apply (subst dist_commute, rule dist_triangle)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   310
done
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   311
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   312
lemma (in bounded_linear) LIMSEQ:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   313
  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   314
unfolding LIMSEQ_conv_tendsto by (rule tendsto)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   315
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   316
lemma (in bounded_bilinear) LIMSEQ:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   317
  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents: 31336
diff changeset
   318
unfolding LIMSEQ_conv_tendsto by (rule tendsto)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   319
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   320
lemma LIMSEQ_mult:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   321
  fixes a b :: "'a::real_normed_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   322
  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   323
by (rule mult.LIMSEQ)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   324
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   325
lemma increasing_LIMSEQ:
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   326
  fixes f :: "nat \<Rightarrow> real"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   327
  assumes inc: "!!n. f n \<le> f (Suc n)"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   328
      and bdd: "!!n. f n \<le> l"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   329
      and en: "!!e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   330
  shows "f ----> l"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   331
proof (auto simp add: LIMSEQ_def)
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   332
  fix e :: real
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   333
  assume e: "0 < e"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   334
  then obtain N where "l \<le> f N + e/2"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   335
    by (metis half_gt_zero e en that)
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   336
  hence N: "l < f N + e" using e
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   337
    by simp
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   338
  { fix k
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   339
    have [simp]: "!!n. \<bar>f n - l\<bar> = l - f n"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   340
      by (simp add: bdd) 
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   341
    have "\<bar>f (N+k) - l\<bar> < e"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   342
    proof (induct k)
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   343
      case 0 show ?case using N
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   344
        by simp   
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   345
    next
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   346
      case (Suc k) thus ?case using N inc [of "N+k"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   347
        by simp
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   348
    qed 
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   349
  } note 1 = this
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   350
  { fix n
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   351
    have "N \<le> n \<Longrightarrow> \<bar>f n - l\<bar> < e" using 1 [of "n-N"]
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   352
      by simp 
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   353
  } note [intro] = this
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   354
  show " \<exists>no. \<forall>n\<ge>no. dist (f n) l < e"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   355
    by (auto simp add: dist_real_def) 
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   356
  qed
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   357
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   358
lemma Bseq_inverse_lemma:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   359
  fixes x :: "'a::real_normed_div_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   360
  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   361
apply (subst nonzero_norm_inverse, clarsimp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   362
apply (erule (1) le_imp_inverse_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   363
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   364
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   365
lemma Bseq_inverse:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   366
  fixes a :: "'a::real_normed_div_algebra"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   367
  shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   368
unfolding LIMSEQ_conv_tendsto Bseq_conv_Bfun
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   369
by (rule Bfun_inverse)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   370
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   371
lemma LIMSEQ_inverse:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   372
  fixes a :: "'a::real_normed_div_algebra"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   373
  shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   374
unfolding LIMSEQ_conv_tendsto
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   375
by (rule tendsto_inverse)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   376
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   377
lemma LIMSEQ_divide:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   378
  fixes a b :: "'a::real_normed_field"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   379
  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   380
by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   381
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   382
lemma LIMSEQ_pow:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30730
diff changeset
   383
  fixes a :: "'a::{power, real_normed_algebra}"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   384
  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   385
by (induct m) (simp_all add: LIMSEQ_const LIMSEQ_mult)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   386
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   387
lemma LIMSEQ_setsum:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   388
  fixes L :: "'a \<Rightarrow> 'b::real_normed_vector"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   389
  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   390
  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31488
diff changeset
   391
using n unfolding LIMSEQ_conv_tendsto by (rule tendsto_setsum)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   392
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   393
lemma LIMSEQ_setprod:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   394
  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   395
  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   396
  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   397
proof (cases "finite S")
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   398
  case True
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   399
  thus ?thesis using n
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   400
  proof (induct)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   401
    case empty
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   402
    show ?case
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   403
      by (simp add: LIMSEQ_const)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   404
  next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   405
    case insert
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   406
    thus ?case
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   407
      by (simp add: LIMSEQ_mult)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   408
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   409
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   410
  case False
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   411
  thus ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   412
    by (simp add: setprod_def LIMSEQ_const)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   413
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   414
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   415
lemma LIMSEQ_add_const:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   416
  fixes a :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   417
  shows "f ----> a ==> (%n.(f n + b)) ----> a + b"
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   418
by (simp add: LIMSEQ_add LIMSEQ_const)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   419
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   420
(* FIXME: delete *)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   421
lemma LIMSEQ_add_minus:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   422
  fixes a b :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   423
  shows "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   424
by (simp only: LIMSEQ_add LIMSEQ_minus)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   425
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   426
lemma LIMSEQ_diff_const:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   427
  fixes a b :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   428
  shows "f ----> a ==> (%n.(f n  - b)) ----> a - b"
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   429
by (simp add: LIMSEQ_diff LIMSEQ_const)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   430
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   431
lemma LIMSEQ_diff_approach_zero:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   432
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   433
  shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   434
by (drule (1) LIMSEQ_add, simp)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   435
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   436
lemma LIMSEQ_diff_approach_zero2:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   437
  fixes L :: "'a::real_normed_vector"
35292
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   438
  shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   439
by (drule (1) LIMSEQ_diff, simp)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   440
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   441
text{*A sequence tends to zero iff its abs does*}
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   442
lemma LIMSEQ_norm_zero:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   443
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   444
  shows "((\<lambda>n. norm (X n)) ----> 0) \<longleftrightarrow> (X ----> 0)"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   445
by (simp add: LIMSEQ_iff)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   446
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   447
lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   448
by (simp add: LIMSEQ_iff)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   449
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   450
lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   451
by (drule LIMSEQ_norm, simp)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   452
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   453
text{*An unbounded sequence's inverse tends to 0*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   454
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   455
lemma LIMSEQ_inverse_zero:
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   456
  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   457
apply (rule LIMSEQ_I)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   458
apply (drule_tac x="inverse r" in spec, safe)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   459
apply (rule_tac x="N" in exI, safe)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   460
apply (drule_tac x="n" in spec, safe)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   461
apply (frule positive_imp_inverse_positive)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   462
apply (frule (1) less_imp_inverse_less)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   463
apply (subgoal_tac "0 < X n", simp)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   464
apply (erule (1) order_less_trans)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   465
done
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   466
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   467
text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   468
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   469
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   470
apply (rule LIMSEQ_inverse_zero, safe)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   471
apply (cut_tac x = r in reals_Archimedean2)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   472
apply (safe, rule_tac x = n in exI)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   473
apply (auto simp add: real_of_nat_Suc)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   474
done
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   475
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   476
text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   477
infinity is now easily proved*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   478
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   479
lemma LIMSEQ_inverse_real_of_nat_add:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   480
     "(%n. r + inverse(real(Suc n))) ----> r"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   481
by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   482
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   483
lemma LIMSEQ_inverse_real_of_nat_add_minus:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   484
     "(%n. r + -inverse(real(Suc n))) ----> r"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   485
by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   486
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   487
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   488
     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   489
by (cut_tac b=1 in
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   490
        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   491
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   492
lemma LIMSEQ_le_const:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   493
  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   494
apply (rule ccontr, simp only: linorder_not_le)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   495
apply (drule_tac r="a - x" in LIMSEQ_D, simp)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   496
apply clarsimp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   497
apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   498
apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   499
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   500
done
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   501
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   502
lemma LIMSEQ_le_const2:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   503
  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   504
apply (subgoal_tac "- a \<le> - x", simp)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   505
apply (rule LIMSEQ_le_const)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   506
apply (erule LIMSEQ_minus)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   507
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   508
done
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   509
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   510
lemma LIMSEQ_le:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   511
  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   512
apply (subgoal_tac "0 \<le> y - x", simp)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   513
apply (rule LIMSEQ_le_const)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   514
apply (erule (1) LIMSEQ_diff)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   515
apply (simp add: le_diff_eq)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   516
done
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   517
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   518
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   519
subsection {* Convergence *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   520
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   521
lemma limI: "X ----> L ==> lim X = L"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   522
apply (simp add: lim_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   523
apply (blast intro: LIMSEQ_unique)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   524
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   525
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   526
lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   527
by (simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   528
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   529
lemma convergentI: "(X ----> L) ==> convergent X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   530
by (auto simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   531
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   532
lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
   533
by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   534
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   535
lemma convergent_minus_iff:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   536
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   537
  shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   538
apply (simp add: convergent_def)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   539
apply (auto dest: LIMSEQ_minus)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   540
apply (drule LIMSEQ_minus, auto)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   541
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   542
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   543
lemma lim_le:
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   544
  fixes x :: real
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   545
  assumes f: "convergent f" and fn_le: "!!n. f n \<le> x"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   546
  shows "lim f \<le> x"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   547
proof (rule classical)
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   548
  assume "\<not> lim f \<le> x"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   549
  hence 0: "0 < lim f - x" by arith
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   550
  have 1: "f----> lim f"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   551
    by (metis convergent_LIMSEQ_iff f) 
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   552
  thus ?thesis
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   553
    proof (simp add: LIMSEQ_iff)
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   554
      assume "\<forall>r>0. \<exists>no. \<forall>n\<ge>no. \<bar>f n - lim f\<bar> < r"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   555
      hence "\<exists>no. \<forall>n\<ge>no. \<bar>f n - lim f\<bar> < lim f - x"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   556
        by (metis 0)
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   557
      from this obtain no where "\<forall>n\<ge>no. \<bar>f n - lim f\<bar> < lim f - x"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   558
        by blast
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   559
      thus "lim f \<le> x"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   560
        by (metis add_cancel_end add_minus_cancel diff_def linorder_linear 
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   561
                  linorder_not_le minus_diff_eq abs_diff_less_iff fn_le) 
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   562
    qed
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   563
qed
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   564
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   565
text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   566
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   567
lemma nat_function_unique: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   568
  unfolding Ex1_def
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   569
  apply (rule_tac x="nat_rec e f" in exI)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   570
  apply (rule conjI)+
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   571
apply (rule def_nat_rec_0, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   572
apply (rule allI, rule def_nat_rec_Suc, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   573
apply (rule allI, rule impI, rule ext)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   574
apply (erule conjE)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   575
apply (induct_tac x)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 33271
diff changeset
   576
apply simp
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   577
apply (erule_tac x="n" in allE)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   578
apply (simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   579
done
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   580
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   581
text{*Subsequence (alternative definition, (e.g. Hoskins)*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   582
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   583
lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   584
apply (simp add: subseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   585
apply (auto dest!: less_imp_Suc_add)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   586
apply (induct_tac k)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   587
apply (auto intro: less_trans)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   588
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   589
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   590
lemma monoseq_Suc:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   591
   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   592
                 | (\<forall>n. X (Suc n) \<le> X n))"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   593
apply (simp add: monoseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   594
apply (auto dest!: le_imp_less_or_eq)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   595
apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   596
apply (induct_tac "ka")
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   597
apply (auto intro: order_trans)
18585
5d379fe2eb74 replaced swap by contrapos_np;
wenzelm
parents: 17439
diff changeset
   598
apply (erule contrapos_np)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   599
apply (induct_tac "k")
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   600
apply (auto intro: order_trans)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   601
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   602
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   603
lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   604
by (simp add: monoseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   605
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   606
lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   607
by (simp add: monoseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   608
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   609
lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   610
by (simp add: monoseq_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   611
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   612
lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   613
by (simp add: monoseq_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   614
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   615
lemma monoseq_minus: assumes "monoseq a"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   616
  shows "monoseq (\<lambda> n. - a n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   617
proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   618
  case True
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   619
  hence "\<forall> m. \<forall> n \<ge> m. - a n \<le> - a m" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   620
  thus ?thesis by (rule monoI2)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   621
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   622
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   623
  hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using `monoseq a`[unfolded monoseq_def] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   624
  thus ?thesis by (rule monoI1)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   625
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   626
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   627
lemma monoseq_le: assumes "monoseq a" and "a ----> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   628
  shows "((\<forall> n. a n \<le> x) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)) \<or> 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   629
         ((\<forall> n. x \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   630
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   631
  { fix x n fix a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   632
    assume "a ----> x" and "\<forall> m. \<forall> n \<ge> m. a m \<le> a n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   633
    hence monotone: "\<And> m n. m \<le> n \<Longrightarrow> a m \<le> a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   634
    have "a n \<le> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   635
    proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   636
      assume "\<not> a n \<le> x" hence "x < a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   637
      hence "0 < a n - x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   638
      from `a ----> x`[THEN LIMSEQ_D, OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   639
      obtain no where "\<And>n'. no \<le> n' \<Longrightarrow> norm (a n' - x) < a n - x" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   640
      hence "norm (a (max no n) - x) < a n - x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   641
      moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   642
      { fix n' have "n \<le> n' \<Longrightarrow> x < a n'" using monotone[where m=n and n=n'] and `x < a n` by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   643
      hence "x < a (max no n)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   644
      ultimately
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   645
      have "a (max no n) < a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   646
      with monotone[where m=n and n="max no n"]
32436
10cd49e0c067 Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
nipkow
parents: 32064
diff changeset
   647
      show False by (auto simp:max_def split:split_if_asm)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   648
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   649
  } note top_down = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   650
  { fix x n m fix a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   651
    assume "a ----> x" and "monoseq a" and "a m < x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   652
    have "a n \<le> x \<and> (\<forall> m. \<forall> n \<ge> m. a m \<le> a n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   653
    proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   654
      case True with top_down and `a ----> x` show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   655
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   656
      case False with `monoseq a`[unfolded monoseq_def] have "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   657
      hence "- a m \<le> - x" using top_down[OF LIMSEQ_minus[OF `a ----> x`]] by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   658
      hence False using `a m < x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   659
      thus ?thesis ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   660
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   661
  } note when_decided = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   662
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   663
  show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   664
  proof (cases "\<exists> m. a m \<noteq> x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   665
    case True then obtain m where "a m \<noteq> x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   666
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   667
    proof (cases "a m < x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   668
      case True with when_decided[OF `a ----> x` `monoseq a`, where m2=m]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   669
      show ?thesis by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   670
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   671
      case False hence "- a m < - x" using `a m \<noteq> x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   672
      with when_decided[OF LIMSEQ_minus[OF `a ----> x`] monoseq_minus[OF `monoseq a`], where m2=m]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   673
      show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   674
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   675
  qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   676
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   677
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   678
text{* for any sequence, there is a mootonic subsequence *}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   679
lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   680
proof-
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   681
  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   682
    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   683
    from nat_function_unique[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   684
    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   685
    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   686
      using H apply - 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   687
      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   688
      unfolding order_le_less by blast 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   689
    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   690
    {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   691
      have "?P (f (Suc n)) (f n)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   692
        unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   693
        using H apply - 
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   694
      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   695
      unfolding order_le_less by blast 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   696
    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   697
  note fSuc = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   698
    {fix p q assume pq: "p \<ge> f q"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   699
      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   700
        by (cases q, simp_all) }
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   701
    note pqth = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   702
    {fix q
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   703
      have "f (Suc q) > f q" apply (induct q) 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   704
        using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   705
    note fss = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   706
    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   707
    {fix a b 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   708
      have "f a \<le> f (a + b)"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   709
      proof(induct b)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   710
        case 0 thus ?case by simp
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   711
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   712
        case (Suc b)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   713
        from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   714
      qed}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   715
    note fmon0 = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   716
    have "monoseq (\<lambda>n. s (f n))" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   717
    proof-
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   718
      {fix n
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   719
        have "s (f n) \<ge> s (f (Suc n))" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   720
        proof(cases n)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   721
          case 0
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   722
          assume n0: "n = 0"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   723
          from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   724
          from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   725
        next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   726
          case (Suc m)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   727
          assume m: "n = Suc m"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   728
          from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   729
          from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   730
        qed}
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   731
      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   732
    qed
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   733
    with th1 have ?thesis by blast}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   734
  moreover
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   735
  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   736
    {fix p assume p: "p \<ge> Suc N" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   737
      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   738
      have "m \<noteq> p" using m(2) by auto 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   739
      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   740
    note th0 = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   741
    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   742
    from nat_function_unique[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   743
    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   744
      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   745
    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   746
      using N apply - 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   747
      apply (erule allE[where x="Suc N"], clarsimp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   748
      apply (rule_tac x="m" in exI)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   749
      apply auto
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   750
      apply (subgoal_tac "Suc N \<noteq> m")
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   751
      apply simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   752
      apply (rule ccontr, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   753
      done
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   754
    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   755
    {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   756
      have "f n > N \<and> ?P (f (Suc n)) (f n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   757
        unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   758
      proof (induct n)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   759
        case 0 thus ?case
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   760
          using f0 N apply auto 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   761
          apply (erule allE[where x="f 0"], clarsimp) 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   762
          apply (rule_tac x="m" in exI, simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   763
          by (subgoal_tac "f 0 \<noteq> m", auto)
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   764
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   765
        case (Suc n)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   766
        from Suc.hyps have Nfn: "N < f n" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   767
        from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   768
        with Nfn have mN: "m > N" by arith
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   769
        note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   770
        
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   771
        from key have th0: "f (Suc n) > N" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   772
        from N[rule_format, OF th0]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   773
        obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   774
        have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   775
        hence "m' > f (Suc n)" using m'(1) by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32877
diff changeset
   776
        with key m'(2) show ?case by auto
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   777
      qed}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   778
    note fSuc = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   779
    {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   780
      have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   781
      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   782
    note thf = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   783
    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   784
    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   785
      apply -
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   786
      apply (rule disjI1)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   787
      apply auto
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   788
      apply (rule order_less_imp_le)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   789
      apply blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   790
      done
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   791
    then have ?thesis  using sqf by blast}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   792
  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   793
qed
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   794
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   795
lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   796
proof(induct n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   797
  case 0 thus ?case by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   798
next
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   799
  case (Suc n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   800
  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   801
  have "n < f (Suc n)" by arith 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   802
  thus ?case by arith
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   803
qed
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   804
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   805
lemma LIMSEQ_subseq_LIMSEQ:
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   806
  "\<lbrakk> X ----> L; subseq f \<rbrakk> \<Longrightarrow> (X o f) ----> L"
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   807
apply (auto simp add: LIMSEQ_def) 
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   808
apply (drule_tac x=r in spec, clarify)  
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   809
apply (rule_tac x=no in exI, clarify) 
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   810
apply (blast intro: seq_suble le_trans dest!: spec) 
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   811
done
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   812
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   813
subsection {* Bounded Monotonic Sequences *}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   814
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   815
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   816
text{*Bounded Sequence*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   817
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   818
lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   819
by (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   820
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   821
lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   822
by (auto simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   823
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   824
lemma lemma_NBseq_def:
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   825
     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   826
      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   827
proof auto
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   828
  fix K :: real
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   829
  from reals_Archimedean2 obtain n :: nat where "K < real n" ..
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   830
  then have "K \<le> real (Suc n)" by auto
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   831
  assume "\<forall>m. norm (X m) \<le> K"
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   832
  have "\<forall>m. norm (X m) \<le> real (Suc n)"
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   833
  proof
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   834
    fix m :: 'a
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   835
    from `\<forall>m. norm (X m) \<le> K` have "norm (X m) \<le> K" ..
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   836
    with `K \<le> real (Suc n)` show "norm (X m) \<le> real (Suc n)" by auto
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   837
  qed
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   838
  then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   839
next
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   840
  fix N :: nat
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   841
  have "real (Suc N) > 0" by (simp add: real_of_nat_Suc)
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   842
  moreover assume "\<forall>n. norm (X n) \<le> real (Suc N)"
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   843
  ultimately show "\<exists>K>0. \<forall>n. norm (X n) \<le> K" by blast
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   844
qed
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   845
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   846
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   847
text{* alternative definition for Bseq *}
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   848
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   849
apply (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   850
apply (simp (no_asm) add: lemma_NBseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   851
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   852
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   853
lemma lemma_NBseq_def2:
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   854
     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   855
apply (subst lemma_NBseq_def, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   856
apply (rule_tac x = "Suc N" in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   857
apply (rule_tac [2] x = N in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   858
apply (auto simp add: real_of_nat_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   859
 prefer 2 apply (blast intro: order_less_imp_le)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   860
apply (drule_tac x = n in spec, simp)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   861
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   862
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   863
(* yet another definition for Bseq *)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   864
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   865
by (simp add: Bseq_def lemma_NBseq_def2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   866
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   867
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   868
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   869
lemma Bseq_isUb:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   870
  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22974
diff changeset
   871
by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   872
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   873
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   874
text{* Use completeness of reals (supremum property)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   875
   to show that any bounded sequence has a least upper bound*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   876
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   877
lemma Bseq_isLub:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   878
  "!!(X::nat=>real). Bseq X ==>
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   879
   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   880
by (blast intro: reals_complete Bseq_isUb)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   881
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   882
subsubsection{*A Bounded and Monotonic Sequence Converges*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   883
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   884
lemma lemma_converg1:
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   885
     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   886
                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   887
               |] ==> \<forall>n \<ge> ma. X n = X ma"