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