src/HOL/SEQ.thy
author wenzelm
Mon, 10 Dec 2012 15:39:20 +0100
changeset 50453 262dc5873f80
parent 50384 b9b967da28e9
child 50937 d249ef928ae1
permissions -rw-r--r--
some clarification for Windows;
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
36822
38a480e0346f minimize imports
huffman
parents: 36776
diff changeset
    13
imports Limits RComplete
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
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    16
subsection {* Monotone sequences and subsequences *}
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    17
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    18
definition
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    19
  monoseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    20
    --{*Definition of monotonicity.
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    21
        The use of disjunction here complicates proofs considerably.
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    22
        One alternative is to add a Boolean argument to indicate the direction.
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    23
        Another is to develop the notions of increasing and decreasing first.*}
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36822
diff changeset
    24
  "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
    25
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    26
definition
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    27
  incseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    28
    --{*Increasing sequence*}
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    29
  "incseq X \<longleftrightarrow> (\<forall>m. \<forall>n\<ge>m. X m \<le> X n)"
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    30
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    31
definition
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    32
  decseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    33
    --{*Decreasing sequence*}
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    34
  "decseq X \<longleftrightarrow> (\<forall>m. \<forall>n\<ge>m. X n \<le> X m)"
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    35
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
    36
definition
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    37
  subseq :: "(nat \<Rightarrow> nat) \<Rightarrow> bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    38
    --{*Definition of subsequence*}
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
    39
  "subseq f \<longleftrightarrow> (\<forall>m. \<forall>n>m. f m < f n)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    40
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    41
lemma incseq_mono: "mono f \<longleftrightarrow> incseq f"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    42
  unfolding mono_def incseq_def by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    43
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    44
lemma incseq_SucI:
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    45
  "(\<And>n. X n \<le> X (Suc n)) \<Longrightarrow> incseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    46
  using lift_Suc_mono_le[of X]
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    47
  by (auto simp: incseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    48
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    49
lemma incseqD: "\<And>i j. incseq f \<Longrightarrow> i \<le> j \<Longrightarrow> f i \<le> f j"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    50
  by (auto simp: incseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    51
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    52
lemma incseq_SucD: "incseq A \<Longrightarrow> A i \<le> A (Suc i)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    53
  using incseqD[of A i "Suc i"] by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    54
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    55
lemma incseq_Suc_iff: "incseq f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    56
  by (auto intro: incseq_SucI dest: incseq_SucD)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    57
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    58
lemma incseq_const[simp, intro]: "incseq (\<lambda>x. k)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    59
  unfolding incseq_def by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    60
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    61
lemma decseq_SucI:
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    62
  "(\<And>n. X (Suc n) \<le> X n) \<Longrightarrow> decseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    63
  using order.lift_Suc_mono_le[OF dual_order, of X]
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    64
  by (auto simp: decseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    65
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    66
lemma decseqD: "\<And>i j. decseq f \<Longrightarrow> i \<le> j \<Longrightarrow> f j \<le> f i"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    67
  by (auto simp: decseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    68
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    69
lemma decseq_SucD: "decseq A \<Longrightarrow> A (Suc i) \<le> A i"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    70
  using decseqD[of A i "Suc i"] by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    71
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    72
lemma decseq_Suc_iff: "decseq f \<longleftrightarrow> (\<forall>n. f (Suc n) \<le> f n)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    73
  by (auto intro: decseq_SucI dest: decseq_SucD)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    74
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    75
lemma decseq_const[simp, intro]: "decseq (\<lambda>x. k)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    76
  unfolding decseq_def by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    77
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    78
lemma monoseq_iff: "monoseq X \<longleftrightarrow> incseq X \<or> decseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    79
  unfolding monoseq_def incseq_def decseq_def ..
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    80
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    81
lemma monoseq_Suc:
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    82
  "monoseq X \<longleftrightarrow> (\<forall>n. X n \<le> X (Suc n)) \<or> (\<forall>n. X (Suc n) \<le> X n)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    83
  unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff ..
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    84
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    85
lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    86
by (simp add: monoseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    87
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    88
lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    89
by (simp add: monoseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    90
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    91
lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    92
by (simp add: monoseq_Suc)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    93
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    94
lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    95
by (simp add: monoseq_Suc)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    96
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    97
lemma monoseq_minus:
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    98
  fixes a :: "nat \<Rightarrow> 'a::ordered_ab_group_add"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
    99
  assumes "monoseq a"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   100
  shows "monoseq (\<lambda> n. - a n)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   101
proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   102
  case True
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   103
  hence "\<forall> m. \<forall> n \<ge> m. - a n \<le> - a m" by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   104
  thus ?thesis by (rule monoI2)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   105
next
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   106
  case False
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   107
  hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using `monoseq a`[unfolded monoseq_def] by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   108
  thus ?thesis by (rule monoI1)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   109
qed
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   110
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   111
text{*Subsequence (alternative definition, (e.g. Hoskins)*}
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   112
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   113
lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   114
apply (simp add: subseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   115
apply (auto dest!: less_imp_Suc_add)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   116
apply (induct_tac k)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   117
apply (auto intro: less_trans)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   118
done
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   119
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   120
text{* for any sequence, there is a monotonic subsequence *}
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   121
lemma seq_monosub:
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   122
  fixes s :: "nat => 'a::linorder"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   123
  shows "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   124
proof cases
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   125
  let "?P p n" = "p > n \<and> (\<forall>m\<ge>p. s m \<le> s p)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   126
  assume *: "\<forall>n. \<exists>p. ?P p n"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   127
  def f \<equiv> "nat_rec (SOME p. ?P p 0) (\<lambda>_ n. SOME p. ?P p n)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   128
  have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   129
  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   130
  have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   131
  have P_Suc: "\<And>i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   132
  then have "subseq f" unfolding subseq_Suc_iff by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   133
  moreover have "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   134
  proof (intro disjI2 allI)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   135
    fix n show "s (f (Suc n)) \<le> s (f n)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   136
    proof (cases n)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   137
      case 0 with P_Suc[of 0] P_0 show ?thesis by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   138
    next
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   139
      case (Suc m)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   140
      from P_Suc[of n] Suc have "f (Suc m) \<le> f (Suc (Suc m))" by simp
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   141
      with P_Suc Suc show ?thesis by simp
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   142
    qed
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   143
  qed
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   144
  ultimately show ?thesis by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   145
next
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   146
  let "?P p m" = "m < p \<and> s m < s p"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   147
  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. s m \<le> s p))"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   148
  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. s p < s m" by (force simp: not_le le_less)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   149
  def f \<equiv> "nat_rec (SOME p. ?P p (Suc N)) (\<lambda>_ n. SOME p. ?P p n)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   150
  have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   151
  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   152
  have P_0: "?P (f 0) (Suc N)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   153
    unfolding f_0 some_eq_ex[of "\<lambda>p. ?P p (Suc N)"] using N[of "Suc N"] by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   154
  { fix i have "N < f i \<Longrightarrow> ?P (f (Suc i)) (f i)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   155
      unfolding f_Suc some_eq_ex[of "\<lambda>p. ?P p (f i)"] using N[of "f i"] . }
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   156
  note P' = this
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   157
  { fix i have "N < f i \<and> ?P (f (Suc i)) (f i)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   158
      by (induct i) (insert P_0 P', auto) }
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   159
  then have "subseq f" "monoseq (\<lambda>x. s (f x))"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   160
    unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   161
  then show ?thesis by auto
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   162
qed
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   163
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   164
lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   165
proof(induct n)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   166
  case 0 thus ?case by simp
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   167
next
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   168
  case (Suc n)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   169
  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   170
  have "n < f (Suc n)" by arith
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   171
  thus ?case by arith
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   172
qed
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   173
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   174
lemma filterlim_subseq: "subseq f \<Longrightarrow> filterlim f sequentially sequentially"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   175
  unfolding filterlim_iff eventually_sequentially by (metis le_trans seq_suble)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   176
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   177
lemma subseq_o: "subseq r \<Longrightarrow> subseq s \<Longrightarrow> subseq (r \<circ> s)"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   178
  unfolding subseq_def by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   179
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   180
lemma subseq_mono: assumes "subseq r" "m < n" shows "r m < r n"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   181
  using assms by (auto simp: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   182
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   183
lemma incseq_imp_monoseq:  "incseq X \<Longrightarrow> monoseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   184
  by (simp add: incseq_def monoseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   185
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   186
lemma decseq_imp_monoseq:  "decseq X \<Longrightarrow> monoseq X"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   187
  by (simp add: decseq_def monoseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   188
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   189
lemma decseq_eq_incseq:
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   190
  fixes X :: "nat \<Rightarrow> 'a::ordered_ab_group_add" shows "decseq X = incseq (\<lambda>n. - X n)" 
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   191
  by (simp add: decseq_def incseq_def)
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   192
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   193
lemma INT_decseq_offset:
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   194
  assumes "decseq F"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   195
  shows "(\<Inter>i. F i) = (\<Inter>i\<in>{n..}. F i)"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   196
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   197
  fix x i assume x: "x \<in> (\<Inter>i\<in>{n..}. F i)"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   198
  show "x \<in> F i"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   199
  proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   200
    from x have "x \<in> F n" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   201
    also assume "i \<le> n" with `decseq F` have "F n \<subseteq> F i"
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   202
      unfolding decseq_def by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   203
    finally show ?thesis .
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   204
  qed (insert x, simp)
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   205
qed auto
635d73673b5e regularity of measures, therefore:
immler
parents: 44714
diff changeset
   206
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   207
subsection {* Defintions of limits *}
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   208
44206
5e4a1664106e locale-ize some constant definitions, so complete_space can inherit from metric_space
huffman
parents: 44205
diff changeset
   209
abbreviation (in topological_space)
5e4a1664106e locale-ize some constant definitions, so complete_space can inherit from metric_space
huffman
parents: 44205
diff changeset
   210
  LIMSEQ :: "[nat \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   211
    ("((_)/ ----> (_))" [60, 60] 60) where
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   212
  "X ----> L \<equiv> (X ---> L) sequentially"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   213
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   214
definition
44205
18da2a87421c generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents: 44194
diff changeset
   215
  lim :: "(nat \<Rightarrow> 'a::t2_space) \<Rightarrow> 'a" where
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   216
    --{*Standard definition of limit using choice operator*}
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   217
  "lim X = (THE L. X ----> L)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   218
44206
5e4a1664106e locale-ize some constant definitions, so complete_space can inherit from metric_space
huffman
parents: 44205
diff changeset
   219
definition (in topological_space) convergent :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
41972
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   220
  "convergent X = (\<exists>L. X ----> L)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   221
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   222
definition
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   223
  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   224
    --{*Standard definition for bounded sequence*}
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   225
  "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
8885ba629692 add lemmas for monotone sequences
hoelzl
parents: 41367
diff changeset
   226
44206
5e4a1664106e locale-ize some constant definitions, so complete_space can inherit from metric_space
huffman
parents: 44205
diff changeset
   227
definition (in metric_space) Cauchy :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36822
diff changeset
   228
  "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
   229
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   230
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   231
subsection {* Bounded Sequences *}
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   232
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
   233
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
   234
unfolding Bseq_def
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   235
proof (intro exI conjI allI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   236
  show "0 < max K 1" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   237
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   238
  fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   239
  have "norm (X n) \<le> K" by (rule K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   240
  thus "norm (X n) \<le> max K 1" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   241
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   242
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   243
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
   244
unfolding Bseq_def by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   245
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
   246
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
   247
proof (rule BseqI')
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   248
  let ?A = "norm ` X ` {..N}"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   249
  have 1: "finite ?A" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   250
  fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   251
  show "norm (X n) \<le> max K (Max ?A)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   252
  proof (cases rule: linorder_le_cases)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   253
    assume "n \<ge> N"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   254
    hence "norm (X n) \<le> K" using K by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   255
    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
   256
  next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   257
    assume "n \<le> N"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   258
    hence "norm (X n) \<in> ?A" by simp
26757
e775accff967 thms Max_ge, Min_le: dropped superfluous premise
haftmann
parents: 26312
diff changeset
   259
    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
   260
    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
   261
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   262
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   263
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   264
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
   265
unfolding Bseq_def by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   266
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   267
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
   268
apply (erule BseqE)
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
   269
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
   270
apply clarify
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   271
apply (drule_tac x="n - k" in spec, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   272
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   273
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   274
lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   275
unfolding Bfun_def eventually_sequentially
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   276
apply (rule iffI)
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   277
apply (simp add: Bseq_def)
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   278
apply (auto intro: BseqI2')
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   279
done
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   280
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   281
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   282
subsection {* Limits of Sequences *}
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   283
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   284
lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   285
  by simp
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   286
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36657
diff changeset
   287
lemma LIMSEQ_def: "X ----> L = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36657
diff changeset
   288
unfolding tendsto_iff eventually_sequentially ..
31392
69570155ddf8 replace filters with filter bases
huffman
parents: 31355
diff changeset
   289
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   290
lemma LIMSEQ_iff:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   291
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   292
  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
   293
unfolding LIMSEQ_def dist_norm ..
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   294
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33042
diff changeset
   295
lemma LIMSEQ_iff_nz: "X ----> L = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36657
diff changeset
   296
  unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33042
diff changeset
   297
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   298
lemma metric_LIMSEQ_I:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   299
  "(\<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
   300
by (simp add: LIMSEQ_def)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   301
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   302
lemma metric_LIMSEQ_D:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   303
  "\<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
   304
by (simp add: LIMSEQ_def)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   305
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   306
lemma LIMSEQ_I:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   307
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   308
  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
   309
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
   310
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   311
lemma LIMSEQ_D:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   312
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   313
  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
   314
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
   315
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   316
lemma LIMSEQ_const_iff:
44205
18da2a87421c generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents: 44194
diff changeset
   317
  fixes k l :: "'a::t2_space"
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   318
  shows "(\<lambda>n. k) ----> l \<longleftrightarrow> k = l"
44205
18da2a87421c generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents: 44194
diff changeset
   319
  using trivial_limit_sequentially by (rule tendsto_const_iff)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   320
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   321
lemma LIMSEQ_ignore_initial_segment:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   322
  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   323
apply (rule topological_tendstoI)
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   324
apply (drule (2) topological_tendstoD)
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   325
apply (simp only: eventually_sequentially)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   326
apply (erule exE, rename_tac N)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   327
apply (rule_tac x=N in exI)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   328
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   329
done
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   330
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   331
lemma LIMSEQ_offset:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   332
  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   333
apply (rule topological_tendstoI)
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   334
apply (drule (2) topological_tendstoD)
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   335
apply (simp only: eventually_sequentially)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   336
apply (erule exE, rename_tac N)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   337
apply (rule_tac x="N + k" in exI)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   338
apply clarify
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   339
apply (drule_tac x="n - k" in spec)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   340
apply (simp add: le_diff_conv2)
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   341
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   342
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   343
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
   344
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
   345
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   346
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
   347
by (rule_tac k="Suc 0" in LIMSEQ_offset, simp)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   348
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   349
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
   350
by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   351
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   352
lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   353
  unfolding tendsto_def eventually_sequentially
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   354
  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
   355
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   356
lemma LIMSEQ_unique:
44205
18da2a87421c generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents: 44194
diff changeset
   357
  fixes a b :: "'a::t2_space"
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   358
  shows "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
44205
18da2a87421c generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents: 44194
diff changeset
   359
  using trivial_limit_sequentially by (rule tendsto_unique)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   360
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   361
lemma increasing_LIMSEQ:
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   362
  fixes f :: "nat \<Rightarrow> real"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   363
  assumes inc: "\<And>n. f n \<le> f (Suc n)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   364
      and bdd: "\<And>n. f n \<le> l"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   365
      and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   366
  shows "f ----> l"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   367
  unfolding LIMSEQ_def
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   368
proof safe
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   369
  fix r :: real assume "0 < r"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   370
  with bdd en[of "r / 2"] obtain n where n: "dist (f n) l \<le> r / 2"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   371
    by (auto simp add: field_simps dist_real_def)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   372
  { fix N assume "n \<le> N"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   373
    then have "dist (f N) l \<le> dist (f n) l"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   374
      using incseq_SucI[of f] inc bdd by (auto dest!: incseqD simp: dist_real_def)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   375
    then have "dist (f N) l < r"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   376
      using `0 < r` n by simp }
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   377
  with `0 < r` show "\<exists>no. \<forall>n\<ge>no. dist (f n) l < r"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   378
    by (auto simp add: LIMSEQ_def field_simps intro!: exI[of _ n])
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   379
qed
32877
6f09346c7c08 New lemmas connected with the reals and infinite series
paulson
parents: 32707
diff changeset
   380
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   381
lemma Bseq_inverse_lemma:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   382
  fixes x :: "'a::real_normed_div_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   383
  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
   384
apply (subst nonzero_norm_inverse, clarsimp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   385
apply (erule (1) le_imp_inverse_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   386
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   387
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   388
lemma Bseq_inverse:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   389
  fixes a :: "'a::real_normed_div_algebra"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   390
  shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36657
diff changeset
   391
unfolding Bseq_conv_Bfun by (rule Bfun_inverse)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   392
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   393
lemma LIMSEQ_diff_approach_zero:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   394
  fixes L :: "'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   395
  shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   396
  by (drule (1) tendsto_add, simp)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   397
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   398
lemma LIMSEQ_diff_approach_zero2:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   399
  fixes L :: "'a::real_normed_vector"
35292
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   400
  shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   401
  by (drule (1) tendsto_diff, simp)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   402
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   403
text{*An unbounded sequence's inverse tends to 0*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   404
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   405
lemma LIMSEQ_inverse_zero:
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   406
  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   407
  apply (rule filterlim_compose[OF tendsto_inverse_0])
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   408
  apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   409
  apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   410
  done
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   411
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   412
text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   413
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   414
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   415
  by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   416
            filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   417
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   418
text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   419
infinity is now easily proved*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   420
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   421
lemma LIMSEQ_inverse_real_of_nat_add:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   422
     "(%n. r + inverse(real(Suc n))) ----> r"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   423
  using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   424
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   425
lemma LIMSEQ_inverse_real_of_nat_add_minus:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   426
     "(%n. r + -inverse(real(Suc n))) ----> r"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   427
  using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]]
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   428
  by auto
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   429
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   430
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   431
     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   432
  using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]]
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   433
  by auto
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   434
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   435
lemma LIMSEQ_le_const:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   436
  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   437
  using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   438
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   439
lemma LIMSEQ_le:
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   440
  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   441
  using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially)
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   442
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   443
lemma LIMSEQ_le_const2:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   444
  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   445
  by (rule LIMSEQ_le[of X x "\<lambda>n. a"]) (auto simp: tendsto_const)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   446
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   447
subsection {* Convergence *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   448
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   449
lemma limI: "X ----> L ==> lim X = L"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   450
apply (simp add: lim_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   451
apply (blast intro: LIMSEQ_unique)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   452
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   453
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   454
lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   455
by (simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   456
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   457
lemma convergentI: "(X ----> L) ==> convergent X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   458
by (auto simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   459
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   460
lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
   461
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
   462
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   463
lemma convergent_const: "convergent (\<lambda>n. c)"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   464
  by (rule convergentI, rule tendsto_const)
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   465
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   466
lemma convergent_add:
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   467
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   468
  assumes "convergent (\<lambda>n. X n)"
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   469
  assumes "convergent (\<lambda>n. Y n)"
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   470
  shows "convergent (\<lambda>n. X n + Y n)"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   471
  using assms unfolding convergent_def by (fast intro: tendsto_add)
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   472
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   473
lemma convergent_setsum:
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   474
  fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
36647
edc381bf7200 Removed unnecessary assumption
hoelzl
parents: 36625
diff changeset
   475
  assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   476
  shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
36647
edc381bf7200 Removed unnecessary assumption
hoelzl
parents: 36625
diff changeset
   477
proof (cases "finite A")
36650
d65f07abfa7c fixed proof (cf. edc381bf7200);
wenzelm
parents: 36647
diff changeset
   478
  case True from this and assms show ?thesis
36647
edc381bf7200 Removed unnecessary assumption
hoelzl
parents: 36625
diff changeset
   479
    by (induct A set: finite) (simp_all add: convergent_const convergent_add)
edc381bf7200 Removed unnecessary assumption
hoelzl
parents: 36625
diff changeset
   480
qed (simp add: convergent_const)
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   481
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   482
lemma (in bounded_linear) convergent:
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   483
  assumes "convergent (\<lambda>n. X n)"
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   484
  shows "convergent (\<lambda>n. f (X n))"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   485
  using assms unfolding convergent_def by (fast intro: tendsto)
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   486
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   487
lemma (in bounded_bilinear) convergent:
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   488
  assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   489
  shows "convergent (\<lambda>n. X n ** Y n)"
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   490
  using assms unfolding convergent_def by (fast intro: tendsto)
36625
2ba6525f9905 add lemmas about convergent
huffman
parents: 35748
diff changeset
   491
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   492
lemma convergent_minus_iff:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   493
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   494
  shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   495
apply (simp add: convergent_def)
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   496
apply (auto dest: tendsto_minus)
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   497
apply (drule tendsto_minus, auto)
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   498
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   499
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   500
lemma lim_le: "convergent f \<Longrightarrow> (\<And>n. f n \<le> (x::real)) \<Longrightarrow> lim f \<le> x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   501
  using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff)
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   502
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
   503
lemma monoseq_le:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   504
  "monoseq a \<Longrightarrow> a ----> (x::real) \<Longrightarrow>
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   505
    ((\<forall> n. a n \<le> x) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)) \<or> ((\<forall> n. x \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m))"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   506
  by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   507
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   508
lemma LIMSEQ_subseq_LIMSEQ:
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   509
  "\<lbrakk> X ----> L; subseq f \<rbrakk> \<Longrightarrow> (X o f) ----> L"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   510
  unfolding comp_def by (rule filterlim_compose[of X, OF _ filterlim_subseq])
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   511
44208
1d2bf1f240ac generalize lemma convergent_subseq_convergent
huffman
parents: 44206
diff changeset
   512
lemma convergent_subseq_convergent:
1d2bf1f240ac generalize lemma convergent_subseq_convergent
huffman
parents: 44206
diff changeset
   513
  "\<lbrakk>convergent X; subseq f\<rbrakk> \<Longrightarrow> convergent (X o f)"
1d2bf1f240ac generalize lemma convergent_subseq_convergent
huffman
parents: 44206
diff changeset
   514
  unfolding convergent_def by (auto intro: LIMSEQ_subseq_LIMSEQ)
1d2bf1f240ac generalize lemma convergent_subseq_convergent
huffman
parents: 44206
diff changeset
   515
1d2bf1f240ac generalize lemma convergent_subseq_convergent
huffman
parents: 44206
diff changeset
   516
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   517
subsection {* Bounded Monotonic Sequences *}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   518
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   519
text{*Bounded Sequence*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   520
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   521
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
   522
by (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   523
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   524
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
   525
by (auto simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   526
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   527
lemma lemma_NBseq_def:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   528
  "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   529
proof safe
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   530
  fix K :: real
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   531
  from reals_Archimedean2 obtain n :: nat where "K < real n" ..
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   532
  then have "K \<le> real (Suc n)" by auto
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   533
  moreover assume "\<forall>m. norm (X m) \<le> K"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   534
  ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   535
    by (blast intro: order_trans)
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   536
  then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   537
qed (force simp add: real_of_nat_Suc)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   538
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   539
text{* alternative definition for Bseq *}
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   540
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
   541
apply (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   542
apply (simp (no_asm) add: lemma_NBseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   543
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   544
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   545
lemma lemma_NBseq_def2:
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   546
     "(\<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
   547
apply (subst lemma_NBseq_def, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   548
apply (rule_tac x = "Suc N" in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   549
apply (rule_tac [2] x = N in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   550
apply (auto simp add: real_of_nat_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   551
 prefer 2 apply (blast intro: order_less_imp_le)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   552
apply (drule_tac x = n in spec, simp)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   553
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   554
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   555
(* yet another definition for Bseq *)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   556
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
   557
by (simp add: Bseq_def lemma_NBseq_def2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   558
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   559
subsubsection{*A Few More Equivalence Theorems for Boundedness*}
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   560
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   561
text{*alternative formulation for boundedness*}
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   562
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   563
apply (unfold Bseq_def, safe)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   564
apply (rule_tac [2] x = "k + norm x" in exI)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   565
apply (rule_tac x = K in exI, simp)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   566
apply (rule exI [where x = 0], auto)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   567
apply (erule order_less_le_trans, simp)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   568
apply (drule_tac x=n in spec, fold diff_minus)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   569
apply (drule order_trans [OF norm_triangle_ineq2])
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   570
apply simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   571
done
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   572
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   573
text{*alternative formulation for boundedness*}
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   574
lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   575
apply safe
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   576
apply (simp add: Bseq_def, safe)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   577
apply (rule_tac x = "K + norm (X N)" in exI)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   578
apply auto
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   579
apply (erule order_less_le_trans, simp)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   580
apply (rule_tac x = N in exI, safe)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   581
apply (drule_tac x = n in spec)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   582
apply (rule order_trans [OF norm_triangle_ineq], simp)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   583
apply (auto simp add: Bseq_iff2)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   584
done
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   585
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   586
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   587
apply (simp add: Bseq_def)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   588
apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   589
apply (drule_tac x = n in spec, arith)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   590
done
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   591
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   592
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   593
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   594
lemma Bseq_isUb:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   595
  "!!(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
   596
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
   597
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   598
text{* Use completeness of reals (supremum property)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   599
   to show that any bounded sequence has a least upper bound*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   600
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   601
lemma Bseq_isLub:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   602
  "!!(X::nat=>real). Bseq X ==>
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   603
   \<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
   604
by (blast intro: reals_complete Bseq_isUb)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   605
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   606
subsubsection{*A Bounded and Monotonic Sequence Converges*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   607
44714
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   608
(* TODO: delete *)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   609
(* FIXME: one use in NSA/HSEQ.thy *)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   610
lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36660
diff changeset
   611
unfolding tendsto_def eventually_sequentially
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   612
apply (rule_tac x = "X m" in exI, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   613
apply (rule_tac x = m in exI, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   614
apply (drule spec, erule impE, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   615
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   616
44714
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   617
text {* A monotone sequence converges to its least upper bound. *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   618
44714
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   619
lemma isLub_mono_imp_LIMSEQ:
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   620
  fixes X :: "nat \<Rightarrow> real"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   621
  assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   622
  assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   623
  shows "X ----> u"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   624
proof (rule LIMSEQ_I)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   625
  have 1: "\<forall>n. X n \<le> u"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   626
    using isLubD2 [OF u] by auto
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   627
  have "\<forall>y. (\<forall>n. X n \<le> y) \<longrightarrow> u \<le> y"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   628
    using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   629
  hence 2: "\<forall>y<u. \<exists>n. y < X n"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   630
    by (metis not_le)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   631
  fix r :: real assume "0 < r"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   632
  hence "u - r < u" by simp
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   633
  hence "\<exists>m. u - r < X m" using 2 by simp
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   634
  then obtain m where "u - r < X m" ..
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   635
  with X have "\<forall>n\<ge>m. u - r < X n"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   636
    by (fast intro: less_le_trans)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   637
  hence "\<exists>m. \<forall>n\<ge>m. u - r < X n" ..
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   638
  thus "\<exists>m. \<forall>n\<ge>m. norm (X n - u) < r"
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   639
    using 1 by (simp add: diff_less_eq add_commute)
a8990b5d7365 simplify proof of Bseq_mono_convergent
huffman
parents: 44710
diff changeset
   640
qed
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   641
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   642
text{*A standard proof of the theorem for monotone increasing sequence*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   643
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   644
lemma Bseq_mono_convergent:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   645
   "Bseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> convergent (X::nat=>real)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   646
  by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   647
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   648
lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   649
  by (simp add: Bseq_def)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   650
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   651
text{*Main monotonicity theorem*}
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   652
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   653
lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   654
  by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   655
            Bseq_mono_convergent)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   656
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   657
subsubsection{*Increasing and Decreasing Series*}
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   658
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   659
lemma incseq_le: "incseq X \<Longrightarrow> X ----> L \<Longrightarrow> X n \<le> (L::real)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   660
  by (metis incseq_def LIMSEQ_le_const)
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   661
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   662
lemma decseq_le: "decseq X \<Longrightarrow> X ----> L \<Longrightarrow> (L::real) \<le> X n"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   663
  by (metis decseq_def LIMSEQ_le_const2)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   664
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   665
subsection {* Cauchy Sequences *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   666
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   667
lemma metric_CauchyI:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   668
  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   669
  by (simp add: Cauchy_def)
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   670
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   671
lemma metric_CauchyD:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   672
  "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   673
  by (simp add: Cauchy_def)
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   674
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   675
lemma Cauchy_iff:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   676
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   677
  shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   678
  unfolding Cauchy_def dist_norm ..
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   679
35292
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   680
lemma Cauchy_iff2:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   681
  "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))"
35292
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   682
apply (simp add: Cauchy_iff, auto)
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   683
apply (drule reals_Archimedean, safe)
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   684
apply (drule_tac x = n in spec, auto)
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   685
apply (rule_tac x = M in exI, auto)
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   686
apply (drule_tac x = m in spec, simp)
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   687
apply (drule_tac x = na in spec, auto)
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   688
done
e4a431b6d9b7 Replaced Integration by Multivariate-Analysis/Real_Integration
hoelzl
parents: 35216
diff changeset
   689
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   690
lemma CauchyI:
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   691
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   692
  shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   693
by (simp add: Cauchy_iff)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   694
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   695
lemma CauchyD:
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   696
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   697
  shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   698
by (simp add: Cauchy_iff)
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   699
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   700
lemma Cauchy_subseq_Cauchy:
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   701
  "\<lbrakk> Cauchy X; subseq f \<rbrakk> \<Longrightarrow> Cauchy (X o f)"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   702
apply (auto simp add: Cauchy_def)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   703
apply (drule_tac x=e in spec, clarify)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   704
apply (rule_tac x=M in exI, clarify)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   705
apply (blast intro: le_trans [OF _ seq_suble] dest!: spec)
30730
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   706
done
4d3565f2cb0e New theorems mostly concerning infinite series.
paulson
parents: 30273
diff changeset
   707
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   708
subsubsection {* Cauchy Sequences are Bounded *}
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   709
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   710
text{*A Cauchy sequence is bounded -- this is the standard
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   711
  proof mechanization rather than the nonstandard proof*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   712
20563
44eda2314aab replace (x + - y) with (x - y)
huffman
parents: 20552
diff changeset
   713
lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   714
          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   715
apply (clarify, drule spec, drule (1) mp)
20563
44eda2314aab replace (x + - y) with (x - y)
huffman
parents: 20552
diff changeset
   716
apply (simp only: norm_minus_commute)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   717
apply (drule order_le_less_trans [OF norm_triangle_ineq2])
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   718
apply simp
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   719
done
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   720
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   721
lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   722
apply (simp add: Cauchy_iff)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   723
apply (drule spec, drule mp, rule zero_less_one, safe)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   724
apply (drule_tac x="M" in spec, simp)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   725
apply (drule lemmaCauchy)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   726
apply (rule_tac k="M" in Bseq_offset)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   727
apply (simp add: Bseq_def)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   728
apply (rule_tac x="1 + norm (X M)" in exI)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   729
apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   730
apply (simp add: order_less_imp_le)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   731
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   732
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   733
subsubsection {* Cauchy Sequences are Convergent *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   734
44206
5e4a1664106e locale-ize some constant definitions, so complete_space can inherit from metric_space
huffman
parents: 44205
diff changeset
   735
class complete_space = metric_space +
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 32960
diff changeset
   736
  assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
20830
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
   737
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 32960
diff changeset
   738
class banach = real_normed_vector + complete_space
31403
0baaad47cef2 class complete_space
huffman
parents: 31392
diff changeset
   739
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   740
theorem LIMSEQ_imp_Cauchy:
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   741
  assumes X: "X ----> a" shows "Cauchy X"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   742
proof (rule metric_CauchyI)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   743
  fix e::real assume "0 < e"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   744
  hence "0 < e/2" by simp
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   745
  with X have "\<exists>N. \<forall>n\<ge>N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   746
  then obtain N where N: "\<forall>n\<ge>N. dist (X n) a < e/2" ..
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   747
  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < e"
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   748
  proof (intro exI allI impI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   749
    fix m assume "N \<le> m"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   750
    hence m: "dist (X m) a < e/2" using N by fast
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   751
    fix n assume "N \<le> n"
31336
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   752
    hence n: "dist (X n) a < e/2" using N by fast
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   753
    have "dist (X m) (X n) \<le> dist (X m) a + dist (X n) a"
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   754
      by (rule dist_triangle2)
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   755
    also from m n have "\<dots> < e" by simp
e17f13cd1280 generalize constants in SEQ.thy to class metric_space
huffman
parents: 31017
diff changeset
   756
    finally show "dist (X m) (X n) < e" .
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   757
  qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   758
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   759
20691
53cbea20e4d9 add lemma convergent_Cauchy
huffman
parents: 20685
diff changeset
   760
lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   761
unfolding convergent_def
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   762
by (erule exE, erule LIMSEQ_imp_Cauchy)
20691
53cbea20e4d9 add lemma convergent_Cauchy
huffman
parents: 20685
diff changeset
   763
31403
0baaad47cef2 class complete_space
huffman
parents: 31392
diff changeset
   764
lemma Cauchy_convergent_iff:
0baaad47cef2 class complete_space
huffman
parents: 31392
diff changeset
   765
  fixes X :: "nat \<Rightarrow> 'a::complete_space"
0baaad47cef2 class complete_space
huffman
parents: 31392
diff changeset
   766
  shows "Cauchy X = convergent X"
0baaad47cef2 class complete_space
huffman
parents: 31392
diff changeset
   767
by (fast intro: Cauchy_convergent convergent_Cauchy)
0baaad47cef2 class complete_space
huffman
parents: 31392
diff changeset
   768
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   769
text {*
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   770
Proof that Cauchy sequences converge based on the one from
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   771
http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   772
*}
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   773
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   774
text {*
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   775
  If sequence @{term "X"} is Cauchy, then its limit is the lub of
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   776
  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   777
*}
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   778
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   779
lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   780
by (simp add: isUbI setleI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   781
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   782
lemma real_Cauchy_convergent:
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   783
  fixes X :: "nat \<Rightarrow> real"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   784
  assumes X: "Cauchy X"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   785
  shows "convergent X"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   786
proof -
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   787
  def S \<equiv> "{x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   788
  then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
27681
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
   789
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   790
  { fix N x assume N: "\<forall>n\<ge>N. X n < x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   791
  have "isUb UNIV S x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   792
  proof (rule isUb_UNIV_I)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   793
  fix y::real assume "y \<in> S"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   794
  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   795
    by (simp add: S_def)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   796
  then obtain M where "\<forall>n\<ge>M. y < X n" ..
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   797
  hence "y < X (max M N)" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   798
  also have "\<dots> < x" using N by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   799
  finally show "y \<le> x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   800
    by (rule order_less_imp_le)
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   801
  qed }
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   802
  note bound_isUb = this 
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   803
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   804
  have "\<exists>u. isLub UNIV S u"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   805
  proof (rule reals_complete)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   806
  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   807
    using CauchyD [OF X zero_less_one] by auto
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   808
  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   809
  show "\<exists>x. x \<in> S"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   810
  proof
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   811
    from N have "\<forall>n\<ge>N. X N - 1 < X n"
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   812
      by (simp add: abs_diff_less_iff)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   813
    thus "X N - 1 \<in> S" by (rule mem_S)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   814
  qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   815
  show "\<exists>u. isUb UNIV S u"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   816
  proof
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   817
    from N have "\<forall>n\<ge>N. X n < X N + 1"
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   818
      by (simp add: abs_diff_less_iff)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   819
    thus "isUb UNIV S (X N + 1)"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   820
      by (rule bound_isUb)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   821
  qed
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   822
  qed
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   823
  then obtain x where x: "isLub UNIV S x" ..
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   824
  have "X ----> x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   825
  proof (rule LIMSEQ_I)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   826
  fix r::real assume "0 < r"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   827
  hence r: "0 < r/2" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   828
  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
32064
53ca12ff305d refinement of lattice classes
haftmann
parents: 31588
diff changeset
   829
    using CauchyD [OF X r] by auto
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   830
  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   831
  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   832
    by (simp only: real_norm_def abs_diff_less_iff)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   833
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   834
  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   835
  hence "X N - r/2 \<in> S" by (rule mem_S)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   836
  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   837
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   838
  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   839
  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   840
  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   841
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   842
  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   843
  proof (intro exI allI impI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   844
    fix n assume n: "N \<le> n"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   845
    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   846
    thus "norm (X n - x) < r" using 1 2
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 32436
diff changeset
   847
      by (simp add: abs_diff_less_iff)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   848
  qed
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   849
  qed
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   850
  then show ?thesis unfolding convergent_def by auto
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   851
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
   852
20830
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
   853
instance real :: banach
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   854
  by intro_classes (rule real_Cauchy_convergent)
20830
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
   855
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   856
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   857
subsection {* Power Sequences *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   858
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   859
text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   860
"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   861
  also fact that bounded and monotonic sequence converges.*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   862
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   863
lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   864
apply (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   865
apply (rule_tac x = 1 in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   866
apply (simp add: power_abs)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   867
apply (auto dest: power_mono)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   868
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   869
41367
1b65137d598c generalized monoseq, decseq and incseq; simplified proof for seq_monosub
hoelzl
parents: 40811
diff changeset
   870
lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   871
apply (clarify intro!: mono_SucI2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   872
apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   873
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   874
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   875
lemma convergent_realpow:
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   876
  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   877
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   878
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   879
lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   880
  by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   881
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   882
lemma LIMSEQ_realpow_zero:
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   883
  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   884
proof cases
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   885
  assume "0 \<le> x" and "x \<noteq> 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   886
  hence x0: "0 < x" by simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   887
  assume x1: "x < 1"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   888
  from x0 x1 have "1 < inverse x"
36776
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36663
diff changeset
   889
    by (rule one_less_inverse)
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   890
  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   891
    by (rule LIMSEQ_inverse_realpow_zero)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
   892
  thus ?thesis by (simp add: power_inverse)
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   893
qed (rule LIMSEQ_imp_Suc, simp add: tendsto_const)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   894
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
   895
lemma LIMSEQ_power_zero:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30730
diff changeset
   896
  fixes x :: "'a::{real_normed_algebra_1}"
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
   897
  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
   898
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
44313
d81d57979771 SEQ.thy: legacy theorem names
huffman
parents: 44282
diff changeset
   899
apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   900
apply (simp add: power_abs norm_power_ineq)
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
   901
done
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
   902
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   903
lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) ----> 0"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   904
  by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   905
15102
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 15085
diff changeset
   906
text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   907
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   908
lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) ----> 0"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   909
  by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   910
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   911
lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) ----> 0"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50087
diff changeset
   912
  by (rule LIMSEQ_power_zero) simp
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   913
50384
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   914
lemma tendsto_at_topI_sequentially:
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   915
  fixes f :: "real \<Rightarrow> real"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   916
  assumes mono: "mono f"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   917
  assumes limseq: "(\<lambda>n. f (real n)) ----> y"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   918
  shows "(f ---> y) at_top"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   919
proof (rule tendstoI)
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   920
  fix e :: real assume "0 < e"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   921
  with limseq obtain N :: nat where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   922
    by (auto simp: LIMSEQ_def dist_real_def)
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   923
  { fix x :: real
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   924
    from ex_le_of_nat[of x] guess n ..
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   925
    note monoD[OF mono this]
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   926
    also have "f (real_of_nat n) \<le> y"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   927
      by (rule LIMSEQ_le_const[OF limseq])
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   928
         (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric])
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   929
    finally have "f x \<le> y" . }
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   930
  note le = this
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   931
  have "eventually (\<lambda>x. real N \<le> x) at_top"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   932
    by (rule eventually_ge_at_top)
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   933
  then show "eventually (\<lambda>x. dist (f x) y < e) at_top"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   934
  proof eventually_elim
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   935
    fix x assume N': "real N \<le> x"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   936
    with N[of N] le have "y - f (real N) < e" by auto
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   937
    moreover note monoD[OF mono N']
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   938
    ultimately show "dist (f x) y < e"
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   939
      using le[of x] by (auto simp: dist_real_def field_simps)
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   940
  qed
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   941
qed
b9b967da28e9 rules for improper Lebesgue integrals (using tendsto at_top)
hoelzl
parents: 50331
diff changeset
   942
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   943
end