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