| author | wenzelm | 
| Wed, 28 Jun 2023 12:25:54 +0200 | |
| changeset 78224 | d85d0d41b2bd | 
| parent 70228 | 2d5b122aa0ff | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 62479 | 1  | 
(* Title: HOL/Nonstandard_Analysis/HSEQ.thy  | 
2  | 
Author: Jacques D. Fleuriot  | 
|
3  | 
Copyright: 1998 University of Cambridge  | 
|
4  | 
||
5  | 
Convergence of sequences and series.  | 
|
6  | 
||
7  | 
Conversion to Isar and new proofs by Lawrence C Paulson, 2004  | 
|
8  | 
Additional contributions by Jeremy Avigad and Brian Huffman.  | 
|
| 27468 | 9  | 
*)  | 
10  | 
||
| 61975 | 11  | 
section \<open>Sequences and Convergence (Nonstandard)\<close>  | 
| 27468 | 12  | 
|
13  | 
theory HSEQ  | 
|
| 67006 | 14  | 
imports Complex_Main NatStar  | 
| 63579 | 15  | 
abbrevs "--->" = "\<longlonglongrightarrow>\<^sub>N\<^sub>S"  | 
| 27468 | 16  | 
begin  | 
17  | 
||
| 64604 | 18  | 
definition NSLIMSEQ :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 61970 | 19  | 
    ("((_)/ \<longlonglongrightarrow>\<^sub>N\<^sub>S (_))" [60, 60] 60) where
 | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67006 
diff
changeset
 | 
20  | 
\<comment> \<open>Nonstandard definition of convergence of sequence\<close>  | 
| 64604 | 21  | 
"X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"  | 
| 27468 | 22  | 
|
| 64604 | 23  | 
definition nslim :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a"  | 
24  | 
where "nslim X = (THE L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"  | 
|
25  | 
\<comment> \<open>Nonstandard definition of limit using choice operator\<close>  | 
|
26  | 
||
| 27468 | 27  | 
|
| 64604 | 28  | 
definition NSconvergent :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"  | 
29  | 
where "NSconvergent X \<longleftrightarrow> (\<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"  | 
|
30  | 
\<comment> \<open>Nonstandard definition of convergence\<close>  | 
|
| 27468 | 31  | 
|
| 64604 | 32  | 
definition NSBseq :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"  | 
33  | 
where "NSBseq X \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite)"  | 
|
34  | 
\<comment> \<open>Nonstandard definition for bounded sequence\<close>  | 
|
35  | 
||
| 27468 | 36  | 
|
| 64604 | 37  | 
definition NSCauchy :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"  | 
38  | 
where "NSCauchy X \<longleftrightarrow> (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)"  | 
|
39  | 
\<comment> \<open>Nonstandard definition\<close>  | 
|
40  | 
||
| 27468 | 41  | 
|
| 61975 | 42  | 
subsection \<open>Limits of Sequences\<close>  | 
| 27468 | 43  | 
|
| 64604 | 44  | 
lemma NSLIMSEQ_I: "(\<And>N. N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L) \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"  | 
45  | 
by (simp add: NSLIMSEQ_def)  | 
|
| 27468 | 46  | 
|
| 64604 | 47  | 
lemma NSLIMSEQ_D: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L"  | 
48  | 
by (simp add: NSLIMSEQ_def)  | 
|
| 27468 | 49  | 
|
| 64604 | 50  | 
lemma NSLIMSEQ_const: "(\<lambda>n. k) \<longlonglongrightarrow>\<^sub>N\<^sub>S k"  | 
51  | 
by (simp add: NSLIMSEQ_def)  | 
|
| 27468 | 52  | 
|
| 64604 | 53  | 
lemma NSLIMSEQ_add: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n + Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"  | 
54  | 
by (auto intro: approx_add simp add: NSLIMSEQ_def)  | 
|
| 27468 | 55  | 
|
| 64604 | 56  | 
lemma NSLIMSEQ_add_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. f n + b) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"  | 
57  | 
by (simp only: NSLIMSEQ_add NSLIMSEQ_const)  | 
|
| 27468 | 58  | 
|
| 64604 | 59  | 
lemma NSLIMSEQ_mult: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n * Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a * b"  | 
60  | 
for a b :: "'a::real_normed_algebra"  | 
|
61  | 
by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def)  | 
|
| 27468 | 62  | 
|
| 64604 | 63  | 
lemma NSLIMSEQ_minus: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. - X n) \<longlonglongrightarrow>\<^sub>N\<^sub>S - a"  | 
64  | 
by (auto simp add: NSLIMSEQ_def)  | 
|
| 27468 | 65  | 
|
| 64604 | 66  | 
lemma NSLIMSEQ_minus_cancel: "(\<lambda>n. - X n) \<longlonglongrightarrow>\<^sub>N\<^sub>S -a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S a"  | 
67  | 
by (drule NSLIMSEQ_minus) simp  | 
|
| 27468 | 68  | 
|
| 64604 | 69  | 
lemma NSLIMSEQ_diff: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n - Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
51525 
diff
changeset
 | 
70  | 
using NSLIMSEQ_add [of X a "- Y" "- b"] by (simp add: NSLIMSEQ_minus fun_Compl_def)  | 
| 
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
51525 
diff
changeset
 | 
71  | 
|
| 64604 | 72  | 
lemma NSLIMSEQ_diff_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. f n - b) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"  | 
73  | 
by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)  | 
|
| 27468 | 74  | 
|
| 64604 | 75  | 
lemma NSLIMSEQ_inverse: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (\<lambda>n. inverse (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S inverse a"  | 
76  | 
for a :: "'a::real_normed_div_algebra"  | 
|
77  | 
by (simp add: NSLIMSEQ_def star_of_approx_inverse)  | 
|
| 27468 | 78  | 
|
| 64604 | 79  | 
lemma NSLIMSEQ_mult_inverse: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> (\<lambda>n. X n / Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a / b"  | 
80  | 
for a b :: "'a::real_normed_field"  | 
|
81  | 
by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse)  | 
|
| 27468 | 82  | 
|
83  | 
lemma starfun_hnorm: "\<And>x. hnorm (( *f* f) x) = ( *f* (\<lambda>x. norm (f x))) x"  | 
|
| 64604 | 84  | 
by transfer simp  | 
| 27468 | 85  | 
|
| 61970 | 86  | 
lemma NSLIMSEQ_norm: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S norm a"  | 
| 64604 | 87  | 
by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm)  | 
| 27468 | 88  | 
|
| 64604 | 89  | 
text \<open>Uniqueness of limit.\<close>  | 
90  | 
lemma NSLIMSEQ_unique: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> a = b"  | 
|
| 
70228
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
91  | 
unfolding NSLIMSEQ_def  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
92  | 
using HNatInfinite_whn approx_trans3 star_of_approx_iff by blast  | 
| 27468 | 93  | 
|
| 64604 | 94  | 
lemma NSLIMSEQ_pow [rule_format]: "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S a) \<longrightarrow> ((\<lambda>n. (X n) ^ m) \<longlonglongrightarrow>\<^sub>N\<^sub>S a ^ m)"  | 
95  | 
  for a :: "'a::{real_normed_algebra,power}"
 | 
|
96  | 
by (induct m) (auto intro: NSLIMSEQ_mult NSLIMSEQ_const)  | 
|
97  | 
||
98  | 
text \<open>We can now try and derive a few properties of sequences,  | 
|
99  | 
starting with the limit comparison property for sequences.\<close>  | 
|
| 27468 | 100  | 
|
| 64604 | 101  | 
lemma NSLIMSEQ_le: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<longlonglongrightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> \<exists>N. \<forall>n \<ge> N. f n \<le> g n \<Longrightarrow> l \<le> m"  | 
102  | 
for l m :: real  | 
|
| 
70228
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
103  | 
unfolding NSLIMSEQ_def  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
104  | 
by (metis HNatInfinite_whn bex_Infinitesimal_iff2 hypnat_of_nat_le_whn hypreal_of_real_le_add_Infininitesimal_cancel2 starfun_le_mono)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
105  | 
|
| 64604 | 106  | 
lemma NSLIMSEQ_le_const: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S r \<Longrightarrow> \<forall>n. a \<le> X n \<Longrightarrow> a \<le> r"  | 
107  | 
for a r :: real  | 
|
108  | 
by (erule NSLIMSEQ_le [OF NSLIMSEQ_const]) auto  | 
|
| 27468 | 109  | 
|
| 64604 | 110  | 
lemma NSLIMSEQ_le_const2: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S r \<Longrightarrow> \<forall>n. X n \<le> a \<Longrightarrow> r \<le> a"  | 
111  | 
for a r :: real  | 
|
112  | 
by (erule NSLIMSEQ_le [OF _ NSLIMSEQ_const]) auto  | 
|
| 27468 | 113  | 
|
| 64604 | 114  | 
text \<open>Shift a convergent series by 1:  | 
| 27468 | 115  | 
By the equivalence between Cauchiness and convergence and because  | 
| 61975 | 116  | 
the successor of an infinite hypernatural is also infinite.\<close>  | 
| 27468 | 117  | 
|
| 
70228
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
118  | 
lemma NSLIMSEQ_Suc_iff: "((\<lambda>n. f (Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l) \<longleftrightarrow> (f \<longlonglongrightarrow>\<^sub>N\<^sub>S l)"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
119  | 
proof  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
120  | 
assume *: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
121  | 
show "(\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
122  | 
proof (rule NSLIMSEQ_I)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
123  | 
fix N  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
124  | 
assume "N \<in> HNatInfinite"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
125  | 
then have "(*f* f) (N + 1) \<approx> star_of l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
126  | 
by (simp add: HNatInfinite_add NSLIMSEQ_D *)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
127  | 
then show "(*f* (\<lambda>n. f (Suc n))) N \<approx> star_of l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
128  | 
by (simp add: starfun_shift_one)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
129  | 
qed  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
130  | 
next  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
131  | 
assume *: "(\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
132  | 
show "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
133  | 
proof (rule NSLIMSEQ_I)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
134  | 
fix N  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
135  | 
assume "N \<in> HNatInfinite"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
136  | 
then have "(*f* (\<lambda>n. f (Suc n))) (N - 1) \<approx> star_of l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
137  | 
using * by (simp add: HNatInfinite_diff NSLIMSEQ_D)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
138  | 
then show "(*f* f) N \<approx> star_of l"  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
139  | 
by (simp add: \<open>N \<in> HNatInfinite\<close> one_le_HNatInfinite starfun_shift_one)  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
140  | 
qed  | 
| 
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
141  | 
qed  | 
| 64604 | 142  | 
|
| 27468 | 143  | 
|
| 69597 | 144  | 
subsubsection \<open>Equivalence of \<^term>\<open>LIMSEQ\<close> and \<^term>\<open>NSLIMSEQ\<close>\<close>  | 
| 27468 | 145  | 
|
146  | 
lemma LIMSEQ_NSLIMSEQ:  | 
|
| 64604 | 147  | 
assumes X: "X \<longlonglongrightarrow> L"  | 
148  | 
shows "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"  | 
|
| 27468 | 149  | 
proof (rule NSLIMSEQ_I)  | 
| 64604 | 150  | 
fix N  | 
151  | 
assume N: "N \<in> HNatInfinite"  | 
|
| 27468 | 152  | 
have "starfun X N - star_of L \<in> Infinitesimal"  | 
153  | 
proof (rule InfinitesimalI2)  | 
|
| 64604 | 154  | 
fix r :: real  | 
155  | 
assume r: "0 < r"  | 
|
156  | 
from LIMSEQ_D [OF X r] obtain no where "\<forall>n\<ge>no. norm (X n - L) < r" ..  | 
|
157  | 
then have "\<forall>n\<ge>star_of no. hnorm (starfun X n - star_of L) < star_of r"  | 
|
| 27468 | 158  | 
by transfer  | 
| 64604 | 159  | 
then show "hnorm (starfun X N - star_of L) < star_of r"  | 
| 27468 | 160  | 
using N by (simp add: star_of_le_HNatInfinite)  | 
161  | 
qed  | 
|
| 64604 | 162  | 
then show "starfun X N \<approx> star_of L"  | 
163  | 
by (simp only: approx_def)  | 
|
| 27468 | 164  | 
qed  | 
165  | 
||
166  | 
lemma NSLIMSEQ_LIMSEQ:  | 
|
| 64604 | 167  | 
assumes X: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"  | 
168  | 
shows "X \<longlonglongrightarrow> L"  | 
|
| 27468 | 169  | 
proof (rule LIMSEQ_I)  | 
| 64604 | 170  | 
fix r :: real  | 
171  | 
assume r: "0 < r"  | 
|
| 27468 | 172  | 
have "\<exists>no. \<forall>n\<ge>no. hnorm (starfun X n - star_of L) < star_of r"  | 
173  | 
proof (intro exI allI impI)  | 
|
| 64604 | 174  | 
fix n  | 
175  | 
assume "whn \<le> n"  | 
|
| 27468 | 176  | 
with HNatInfinite_whn have "n \<in> HNatInfinite"  | 
177  | 
by (rule HNatInfinite_upward_closed)  | 
|
178  | 
with X have "starfun X n \<approx> star_of L"  | 
|
179  | 
by (rule NSLIMSEQ_D)  | 
|
| 64604 | 180  | 
then have "starfun X n - star_of L \<in> Infinitesimal"  | 
181  | 
by (simp only: approx_def)  | 
|
182  | 
then show "hnorm (starfun X n - star_of L) < star_of r"  | 
|
| 27468 | 183  | 
using r by (rule InfinitesimalD2)  | 
184  | 
qed  | 
|
| 64604 | 185  | 
then show "\<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"  | 
| 27468 | 186  | 
by transfer  | 
187  | 
qed  | 
|
188  | 
||
| 64604 | 189  | 
theorem LIMSEQ_NSLIMSEQ_iff: "f \<longlonglongrightarrow> L \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S L"  | 
190  | 
by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ)  | 
|
191  | 
||
| 27468 | 192  | 
|
| 69597 | 193  | 
subsubsection \<open>Derived theorems about \<^term>\<open>NSLIMSEQ\<close>\<close>  | 
| 27468 | 194  | 
|
| 64604 | 195  | 
text \<open>We prove the NS version from the standard one, since the NS proof  | 
196  | 
seems more complicated than the standard one above!\<close>  | 
|
197  | 
lemma NSLIMSEQ_norm_zero: "(\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0 \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"  | 
|
198  | 
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_norm_zero_iff)  | 
|
| 27468 | 199  | 
|
| 64604 | 200  | 
lemma NSLIMSEQ_rabs_zero: "(\<lambda>n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0 \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S (0::real)"  | 
201  | 
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_rabs_zero_iff)  | 
|
| 27468 | 202  | 
|
| 64604 | 203  | 
text \<open>Generalization to other limits.\<close>  | 
204  | 
lemma NSLIMSEQ_imp_rabs: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> (\<lambda>n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S \<bar>l\<bar>"  | 
|
205  | 
for l :: real  | 
|
206  | 
by (simp add: NSLIMSEQ_def) (auto intro: approx_hrabs simp add: starfun_abs)  | 
|
207  | 
||
208  | 
lemma NSLIMSEQ_inverse_zero: "\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f n \<Longrightarrow> (\<lambda>n. inverse (f n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"  | 
|
209  | 
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero)  | 
|
| 27468 | 210  | 
|
| 64604 | 211  | 
lemma NSLIMSEQ_inverse_real_of_nat: "(\<lambda>n. inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"  | 
212  | 
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat del: of_nat_Suc)  | 
|
| 27468 | 213  | 
|
| 64604 | 214  | 
lemma NSLIMSEQ_inverse_real_of_nat_add: "(\<lambda>n. r + inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"  | 
215  | 
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add del: of_nat_Suc)  | 
|
| 27468 | 216  | 
|
| 64604 | 217  | 
lemma NSLIMSEQ_inverse_real_of_nat_add_minus: "(\<lambda>n. r + - inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
51525 
diff
changeset
 | 
218  | 
using LIMSEQ_inverse_real_of_nat_add_minus by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])  | 
| 27468 | 219  | 
|
220  | 
lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:  | 
|
| 64604 | 221  | 
"(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"  | 
222  | 
using LIMSEQ_inverse_real_of_nat_add_minus_mult  | 
|
223  | 
by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])  | 
|
| 27468 | 224  | 
|
225  | 
||
| 61975 | 226  | 
subsection \<open>Convergence\<close>  | 
| 27468 | 227  | 
|
| 64604 | 228  | 
lemma nslimI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> nslim X = L"  | 
229  | 
by (simp add: nslim_def) (blast intro: NSLIMSEQ_unique)  | 
|
| 27468 | 230  | 
|
231  | 
lemma lim_nslim_iff: "lim X = nslim X"  | 
|
| 64604 | 232  | 
by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff)  | 
| 27468 | 233  | 
|
| 64604 | 234  | 
lemma NSconvergentD: "NSconvergent X \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"  | 
235  | 
by (simp add: NSconvergent_def)  | 
|
| 27468 | 236  | 
|
| 64604 | 237  | 
lemma NSconvergentI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> NSconvergent X"  | 
238  | 
by (auto simp add: NSconvergent_def)  | 
|
| 27468 | 239  | 
|
240  | 
lemma convergent_NSconvergent_iff: "convergent X = NSconvergent X"  | 
|
| 64604 | 241  | 
by (simp add: convergent_def NSconvergent_def LIMSEQ_NSLIMSEQ_iff)  | 
| 27468 | 242  | 
|
| 64604 | 243  | 
lemma NSconvergent_NSLIMSEQ_iff: "NSconvergent X \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S nslim X"  | 
244  | 
by (auto intro: theI NSLIMSEQ_unique simp add: NSconvergent_def nslim_def)  | 
|
| 27468 | 245  | 
|
246  | 
||
| 61975 | 247  | 
subsection \<open>Bounded Monotonic Sequences\<close>  | 
| 27468 | 248  | 
|
| 64604 | 249  | 
lemma NSBseqD: "NSBseq X \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> ( *f* X) N \<in> HFinite"  | 
250  | 
by (simp add: NSBseq_def)  | 
|
| 27468 | 251  | 
|
252  | 
lemma Standard_subset_HFinite: "Standard \<subseteq> HFinite"  | 
|
| 64604 | 253  | 
by (auto simp: Standard_def)  | 
| 27468 | 254  | 
|
255  | 
lemma NSBseqD2: "NSBseq X \<Longrightarrow> ( *f* X) N \<in> HFinite"  | 
|
| 
70228
 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
256  | 
using HNatInfinite_def NSBseq_def Nats_eq_Standard Standard_starfun Standard_subset_HFinite by blast  | 
| 27468 | 257  | 
|
| 64604 | 258  | 
lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite \<Longrightarrow> NSBseq X"  | 
259  | 
by (simp add: NSBseq_def)  | 
|
| 27468 | 260  | 
|
| 64604 | 261  | 
text \<open>The standard definition implies the nonstandard definition.\<close>  | 
262  | 
lemma Bseq_NSBseq: "Bseq X \<Longrightarrow> NSBseq X"  | 
|
263  | 
unfolding NSBseq_def  | 
|
264  | 
proof safe  | 
|
| 27468 | 265  | 
assume X: "Bseq X"  | 
| 64604 | 266  | 
fix N  | 
267  | 
assume N: "N \<in> HNatInfinite"  | 
|
268  | 
from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K"  | 
|
269  | 
by fast  | 
|
270  | 
then have "\<forall>N. hnorm (starfun X N) \<le> star_of K"  | 
|
271  | 
by transfer  | 
|
272  | 
then have "hnorm (starfun X N) \<le> star_of K"  | 
|
273  | 
by simp  | 
|
274  | 
also have "star_of K < star_of (K + 1)"  | 
|
275  | 
by simp  | 
|
276  | 
finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x"  | 
|
277  | 
by (rule bexI) simp  | 
|
278  | 
then show "starfun X N \<in> HFinite"  | 
|
279  | 
by (simp add: HFinite_def)  | 
|
| 27468 | 280  | 
qed  | 
281  | 
||
| 64604 | 282  | 
text \<open>The nonstandard definition implies the standard definition.\<close>  | 
| 27468 | 283  | 
lemma SReal_less_omega: "r \<in> \<real> \<Longrightarrow> r < \<omega>"  | 
| 64604 | 284  | 
using HInfinite_omega  | 
285  | 
by (simp add: HInfinite_def) (simp add: order_less_imp_le)  | 
|
| 27468 | 286  | 
|
287  | 
lemma NSBseq_Bseq: "NSBseq X \<Longrightarrow> Bseq X"  | 
|
288  | 
proof (rule ccontr)  | 
|
289  | 
let ?n = "\<lambda>K. LEAST n. K < norm (X n)"  | 
|
290  | 
assume "NSBseq X"  | 
|
| 64604 | 291  | 
then have finite: "( *f* X) (( *f* ?n) \<omega>) \<in> HFinite"  | 
| 27468 | 292  | 
by (rule NSBseqD2)  | 
293  | 
assume "\<not> Bseq X"  | 
|
| 64604 | 294  | 
then have "\<forall>K>0. \<exists>n. K < norm (X n)"  | 
| 27468 | 295  | 
by (simp add: Bseq_def linorder_not_le)  | 
| 64604 | 296  | 
then have "\<forall>K>0. K < norm (X (?n K))"  | 
| 27468 | 297  | 
by (auto intro: LeastI_ex)  | 
| 64604 | 298  | 
then have "\<forall>K>0. K < hnorm (( *f* X) (( *f* ?n) K))"  | 
| 27468 | 299  | 
by transfer  | 
| 64604 | 300  | 
then have "\<omega> < hnorm (( *f* X) (( *f* ?n) \<omega>))"  | 
| 27468 | 301  | 
by simp  | 
| 64604 | 302  | 
then have "\<forall>r\<in>\<real>. r < hnorm (( *f* X) (( *f* ?n) \<omega>))"  | 
| 27468 | 303  | 
by (simp add: order_less_trans [OF SReal_less_omega])  | 
| 64604 | 304  | 
then have "( *f* X) (( *f* ?n) \<omega>) \<in> HInfinite"  | 
| 27468 | 305  | 
by (simp add: HInfinite_def)  | 
306  | 
with finite show "False"  | 
|
307  | 
by (simp add: HFinite_HInfinite_iff)  | 
|
308  | 
qed  | 
|
309  | 
||
| 64604 | 310  | 
text \<open>Equivalence of nonstandard and standard definitions for a bounded sequence.\<close>  | 
311  | 
lemma Bseq_NSBseq_iff: "Bseq X = NSBseq X"  | 
|
312  | 
by (blast intro!: NSBseq_Bseq Bseq_NSBseq)  | 
|
| 27468 | 313  | 
|
| 64604 | 314  | 
text \<open>A convergent sequence is bounded:  | 
315  | 
Boundedness as a necessary condition for convergence.  | 
|
316  | 
The nonstandard version has no existential, as usual.\<close>  | 
|
317  | 
lemma NSconvergent_NSBseq: "NSconvergent X \<Longrightarrow> NSBseq X"  | 
|
318  | 
by (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def)  | 
|
319  | 
(blast intro: HFinite_star_of approx_sym approx_HFinite)  | 
|
| 27468 | 320  | 
|
| 64604 | 321  | 
text \<open>Standard Version: easily now proved using equivalence of NS and  | 
322  | 
standard definitions.\<close>  | 
|
| 27468 | 323  | 
|
| 64604 | 324  | 
lemma convergent_Bseq: "convergent X \<Longrightarrow> Bseq X"  | 
325  | 
for X :: "nat \<Rightarrow> 'b::real_normed_vector"  | 
|
326  | 
by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff)  | 
|
| 27468 | 327  | 
|
328  | 
||
| 64604 | 329  | 
subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close>  | 
| 27468 | 330  | 
|
| 64604 | 331  | 
lemma NSBseq_isUb: "NSBseq X \<Longrightarrow> \<exists>U::real. isUb UNIV {x. \<exists>n. X n = x} U"
 | 
332  | 
by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb)  | 
|
| 27468 | 333  | 
|
| 64604 | 334  | 
lemma NSBseq_isLub: "NSBseq X \<Longrightarrow> \<exists>U::real. isLub UNIV {x. \<exists>n. X n = x} U"
 | 
335  | 
by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub)  | 
|
336  | 
||
| 27468 | 337  | 
|
| 64604 | 338  | 
subsubsection \<open>A Bounded and Monotonic Sequence Converges\<close>  | 
| 27468 | 339  | 
|
| 64604 | 340  | 
text \<open>The best of both worlds: Easier to prove this result as a standard  | 
| 27468 | 341  | 
theorem and then use equivalence to "transfer" it into the  | 
| 61975 | 342  | 
equivalent nonstandard form if needed!\<close>  | 
| 27468 | 343  | 
|
| 68614 | 344  | 
lemma Bmonoseq_NSLIMSEQ: "\<forall>\<^sub>F k in sequentially. X k = X m \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S X m"  | 
345  | 
unfolding LIMSEQ_NSLIMSEQ_iff[symmetric]  | 
|
346  | 
by (simp add: eventually_mono eventually_nhds_x_imp_x filterlim_iff)  | 
|
| 27468 | 347  | 
|
| 64604 | 348  | 
lemma NSBseq_mono_NSconvergent: "NSBseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> NSconvergent X"  | 
349  | 
for X :: "nat \<Rightarrow> real"  | 
|
350  | 
by (auto intro: Bseq_mono_convergent  | 
|
351  | 
simp: convergent_NSconvergent_iff [symmetric] Bseq_NSBseq_iff [symmetric])  | 
|
| 27468 | 352  | 
|
353  | 
||
| 61975 | 354  | 
subsection \<open>Cauchy Sequences\<close>  | 
| 27468 | 355  | 
|
356  | 
lemma NSCauchyI:  | 
|
| 64604 | 357  | 
"(\<And>M N. M \<in> HNatInfinite \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X M \<approx> starfun X N) \<Longrightarrow> NSCauchy X"  | 
358  | 
by (simp add: NSCauchy_def)  | 
|
| 27468 | 359  | 
|
360  | 
lemma NSCauchyD:  | 
|
| 64604 | 361  | 
"NSCauchy X \<Longrightarrow> M \<in> HNatInfinite \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X M \<approx> starfun X N"  | 
362  | 
by (simp add: NSCauchy_def)  | 
|
| 27468 | 363  | 
|
| 64604 | 364  | 
|
365  | 
subsubsection \<open>Equivalence Between NS and Standard\<close>  | 
|
| 27468 | 366  | 
|
367  | 
lemma Cauchy_NSCauchy:  | 
|
| 64604 | 368  | 
assumes X: "Cauchy X"  | 
369  | 
shows "NSCauchy X"  | 
|
| 27468 | 370  | 
proof (rule NSCauchyI)  | 
| 64604 | 371  | 
fix M  | 
372  | 
assume M: "M \<in> HNatInfinite"  | 
|
373  | 
fix N  | 
|
374  | 
assume N: "N \<in> HNatInfinite"  | 
|
| 27468 | 375  | 
have "starfun X M - starfun X N \<in> Infinitesimal"  | 
376  | 
proof (rule InfinitesimalI2)  | 
|
| 64604 | 377  | 
fix r :: real  | 
378  | 
assume r: "0 < r"  | 
|
379  | 
from CauchyD [OF X r] obtain k where "\<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" ..  | 
|
380  | 
then have "\<forall>m\<ge>star_of k. \<forall>n\<ge>star_of k. hnorm (starfun X m - starfun X n) < star_of r"  | 
|
| 27468 | 381  | 
by transfer  | 
| 64604 | 382  | 
then show "hnorm (starfun X M - starfun X N) < star_of r"  | 
| 27468 | 383  | 
using M N by (simp add: star_of_le_HNatInfinite)  | 
384  | 
qed  | 
|
| 64604 | 385  | 
then show "starfun X M \<approx> starfun X N"  | 
386  | 
by (simp only: approx_def)  | 
|
| 27468 | 387  | 
qed  | 
388  | 
||
389  | 
lemma NSCauchy_Cauchy:  | 
|
| 64604 | 390  | 
assumes X: "NSCauchy X"  | 
391  | 
shows "Cauchy X"  | 
|
| 27468 | 392  | 
proof (rule CauchyI)  | 
| 64604 | 393  | 
fix r :: real  | 
394  | 
assume r: "0 < r"  | 
|
| 27468 | 395  | 
have "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. hnorm (starfun X m - starfun X n) < star_of r"  | 
396  | 
proof (intro exI allI impI)  | 
|
| 64604 | 397  | 
fix M  | 
398  | 
assume "whn \<le> M"  | 
|
| 27468 | 399  | 
with HNatInfinite_whn have M: "M \<in> HNatInfinite"  | 
400  | 
by (rule HNatInfinite_upward_closed)  | 
|
| 64604 | 401  | 
fix N  | 
402  | 
assume "whn \<le> N"  | 
|
| 27468 | 403  | 
with HNatInfinite_whn have N: "N \<in> HNatInfinite"  | 
404  | 
by (rule HNatInfinite_upward_closed)  | 
|
405  | 
from X M N have "starfun X M \<approx> starfun X N"  | 
|
406  | 
by (rule NSCauchyD)  | 
|
| 64604 | 407  | 
then have "starfun X M - starfun X N \<in> Infinitesimal"  | 
408  | 
by (simp only: approx_def)  | 
|
409  | 
then show "hnorm (starfun X M - starfun X N) < star_of r"  | 
|
| 27468 | 410  | 
using r by (rule InfinitesimalD2)  | 
411  | 
qed  | 
|
| 64604 | 412  | 
then show "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r"  | 
| 27468 | 413  | 
by transfer  | 
414  | 
qed  | 
|
415  | 
||
416  | 
theorem NSCauchy_Cauchy_iff: "NSCauchy X = Cauchy X"  | 
|
| 64604 | 417  | 
by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy)  | 
418  | 
||
| 27468 | 419  | 
|
| 61975 | 420  | 
subsubsection \<open>Cauchy Sequences are Bounded\<close>  | 
| 27468 | 421  | 
|
| 64604 | 422  | 
text \<open>A Cauchy sequence is bounded -- nonstandard version.\<close>  | 
| 27468 | 423  | 
|
| 64604 | 424  | 
lemma NSCauchy_NSBseq: "NSCauchy X \<Longrightarrow> NSBseq X"  | 
425  | 
by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff)  | 
|
426  | 
||
| 27468 | 427  | 
|
| 61975 | 428  | 
subsubsection \<open>Cauchy Sequences are Convergent\<close>  | 
| 27468 | 429  | 
|
| 64604 | 430  | 
text \<open>Equivalence of Cauchy criterion and convergence:  | 
| 27468 | 431  | 
We will prove this using our NS formulation which provides a  | 
432  | 
much easier proof than using the standard definition. We do not  | 
|
433  | 
need to use properties of subsequences such as boundedness,  | 
|
434  | 
monotonicity etc... Compare with Harrison's corresponding proof  | 
|
435  | 
in HOL which is much longer and more complicated. Of course, we do  | 
|
436  | 
not have problems which he encountered with guessing the right  | 
|
437  | 
instantiations for his 'espsilon-delta' proof(s) in this case  | 
|
| 61975 | 438  | 
since the NS formulations do not involve existential quantifiers.\<close>  | 
| 27468 | 439  | 
|
440  | 
lemma NSconvergent_NSCauchy: "NSconvergent X \<Longrightarrow> NSCauchy X"  | 
|
| 64604 | 441  | 
by (simp add: NSconvergent_def NSLIMSEQ_def NSCauchy_def) (auto intro: approx_trans2)  | 
| 27468 | 442  | 
|
| 68614 | 443  | 
lemma real_NSCauchy_NSconvergent:  | 
444  | 
fixes X :: "nat \<Rightarrow> real"  | 
|
445  | 
assumes "NSCauchy X" shows "NSconvergent X"  | 
|
446  | 
unfolding NSconvergent_def NSLIMSEQ_def  | 
|
447  | 
proof -  | 
|
448  | 
have "( *f* X) whn \<in> HFinite"  | 
|
449  | 
by (simp add: NSBseqD2 NSCauchy_NSBseq assms)  | 
|
450  | 
moreover have "\<forall>N\<in>HNatInfinite. ( *f* X) whn \<approx> ( *f* X) N"  | 
|
451  | 
using HNatInfinite_whn NSCauchy_def assms by blast  | 
|
452  | 
ultimately show "\<exists>L. \<forall>N\<in>HNatInfinite. ( *f* X) N \<approx> hypreal_of_real L"  | 
|
453  | 
by (force dest!: st_part_Ex simp add: SReal_iff intro: approx_trans3)  | 
|
454  | 
qed  | 
|
| 27468 | 455  | 
|
| 64604 | 456  | 
lemma NSCauchy_NSconvergent: "NSCauchy X \<Longrightarrow> NSconvergent X"  | 
457  | 
for X :: "nat \<Rightarrow> 'a::banach"  | 
|
| 68614 | 458  | 
using Cauchy_convergent NSCauchy_Cauchy convergent_NSconvergent_iff by auto  | 
| 27468 | 459  | 
|
| 64604 | 460  | 
lemma NSCauchy_NSconvergent_iff: "NSCauchy X = NSconvergent X"  | 
461  | 
for X :: "nat \<Rightarrow> 'a::banach"  | 
|
462  | 
by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy)  | 
|
| 27468 | 463  | 
|
464  | 
||
| 61975 | 465  | 
subsection \<open>Power Sequences\<close>  | 
| 27468 | 466  | 
|
| 69597 | 467  | 
text \<open>The sequence \<^term>\<open>x^n\<close> tends to 0 if \<^term>\<open>0\<le>x\<close> and \<^term>\<open>x<1\<close>. Proof will use (NS) Cauchy equivalence for convergence and  | 
| 61975 | 468  | 
also fact that bounded and monotonic sequence converges.\<close>  | 
| 27468 | 469  | 
|
| 64604 | 470  | 
text \<open>We now use NS criterion to bring proof of theorem through.\<close>  | 
| 68614 | 471  | 
lemma NSLIMSEQ_realpow_zero:  | 
472  | 
fixes x :: real  | 
|
473  | 
assumes "0 \<le> x" "x < 1" shows "(\<lambda>n. x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"  | 
|
474  | 
proof -  | 
|
475  | 
have "( *f* (^) x) N \<approx> 0"  | 
|
476  | 
if N: "N \<in> HNatInfinite" and x: "NSconvergent ((^) x)" for N  | 
|
477  | 
proof -  | 
|
478  | 
have "hypreal_of_real x pow N \<approx> hypreal_of_real x pow (N + 1)"  | 
|
479  | 
by (metis HNatInfinite_add N NSCauchy_NSconvergent_iff NSCauchy_def starfun_pow x)  | 
|
480  | 
moreover obtain L where L: "hypreal_of_real x pow N \<approx> hypreal_of_real L"  | 
|
481  | 
using NSconvergentD [OF x] N by (auto simp add: NSLIMSEQ_def starfun_pow)  | 
|
482  | 
ultimately have "hypreal_of_real x pow N \<approx> hypreal_of_real L * hypreal_of_real x"  | 
|
483  | 
by (simp add: approx_mult_subst_star_of hyperpow_add)  | 
|
484  | 
then have "hypreal_of_real L \<approx> hypreal_of_real L * hypreal_of_real x"  | 
|
485  | 
using L approx_trans3 by blast  | 
|
486  | 
then show ?thesis  | 
|
487  | 
by (metis L \<open>x < 1\<close> hyperpow_def less_irrefl mult.right_neutral mult_left_cancel star_of_approx_iff star_of_mult star_of_simps(9) starfun2_star_of)  | 
|
488  | 
qed  | 
|
489  | 
with assms show ?thesis  | 
|
490  | 
by (force dest!: convergent_realpow simp add: NSLIMSEQ_def convergent_NSconvergent_iff)  | 
|
491  | 
qed  | 
|
| 27468 | 492  | 
|
| 68614 | 493  | 
lemma NSLIMSEQ_abs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"  | 
| 64604 | 494  | 
for c :: real  | 
| 68614 | 495  | 
by (simp add: LIMSEQ_abs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric])  | 
| 27468 | 496  | 
|
| 68614 | 497  | 
lemma NSLIMSEQ_abs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"  | 
498  | 
for c :: real  | 
|
499  | 
by (simp add: LIMSEQ_abs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric])  | 
|
| 27468 | 500  | 
|
501  | 
end  |