author | wenzelm |
Fri, 17 Nov 2006 02:20:03 +0100 | |
changeset 21404 | eb85850d3eb7 |
parent 21139 | c957e02e7a36 |
child 21810 | b2d23672b003 |
permissions | -rw-r--r-- |
10751 | 1 |
(* Title : SEQ.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 1998 University of Cambridge |
|
4 |
Description : Convergence of sequences and series |
|
15082 | 5 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
16819 | 6 |
Additional contributions by Jeremy Avigad |
15082 | 7 |
*) |
10751 | 8 |
|
17439 | 9 |
header {* Sequences and Series *} |
10 |
||
15131 | 11 |
theory SEQ |
15236
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
nipkow
parents:
15229
diff
changeset
|
12 |
imports NatStar |
15131 | 13 |
begin |
10751 | 14 |
|
19765 | 15 |
definition |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
16 |
LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
17 |
("((_)/ ----> (_))" [60, 60] 60) where |
15082 | 18 |
--{*Standard definition of convergence of sequence*} |
20563 | 19 |
"X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))" |
10751 | 20 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
21 |
definition |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
22 |
NSLIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
23 |
("((_)/ ----NS> (_))" [60, 60] 60) where |
15082 | 24 |
--{*Nonstandard definition of convergence of sequence*} |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
25 |
"X ----NS> L = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)" |
15082 | 26 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
27 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
28 |
lim :: "(nat => 'a::real_normed_vector) => 'a" where |
15082 | 29 |
--{*Standard definition of limit using choice operator*} |
20682 | 30 |
"lim X = (THE L. X ----> L)" |
10751 | 31 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
32 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
33 |
nslim :: "(nat => 'a::real_normed_vector) => 'a" where |
15082 | 34 |
--{*Nonstandard definition of limit using choice operator*} |
20682 | 35 |
"nslim X = (THE L. X ----NS> L)" |
10751 | 36 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
37 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
38 |
convergent :: "(nat => 'a::real_normed_vector) => bool" where |
15082 | 39 |
--{*Standard definition of convergence*} |
20682 | 40 |
"convergent X = (\<exists>L. X ----> L)" |
10751 | 41 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
42 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
43 |
NSconvergent :: "(nat => 'a::real_normed_vector) => bool" where |
15082 | 44 |
--{*Nonstandard definition of convergence*} |
20682 | 45 |
"NSconvergent X = (\<exists>L. X ----NS> L)" |
15082 | 46 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
47 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
48 |
Bseq :: "(nat => 'a::real_normed_vector) => bool" where |
15082 | 49 |
--{*Standard definition for bounded sequence*} |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
50 |
"Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)" |
10751 | 51 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
52 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
53 |
NSBseq :: "(nat => 'a::real_normed_vector) => bool" where |
15082 | 54 |
--{*Nonstandard definition for bounded sequence*} |
19765 | 55 |
"NSBseq X = (\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite)" |
15082 | 56 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
57 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
58 |
monoseq :: "(nat=>real)=>bool" where |
15082 | 59 |
--{*Definition for monotonicity*} |
19765 | 60 |
"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 | 61 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
62 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
63 |
subseq :: "(nat => nat) => bool" where |
15082 | 64 |
--{*Definition of subsequence*} |
19765 | 65 |
"subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))" |
10751 | 66 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
67 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
68 |
Cauchy :: "(nat => 'a::real_normed_vector) => bool" where |
15082 | 69 |
--{*Standard definition of the Cauchy condition*} |
20563 | 70 |
"Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)" |
10751 | 71 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
72 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21139
diff
changeset
|
73 |
NSCauchy :: "(nat => 'a::real_normed_vector) => bool" where |
15082 | 74 |
--{*Nonstandard definition*} |
19765 | 75 |
"NSCauchy X = (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)" |
15082 | 76 |
|
77 |
||
20696 | 78 |
subsection {* Limits of Sequences *} |
79 |
||
80 |
subsubsection {* Purely standard proofs *} |
|
15082 | 81 |
|
82 |
lemma LIMSEQ_iff: |
|
20563 | 83 |
"(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)" |
15082 | 84 |
by (simp add: LIMSEQ_def) |
85 |
||
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
86 |
lemma LIMSEQ_I: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
87 |
"(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
88 |
by (simp add: LIMSEQ_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
89 |
|
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
90 |
lemma LIMSEQ_D: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
91 |
"\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
92 |
by (simp add: LIMSEQ_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
93 |
|
20696 | 94 |
lemma LIMSEQ_const: "(%n. k) ----> k" |
95 |
by (simp add: LIMSEQ_def) |
|
96 |
||
97 |
lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a" |
|
98 |
apply (simp add: LIMSEQ_def, safe) |
|
99 |
apply (drule_tac x="r" in spec, safe) |
|
100 |
apply (rule_tac x="no" in exI, safe) |
|
101 |
apply (drule_tac x="n" in spec, safe) |
|
102 |
apply (erule order_le_less_trans [OF norm_triangle_ineq3]) |
|
103 |
done |
|
104 |
||
105 |
lemma LIMSEQ_ignore_initial_segment: "f ----> a |
|
106 |
==> (%n. f(n + k)) ----> a" |
|
107 |
apply (unfold LIMSEQ_def) |
|
108 |
apply (clarify) |
|
109 |
apply (drule_tac x = r in spec) |
|
110 |
apply (clarify) |
|
111 |
apply (rule_tac x = "no + k" in exI) |
|
112 |
by auto |
|
113 |
||
114 |
lemma LIMSEQ_offset: "(%x. f (x + k)) ----> a ==> |
|
115 |
f ----> a" |
|
116 |
apply (unfold LIMSEQ_def) |
|
117 |
apply clarsimp |
|
118 |
apply (drule_tac x = r in spec) |
|
119 |
apply clarsimp |
|
120 |
apply (rule_tac x = "no + k" in exI) |
|
121 |
apply clarsimp |
|
122 |
apply (drule_tac x = "n - k" in spec) |
|
123 |
apply (frule mp) |
|
124 |
apply arith |
|
125 |
apply simp |
|
126 |
done |
|
127 |
||
128 |
subsubsection {* Purely nonstandard proofs *} |
|
129 |
||
15082 | 130 |
lemma NSLIMSEQ_iff: |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
131 |
"(X ----NS> L) = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)" |
15082 | 132 |
by (simp add: NSLIMSEQ_def) |
133 |
||
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
134 |
lemma NSLIMSEQ_I: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
135 |
"(\<And>N. N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L) \<Longrightarrow> X ----NS> L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
136 |
by (simp add: NSLIMSEQ_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
137 |
|
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
138 |
lemma NSLIMSEQ_D: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
139 |
"\<lbrakk>X ----NS> L; N \<in> HNatInfinite\<rbrakk> \<Longrightarrow> starfun X N \<approx> star_of L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
140 |
by (simp add: NSLIMSEQ_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
141 |
|
20696 | 142 |
lemma NSLIMSEQ_const: "(%n. k) ----NS> k" |
143 |
by (simp add: NSLIMSEQ_def) |
|
144 |
||
145 |
lemma NSLIMSEQ_add: |
|
146 |
"[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + Y n) ----NS> a + b" |
|
147 |
by (auto intro: approx_add simp add: NSLIMSEQ_def starfun_add [symmetric]) |
|
148 |
||
149 |
lemma NSLIMSEQ_add_const: "f ----NS> a ==> (%n.(f n + b)) ----NS> a + b" |
|
150 |
by (simp only: NSLIMSEQ_add NSLIMSEQ_const) |
|
151 |
||
152 |
lemma NSLIMSEQ_mult: |
|
153 |
fixes a b :: "'a::real_normed_algebra" |
|
154 |
shows "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n * Y n) ----NS> a * b" |
|
155 |
by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def) |
|
156 |
||
157 |
lemma NSLIMSEQ_minus: "X ----NS> a ==> (%n. -(X n)) ----NS> -a" |
|
158 |
by (auto simp add: NSLIMSEQ_def) |
|
159 |
||
160 |
lemma NSLIMSEQ_minus_cancel: "(%n. -(X n)) ----NS> -a ==> X ----NS> a" |
|
161 |
by (drule NSLIMSEQ_minus, simp) |
|
162 |
||
163 |
(* FIXME: delete *) |
|
164 |
lemma NSLIMSEQ_add_minus: |
|
165 |
"[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + -Y n) ----NS> a + -b" |
|
166 |
by (simp add: NSLIMSEQ_add NSLIMSEQ_minus) |
|
167 |
||
168 |
lemma NSLIMSEQ_diff: |
|
169 |
"[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b" |
|
170 |
by (simp add: diff_minus NSLIMSEQ_add NSLIMSEQ_minus) |
|
171 |
||
172 |
lemma NSLIMSEQ_diff_const: "f ----NS> a ==> (%n.(f n - b)) ----NS> a - b" |
|
173 |
by (simp add: NSLIMSEQ_diff NSLIMSEQ_const) |
|
174 |
||
175 |
lemma NSLIMSEQ_inverse: |
|
176 |
fixes a :: "'a::real_normed_div_algebra" |
|
177 |
shows "[| X ----NS> a; a ~= 0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)" |
|
178 |
by (simp add: NSLIMSEQ_def star_of_approx_inverse) |
|
179 |
||
180 |
lemma NSLIMSEQ_mult_inverse: |
|
181 |
fixes a b :: "'a::real_normed_field" |
|
182 |
shows |
|
183 |
"[| X ----NS> a; Y ----NS> b; b ~= 0 |] ==> (%n. X n / Y n) ----NS> a/b" |
|
184 |
by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse) |
|
185 |
||
186 |
lemma starfun_hnorm: "\<And>x. hnorm (( *f* f) x) = ( *f* (\<lambda>x. norm (f x))) x" |
|
187 |
by transfer simp |
|
188 |
||
189 |
lemma NSLIMSEQ_norm: "X ----NS> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----NS> norm a" |
|
190 |
by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm) |
|
191 |
||
192 |
text{*Uniqueness of limit*} |
|
193 |
lemma NSLIMSEQ_unique: "[| X ----NS> a; X ----NS> b |] ==> a = b" |
|
194 |
apply (simp add: NSLIMSEQ_def) |
|
195 |
apply (drule HNatInfinite_whn [THEN [2] bspec])+ |
|
196 |
apply (auto dest: approx_trans3) |
|
197 |
done |
|
198 |
||
199 |
lemma NSLIMSEQ_pow [rule_format]: |
|
200 |
fixes a :: "'a::{real_normed_algebra,recpower}" |
|
201 |
shows "(X ----NS> a) --> ((%n. (X n) ^ m) ----NS> a ^ m)" |
|
202 |
apply (induct "m") |
|
203 |
apply (auto simp add: power_Suc intro: NSLIMSEQ_mult NSLIMSEQ_const) |
|
204 |
done |
|
205 |
||
206 |
subsubsection {* Equivalence of @{term LIMSEQ} and @{term NSLIMSEQ} *} |
|
15082 | 207 |
|
208 |
lemma LIMSEQ_NSLIMSEQ: |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
209 |
assumes X: "X ----> L" shows "X ----NS> L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
210 |
proof (rule NSLIMSEQ_I) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
211 |
fix N assume N: "N \<in> HNatInfinite" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
212 |
have "starfun X N - star_of L \<in> Infinitesimal" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
213 |
proof (rule InfinitesimalI2) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
214 |
fix r::real assume r: "0 < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
215 |
from LIMSEQ_D [OF X r] |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
216 |
obtain no where "\<forall>n\<ge>no. norm (X n - L) < r" .. |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
217 |
hence "\<forall>n\<ge>star_of no. hnorm (starfun X n - star_of L) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
218 |
by transfer |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
219 |
thus "hnorm (starfun X N - star_of L) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
220 |
using N by (simp add: star_of_le_HNatInfinite) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
221 |
qed |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
222 |
thus "starfun X N \<approx> star_of L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
223 |
by (unfold approx_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
224 |
qed |
15082 | 225 |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
226 |
lemma NSLIMSEQ_LIMSEQ: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
227 |
assumes X: "X ----NS> L" shows "X ----> L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
228 |
proof (rule LIMSEQ_I) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
229 |
fix r::real assume r: "0 < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
230 |
have "\<exists>no. \<forall>n\<ge>no. hnorm (starfun X n - star_of L) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
231 |
proof (intro exI allI impI) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
232 |
fix n assume "whn \<le> n" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
233 |
with HNatInfinite_whn have "n \<in> HNatInfinite" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
234 |
by (rule HNatInfinite_upward_closed) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
235 |
with X have "starfun X n \<approx> star_of L" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
236 |
by (rule NSLIMSEQ_D) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
237 |
hence "starfun X n - star_of L \<in> Infinitesimal" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
238 |
by (unfold approx_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
239 |
thus "hnorm (starfun X n - star_of L) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
240 |
using r by (rule InfinitesimalD2) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
241 |
qed |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
242 |
thus "\<exists>no. \<forall>n\<ge>no. norm (X n - L) < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
243 |
by transfer |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
244 |
qed |
15082 | 245 |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
246 |
theorem LIMSEQ_NSLIMSEQ_iff: "(f ----> L) = (f ----NS> L)" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
247 |
by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ) |
15082 | 248 |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
249 |
(* Used once by Integration/Rats.thy in AFP *) |
15082 | 250 |
lemma NSLIMSEQ_finite_set: |
251 |
"!!(f::nat=>nat). \<forall>n. n \<le> f n ==> finite {n. f n \<le> u}" |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
252 |
by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) |
15082 | 253 |
|
20696 | 254 |
subsubsection {* Derived theorems about @{term LIMSEQ} *} |
15082 | 255 |
|
256 |
lemma LIMSEQ_add: "[| X ----> a; Y ----> b |] ==> (%n. X n + Y n) ----> a + b" |
|
257 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_add) |
|
258 |
||
16819 | 259 |
lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b" |
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
260 |
by (simp add: LIMSEQ_add LIMSEQ_const) |
16819 | 261 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
262 |
lemma LIMSEQ_mult: |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
263 |
fixes a b :: "'a::real_normed_algebra" |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
264 |
shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b" |
15082 | 265 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_mult) |
266 |
||
267 |
lemma LIMSEQ_minus: "X ----> a ==> (%n. -(X n)) ----> -a" |
|
268 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_minus) |
|
269 |
||
270 |
lemma LIMSEQ_minus_cancel: "(%n. -(X n)) ----> -a ==> X ----> a" |
|
271 |
by (drule LIMSEQ_minus, simp) |
|
272 |
||
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
273 |
(* FIXME: delete *) |
15082 | 274 |
lemma LIMSEQ_add_minus: |
275 |
"[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b" |
|
276 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_add_minus) |
|
277 |
||
278 |
lemma LIMSEQ_diff: "[| X ----> a; Y ----> b |] ==> (%n. X n - Y n) ----> a - b" |
|
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
279 |
by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus) |
15082 | 280 |
|
16819 | 281 |
lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n - b)) ----> a - b" |
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
282 |
by (simp add: LIMSEQ_diff LIMSEQ_const) |
16819 | 283 |
|
15082 | 284 |
lemma LIMSEQ_inverse: |
20653
24cda2c5fd40
removed division_by_zero class requirements from several lemmas
huffman
parents:
20563
diff
changeset
|
285 |
fixes a :: "'a::real_normed_div_algebra" |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
286 |
shows "[| X ----> a; a ~= 0 |] ==> (%n. inverse(X n)) ----> inverse(a)" |
15082 | 287 |
by (simp add: NSLIMSEQ_inverse LIMSEQ_NSLIMSEQ_iff) |
288 |
||
289 |
lemma LIMSEQ_divide: |
|
20696 | 290 |
fixes a b :: "'a::real_normed_field" |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
291 |
shows "[| X ----> a; Y ----> b; b ~= 0 |] ==> (%n. X n / Y n) ----> a/b" |
15082 | 292 |
by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse) |
293 |
||
294 |
lemma LIMSEQ_unique: "[| X ----> a; X ----> b |] ==> a = b" |
|
295 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_unique) |
|
296 |
||
20696 | 297 |
lemma LIMSEQ_pow: |
298 |
fixes a :: "'a::{real_normed_algebra,recpower}" |
|
299 |
shows "X ----> a ==> (%n. (X n) ^ m) ----> a ^ m" |
|
300 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_pow) |
|
301 |
||
15312 | 302 |
lemma LIMSEQ_setsum: |
303 |
assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n" |
|
304 |
shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)" |
|
305 |
proof (cases "finite S") |
|
306 |
case True |
|
307 |
thus ?thesis using n |
|
308 |
proof (induct) |
|
309 |
case empty |
|
310 |
show ?case |
|
311 |
by (simp add: LIMSEQ_const) |
|
312 |
next |
|
313 |
case insert |
|
314 |
thus ?case |
|
315 |
by (simp add: LIMSEQ_add) |
|
316 |
qed |
|
317 |
next |
|
318 |
case False |
|
319 |
thus ?thesis |
|
320 |
by (simp add: setsum_def LIMSEQ_const) |
|
321 |
qed |
|
322 |
||
16819 | 323 |
lemma LIMSEQ_setprod: |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
324 |
fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}" |
16819 | 325 |
assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n" |
326 |
shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)" |
|
327 |
proof (cases "finite S") |
|
328 |
case True |
|
329 |
thus ?thesis using n |
|
330 |
proof (induct) |
|
331 |
case empty |
|
332 |
show ?case |
|
333 |
by (simp add: LIMSEQ_const) |
|
334 |
next |
|
335 |
case insert |
|
336 |
thus ?case |
|
337 |
by (simp add: LIMSEQ_mult) |
|
338 |
qed |
|
339 |
next |
|
340 |
case False |
|
341 |
thus ?thesis |
|
342 |
by (simp add: setprod_def LIMSEQ_const) |
|
343 |
qed |
|
344 |
||
345 |
lemma LIMSEQ_diff_approach_zero: |
|
346 |
"g ----> L ==> (%x. f x - g x) ----> 0 ==> |
|
347 |
f ----> L" |
|
348 |
apply (drule LIMSEQ_add) |
|
349 |
apply assumption |
|
350 |
apply simp |
|
351 |
done |
|
352 |
||
353 |
lemma LIMSEQ_diff_approach_zero2: |
|
354 |
"f ----> L ==> (%x. f x - g x) ----> 0 ==> |
|
355 |
g ----> L"; |
|
356 |
apply (drule LIMSEQ_diff) |
|
357 |
apply assumption |
|
358 |
apply simp |
|
359 |
done |
|
360 |
||
15082 | 361 |
|
20696 | 362 |
subsection {* Convergence *} |
15082 | 363 |
|
364 |
lemma limI: "X ----> L ==> lim X = L" |
|
365 |
apply (simp add: lim_def) |
|
366 |
apply (blast intro: LIMSEQ_unique) |
|
367 |
done |
|
368 |
||
369 |
lemma nslimI: "X ----NS> L ==> nslim X = L" |
|
370 |
apply (simp add: nslim_def) |
|
371 |
apply (blast intro: NSLIMSEQ_unique) |
|
372 |
done |
|
373 |
||
374 |
lemma lim_nslim_iff: "lim X = nslim X" |
|
375 |
by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff) |
|
376 |
||
377 |
lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)" |
|
378 |
by (simp add: convergent_def) |
|
379 |
||
380 |
lemma convergentI: "(X ----> L) ==> convergent X" |
|
381 |
by (auto simp add: convergent_def) |
|
382 |
||
383 |
lemma NSconvergentD: "NSconvergent X ==> \<exists>L. (X ----NS> L)" |
|
384 |
by (simp add: NSconvergent_def) |
|
385 |
||
386 |
lemma NSconvergentI: "(X ----NS> L) ==> NSconvergent X" |
|
387 |
by (auto simp add: NSconvergent_def) |
|
388 |
||
389 |
lemma convergent_NSconvergent_iff: "convergent X = NSconvergent X" |
|
390 |
by (simp add: convergent_def NSconvergent_def LIMSEQ_NSLIMSEQ_iff) |
|
391 |
||
392 |
lemma NSconvergent_NSLIMSEQ_iff: "NSconvergent X = (X ----NS> nslim X)" |
|
20682 | 393 |
by (auto intro: theI NSLIMSEQ_unique simp add: NSconvergent_def nslim_def) |
15082 | 394 |
|
395 |
lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)" |
|
20682 | 396 |
by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) |
15082 | 397 |
|
20696 | 398 |
lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))" |
399 |
apply (simp add: convergent_def) |
|
400 |
apply (auto dest: LIMSEQ_minus) |
|
401 |
apply (drule LIMSEQ_minus, auto) |
|
402 |
done |
|
403 |
||
404 |
||
405 |
subsection {* Bounded Monotonic Sequences *} |
|
406 |
||
15082 | 407 |
text{*Subsequence (alternative definition, (e.g. Hoskins)*} |
408 |
||
409 |
lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))" |
|
410 |
apply (simp add: subseq_def) |
|
411 |
apply (auto dest!: less_imp_Suc_add) |
|
412 |
apply (induct_tac k) |
|
413 |
apply (auto intro: less_trans) |
|
414 |
done |
|
415 |
||
416 |
lemma monoseq_Suc: |
|
417 |
"monoseq X = ((\<forall>n. X n \<le> X (Suc n)) |
|
418 |
| (\<forall>n. X (Suc n) \<le> X n))" |
|
419 |
apply (simp add: monoseq_def) |
|
420 |
apply (auto dest!: le_imp_less_or_eq) |
|
421 |
apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add) |
|
422 |
apply (induct_tac "ka") |
|
423 |
apply (auto intro: order_trans) |
|
18585 | 424 |
apply (erule contrapos_np) |
15082 | 425 |
apply (induct_tac "k") |
426 |
apply (auto intro: order_trans) |
|
427 |
done |
|
428 |
||
15360 | 429 |
lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X" |
15082 | 430 |
by (simp add: monoseq_def) |
431 |
||
15360 | 432 |
lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X" |
15082 | 433 |
by (simp add: monoseq_def) |
434 |
||
435 |
lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X" |
|
436 |
by (simp add: monoseq_Suc) |
|
437 |
||
438 |
lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X" |
|
439 |
by (simp add: monoseq_Suc) |
|
440 |
||
20696 | 441 |
text{*Bounded Sequence*} |
15082 | 442 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
443 |
lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)" |
15082 | 444 |
by (simp add: Bseq_def) |
445 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
446 |
lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X" |
15082 | 447 |
by (auto simp add: Bseq_def) |
448 |
||
449 |
lemma lemma_NBseq_def: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
450 |
"(\<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
|
451 |
(\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
15082 | 452 |
apply auto |
453 |
prefer 2 apply force |
|
454 |
apply (cut_tac x = K in reals_Archimedean2, clarify) |
|
455 |
apply (rule_tac x = n in exI, clarify) |
|
456 |
apply (drule_tac x = na in spec) |
|
457 |
apply (auto simp add: real_of_nat_Suc) |
|
458 |
done |
|
459 |
||
460 |
text{* alternative definition for Bseq *} |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
461 |
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
15082 | 462 |
apply (simp add: Bseq_def) |
463 |
apply (simp (no_asm) add: lemma_NBseq_def) |
|
464 |
done |
|
465 |
||
466 |
lemma lemma_NBseq_def2: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
467 |
"(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
15082 | 468 |
apply (subst lemma_NBseq_def, auto) |
469 |
apply (rule_tac x = "Suc N" in exI) |
|
470 |
apply (rule_tac [2] x = N in exI) |
|
471 |
apply (auto simp add: real_of_nat_Suc) |
|
472 |
prefer 2 apply (blast intro: order_less_imp_le) |
|
473 |
apply (drule_tac x = n in spec, simp) |
|
474 |
done |
|
475 |
||
476 |
(* yet another definition for Bseq *) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
477 |
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
15082 | 478 |
by (simp add: Bseq_def lemma_NBseq_def2) |
479 |
||
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
480 |
lemma NSBseqD: "[| NSBseq X; N: HNatInfinite |] ==> ( *f* X) N : HFinite" |
15082 | 481 |
by (simp add: NSBseq_def) |
482 |
||
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
483 |
lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite ==> NSBseq X" |
15082 | 484 |
by (simp add: NSBseq_def) |
485 |
||
486 |
text{*The standard definition implies the nonstandard definition*} |
|
487 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
488 |
lemma lemma_Bseq: "\<forall>n. norm (X n) \<le> K ==> \<forall>n. norm(X((f::nat=>nat) n)) \<le> K" |
15082 | 489 |
by auto |
490 |
||
491 |
lemma Bseq_NSBseq: "Bseq X ==> NSBseq X" |
|
21139 | 492 |
proof (unfold NSBseq_def, safe) |
493 |
assume X: "Bseq X" |
|
494 |
fix N assume N: "N \<in> HNatInfinite" |
|
495 |
from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K" by fast |
|
496 |
hence "\<forall>N. hnorm (starfun X N) \<le> star_of K" by transfer |
|
497 |
hence "hnorm (starfun X N) \<le> star_of K" by simp |
|
498 |
also have "star_of K < star_of (K + 1)" by simp |
|
499 |
finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x" by (rule bexI, simp) |
|
500 |
thus "starfun X N \<in> HFinite" by (simp add: HFinite_def) |
|
501 |
qed |
|
15082 | 502 |
|
503 |
text{*The nonstandard definition implies the standard definition*} |
|
504 |
||
505 |
(* similar to NSLIM proof in REALTOPOS *) |
|
506 |
||
507 |
text{* We need to get rid of the real variable and do so by proving the |
|
508 |
following, which relies on the Archimedean property of the reals. |
|
509 |
When we skolemize we then get the required function @{term "f::nat=>nat"}. |
|
510 |
Otherwise, we would be stuck with a skolem function @{term "f::real=>nat"} |
|
511 |
which woulid be useless.*} |
|
512 |
||
513 |
lemma lemmaNSBseq: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
514 |
"\<forall>K > 0. \<exists>n. K < norm (X n) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
515 |
==> \<forall>N. \<exists>n. real(Suc N) < norm (X n)" |
15082 | 516 |
apply safe |
517 |
apply (cut_tac n = N in real_of_nat_Suc_gt_zero, blast) |
|
518 |
done |
|
519 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
520 |
lemma lemmaNSBseq2: "\<forall>K > 0. \<exists>n::nat. K < norm (X n) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
521 |
==> \<exists>f. \<forall>N. real(Suc N) < norm (X (f N))" |
15082 | 522 |
apply (drule lemmaNSBseq) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
523 |
apply (drule no_choice, blast) |
15082 | 524 |
done |
525 |
||
526 |
lemma real_seq_to_hypreal_HInfinite: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
527 |
"\<forall>N. real(Suc N) < norm (X (f N)) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
528 |
==> star_n (X o f) : HInfinite" |
15082 | 529 |
apply (auto simp add: HInfinite_FreeUltrafilterNat_iff o_def) |
530 |
apply (cut_tac u = u in FreeUltrafilterNat_nat_gt_real) |
|
531 |
apply (drule FreeUltrafilterNat_all) |
|
532 |
apply (erule FreeUltrafilterNat_Int [THEN FreeUltrafilterNat_subset]) |
|
533 |
apply (auto simp add: real_of_nat_Suc) |
|
534 |
done |
|
535 |
||
536 |
text{* Now prove that we can get out an infinite hypernatural as well |
|
537 |
defined using the skolem function @{term "f::nat=>nat"} above*} |
|
538 |
||
539 |
lemma lemma_finite_NSBseq: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
540 |
"{n. f n \<le> Suc u & real(Suc n) < norm (X (f n))} \<le> |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
541 |
{n. f n \<le> u & real(Suc n) < norm (X (f n))} Un |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
542 |
{n. real(Suc n) < norm (X (Suc u))}" |
15082 | 543 |
by (auto dest!: le_imp_less_or_eq) |
544 |
||
545 |
lemma lemma_finite_NSBseq2: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
546 |
"finite {n. f n \<le> (u::nat) & real(Suc n) < norm (X(f n))}" |
15251 | 547 |
apply (induct "u") |
15082 | 548 |
apply (rule_tac [2] lemma_finite_NSBseq [THEN finite_subset]) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
549 |
apply (rule_tac B = "{n. real (Suc n) < norm (X 0) }" in finite_subset) |
15082 | 550 |
apply (auto intro: finite_real_of_nat_less_real |
551 |
simp add: real_of_nat_Suc less_diff_eq [symmetric]) |
|
552 |
done |
|
553 |
||
554 |
lemma HNatInfinite_skolem_f: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
555 |
"\<forall>N. real(Suc N) < norm (X (f N)) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
556 |
==> star_n f : HNatInfinite" |
15082 | 557 |
apply (auto simp add: HNatInfinite_FreeUltrafilterNat_iff) |
558 |
apply (rule ccontr, drule FreeUltrafilterNat_Compl_mem) |
|
559 |
apply (rule lemma_finite_NSBseq2 [THEN FreeUltrafilterNat_finite, THEN notE]) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
560 |
apply (subgoal_tac "{n. f n \<le> u & real (Suc n) < norm (X (f n))} = |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
561 |
{n. f n \<le> u} \<inter> {N. real (Suc N) < norm (X (f N))}") |
15082 | 562 |
apply (erule ssubst) |
563 |
apply (auto simp add: linorder_not_less Compl_def) |
|
564 |
done |
|
565 |
||
566 |
lemma NSBseq_Bseq: "NSBseq X ==> Bseq X" |
|
567 |
apply (simp add: Bseq_def NSBseq_def) |
|
568 |
apply (rule ccontr) |
|
569 |
apply (auto simp add: linorder_not_less [symmetric]) |
|
570 |
apply (drule lemmaNSBseq2, safe) |
|
571 |
apply (frule_tac X = X and f = f in real_seq_to_hypreal_HInfinite) |
|
572 |
apply (drule HNatInfinite_skolem_f [THEN [2] bspec]) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
573 |
apply (auto simp add: starfun o_def HFinite_HInfinite_iff) |
15082 | 574 |
done |
575 |
||
576 |
text{* Equivalence of nonstandard and standard definitions |
|
577 |
for a bounded sequence*} |
|
578 |
lemma Bseq_NSBseq_iff: "(Bseq X) = (NSBseq X)" |
|
579 |
by (blast intro!: NSBseq_Bseq Bseq_NSBseq) |
|
580 |
||
581 |
text{*A convergent sequence is bounded: |
|
582 |
Boundedness as a necessary condition for convergence. |
|
583 |
The nonstandard version has no existential, as usual *} |
|
584 |
||
585 |
lemma NSconvergent_NSBseq: "NSconvergent X ==> NSBseq X" |
|
586 |
apply (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
587 |
apply (blast intro: HFinite_star_of approx_sym approx_HFinite) |
15082 | 588 |
done |
589 |
||
590 |
text{*Standard Version: easily now proved using equivalence of NS and |
|
591 |
standard definitions *} |
|
592 |
lemma convergent_Bseq: "convergent X ==> Bseq X" |
|
593 |
by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff) |
|
594 |
||
20696 | 595 |
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*} |
15082 | 596 |
|
597 |
lemma Bseq_isUb: |
|
598 |
"!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U" |
|
599 |
by (auto intro: isUbI setleI simp add: Bseq_def abs_le_interval_iff) |
|
600 |
||
601 |
||
602 |
text{* Use completeness of reals (supremum property) |
|
603 |
to show that any bounded sequence has a least upper bound*} |
|
604 |
||
605 |
lemma Bseq_isLub: |
|
606 |
"!!(X::nat=>real). Bseq X ==> |
|
607 |
\<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U" |
|
608 |
by (blast intro: reals_complete Bseq_isUb) |
|
609 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
610 |
lemma NSBseq_isUb: "NSBseq X ==> \<exists>U::real. isUb UNIV {x. \<exists>n. X n = x} U" |
15082 | 611 |
by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb) |
612 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
613 |
lemma NSBseq_isLub: "NSBseq X ==> \<exists>U::real. isLub UNIV {x. \<exists>n. X n = x} U" |
15082 | 614 |
by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub) |
615 |
||
616 |
||
20696 | 617 |
subsubsection{*A Bounded and Monotonic Sequence Converges*} |
15082 | 618 |
|
619 |
lemma lemma_converg1: |
|
15360 | 620 |
"!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n; |
15082 | 621 |
isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma) |
15360 | 622 |
|] ==> \<forall>n \<ge> ma. X n = X ma" |
15082 | 623 |
apply safe |
624 |
apply (drule_tac y = "X n" in isLubD2) |
|
625 |
apply (blast dest: order_antisym)+ |
|
626 |
done |
|
627 |
||
628 |
text{* The best of both worlds: Easier to prove this result as a standard |
|
629 |
theorem and then use equivalence to "transfer" it into the |
|
630 |
equivalent nonstandard form if needed!*} |
|
631 |
||
632 |
lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)" |
|
633 |
apply (simp add: LIMSEQ_def) |
|
634 |
apply (rule_tac x = "X m" in exI, safe) |
|
635 |
apply (rule_tac x = m in exI, safe) |
|
636 |
apply (drule spec, erule impE, auto) |
|
637 |
done |
|
638 |
||
639 |
text{*Now, the same theorem in terms of NS limit *} |
|
15360 | 640 |
lemma Bmonoseq_NSLIMSEQ: "\<forall>n \<ge> m. X n = X m ==> \<exists>L. (X ----NS> L)" |
15082 | 641 |
by (auto dest!: Bmonoseq_LIMSEQ simp add: LIMSEQ_NSLIMSEQ_iff) |
642 |
||
643 |
lemma lemma_converg2: |
|
644 |
"!!(X::nat=>real). |
|
645 |
[| \<forall>m. X m ~= U; isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U" |
|
646 |
apply safe |
|
647 |
apply (drule_tac y = "X m" in isLubD2) |
|
648 |
apply (auto dest!: order_le_imp_less_or_eq) |
|
649 |
done |
|
650 |
||
651 |
lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U" |
|
652 |
by (rule setleI [THEN isUbI], auto) |
|
653 |
||
654 |
text{* FIXME: @{term "U - T < U"} is redundant *} |
|
655 |
lemma lemma_converg4: "!!(X::nat=> real). |
|
656 |
[| \<forall>m. X m ~= U; |
|
657 |
isLub UNIV {x. \<exists>n. X n = x} U; |
|
658 |
0 < T; |
|
659 |
U + - T < U |
|
660 |
|] ==> \<exists>m. U + -T < X m & X m < U" |
|
661 |
apply (drule lemma_converg2, assumption) |
|
662 |
apply (rule ccontr, simp) |
|
663 |
apply (simp add: linorder_not_less) |
|
664 |
apply (drule lemma_converg3) |
|
665 |
apply (drule isLub_le_isUb, assumption) |
|
666 |
apply (auto dest: order_less_le_trans) |
|
667 |
done |
|
668 |
||
669 |
text{*A standard proof of the theorem for monotone increasing sequence*} |
|
670 |
||
671 |
lemma Bseq_mono_convergent: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
672 |
"[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)" |
15082 | 673 |
apply (simp add: convergent_def) |
674 |
apply (frule Bseq_isLub, safe) |
|
675 |
apply (case_tac "\<exists>m. X m = U", auto) |
|
676 |
apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ) |
|
677 |
(* second case *) |
|
678 |
apply (rule_tac x = U in exI) |
|
679 |
apply (subst LIMSEQ_iff, safe) |
|
680 |
apply (frule lemma_converg2, assumption) |
|
681 |
apply (drule lemma_converg4, auto) |
|
682 |
apply (rule_tac x = m in exI, safe) |
|
683 |
apply (subgoal_tac "X m \<le> X n") |
|
684 |
prefer 2 apply blast |
|
685 |
apply (drule_tac x=n and P="%m. X m < U" in spec, arith) |
|
686 |
done |
|
687 |
||
688 |
text{*Nonstandard version of the theorem*} |
|
689 |
||
690 |
lemma NSBseq_mono_NSconvergent: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
691 |
"[| NSBseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> NSconvergent (X::nat=>real)" |
15082 | 692 |
by (auto intro: Bseq_mono_convergent |
693 |
simp add: convergent_NSconvergent_iff [symmetric] |
|
694 |
Bseq_NSBseq_iff [symmetric]) |
|
695 |
||
696 |
lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X" |
|
697 |
by (simp add: Bseq_def) |
|
698 |
||
699 |
text{*Main monotonicity theorem*} |
|
700 |
lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X" |
|
701 |
apply (simp add: monoseq_def, safe) |
|
702 |
apply (rule_tac [2] convergent_minus_iff [THEN ssubst]) |
|
703 |
apply (drule_tac [2] Bseq_minus_iff [THEN ssubst]) |
|
704 |
apply (auto intro!: Bseq_mono_convergent) |
|
705 |
done |
|
706 |
||
20696 | 707 |
subsubsection{*A Few More Equivalence Theorems for Boundedness*} |
15082 | 708 |
|
709 |
text{*alternative formulation for boundedness*} |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
710 |
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)" |
15082 | 711 |
apply (unfold Bseq_def, safe) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
712 |
apply (rule_tac [2] x = "k + norm x" in exI) |
15360 | 713 |
apply (rule_tac x = K in exI, simp) |
15221 | 714 |
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
|
715 |
apply (erule order_less_le_trans, simp) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
716 |
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
|
717 |
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
|
718 |
apply simp |
15082 | 719 |
done |
720 |
||
721 |
text{*alternative formulation for boundedness*} |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
722 |
lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)" |
15082 | 723 |
apply safe |
724 |
apply (simp add: Bseq_def, safe) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
725 |
apply (rule_tac x = "K + norm (X N)" in exI) |
15082 | 726 |
apply auto |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
727 |
apply (erule order_less_le_trans, simp) |
15082 | 728 |
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
|
729 |
apply (drule_tac x = n in spec) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
730 |
apply (rule order_trans [OF norm_triangle_ineq], simp) |
15082 | 731 |
apply (auto simp add: Bseq_iff2) |
732 |
done |
|
733 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
734 |
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f" |
15082 | 735 |
apply (simp add: Bseq_def) |
15221 | 736 |
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
|
737 |
apply (drule_tac x = n in spec, arith) |
15082 | 738 |
done |
739 |
||
740 |
||
20696 | 741 |
subsection {* Cauchy Sequences *} |
15082 | 742 |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
743 |
lemma CauchyI: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
744 |
"(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
745 |
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
|
746 |
|
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
747 |
lemma CauchyD: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
748 |
"\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
749 |
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
|
750 |
|
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
751 |
lemma NSCauchyI: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
752 |
"(\<And>M N. \<lbrakk>M \<in> HNatInfinite; N \<in> HNatInfinite\<rbrakk> \<Longrightarrow> starfun X M \<approx> starfun X N) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
753 |
\<Longrightarrow> NSCauchy X" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
754 |
by (simp add: NSCauchy_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
755 |
|
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
756 |
lemma NSCauchyD: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
757 |
"\<lbrakk>NSCauchy X; M \<in> HNatInfinite; N \<in> HNatInfinite\<rbrakk> |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
758 |
\<Longrightarrow> starfun X M \<approx> starfun X N" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
759 |
by (simp add: NSCauchy_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
760 |
|
20696 | 761 |
subsubsection{*Equivalence Between NS and Standard*} |
762 |
||
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
763 |
lemma Cauchy_NSCauchy: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
764 |
assumes X: "Cauchy X" shows "NSCauchy X" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
765 |
proof (rule NSCauchyI) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
766 |
fix M assume M: "M \<in> HNatInfinite" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
767 |
fix N assume N: "N \<in> HNatInfinite" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
768 |
have "starfun X M - starfun X N \<in> Infinitesimal" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
769 |
proof (rule InfinitesimalI2) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
770 |
fix r :: real assume r: "0 < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
771 |
from CauchyD [OF X r] |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
772 |
obtain k where "\<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" .. |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
773 |
hence "\<forall>m\<ge>star_of k. \<forall>n\<ge>star_of k. |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
774 |
hnorm (starfun X m - starfun X n) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
775 |
by transfer |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
776 |
thus "hnorm (starfun X M - starfun X N) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
777 |
using M N by (simp add: star_of_le_HNatInfinite) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
778 |
qed |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
779 |
thus "starfun X M \<approx> starfun X N" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
780 |
by (unfold approx_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
781 |
qed |
15082 | 782 |
|
20751
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
783 |
lemma NSCauchy_Cauchy: |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
784 |
assumes X: "NSCauchy X" shows "Cauchy X" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
785 |
proof (rule CauchyI) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
786 |
fix r::real assume r: "0 < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
787 |
have "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. hnorm (starfun X m - starfun X n) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
788 |
proof (intro exI allI impI) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
789 |
fix M assume "whn \<le> M" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
790 |
with HNatInfinite_whn have M: "M \<in> HNatInfinite" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
791 |
by (rule HNatInfinite_upward_closed) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
792 |
fix N assume "whn \<le> N" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
793 |
with HNatInfinite_whn have N: "N \<in> HNatInfinite" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
794 |
by (rule HNatInfinite_upward_closed) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
795 |
from X M N have "starfun X M \<approx> starfun X N" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
796 |
by (rule NSCauchyD) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
797 |
hence "starfun X M - starfun X N \<in> Infinitesimal" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
798 |
by (unfold approx_def) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
799 |
thus "hnorm (starfun X M - starfun X N) < star_of r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
800 |
using r by (rule InfinitesimalD2) |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
801 |
qed |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
802 |
thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
803 |
by transfer |
93271c59d211
add intro/dest rules for (NS)LIMSEQ and (NS)Cauchy; rewrite equivalence proofs using transfer
huffman
parents:
20740
diff
changeset
|
804 |
qed |
15082 | 805 |
|
806 |
theorem NSCauchy_Cauchy_iff: "NSCauchy X = Cauchy X" |
|
807 |
by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy) |
|
808 |
||
20696 | 809 |
subsubsection {* Cauchy Sequences are Bounded *} |
810 |
||
15082 | 811 |
text{*A Cauchy sequence is bounded -- this is the standard |
812 |
proof mechanization rather than the nonstandard proof*} |
|
813 |
||
20563 | 814 |
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
|
815 |
==> \<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
|
816 |
apply (clarify, drule spec, drule (1) mp) |
20563 | 817 |
apply (simp only: norm_minus_commute) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
818 |
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
|
819 |
apply simp |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
820 |
done |
15082 | 821 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
822 |
lemma Bseq_Suc_imp_Bseq: "Bseq (\<lambda>n. X (Suc n)) \<Longrightarrow> Bseq X" |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
823 |
apply (unfold Bseq_def, clarify) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
824 |
apply (rule_tac x="max K (norm (X 0))" in exI) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
825 |
apply (simp add: order_less_le_trans [OF _ le_maxI1]) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
826 |
apply (clarify, case_tac "n", simp) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
827 |
apply (simp add: order_trans [OF _ le_maxI1]) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
828 |
done |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
829 |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
830 |
lemma Bseq_shift_imp_Bseq: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X" |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
831 |
apply (induct k, simp_all) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
832 |
apply (subgoal_tac "Bseq (\<lambda>n. X (n + k))", simp) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
833 |
apply (rule Bseq_Suc_imp_Bseq, simp) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
834 |
done |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
835 |
|
15082 | 836 |
lemma Cauchy_Bseq: "Cauchy X ==> Bseq X" |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
837 |
apply (simp add: Cauchy_def) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
838 |
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
|
839 |
apply (drule_tac x="M" in spec, simp) |
15082 | 840 |
apply (drule lemmaCauchy) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
841 |
apply (rule_tac k="M" in Bseq_shift_imp_Bseq) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
842 |
apply (simp add: Bseq_def) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
843 |
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
|
844 |
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
|
845 |
apply (simp add: order_less_imp_le) |
15082 | 846 |
done |
847 |
||
848 |
text{*A Cauchy sequence is bounded -- nonstandard version*} |
|
849 |
||
850 |
lemma NSCauchy_NSBseq: "NSCauchy X ==> NSBseq X" |
|
851 |
by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff) |
|
852 |
||
20696 | 853 |
subsubsection {* Cauchy Sequences are Convergent *} |
15082 | 854 |
|
20830
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
855 |
axclass banach \<subseteq> real_normed_vector |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
856 |
Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
857 |
|
15082 | 858 |
text{*Equivalence of Cauchy criterion and convergence: |
859 |
We will prove this using our NS formulation which provides a |
|
860 |
much easier proof than using the standard definition. We do not |
|
861 |
need to use properties of subsequences such as boundedness, |
|
862 |
monotonicity etc... Compare with Harrison's corresponding proof |
|
863 |
in HOL which is much longer and more complicated. Of course, we do |
|
864 |
not have problems which he encountered with guessing the right |
|
865 |
instantiations for his 'espsilon-delta' proof(s) in this case |
|
866 |
since the NS formulations do not involve existential quantifiers.*} |
|
867 |
||
20691 | 868 |
lemma NSconvergent_NSCauchy: "NSconvergent X \<Longrightarrow> NSCauchy X" |
869 |
apply (simp add: NSconvergent_def NSLIMSEQ_def NSCauchy_def, safe) |
|
870 |
apply (auto intro: approx_trans2) |
|
871 |
done |
|
872 |
||
873 |
lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X" |
|
874 |
apply (rule NSconvergent_NSCauchy [THEN NSCauchy_Cauchy]) |
|
875 |
apply (simp add: convergent_NSconvergent_iff) |
|
876 |
done |
|
877 |
||
20830
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
878 |
lemma real_NSCauchy_NSconvergent: |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
879 |
fixes X :: "nat \<Rightarrow> real" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
880 |
shows "NSCauchy X \<Longrightarrow> NSconvergent X" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
881 |
apply (simp add: NSconvergent_def NSLIMSEQ_def) |
15082 | 882 |
apply (frule NSCauchy_NSBseq) |
20830
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
883 |
apply (simp add: NSBseq_def NSCauchy_def) |
15082 | 884 |
apply (drule HNatInfinite_whn [THEN [2] bspec]) |
885 |
apply (drule HNatInfinite_whn [THEN [2] bspec]) |
|
886 |
apply (auto dest!: st_part_Ex simp add: SReal_iff) |
|
887 |
apply (blast intro: approx_trans3) |
|
888 |
done |
|
889 |
||
890 |
text{*Standard proof for free*} |
|
20830
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
891 |
lemma real_Cauchy_convergent: |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
892 |
fixes X :: "nat \<Rightarrow> real" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
893 |
shows "Cauchy X \<Longrightarrow> convergent X" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
894 |
apply (drule Cauchy_NSCauchy [THEN real_NSCauchy_NSconvergent]) |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
895 |
apply (erule convergent_NSconvergent_iff [THEN iffD2]) |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
896 |
done |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
897 |
|
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
898 |
instance real :: banach |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
899 |
by intro_classes (rule real_Cauchy_convergent) |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
900 |
|
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
901 |
lemma NSCauchy_NSconvergent: |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
902 |
fixes X :: "nat \<Rightarrow> 'a::banach" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
903 |
shows "NSCauchy X \<Longrightarrow> NSconvergent X" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
904 |
apply (drule NSCauchy_Cauchy [THEN Cauchy_convergent]) |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
905 |
apply (erule convergent_NSconvergent_iff [THEN iffD1]) |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
906 |
done |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
907 |
|
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
908 |
lemma NSCauchy_NSconvergent_iff: |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
909 |
fixes X :: "nat \<Rightarrow> 'a::banach" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
910 |
shows "NSCauchy X = NSconvergent X" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
911 |
by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy) |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
912 |
|
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
913 |
lemma Cauchy_convergent_iff: |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
914 |
fixes X :: "nat \<Rightarrow> 'a::banach" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
915 |
shows "Cauchy X = convergent X" |
65ba80cae6df
add axclass banach for complete normed vector spaces
huffman
parents:
20829
diff
changeset
|
916 |
by (fast intro: Cauchy_convergent convergent_Cauchy) |
15082 | 917 |
|
918 |
||
20696 | 919 |
subsection {* More Properties of Sequences *} |
920 |
||
15082 | 921 |
text{*We can now try and derive a few properties of sequences, |
922 |
starting with the limit comparison property for sequences.*} |
|
923 |
||
924 |
lemma NSLIMSEQ_le: |
|
925 |
"[| f ----NS> l; g ----NS> m; |
|
15360 | 926 |
\<exists>N. \<forall>n \<ge> N. f(n) \<le> g(n) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
927 |
|] ==> l \<le> (m::real)" |
15082 | 928 |
apply (simp add: NSLIMSEQ_def, safe) |
929 |
apply (drule starfun_le_mono) |
|
930 |
apply (drule HNatInfinite_whn [THEN [2] bspec])+ |
|
931 |
apply (drule_tac x = whn in spec) |
|
932 |
apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ |
|
933 |
apply clarify |
|
934 |
apply (auto intro: hypreal_of_real_le_add_Infininitesimal_cancel2) |
|
935 |
done |
|
936 |
||
937 |
(* standard version *) |
|
938 |
lemma LIMSEQ_le: |
|
15360 | 939 |
"[| f ----> l; g ----> m; \<exists>N. \<forall>n \<ge> N. f(n) \<le> g(n) |] |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
940 |
==> l \<le> (m::real)" |
15082 | 941 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_le) |
942 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
943 |
lemma LIMSEQ_le_const: "[| X ----> (r::real); \<forall>n. a \<le> X n |] ==> a \<le> r" |
15082 | 944 |
apply (rule LIMSEQ_le) |
945 |
apply (rule LIMSEQ_const, auto) |
|
946 |
done |
|
947 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
948 |
lemma NSLIMSEQ_le_const: "[| X ----NS> (r::real); \<forall>n. a \<le> X n |] ==> a \<le> r" |
15082 | 949 |
by (simp add: LIMSEQ_NSLIMSEQ_iff LIMSEQ_le_const) |
950 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
951 |
lemma LIMSEQ_le_const2: "[| X ----> (r::real); \<forall>n. X n \<le> a |] ==> r \<le> a" |
15082 | 952 |
apply (rule LIMSEQ_le) |
953 |
apply (rule_tac [2] LIMSEQ_const, auto) |
|
954 |
done |
|
955 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
956 |
lemma NSLIMSEQ_le_const2: "[| X ----NS> (r::real); \<forall>n. X n \<le> a |] ==> r \<le> a" |
15082 | 957 |
by (simp add: LIMSEQ_NSLIMSEQ_iff LIMSEQ_le_const2) |
958 |
||
959 |
text{*Shift a convergent series by 1: |
|
960 |
By the equivalence between Cauchiness and convergence and because |
|
961 |
the successor of an infinite hypernatural is also infinite.*} |
|
962 |
||
963 |
lemma NSLIMSEQ_Suc: "f ----NS> l ==> (%n. f(Suc n)) ----NS> l" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
964 |
apply (unfold NSLIMSEQ_def, safe) |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
965 |
apply (drule_tac x="N + 1" in bspec) |
20740
5a103b43da5a
reorganized HNatInfinite proofs; simplified and renamed some lemmas
huffman
parents:
20728
diff
changeset
|
966 |
apply (erule HNatInfinite_add) |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
967 |
apply (simp add: starfun_shift_one) |
15082 | 968 |
done |
969 |
||
970 |
text{* standard version *} |
|
971 |
lemma LIMSEQ_Suc: "f ----> l ==> (%n. f(Suc n)) ----> l" |
|
972 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_Suc) |
|
973 |
||
974 |
lemma NSLIMSEQ_imp_Suc: "(%n. f(Suc n)) ----NS> l ==> f ----NS> l" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
975 |
apply (unfold NSLIMSEQ_def, safe) |
15082 | 976 |
apply (drule_tac x="N - 1" in bspec) |
20740
5a103b43da5a
reorganized HNatInfinite proofs; simplified and renamed some lemmas
huffman
parents:
20728
diff
changeset
|
977 |
apply (erule Nats_1 [THEN [2] HNatInfinite_diff]) |
5a103b43da5a
reorganized HNatInfinite proofs; simplified and renamed some lemmas
huffman
parents:
20728
diff
changeset
|
978 |
apply (simp add: starfun_shift_one one_le_HNatInfinite) |
15082 | 979 |
done |
980 |
||
981 |
lemma LIMSEQ_imp_Suc: "(%n. f(Suc n)) ----> l ==> f ----> l" |
|
982 |
apply (simp add: LIMSEQ_NSLIMSEQ_iff) |
|
983 |
apply (erule NSLIMSEQ_imp_Suc) |
|
984 |
done |
|
985 |
||
986 |
lemma LIMSEQ_Suc_iff: "((%n. f(Suc n)) ----> l) = (f ----> l)" |
|
987 |
by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) |
|
988 |
||
989 |
lemma NSLIMSEQ_Suc_iff: "((%n. f(Suc n)) ----NS> l) = (f ----NS> l)" |
|
990 |
by (blast intro: NSLIMSEQ_imp_Suc NSLIMSEQ_Suc) |
|
991 |
||
992 |
text{*A sequence tends to zero iff its abs does*} |
|
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
993 |
lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)" |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
994 |
by (simp add: LIMSEQ_def) |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
995 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
996 |
lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))" |
15082 | 997 |
by (simp add: LIMSEQ_def) |
998 |
||
999 |
text{*We prove the NS version from the standard one, since the NS proof |
|
1000 |
seems more complicated than the standard one above!*} |
|
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1001 |
lemma NSLIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----NS> 0) = (X ----NS> 0)" |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1002 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_norm_zero) |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1003 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1004 |
lemma NSLIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----NS> 0) = (f ----NS> (0::real))" |
15082 | 1005 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_rabs_zero) |
1006 |
||
1007 |
text{*Generalization to other limits*} |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1008 |
lemma NSLIMSEQ_imp_rabs: "f ----NS> (l::real) ==> (%n. \<bar>f n\<bar>) ----NS> \<bar>l\<bar>" |
15082 | 1009 |
apply (simp add: NSLIMSEQ_def) |
1010 |
apply (auto intro: approx_hrabs |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
1011 |
simp add: starfun_abs hypreal_of_real_hrabs [symmetric]) |
15082 | 1012 |
done |
1013 |
||
1014 |
text{* standard version *} |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1015 |
lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>" |
15082 | 1016 |
by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_imp_rabs) |
1017 |
||
1018 |
text{*An unbounded sequence's inverse tends to 0*} |
|
1019 |
||
1020 |
text{* standard proof seems easier *} |
|
1021 |
lemma LIMSEQ_inverse_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1022 |
"\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f(n) ==> (%n. inverse(f n)) ----> 0" |
15082 | 1023 |
apply (simp add: LIMSEQ_def, safe) |
1024 |
apply (drule_tac x = "inverse r" in spec, safe) |
|
1025 |
apply (rule_tac x = N in exI, safe) |
|
1026 |
apply (drule spec, auto) |
|
1027 |
apply (frule positive_imp_inverse_positive) |
|
1028 |
apply (frule order_less_trans, assumption) |
|
1029 |
apply (frule_tac a = "f n" in positive_imp_inverse_positive) |
|
1030 |
apply (simp add: abs_if) |
|
1031 |
apply (rule_tac t = r in inverse_inverse_eq [THEN subst]) |
|
1032 |
apply (auto intro: inverse_less_iff_less [THEN iffD2] |
|
1033 |
simp del: inverse_inverse_eq) |
|
1034 |
done |
|
1035 |
||
1036 |
lemma NSLIMSEQ_inverse_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1037 |
"\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f(n) |
15082 | 1038 |
==> (%n. inverse(f n)) ----NS> 0" |
1039 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero) |
|
1040 |
||
1041 |
text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*} |
|
1042 |
||
1043 |
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0" |
|
1044 |
apply (rule LIMSEQ_inverse_zero, safe) |
|
1045 |
apply (cut_tac x = y in reals_Archimedean2) |
|
1046 |
apply (safe, rule_tac x = n in exI) |
|
1047 |
apply (auto simp add: real_of_nat_Suc) |
|
1048 |
done |
|
1049 |
||
1050 |
lemma NSLIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----NS> 0" |
|
1051 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat) |
|
1052 |
||
1053 |
text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to |
|
1054 |
infinity is now easily proved*} |
|
1055 |
||
1056 |
lemma LIMSEQ_inverse_real_of_nat_add: |
|
1057 |
"(%n. r + inverse(real(Suc n))) ----> r" |
|
1058 |
by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto) |
|
1059 |
||
1060 |
lemma NSLIMSEQ_inverse_real_of_nat_add: |
|
1061 |
"(%n. r + inverse(real(Suc n))) ----NS> r" |
|
1062 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add) |
|
1063 |
||
1064 |
lemma LIMSEQ_inverse_real_of_nat_add_minus: |
|
1065 |
"(%n. r + -inverse(real(Suc n))) ----> r" |
|
1066 |
by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto) |
|
1067 |
||
1068 |
lemma NSLIMSEQ_inverse_real_of_nat_add_minus: |
|
1069 |
"(%n. r + -inverse(real(Suc n))) ----NS> r" |
|
1070 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus) |
|
1071 |
||
1072 |
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: |
|
1073 |
"(%n. r*( 1 + -inverse(real(Suc n)))) ----> r" |
|
1074 |
by (cut_tac b=1 in |
|
1075 |
LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto) |
|
1076 |
||
1077 |
lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult: |
|
1078 |
"(%n. r*( 1 + -inverse(real(Suc n)))) ----NS> r" |
|
1079 |
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus_mult) |
|
1080 |
||
1081 |
||
20696 | 1082 |
subsection {* Power Sequences *} |
15082 | 1083 |
|
1084 |
text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term |
|
1085 |
"x<1"}. Proof will use (NS) Cauchy equivalence for convergence and |
|
1086 |
also fact that bounded and monotonic sequence converges.*} |
|
1087 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1088 |
lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)" |
15082 | 1089 |
apply (simp add: Bseq_def) |
1090 |
apply (rule_tac x = 1 in exI) |
|
1091 |
apply (simp add: power_abs) |
|
1092 |
apply (auto dest: power_mono intro: order_less_imp_le simp add: abs_if) |
|
1093 |
done |
|
1094 |
||
1095 |
lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)" |
|
1096 |
apply (clarify intro!: mono_SucI2) |
|
1097 |
apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto) |
|
1098 |
done |
|
1099 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1100 |
lemma convergent_realpow: |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1101 |
"[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)" |
15082 | 1102 |
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) |
1103 |
||
1104 |
text{* We now use NS criterion to bring proof of theorem through *} |
|
1105 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1106 |
lemma NSLIMSEQ_realpow_zero: |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1107 |
"[| 0 \<le> (x::real); x < 1 |] ==> (%n. x ^ n) ----NS> 0" |
15082 | 1108 |
apply (simp add: NSLIMSEQ_def) |
1109 |
apply (auto dest!: convergent_realpow simp add: convergent_NSconvergent_iff) |
|
1110 |
apply (frule NSconvergentD) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
1111 |
apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfun_pow) |
15082 | 1112 |
apply (frule HNatInfinite_add_one) |
1113 |
apply (drule bspec, assumption) |
|
1114 |
apply (drule bspec, assumption) |
|
1115 |
apply (drule_tac x = "N + (1::hypnat) " in bspec, assumption) |
|
1116 |
apply (simp add: hyperpow_add) |
|
1117 |
apply (drule approx_mult_subst_SReal, assumption) |
|
1118 |
apply (drule approx_trans3, assumption) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
1119 |
apply (auto simp del: star_of_mult simp add: star_of_mult [symmetric]) |
15082 | 1120 |
done |
1121 |
||
1122 |
text{* standard version *} |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1123 |
lemma LIMSEQ_realpow_zero: |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1124 |
"[| 0 \<le> (x::real); x < 1 |] ==> (%n. x ^ n) ----> 0" |
15082 | 1125 |
by (simp add: NSLIMSEQ_realpow_zero LIMSEQ_NSLIMSEQ_iff) |
1126 |
||
20685
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1127 |
lemma LIMSEQ_power_zero: |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1128 |
fixes x :: "'a::{real_normed_div_algebra,recpower}" |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1129 |
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
|
1130 |
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1131 |
apply (simp add: norm_power [symmetric] LIMSEQ_norm_zero) |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1132 |
done |
fee8c75e3b5d
added lemmas about LIMSEQ and norm; simplified some proofs
huffman
parents:
20682
diff
changeset
|
1133 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1134 |
lemma LIMSEQ_divide_realpow_zero: |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1135 |
"1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0" |
15082 | 1136 |
apply (cut_tac a = a and x1 = "inverse x" in |
1137 |
LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero]) |
|
1138 |
apply (auto simp add: divide_inverse power_inverse) |
|
1139 |
apply (simp add: inverse_eq_divide pos_divide_less_eq) |
|
1140 |
done |
|
1141 |
||
15102 | 1142 |
text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*} |
15082 | 1143 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1144 |
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
|
1145 |
by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) |
15082 | 1146 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1147 |
lemma NSLIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----NS> 0" |
15082 | 1148 |
by (simp add: LIMSEQ_rabs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric]) |
1149 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1150 |
lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0" |
15082 | 1151 |
apply (rule LIMSEQ_rabs_zero [THEN iffD1]) |
1152 |
apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs) |
|
1153 |
done |
|
1154 |
||
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20408
diff
changeset
|
1155 |
lemma NSLIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----NS> 0" |
15082 | 1156 |
by (simp add: LIMSEQ_rabs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric]) |
1157 |
||
1158 |
subsection{*Hyperreals and Sequences*} |
|
1159 |
||
1160 |
text{*A bounded sequence is a finite hyperreal*} |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
1161 |
lemma NSBseq_HFinite_hypreal: "NSBseq X ==> star_n X : HFinite" |
17298 | 1162 |
by (auto intro!: bexI lemma_starrel_refl |
15082 | 1163 |
intro: FreeUltrafilterNat_all [THEN FreeUltrafilterNat_subset] |
1164 |
simp add: HFinite_FreeUltrafilterNat_iff Bseq_NSBseq_iff [symmetric] |
|
1165 |
Bseq_iff1a) |
|
1166 |
||
1167 |
text{*A sequence converging to zero defines an infinitesimal*} |
|
1168 |
lemma NSLIMSEQ_zero_Infinitesimal_hypreal: |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
1169 |
"X ----NS> 0 ==> star_n X : Infinitesimal" |
15082 | 1170 |
apply (simp add: NSLIMSEQ_def) |
1171 |
apply (drule_tac x = whn in bspec) |
|
1172 |
apply (simp add: HNatInfinite_whn) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17299
diff
changeset
|
1173 |
apply (auto simp add: hypnat_omega_def mem_infmal_iff [symmetric] starfun) |
15082 | 1174 |
done |
1175 |
||
1176 |
(***--------------------------------------------------------------- |
|
1177 |
Theorems proved by Harrison in HOL that we do not need |
|
1178 |
in order to prove equivalence between Cauchy criterion |
|
1179 |
and convergence: |
|
1180 |
-- Show that every sequence contains a monotonic subsequence |
|
1181 |
Goal "\<exists>f. subseq f & monoseq (%n. s (f n))" |
|
1182 |
-- Show that a subsequence of a bounded sequence is bounded |
|
1183 |
Goal "Bseq X ==> Bseq (%n. X (f n))"; |
|
1184 |
-- Show we can take subsequential terms arbitrarily far |
|
1185 |
up a sequence |
|
1186 |
Goal "subseq f ==> n \<le> f(n)"; |
|
1187 |
Goal "subseq f ==> \<exists>n. N1 \<le> n & N2 \<le> f(n)"; |
|
1188 |
---------------------------------------------------------------***) |
|
1189 |
||
10751 | 1190 |
end |