src/HOL/SEQ.thy
author wenzelm
Wed, 11 Mar 2009 20:36:20 +0100
changeset 30458 804de935c328
parent 30273 ecd6f0ca62ea
child 30730 4d3565f2cb0e
permissions -rw-r--r--
delete unused generated files;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : SEQ.thy
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     4
    Description : Convergence of sequences and series
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
     5
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
     6
    Additional contributions by Jeremy Avigad and Brian Huffman
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
     7
*)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     8
22631
7ae5a6ab7bd6 moved nonstandard stuff from SEQ.thy into new file HSEQ.thy
huffman
parents: 22629
diff changeset
     9
header {* Sequences and Convergence *}
17439
a358da1a0218 add header
huffman
parents: 17429
diff changeset
    10
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15102
diff changeset
    11
theory SEQ
29197
6d4cb27ed19c adapted HOL source structure to distribution layout
haftmann
parents: 28952
diff changeset
    12
imports RealVector RComplete
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15102
diff changeset
    13
begin
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    14
19765
dfe940911617 misc cleanup;
wenzelm
parents: 18585
diff changeset
    15
definition
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    16
  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    17
    --{*Standard definition of sequence converging to zero*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    18
  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    19
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    20
definition
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
    21
  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    22
    ("((_)/ ----> (_))" [60, 60] 60) where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    23
    --{*Standard definition of convergence of sequence*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    24
  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
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
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    27
  lim :: "(nat => 'a::real_normed_vector) => 'a" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    28
    --{*Standard definition of limit using choice operator*}
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
    29
  "lim X = (THE L. X ----> L)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    30
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    31
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    32
  convergent :: "(nat => 'a::real_normed_vector) => bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    33
    --{*Standard definition of convergence*}
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
    34
  "convergent X = (\<exists>L. X ----> L)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    35
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    36
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    37
  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    38
    --{*Standard definition for bounded sequence*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    39
  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    40
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    41
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    42
  monoseq :: "(nat=>real)=>bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    43
    --{*Definition for monotonicity*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    44
  [code del]: "monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) | (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    45
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    46
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    47
  subseq :: "(nat => nat) => bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    48
    --{*Definition of subsequence*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    49
  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    50
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    51
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21139
diff changeset
    52
  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    53
    --{*Standard definition of the Cauchy condition*}
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27681
diff changeset
    54
  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    55
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
    56
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    57
subsection {* Bounded Sequences *}
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    58
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
    59
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
    60
unfolding Bseq_def
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    61
proof (intro exI conjI allI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    62
  show "0 < max K 1" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    63
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    64
  fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    65
  have "norm (X n) \<le> K" by (rule K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    66
  thus "norm (X n) \<le> max K 1" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    67
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    68
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    69
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
    70
unfolding Bseq_def by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    71
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
    72
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
    73
proof (rule BseqI')
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    74
  let ?A = "norm ` X ` {..N}"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    75
  have 1: "finite ?A" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    76
  fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    77
  show "norm (X n) \<le> max K (Max ?A)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    78
  proof (cases rule: linorder_le_cases)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    79
    assume "n \<ge> N"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    80
    hence "norm (X n) \<le> K" using K by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    81
    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
    82
  next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    83
    assume "n \<le> N"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    84
    hence "norm (X n) \<in> ?A" by simp
26757
e775accff967 thms Max_ge, Min_le: dropped superfluous premise
haftmann
parents: 26312
diff changeset
    85
    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
    86
    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
    87
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    88
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    89
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    90
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
    91
unfolding Bseq_def by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    92
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    93
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
    94
apply (erule BseqE)
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
    95
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
    96
apply clarify
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    97
apply (drule_tac x="n - k" in spec, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    98
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
    99
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   100
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   101
subsection {* Sequences That Converge to Zero *}
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   102
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   103
lemma ZseqI:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   104
  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   105
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   106
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   107
lemma ZseqD:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   108
  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   109
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   110
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   111
lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   112
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   113
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   114
lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   115
unfolding Zseq_def by force
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   116
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   117
lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   118
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   119
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   120
lemma Zseq_imp_Zseq:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   121
  assumes X: "Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   122
  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   123
  shows "Zseq (\<lambda>n. Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   124
proof (cases)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   125
  assume K: "0 < K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   126
  show ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   127
  proof (rule ZseqI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   128
    fix r::real assume "0 < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   129
    hence "0 < r / K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   130
      using K by (rule divide_pos_pos)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   131
    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   132
      using ZseqD [OF X] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   133
    hence "\<forall>n\<ge>N. norm (X n) * K < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   134
      by (simp add: pos_less_divide_eq K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   135
    hence "\<forall>n\<ge>N. norm (Y n) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   136
      by (simp add: order_le_less_trans [OF Y])
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   137
    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   138
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   139
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   140
  assume "\<not> 0 < K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   141
  hence K: "K \<le> 0" by (simp only: linorder_not_less)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   142
  {
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   143
    fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   144
    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   145
    also have "\<dots> \<le> norm (X n) * 0"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   146
      using K norm_ge_zero by (rule mult_left_mono)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   147
    finally have "norm (Y n) = 0" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   148
  }
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   149
  thus ?thesis by (simp add: Zseq_zero)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   150
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   151
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   152
lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   153
by (erule_tac K="1" in Zseq_imp_Zseq, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   154
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   155
lemma Zseq_add:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   156
  assumes X: "Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   157
  assumes Y: "Zseq Y"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   158
  shows "Zseq (\<lambda>n. X n + Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   159
proof (rule ZseqI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   160
  fix r::real assume "0 < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   161
  hence r: "0 < r / 2" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   162
  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   163
    using ZseqD [OF X r] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   164
  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   165
    using ZseqD [OF Y r] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   166
  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   167
  proof (intro exI allI impI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   168
    fix n assume n: "max M N \<le> n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   169
    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   170
      by (rule norm_triangle_ineq)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   171
    also have "\<dots> < r/2 + r/2"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   172
    proof (rule add_strict_mono)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   173
      from M n show "norm (X n) < r/2" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   174
      from N n show "norm (Y n) < r/2" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   175
    qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   176
    finally show "norm (X n + Y n) < r" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   177
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   178
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   179
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   180
lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   181
unfolding Zseq_def by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   182
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   183
lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   184
by (simp only: diff_minus Zseq_add Zseq_minus)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   185
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   186
lemma (in bounded_linear) Zseq:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   187
  assumes X: "Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   188
  shows "Zseq (\<lambda>n. f (X n))"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   189
proof -
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   190
  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   191
    using bounded by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   192
  with X show ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   193
    by (rule Zseq_imp_Zseq)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   194
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   195
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   196
lemma (in bounded_bilinear) Zseq:
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   197
  assumes X: "Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   198
  assumes Y: "Zseq Y"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   199
  shows "Zseq (\<lambda>n. X n ** Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   200
proof (rule ZseqI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   201
  fix r::real assume r: "0 < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   202
  obtain K where K: "0 < K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   203
    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   204
    using pos_bounded by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   205
  from K have K': "0 < inverse K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   206
    by (rule positive_imp_inverse_positive)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   207
  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   208
    using ZseqD [OF X r] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   209
  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   210
    using ZseqD [OF Y K'] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   211
  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   212
  proof (intro exI allI impI)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   213
    fix n assume n: "max M N \<le> n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   214
    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   215
      by (rule norm_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   216
    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   217
    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   218
      from M n show Xn: "norm (X n) < r" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   219
      from N n show Yn: "norm (Y n) < inverse K" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   220
    qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   221
    also from K have "r * inverse K * K = r" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   222
    finally show "norm (X n ** Y n) < r" .
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   223
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   224
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   225
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   226
lemma (in bounded_bilinear) Zseq_prod_Bseq:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   227
  assumes X: "Zseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   228
  assumes Y: "Bseq Y"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   229
  shows "Zseq (\<lambda>n. X n ** Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   230
proof -
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   231
  obtain K where K: "0 \<le> K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   232
    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   233
    using nonneg_bounded by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   234
  obtain B where B: "0 < B"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   235
    and norm_Y: "\<And>n. norm (Y n) \<le> B"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   236
    using Y [unfolded Bseq_def] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   237
  from X show ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   238
  proof (rule Zseq_imp_Zseq)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   239
    fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   240
    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   241
      by (rule norm_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   242
    also have "\<dots> \<le> norm (X n) * B * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   243
      by (intro mult_mono' order_refl norm_Y norm_ge_zero
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   244
                mult_nonneg_nonneg K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   245
    also have "\<dots> = norm (X n) * (B * K)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   246
      by (rule mult_assoc)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   247
    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   248
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   249
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   250
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   251
lemma (in bounded_bilinear) Bseq_prod_Zseq:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   252
  assumes X: "Bseq X"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   253
  assumes Y: "Zseq Y"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   254
  shows "Zseq (\<lambda>n. X n ** Y n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   255
proof -
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   256
  obtain K where K: "0 \<le> K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   257
    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   258
    using nonneg_bounded by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   259
  obtain B where B: "0 < B"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   260
    and norm_X: "\<And>n. norm (X n) \<le> B"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   261
    using X [unfolded Bseq_def] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   262
  from Y show ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   263
  proof (rule Zseq_imp_Zseq)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   264
    fix n::nat
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   265
    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   266
      by (rule norm_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   267
    also have "\<dots> \<le> B * norm (Y n) * K"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   268
      by (intro mult_mono' order_refl norm_X norm_ge_zero
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   269
                mult_nonneg_nonneg K)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   270
    also have "\<dots> = norm (Y n) * (B * K)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   271
      by (simp only: mult_ac)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   272
    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   273
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   274
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   275
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   276
lemma (in bounded_bilinear) Zseq_left:
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   277
  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   278
by (rule bounded_linear_left [THEN bounded_linear.Zseq])
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   279
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   280
lemma (in bounded_bilinear) Zseq_right:
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   281
  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   282
by (rule bounded_linear_right [THEN bounded_linear.Zseq])
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   283
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   284
lemmas Zseq_mult = mult.Zseq
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   285
lemmas Zseq_mult_right = mult.Zseq_right
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   286
lemmas Zseq_mult_left = mult.Zseq_left
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   287
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   288
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   289
subsection {* Limits of Sequences *}
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   290
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   291
lemma LIMSEQ_iff:
20563
44eda2314aab replace (x + - y) with (x - y)
huffman
parents: 20552
diff changeset
   292
      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   293
by (rule LIMSEQ_def)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   294
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   295
lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   296
by (simp only: LIMSEQ_def Zseq_def)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   297
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   298
lemma LIMSEQ_I:
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   299
  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   300
by (simp add: LIMSEQ_def)
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   301
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   302
lemma LIMSEQ_D:
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   303
  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   304
by (simp add: LIMSEQ_def)
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
   305
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   306
lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   307
by (simp add: LIMSEQ_def)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   308
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   309
lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   310
by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   311
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   312
lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   313
apply (simp add: LIMSEQ_def, safe)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   314
apply (drule_tac x="r" in spec, safe)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   315
apply (rule_tac x="no" in exI, safe)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   316
apply (drule_tac x="n" in spec, safe)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   317
apply (erule order_le_less_trans [OF norm_triangle_ineq3])
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   318
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   319
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   320
lemma LIMSEQ_ignore_initial_segment:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   321
  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   322
apply (rule LIMSEQ_I)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   323
apply (drule (1) LIMSEQ_D)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   324
apply (erule exE, rename_tac N)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   325
apply (rule_tac x=N in exI)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   326
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   327
done
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   328
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   329
lemma LIMSEQ_offset:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   330
  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   331
apply (rule LIMSEQ_I)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   332
apply (drule (1) LIMSEQ_D)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   333
apply (erule exE, rename_tac N)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   334
apply (rule_tac x="N + k" in exI)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   335
apply clarify
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   336
apply (drule_tac x="n - k" in spec)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   337
apply (simp add: le_diff_conv2)
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   338
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   339
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   340
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
   341
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
   342
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   343
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
   344
by (rule_tac k="Suc 0" in LIMSEQ_offset, 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_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   347
by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   348
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   349
lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   350
  unfolding LIMSEQ_def
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   351
  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
   352
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   353
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   354
lemma add_diff_add:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   355
  fixes a b c d :: "'a::ab_group_add"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   356
  shows "(a + c) - (b + d) = (a - b) + (c - d)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   357
by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   358
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   359
lemma minus_diff_minus:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   360
  fixes a b :: "'a::ab_group_add"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   361
  shows "(- a) - (- b) = - (a - b)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   362
by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   363
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   364
lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   365
by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   366
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   367
lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   368
by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   369
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   370
lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   371
by (drule LIMSEQ_minus, simp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   372
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   373
lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   374
by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   375
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   376
lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   377
by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   378
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   379
lemma (in bounded_linear) LIMSEQ:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   380
  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   381
by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   382
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   383
lemma (in bounded_bilinear) LIMSEQ:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   384
  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   385
by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   386
               Zseq_add Zseq Zseq_left Zseq_right)
22608
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 LIMSEQ_mult:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   389
  fixes a b :: "'a::real_normed_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   390
  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   391
by (rule mult.LIMSEQ)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   392
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   393
lemma inverse_diff_inverse:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   394
  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   395
   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29197
diff changeset
   396
by (simp add: algebra_simps)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   397
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   398
lemma Bseq_inverse_lemma:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   399
  fixes x :: "'a::real_normed_div_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   400
  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
   401
apply (subst nonzero_norm_inverse, clarsimp)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   402
apply (erule (1) le_imp_inverse_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   403
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   404
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   405
lemma Bseq_inverse:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   406
  fixes a :: "'a::real_normed_div_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   407
  assumes X: "X ----> a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   408
  assumes a: "a \<noteq> 0"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   409
  shows "Bseq (\<lambda>n. inverse (X n))"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   410
proof -
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   411
  from a have "0 < norm a" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   412
  hence "\<exists>r>0. r < norm a" by (rule dense)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   413
  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   414
  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   415
    using LIMSEQ_D [OF X r1] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   416
  show ?thesis
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 23482
diff changeset
   417
  proof (rule BseqI2' [rule_format])
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   418
    fix n assume n: "N \<le> n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   419
    hence 1: "norm (X n - a) < r" by (rule N)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   420
    hence 2: "X n \<noteq> 0" using r2 by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   421
    hence "norm (inverse (X n)) = inverse (norm (X n))"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   422
      by (rule nonzero_norm_inverse)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   423
    also have "\<dots> \<le> inverse (norm a - r)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   424
    proof (rule le_imp_inverse_le)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   425
      show "0 < norm a - r" using r2 by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   426
    next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   427
      have "norm a - norm (X n) \<le> norm (a - X n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   428
        by (rule norm_triangle_ineq2)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   429
      also have "\<dots> = norm (X n - a)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   430
        by (rule norm_minus_commute)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   431
      also have "\<dots> < r" using 1 .
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   432
      finally show "norm a - r \<le> norm (X n)" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   433
    qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   434
    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   435
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   436
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   437
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   438
lemma LIMSEQ_inverse_lemma:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   439
  fixes a :: "'a::real_normed_div_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   440
  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   441
         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   442
apply (subst LIMSEQ_Zseq_iff)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   443
apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   444
apply (rule Zseq_minus)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   445
apply (rule Zseq_mult_left)
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 22998
diff changeset
   446
apply (rule mult.Bseq_prod_Zseq)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   447
apply (erule (1) Bseq_inverse)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   448
apply (simp add: LIMSEQ_Zseq_iff)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   449
done
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   450
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   451
lemma LIMSEQ_inverse:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   452
  fixes a :: "'a::real_normed_div_algebra"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   453
  assumes X: "X ----> a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   454
  assumes a: "a \<noteq> 0"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   455
  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   456
proof -
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   457
  from a have "0 < norm a" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   458
  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   459
    using LIMSEQ_D [OF X] by fast
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   460
  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   461
  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   462
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   463
  from X have "(\<lambda>n. X (n + k)) ----> a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   464
    by (rule LIMSEQ_ignore_initial_segment)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   465
  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   466
    using a k by (rule LIMSEQ_inverse_lemma)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   467
  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   468
    by (rule LIMSEQ_offset)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   469
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   470
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   471
lemma LIMSEQ_divide:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   472
  fixes a b :: "'a::real_normed_field"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   473
  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   474
by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   475
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   476
lemma LIMSEQ_pow:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   477
  fixes a :: "'a::{real_normed_algebra,recpower}"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   478
  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   479
by (induct m) (simp_all add: LIMSEQ_const LIMSEQ_mult)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   480
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   481
lemma LIMSEQ_setsum:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   482
  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   483
  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   484
proof (cases "finite S")
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   485
  case True
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   486
  thus ?thesis using n
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   487
  proof (induct)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   488
    case empty
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   489
    show ?case
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   490
      by (simp add: LIMSEQ_const)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   491
  next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   492
    case insert
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   493
    thus ?case
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   494
      by (simp add: LIMSEQ_add)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   495
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   496
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   497
  case False
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   498
  thus ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   499
    by (simp add: LIMSEQ_const)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   500
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   501
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   502
lemma LIMSEQ_setprod:
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   503
  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   504
  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   505
  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   506
proof (cases "finite S")
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   507
  case True
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   508
  thus ?thesis using n
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   509
  proof (induct)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   510
    case empty
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   511
    show ?case
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   512
      by (simp add: LIMSEQ_const)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   513
  next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   514
    case insert
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   515
    thus ?case
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   516
      by (simp add: LIMSEQ_mult)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   517
  qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   518
next
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   519
  case False
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   520
  thus ?thesis
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   521
    by (simp add: setprod_def LIMSEQ_const)
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   522
qed
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
   523
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   524
lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   525
by (simp add: LIMSEQ_add LIMSEQ_const)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   526
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   527
(* FIXME: delete *)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   528
lemma LIMSEQ_add_minus:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   529
     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   530
by (simp only: LIMSEQ_add LIMSEQ_minus)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   531
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   532
lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   533
by (simp add: LIMSEQ_diff LIMSEQ_const)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   534
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   535
lemma LIMSEQ_diff_approach_zero: 
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   536
  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   537
     f ----> L"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   538
  apply (drule LIMSEQ_add)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   539
  apply assumption
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   540
  apply simp
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   541
done
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   542
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   543
lemma LIMSEQ_diff_approach_zero2: 
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   544
  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   545
     g ----> L";
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   546
  apply (drule LIMSEQ_diff)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   547
  apply assumption
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   548
  apply simp
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   549
done
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   550
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   551
text{*A sequence tends to zero iff its abs does*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   552
lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   553
by (simp add: LIMSEQ_def)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   554
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   555
lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   556
by (simp add: LIMSEQ_def)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   557
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   558
lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   559
by (drule LIMSEQ_norm, simp)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   560
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   561
text{*An unbounded sequence's inverse tends to 0*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   562
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   563
lemma LIMSEQ_inverse_zero:
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   564
  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   565
apply (rule LIMSEQ_I)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   566
apply (drule_tac x="inverse r" in spec, safe)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   567
apply (rule_tac x="N" in exI, safe)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   568
apply (drule_tac x="n" in spec, safe)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   569
apply (frule positive_imp_inverse_positive)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   570
apply (frule (1) less_imp_inverse_less)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   571
apply (subgoal_tac "0 < X n", simp)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   572
apply (erule (1) order_less_trans)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   573
done
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   574
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   575
text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   576
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   577
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   578
apply (rule LIMSEQ_inverse_zero, safe)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
   579
apply (cut_tac x = r in reals_Archimedean2)
22614
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   580
apply (safe, rule_tac x = n in exI)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   581
apply (auto simp add: real_of_nat_Suc)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   582
done
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   583
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   584
text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   585
infinity is now easily proved*}
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   586
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   587
lemma LIMSEQ_inverse_real_of_nat_add:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   588
     "(%n. r + inverse(real(Suc n))) ----> r"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   589
by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   590
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   591
lemma LIMSEQ_inverse_real_of_nat_add_minus:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   592
     "(%n. r + -inverse(real(Suc n))) ----> r"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   593
by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   594
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   595
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   596
     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   597
by (cut_tac b=1 in
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   598
        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
17644bc9cee4 rearranged sections
huffman
parents: 22608
diff changeset
   599
22615
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   600
lemma LIMSEQ_le_const:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   601
  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   602
apply (rule ccontr, simp only: linorder_not_le)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   603
apply (drule_tac r="a - x" in LIMSEQ_D, simp)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   604
apply clarsimp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   605
apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   606
apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   607
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   608
done
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   609
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   610
lemma LIMSEQ_le_const2:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   611
  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   612
apply (subgoal_tac "- a \<le> - x", simp)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   613
apply (rule LIMSEQ_le_const)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   614
apply (erule LIMSEQ_minus)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   615
apply simp
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   616
done
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   617
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   618
lemma LIMSEQ_le:
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   619
  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   620
apply (subgoal_tac "0 \<le> y - x", simp)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   621
apply (rule LIMSEQ_le_const)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   622
apply (erule (1) LIMSEQ_diff)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   623
apply (simp add: le_diff_eq)
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   624
done
d650e51b5970 new standard proofs of some LIMSEQ lemmas
huffman
parents: 22614
diff changeset
   625
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   626
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   627
subsection {* Convergence *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   628
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   629
lemma limI: "X ----> L ==> lim X = L"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   630
apply (simp add: lim_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   631
apply (blast intro: LIMSEQ_unique)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   632
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   633
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   634
lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   635
by (simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   636
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   637
lemma convergentI: "(X ----> L) ==> convergent X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   638
by (auto simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   639
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   640
lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
20682
cecff1f51431 define constants with THE instead of SOME
huffman
parents: 20653
diff changeset
   641
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
   642
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   643
lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   644
apply (simp add: convergent_def)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   645
apply (auto dest: LIMSEQ_minus)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   646
apply (drule LIMSEQ_minus, auto)
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   647
done
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   648
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   649
text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   650
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   651
lemma nat_function_unique: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   652
  unfolding Ex1_def
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   653
  apply (rule_tac x="nat_rec e f" in exI)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   654
  apply (rule conjI)+
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   655
apply (rule def_nat_rec_0, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   656
apply (rule allI, rule def_nat_rec_Suc, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   657
apply (rule allI, rule impI, rule ext)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   658
apply (erule conjE)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   659
apply (induct_tac x)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   660
apply (simp add: nat_rec_0)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   661
apply (erule_tac x="n" in allE)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   662
apply (simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   663
done
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   664
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   665
text{*Subsequence (alternative definition, (e.g. Hoskins)*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   666
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   667
lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   668
apply (simp add: subseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   669
apply (auto dest!: less_imp_Suc_add)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   670
apply (induct_tac k)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   671
apply (auto intro: less_trans)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   672
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   673
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   674
lemma monoseq_Suc:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   675
   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   676
                 | (\<forall>n. X (Suc n) \<le> X n))"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   677
apply (simp add: monoseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   678
apply (auto dest!: le_imp_less_or_eq)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   679
apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   680
apply (induct_tac "ka")
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   681
apply (auto intro: order_trans)
18585
5d379fe2eb74 replaced swap by contrapos_np;
wenzelm
parents: 17439
diff changeset
   682
apply (erule contrapos_np)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   683
apply (induct_tac "k")
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   684
apply (auto intro: order_trans)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   685
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   686
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   687
lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   688
by (simp add: monoseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   689
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   690
lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   691
by (simp add: monoseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   692
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   693
lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   694
by (simp add: monoseq_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   695
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   696
lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   697
by (simp add: monoseq_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   698
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   699
lemma monoseq_minus: assumes "monoseq a"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   700
  shows "monoseq (\<lambda> n. - a n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   701
proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   702
  case True
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   703
  hence "\<forall> m. \<forall> n \<ge> m. - a n \<le> - a m" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   704
  thus ?thesis by (rule monoI2)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   705
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   706
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   707
  hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using `monoseq a`[unfolded monoseq_def] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   708
  thus ?thesis by (rule monoI1)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   709
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   710
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   711
lemma monoseq_le: assumes "monoseq a" and "a ----> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   712
  shows "((\<forall> n. a n \<le> x) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)) \<or> 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   713
         ((\<forall> n. x \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   714
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   715
  { fix x n fix a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   716
    assume "a ----> x" and "\<forall> m. \<forall> n \<ge> m. a m \<le> a n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   717
    hence monotone: "\<And> m n. m \<le> n \<Longrightarrow> a m \<le> a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   718
    have "a n \<le> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   719
    proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   720
      assume "\<not> a n \<le> x" hence "x < a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   721
      hence "0 < a n - x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   722
      from `a ----> x`[THEN LIMSEQ_D, OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   723
      obtain no where "\<And>n'. no \<le> n' \<Longrightarrow> norm (a n' - x) < a n - x" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   724
      hence "norm (a (max no n) - x) < a n - x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   725
      moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   726
      { fix n' have "n \<le> n' \<Longrightarrow> x < a n'" using monotone[where m=n and n=n'] and `x < a n` by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   727
      hence "x < a (max no n)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   728
      ultimately
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   729
      have "a (max no n) < a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   730
      with monotone[where m=n and n="max no n"]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   731
      show False by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   732
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   733
  } note top_down = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   734
  { fix x n m fix a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   735
    assume "a ----> x" and "monoseq a" and "a m < x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   736
    have "a n \<le> x \<and> (\<forall> m. \<forall> n \<ge> m. a m \<le> a n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   737
    proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   738
      case True with top_down and `a ----> x` show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   739
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   740
      case False with `monoseq a`[unfolded monoseq_def] have "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   741
      hence "- a m \<le> - x" using top_down[OF LIMSEQ_minus[OF `a ----> x`]] by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   742
      hence False using `a m < x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   743
      thus ?thesis ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   744
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   745
  } note when_decided = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   746
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   747
  show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   748
  proof (cases "\<exists> m. a m \<noteq> x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   749
    case True then obtain m where "a m \<noteq> x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   750
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   751
    proof (cases "a m < x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   752
      case True with when_decided[OF `a ----> x` `monoseq a`, where m2=m]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   753
      show ?thesis by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   754
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   755
      case False hence "- a m < - x" using `a m \<noteq> x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   756
      with when_decided[OF LIMSEQ_minus[OF `a ----> x`] monoseq_minus[OF `monoseq a`], where m2=m]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   757
      show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   758
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   759
  qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   760
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   761
30196
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   762
text{* for any sequence, there is a mootonic subsequence *}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   763
lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   764
proof-
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   765
  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   766
    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   767
    from nat_function_unique[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   768
    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   769
    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   770
      using H apply - 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   771
      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   772
      unfolding order_le_less by blast 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   773
    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   774
    {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   775
      have "?P (f (Suc n)) (f n)" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   776
	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   777
	using H apply - 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   778
      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   779
      unfolding order_le_less by blast 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   780
    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   781
  note fSuc = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   782
    {fix p q assume pq: "p \<ge> f q"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   783
      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   784
	by (cases q, simp_all) }
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   785
    note pqth = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   786
    {fix q
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   787
      have "f (Suc q) > f q" apply (induct q) 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   788
	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   789
    note fss = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   790
    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   791
    {fix a b 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   792
      have "f a \<le> f (a + b)"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   793
      proof(induct b)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   794
	case 0 thus ?case by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   795
      next
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   796
	case (Suc b)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   797
	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   798
      qed}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   799
    note fmon0 = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   800
    have "monoseq (\<lambda>n. s (f n))" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   801
    proof-
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   802
      {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   803
	have "s (f n) \<ge> s (f (Suc n))" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   804
	proof(cases n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   805
	  case 0
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   806
	  assume n0: "n = 0"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   807
	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   808
	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   809
	next
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   810
	  case (Suc m)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   811
	  assume m: "n = Suc m"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   812
	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   813
	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   814
	qed}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   815
      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   816
    qed
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   817
    with th1 have ?thesis by blast}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   818
  moreover
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   819
  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   820
    {fix p assume p: "p \<ge> Suc N" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   821
      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   822
      have "m \<noteq> p" using m(2) by auto 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   823
      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   824
    note th0 = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   825
    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   826
    from nat_function_unique[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   827
    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   828
      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   829
    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   830
      using N apply - 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   831
      apply (erule allE[where x="Suc N"], clarsimp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   832
      apply (rule_tac x="m" in exI)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   833
      apply auto
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   834
      apply (subgoal_tac "Suc N \<noteq> m")
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   835
      apply simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   836
      apply (rule ccontr, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   837
      done
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   838
    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   839
    {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   840
      have "f n > N \<and> ?P (f (Suc n)) (f n)"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   841
	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   842
      proof (induct n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   843
	case 0 thus ?case
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   844
	  using f0 N apply auto 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   845
	  apply (erule allE[where x="f 0"], clarsimp) 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   846
	  apply (rule_tac x="m" in exI, simp)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   847
	  by (subgoal_tac "f 0 \<noteq> m", auto)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   848
      next
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   849
	case (Suc n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   850
	from Suc.hyps have Nfn: "N < f n" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   851
	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   852
	with Nfn have mN: "m > N" by arith
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   853
	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   854
	
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   855
	from key have th0: "f (Suc n) > N" by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   856
	from N[rule_format, OF th0]
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   857
	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   858
	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   859
	hence "m' > f (Suc n)" using m'(1) by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   860
	with key m'(2) show ?case by auto
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   861
      qed}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   862
    note fSuc = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   863
    {fix n
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   864
      have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   865
      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   866
    note thf = this
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   867
    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   868
    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   869
      apply -
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   870
      apply (rule disjI1)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   871
      apply auto
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   872
      apply (rule order_less_imp_le)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   873
      apply blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   874
      done
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   875
    then have ?thesis  using sqf by blast}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   876
  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   877
qed
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   878
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   879
lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   880
proof(induct n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   881
  case 0 thus ?case by simp
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   882
next
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   883
  case (Suc n)
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   884
  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   885
  have "n < f (Suc n)" by arith 
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   886
  thus ?case by arith
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   887
qed
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   888
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   889
subsection {* Bounded Monotonic Sequences *}
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   890
6ffaa79c352c Moved a few theorems about monotonic sequences from Fundamental_Theorem_Algebra to SEQ.thy
chaieb
parents: 30082
diff changeset
   891
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   892
text{*Bounded Sequence*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   893
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   894
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
   895
by (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   896
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   897
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
   898
by (auto simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   899
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   900
lemma lemma_NBseq_def:
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   901
     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   902
      (\<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
   903
apply auto
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   904
 prefer 2 apply force
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   905
apply (cut_tac x = K in reals_Archimedean2, clarify)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   906
apply (rule_tac x = n in exI, clarify)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   907
apply (drule_tac x = na in spec)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   908
apply (auto simp add: real_of_nat_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   909
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   910
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   911
text{* alternative definition for Bseq *}
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   912
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
   913
apply (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   914
apply (simp (no_asm) add: lemma_NBseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   915
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   916
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   917
lemma lemma_NBseq_def2:
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   918
     "(\<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
   919
apply (subst lemma_NBseq_def, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   920
apply (rule_tac x = "Suc N" in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   921
apply (rule_tac [2] x = N in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   922
apply (auto simp add: real_of_nat_Suc)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   923
 prefer 2 apply (blast intro: order_less_imp_le)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   924
apply (drule_tac x = n in spec, simp)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   925
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   926
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   927
(* yet another definition for Bseq *)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   928
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
   929
by (simp add: Bseq_def lemma_NBseq_def2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   930
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   931
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   932
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   933
lemma Bseq_isUb:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   934
  "!!(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
   935
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
   936
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   937
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   938
text{* Use completeness of reals (supremum property)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   939
   to show that any bounded sequence has a least upper bound*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   940
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   941
lemma Bseq_isLub:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   942
  "!!(X::nat=>real). Bseq X ==>
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   943
   \<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
   944
by (blast intro: reals_complete Bseq_isUb)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   945
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
   946
subsubsection{*A Bounded and Monotonic Sequence Converges*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   947
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   948
lemma lemma_converg1:
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   949
     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   950
                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
   951
               |] ==> \<forall>n \<ge> ma. X n = X ma"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   952
apply safe
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   953
apply (drule_tac y = "X n" in isLubD2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   954
apply (blast dest: order_antisym)+
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   955
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   956
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   957
text{* The best of both worlds: Easier to prove this result as a standard
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   958
   theorem and then use equivalence to "transfer" it into the
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   959
   equivalent nonstandard form if needed!*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   960
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   961
lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   962
apply (simp add: LIMSEQ_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   963
apply (rule_tac x = "X m" in exI, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   964
apply (rule_tac x = m in exI, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   965
apply (drule spec, erule impE, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   966
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   967
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   968
lemma lemma_converg2:
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   969
   "!!(X::nat=>real).
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   970
    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   971
apply safe
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   972
apply (drule_tac y = "X m" in isLubD2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   973
apply (auto dest!: order_le_imp_less_or_eq)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   974
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   975
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   976
lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   977
by (rule setleI [THEN isUbI], auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   978
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   979
text{* FIXME: @{term "U - T < U"} is redundant *}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   980
lemma lemma_converg4: "!!(X::nat=> real).
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   981
               [| \<forall>m. X m ~= U;
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   982
                  isLub UNIV {x. \<exists>n. X n = x} U;
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   983
                  0 < T;
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   984
                  U + - T < U
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   985
               |] ==> \<exists>m. U + -T < X m & X m < U"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   986
apply (drule lemma_converg2, assumption)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   987
apply (rule ccontr, simp)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   988
apply (simp add: linorder_not_less)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   989
apply (drule lemma_converg3)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   990
apply (drule isLub_le_isUb, assumption)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   991
apply (auto dest: order_less_le_trans)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   992
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   993
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   994
text{*A standard proof of the theorem for monotone increasing sequence*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   995
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   996
lemma Bseq_mono_convergent:
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
   997
     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   998
apply (simp add: convergent_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
   999
apply (frule Bseq_isLub, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1000
apply (case_tac "\<exists>m. X m = U", auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1001
apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1002
(* second case *)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1003
apply (rule_tac x = U in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1004
apply (subst LIMSEQ_iff, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1005
apply (frule lemma_converg2, assumption)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1006
apply (drule lemma_converg4, auto)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1007
apply (rule_tac x = m in exI, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1008
apply (subgoal_tac "X m \<le> X n")
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1009
 prefer 2 apply blast
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1010
apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1011
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1012
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1013
lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1014
by (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1015
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1016
text{*Main monotonicity theorem*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1017
lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1018
apply (simp add: monoseq_def, safe)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1019
apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1020
apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1021
apply (auto intro!: Bseq_mono_convergent)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1022
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1023
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
  1024
subsubsection{*A Few More Equivalence Theorems for Boundedness*}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1025
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1026
text{*alternative formulation for boundedness*}
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1027
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1028
apply (unfold Bseq_def, safe)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1029
apply (rule_tac [2] x = "k + norm x" in exI)
15360
300e09825d8b Added "ALL x > y" and relatives.
nipkow
parents: 15312
diff changeset
  1030
apply (rule_tac x = K in exI, simp)
15221
8412cfdf3287 tweaking of arithmetic proofs
paulson
parents: 15140
diff changeset
  1031
apply (rule exI [where x = 0], auto)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1032
apply (erule order_less_le_trans, simp)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1033
apply (drule_tac x=n in spec, fold diff_def)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1034
apply (drule order_trans [OF norm_triangle_ineq2])
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1035
apply simp
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1036
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1037
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1038
text{*alternative formulation for boundedness*}
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1039
lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1040
apply safe
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1041
apply (simp add: Bseq_def, safe)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1042
apply (rule_tac x = "K + norm (X N)" in exI)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1043
apply auto
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1044
apply (erule order_less_le_trans, simp)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1045
apply (rule_tac x = N in exI, safe)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1046
apply (drule_tac x = n in spec)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1047
apply (rule order_trans [OF norm_triangle_ineq], simp)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1048
apply (auto simp add: Bseq_iff2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1049
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1050
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1051
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1052
apply (simp add: Bseq_def)
15221
8412cfdf3287 tweaking of arithmetic proofs
paulson
parents: 15140
diff changeset
  1053
apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
  1054
apply (drule_tac x = n in spec, arith)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1055
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1056
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1057
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
  1058
subsection {* Cauchy Sequences *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1059
20751
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1060
lemma CauchyI:
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1061
  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1062
by (simp add: Cauchy_def)
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1063
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1064
lemma CauchyD:
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1065
  "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1066
by (simp add: Cauchy_def)
93271c59d211 add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents: 20740
diff changeset
  1067
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
  1068
subsubsection {* Cauchy Sequences are Bounded *}
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
  1069
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1070
text{*A Cauchy sequence is bounded -- this is the standard
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1071
  proof mechanization rather than the nonstandard proof*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1072
20563
44eda2314aab replace (x + - y) with (x - y)
huffman
parents: 20552
diff changeset
  1073
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
  1074
          ==>  \<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
  1075
apply (clarify, drule spec, drule (1) mp)
20563
44eda2314aab replace (x + - y) with (x - y)
huffman
parents: 20552
diff changeset
  1076
apply (simp only: norm_minus_commute)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1077
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
  1078
apply simp
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1079
done
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1080
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1081
lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1082
apply (simp add: Cauchy_def)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1083
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
  1084
apply (drule_tac x="M" in spec, simp)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1085
apply (drule lemmaCauchy)
22608
092a3353241e add new standard proofs for limits of sequences
huffman
parents: 21842
diff changeset
  1086
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
  1087
apply (simp add: Bseq_def)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1088
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
  1089
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
  1090
apply (simp add: order_less_imp_le)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1091
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1092
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
  1093
subsubsection {* Cauchy Sequences are Convergent *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1094
20830
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1095
axclass banach \<subseteq> real_normed_vector
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1096
  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1097
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1098
theorem LIMSEQ_imp_Cauchy:
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1099
  assumes X: "X ----> a" shows "Cauchy X"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1100
proof (rule CauchyI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1101
  fix e::real assume "0 < e"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1102
  hence "0 < e/2" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1103
  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1104
  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1105
  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1106
  proof (intro exI allI impI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1107
    fix m assume "N \<le> m"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1108
    hence m: "norm (X m - a) < e/2" using N by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1109
    fix n assume "N \<le> n"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1110
    hence n: "norm (X n - a) < e/2" using N by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1111
    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1112
    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1113
      by (rule norm_triangle_ineq4)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1114
    also from m n have "\<dots> < e" by(simp add:field_simps)
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1115
    finally show "norm (X m - X n) < e" .
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1116
  qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1117
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1118
20691
53cbea20e4d9 add lemma convergent_Cauchy
huffman
parents: 20685
diff changeset
  1119
lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1120
unfolding convergent_def
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1121
by (erule exE, erule LIMSEQ_imp_Cauchy)
20691
53cbea20e4d9 add lemma convergent_Cauchy
huffman
parents: 20685
diff changeset
  1122
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1123
text {*
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1124
Proof that Cauchy sequences converge based on the one from
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1125
http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1126
*}
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1127
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1128
text {*
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1129
  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
  1130
  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1131
*}
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1132
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1133
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
  1134
by (simp add: isUbI setleI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1135
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1136
lemma real_abs_diff_less_iff:
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1137
  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1138
by auto
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1139
27681
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1140
locale real_Cauchy =
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1141
  fixes X :: "nat \<Rightarrow> real"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1142
  assumes X: "Cauchy X"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1143
  fixes S :: "real set"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1144
  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1145
27681
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1146
lemma real_CauchyI:
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1147
  assumes "Cauchy X"
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1148
  shows "real_Cauchy X"
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28562
diff changeset
  1149
  proof qed (fact assms)
27681
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1150
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1151
lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1152
by (unfold S_def, auto)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1153
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1154
lemma (in real_Cauchy) bound_isUb:
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1155
  assumes N: "\<forall>n\<ge>N. X n < x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1156
  shows "isUb UNIV S x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1157
proof (rule isUb_UNIV_I)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1158
  fix y::real assume "y \<in> S"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1159
  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1160
    by (simp add: S_def)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1161
  then obtain M where "\<forall>n\<ge>M. y < X n" ..
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1162
  hence "y < X (max M N)" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1163
  also have "\<dots> < x" using N by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1164
  finally show "y \<le> x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1165
    by (rule order_less_imp_le)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1166
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1167
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1168
lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1169
proof (rule reals_complete)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1170
  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1171
    using CauchyD [OF X zero_less_one] by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1172
  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
  1173
  show "\<exists>x. x \<in> S"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1174
  proof
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1175
    from N have "\<forall>n\<ge>N. X N - 1 < X n"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1176
      by (simp add: real_abs_diff_less_iff)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1177
    thus "X N - 1 \<in> S" by (rule mem_S)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1178
  qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1179
  show "\<exists>u. isUb UNIV S u"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1180
  proof
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1181
    from N have "\<forall>n\<ge>N. X n < X N + 1"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1182
      by (simp add: real_abs_diff_less_iff)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1183
    thus "isUb UNIV S (X N + 1)"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1184
      by (rule bound_isUb)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1185
  qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1186
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1187
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1188
lemma (in real_Cauchy) isLub_imp_LIMSEQ:
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1189
  assumes x: "isLub UNIV S x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1190
  shows "X ----> x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1191
proof (rule LIMSEQ_I)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1192
  fix r::real assume "0 < r"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1193
  hence r: "0 < r/2" by simp
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1194
  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1195
    using CauchyD [OF X r] by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1196
  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
  1197
  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1198
    by (simp only: real_norm_def real_abs_diff_less_iff)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1199
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1200
  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
  1201
  hence "X N - r/2 \<in> S" by (rule mem_S)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1202
  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
  1203
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1204
  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
  1205
  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1206
  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
  1207
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1208
  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1209
  proof (intro exI allI impI)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1210
    fix n assume n: "N \<le> n"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1211
    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
  1212
    thus "norm (X n - x) < r" using 1 2
22629
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1213
      by (simp add: real_abs_diff_less_iff)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1214
  qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1215
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1216
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1217
lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1218
proof -
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1219
  obtain x where "isLub UNIV S x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1220
    using isLub_ex by fast
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1221
  hence "X ----> x"
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1222
    by (rule isLub_imp_LIMSEQ)
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1223
  thus ?thesis ..
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1224
qed
73771f454861 new standard proof of convergent = Cauchy
huffman
parents: 22628
diff changeset
  1225
20830
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1226
lemma real_Cauchy_convergent:
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1227
  fixes X :: "nat \<Rightarrow> real"
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1228
  shows "Cauchy X \<Longrightarrow> convergent X"
27681
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1229
unfolding convergent_def
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1230
by (rule real_Cauchy.LIMSEQ_ex)
8cedebf55539 dropped locale (open)
haftmann
parents: 27543
diff changeset
  1231
 (rule real_CauchyI)
20830
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1232
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1233
instance real :: banach
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1234
by intro_classes (rule real_Cauchy_convergent)
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1235
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1236
lemma Cauchy_convergent_iff:
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1237
  fixes X :: "nat \<Rightarrow> 'a::banach"
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1238
  shows "Cauchy X = convergent X"
65ba80cae6df add axclass banach for complete normed vector spaces
huffman
parents: 20829
diff changeset
  1239
by (fast intro: Cauchy_convergent convergent_Cauchy)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1240
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1241
20696
3b887ad7d196 reorganized subsection headings
huffman
parents: 20695
diff changeset
  1242
subsection {* Power Sequences *}
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1243
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1244
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
  1245
"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1246
  also fact that bounded and monotonic sequence converges.*}
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1247
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1248
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
  1249
apply (simp add: Bseq_def)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1250
apply (rule_tac x = 1 in exI)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1251
apply (simp add: power_abs)
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
  1252
apply (auto dest: power_mono)
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1253
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1254
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1255
lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1256
apply (clarify intro!: mono_SucI2)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1257
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
  1258
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1259
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1260
lemma convergent_realpow:
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1261
  "[| 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
  1262
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1263
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1264
lemma LIMSEQ_inverse_realpow_zero_lemma:
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1265
  fixes x :: real
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1266
  assumes x: "0 \<le> x"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1267
  shows "real n * x + 1 \<le> (x + 1) ^ n"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1268
apply (induct n)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1269
apply simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1270
apply simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1271
apply (rule order_trans)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1272
prefer 2
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1273
apply (erule mult_left_mono)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1274
apply (rule add_increasing [OF x], simp)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1275
apply (simp add: real_of_nat_Suc)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23127
diff changeset
  1276
apply (simp add: ring_distribs)
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1277
apply (simp add: mult_nonneg_nonneg x)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1278
done
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1279
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1280
lemma LIMSEQ_inverse_realpow_zero:
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1281
  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1282
proof (rule LIMSEQ_inverse_zero [rule_format])
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1283
  fix y :: real
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1284
  assume x: "1 < x"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1285
  hence "0 < x - 1" by simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1286
  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1287
    by (rule reals_Archimedean3)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1288
  hence "\<exists>N::nat. y < real N * (x - 1)" ..
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1289
  then obtain N::nat where "y < real N * (x - 1)" ..
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1290
  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1291
  also have "\<dots> \<le> (x - 1 + 1) ^ N"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1292
    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1293
  also have "\<dots> = x ^ N" by simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1294
  finally have "y < x ^ N" .
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1295
  hence "\<forall>n\<ge>N. y < x ^ n"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1296
    apply clarify
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1297
    apply (erule order_less_le_trans)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1298
    apply (erule power_increasing)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1299
    apply (rule order_less_imp_le [OF x])
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1300
    done
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1301
  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1302
qed
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1303
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1304
lemma LIMSEQ_realpow_zero:
22628
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1305
  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1306
proof (cases)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1307
  assume "x = 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1308
  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1309
  thus ?thesis by (rule LIMSEQ_imp_Suc)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1310
next
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1311
  assume "0 \<le> x" and "x \<noteq> 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1312
  hence x0: "0 < x" by simp
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1313
  assume x1: "x < 1"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1314
  from x0 x1 have "1 < inverse x"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1315
    by (rule real_inverse_gt_one)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1316
  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1317
    by (rule LIMSEQ_inverse_realpow_zero)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1318
  thus ?thesis by (simp add: power_inverse)
0e5ac9503d7e new standard proof of LIMSEQ_realpow_zero
huffman
parents: 22615
diff changeset
  1319
qed
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1320
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
  1321
lemma LIMSEQ_power_zero:
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
  1322
  fixes x :: "'a::{real_normed_algebra_1,recpower}"
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
  1323
  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
  1324
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
22974
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
  1325
apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
08b0fa905ea0 tuned proofs
huffman
parents: 22631
diff changeset
  1326
apply (simp add: power_abs norm_power_ineq)
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
  1327
done
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
  1328
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1329
lemma LIMSEQ_divide_realpow_zero:
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1330
  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1331
apply (cut_tac a = a and x1 = "inverse x" in
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1332
        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1333
apply (auto simp add: divide_inverse power_inverse)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1334
apply (simp add: inverse_eq_divide pos_divide_less_eq)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1335
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1336
15102
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 15085
diff changeset
  1337
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
  1338
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1339
lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
20685
fee8c75e3b5d added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents: 20682
diff changeset
  1340
by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1341
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20408
diff changeset
  1342
lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
15082
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1343
apply (rule LIMSEQ_rabs_zero [THEN iffD1])
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1344
apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1345
done
6c3276a2735b conversion of SEQ.ML to Isar script
paulson
parents: 13810
diff changeset
  1346
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
  1347
end