| author | nipkow | 
| Sun, 21 Jan 2007 13:27:41 +0100 | |
| changeset 22143 | cf58486ca11b | 
| parent 21810 | b2d23672b003 | 
| child 22442 | 15d9ed9b5051 | 
| permissions | -rw-r--r-- | 
| 10751 | 1  | 
(* Title : Lim.thy  | 
| 14477 | 2  | 
ID : $Id$  | 
| 10751 | 3  | 
Author : Jacques D. Fleuriot  | 
4  | 
Copyright : 1998 University of Cambridge  | 
|
| 14477 | 5  | 
Conversion to Isar and new proofs by Lawrence C Paulson, 2004  | 
| 10751 | 6  | 
*)  | 
7  | 
||
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
8  | 
header{* Limits and Continuity *}
 | 
| 10751 | 9  | 
|
| 15131 | 10  | 
theory Lim  | 
| 15360 | 11  | 
imports SEQ  | 
| 15131 | 12  | 
begin  | 
| 14477 | 13  | 
|
14  | 
text{*Standard and Nonstandard Definitions*}
 | 
|
| 10751 | 15  | 
|
| 19765 | 16  | 
definition  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
17  | 
LIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool"  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
18  | 
        ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where
 | 
| 19765 | 19  | 
"f -- a --> L =  | 
| 20563 | 20  | 
(\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & norm (x - a) < s  | 
21  | 
--> norm (f x - L) < r)"  | 
|
| 10751 | 22  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
23  | 
definition  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
24  | 
NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool"  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
25  | 
            ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) where
 | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
26  | 
"f -- a --NS> L =  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
27  | 
(\<forall>x. (x \<noteq> star_of a & x @= star_of a --> ( *f* f) x @= star_of L))"  | 
| 10751 | 28  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
29  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
30  | 
isCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where  | 
| 19765 | 31  | 
"isCont f a = (f -- a --> (f a))"  | 
| 10751 | 32  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
33  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
34  | 
isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where  | 
| 15228 | 35  | 
    --{*NS definition dispenses with limit notions*}
 | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
36  | 
"isNSCont f a = (\<forall>y. y @= star_of a -->  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
37  | 
( *f* f) y @= star_of (f a))"  | 
| 10751 | 38  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
39  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
40  | 
isUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where  | 
| 20752 | 41  | 
"isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r)"  | 
| 10751 | 42  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
43  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
44  | 
isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where  | 
| 19765 | 45  | 
"isNSUCont f = (\<forall>x y. x @= y --> ( *f* f) x @= ( *f* f) y)"  | 
| 10751 | 46  | 
|
47  | 
||
| 20755 | 48  | 
subsection {* Limits of Functions *}
 | 
| 14477 | 49  | 
|
| 20755 | 50  | 
subsubsection {* Purely standard proofs *}
 | 
| 14477 | 51  | 
|
52  | 
lemma LIM_eq:  | 
|
53  | 
"f -- a --> L =  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
54  | 
(\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"  | 
| 14477 | 55  | 
by (simp add: LIM_def diff_minus)  | 
56  | 
||
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
57  | 
lemma LIM_I:  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
58  | 
"(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
59  | 
==> f -- a --> L"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
60  | 
by (simp add: LIM_eq)  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
61  | 
|
| 14477 | 62  | 
lemma LIM_D:  | 
63  | 
"[| f -- a --> L; 0<r |]  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
64  | 
==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"  | 
| 14477 | 65  | 
by (simp add: LIM_eq)  | 
66  | 
||
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
67  | 
lemma LIM_offset: "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"  | 
| 20756 | 68  | 
apply (rule LIM_I)  | 
69  | 
apply (drule_tac r="r" in LIM_D, safe)  | 
|
70  | 
apply (rule_tac x="s" in exI, safe)  | 
|
71  | 
apply (drule_tac x="x + k" in spec)  | 
|
72  | 
apply (simp add: compare_rls)  | 
|
73  | 
done  | 
|
74  | 
||
| 21239 | 75  | 
lemma LIM_offset_zero: "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"  | 
76  | 
by (drule_tac k="a" in LIM_offset, simp add: add_commute)  | 
|
77  | 
||
78  | 
lemma LIM_offset_zero_cancel: "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"  | 
|
79  | 
by (drule_tac k="- a" in LIM_offset, simp)  | 
|
80  | 
||
| 15228 | 81  | 
lemma LIM_const [simp]: "(%x. k) -- x --> k"  | 
| 14477 | 82  | 
by (simp add: LIM_def)  | 
83  | 
||
84  | 
lemma LIM_add:  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
85  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"  | 
| 14477 | 86  | 
assumes f: "f -- a --> L" and g: "g -- a --> M"  | 
87  | 
shows "(%x. f x + g(x)) -- a --> (L + M)"  | 
|
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
88  | 
proof (rule LIM_I)  | 
| 14477 | 89  | 
fix r :: real  | 
| 20409 | 90  | 
assume r: "0 < r"  | 
| 14477 | 91  | 
from LIM_D [OF f half_gt_zero [OF r]]  | 
92  | 
obtain fs  | 
|
93  | 
where fs: "0 < fs"  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
94  | 
and fs_lt: "\<forall>x. x \<noteq> a & norm (x-a) < fs --> norm (f x - L) < r/2"  | 
| 14477 | 95  | 
by blast  | 
96  | 
from LIM_D [OF g half_gt_zero [OF r]]  | 
|
97  | 
obtain gs  | 
|
98  | 
where gs: "0 < gs"  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
99  | 
and gs_lt: "\<forall>x. x \<noteq> a & norm (x-a) < gs --> norm (g x - M) < r/2"  | 
| 14477 | 100  | 
by blast  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
101  | 
show "\<exists>s>0.\<forall>x. x \<noteq> a \<and> norm (x-a) < s \<longrightarrow> norm (f x + g x - (L + M)) < r"  | 
| 14477 | 102  | 
proof (intro exI conjI strip)  | 
103  | 
show "0 < min fs gs" by (simp add: fs gs)  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
104  | 
fix x :: 'a  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
105  | 
assume "x \<noteq> a \<and> norm (x-a) < min fs gs"  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
106  | 
hence "x \<noteq> a \<and> norm (x-a) < fs \<and> norm (x-a) < gs" by simp  | 
| 14477 | 107  | 
with fs_lt gs_lt  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
108  | 
have "norm (f x - L) < r/2" and "norm (g x - M) < r/2" by blast+  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
109  | 
hence "norm (f x - L) + norm (g x - M) < r" by arith  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
110  | 
thus "norm (f x + g x - (L + M)) < r"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
111  | 
by (blast intro: norm_diff_triangle_ineq order_le_less_trans)  | 
| 14477 | 112  | 
qed  | 
113  | 
qed  | 
|
114  | 
||
| 21257 | 115  | 
lemma LIM_add_zero:  | 
116  | 
"\<lbrakk>f -- a --> 0; g -- a --> 0\<rbrakk> \<Longrightarrow> (\<lambda>x. f x + g x) -- a --> 0"  | 
|
117  | 
by (drule (1) LIM_add, simp)  | 
|
118  | 
||
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
119  | 
lemma minus_diff_minus:  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
120  | 
fixes a b :: "'a::ab_group_add"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
121  | 
shows "(- a) - (- b) = - (a - b)"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
122  | 
by simp  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
123  | 
|
| 14477 | 124  | 
lemma LIM_minus: "f -- a --> L ==> (%x. -f(x)) -- a --> -L"  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
125  | 
by (simp only: LIM_eq minus_diff_minus norm_minus_cancel)  | 
| 14477 | 126  | 
|
127  | 
lemma LIM_add_minus:  | 
|
128  | 
"[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)"  | 
|
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
129  | 
by (intro LIM_add LIM_minus)  | 
| 14477 | 130  | 
|
131  | 
lemma LIM_diff:  | 
|
132  | 
"[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) - g(x)) -- x --> l-m"  | 
|
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
133  | 
by (simp only: diff_minus LIM_add LIM_minus)  | 
| 14477 | 134  | 
|
| 21239 | 135  | 
lemma LIM_zero: "f -- a --> l \<Longrightarrow> (\<lambda>x. f x - l) -- a --> 0"  | 
136  | 
by (simp add: LIM_def)  | 
|
137  | 
||
138  | 
lemma LIM_zero_cancel: "(\<lambda>x. f x - l) -- a --> 0 \<Longrightarrow> f -- a --> l"  | 
|
139  | 
by (simp add: LIM_def)  | 
|
140  | 
||
| 21399 | 141  | 
lemma LIM_zero_iff: "(\<lambda>x. f x - l) -- a --> 0 = f -- a --> l"  | 
142  | 
by (simp add: LIM_def)  | 
|
143  | 
||
| 21257 | 144  | 
lemma LIM_imp_LIM:  | 
145  | 
assumes f: "f -- a --> l"  | 
|
146  | 
assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"  | 
|
147  | 
shows "g -- a --> m"  | 
|
148  | 
apply (rule LIM_I, drule LIM_D [OF f], safe)  | 
|
149  | 
apply (rule_tac x="s" in exI, safe)  | 
|
150  | 
apply (drule_tac x="x" in spec, safe)  | 
|
151  | 
apply (erule (1) order_le_less_trans [OF le])  | 
|
152  | 
done  | 
|
153  | 
||
154  | 
lemma LIM_norm: "f -- a --> l \<Longrightarrow> (\<lambda>x. norm (f x)) -- a --> norm l"  | 
|
155  | 
by (erule LIM_imp_LIM, simp add: norm_triangle_ineq3)  | 
|
156  | 
||
157  | 
lemma LIM_norm_zero: "f -- a --> 0 \<Longrightarrow> (\<lambda>x. norm (f x)) -- a --> 0"  | 
|
158  | 
by (drule LIM_norm, simp)  | 
|
159  | 
||
160  | 
lemma LIM_norm_zero_cancel: "(\<lambda>x. norm (f x)) -- a --> 0 \<Longrightarrow> f -- a --> 0"  | 
|
161  | 
by (erule LIM_imp_LIM, simp)  | 
|
162  | 
||
| 21399 | 163  | 
lemma LIM_norm_zero_iff: "(\<lambda>x. norm (f x)) -- a --> 0 = f -- a --> 0"  | 
164  | 
by (rule iffI [OF LIM_norm_zero_cancel LIM_norm_zero])  | 
|
165  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
166  | 
lemma LIM_const_not_eq:  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
167  | 
fixes a :: "'a::real_normed_div_algebra"  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
168  | 
shows "k \<noteq> L ==> ~ ((%x. k) -- a --> L)"  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
169  | 
apply (simp add: LIM_eq)  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
170  | 
apply (rule_tac x="norm (k - L)" in exI, simp, safe)  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
171  | 
apply (rule_tac x="a + of_real (s/2)" in exI, simp add: norm_of_real)  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
172  | 
done  | 
| 14477 | 173  | 
|
| 21786 | 174  | 
lemmas LIM_not_zero = LIM_const_not_eq [where L = 0]  | 
175  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
176  | 
lemma LIM_const_eq:  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
177  | 
fixes a :: "'a::real_normed_div_algebra"  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
178  | 
shows "(%x. k) -- a --> L ==> k = L"  | 
| 14477 | 179  | 
apply (rule ccontr)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
180  | 
apply (blast dest: LIM_const_not_eq)  | 
| 14477 | 181  | 
done  | 
182  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
183  | 
lemma LIM_unique:  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
184  | 
fixes a :: "'a::real_normed_div_algebra"  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
185  | 
shows "[| f -- a --> L; f -- a --> M |] ==> L = M"  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
186  | 
apply (drule LIM_diff, assumption)  | 
| 14477 | 187  | 
apply (auto dest!: LIM_const_eq)  | 
188  | 
done  | 
|
189  | 
||
190  | 
lemma LIM_mult_zero:  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
191  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"  | 
| 14477 | 192  | 
assumes f: "f -- a --> 0" and g: "g -- a --> 0"  | 
193  | 
shows "(%x. f(x) * g(x)) -- a --> 0"  | 
|
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
194  | 
proof (rule LIM_I, simp)  | 
| 14477 | 195  | 
fix r :: real  | 
196  | 
assume r: "0<r"  | 
|
197  | 
from LIM_D [OF f zero_less_one]  | 
|
198  | 
obtain fs  | 
|
199  | 
where fs: "0 < fs"  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
200  | 
and fs_lt: "\<forall>x. x \<noteq> a & norm (x-a) < fs --> norm (f x) < 1"  | 
| 14477 | 201  | 
by auto  | 
202  | 
from LIM_D [OF g r]  | 
|
203  | 
obtain gs  | 
|
204  | 
where gs: "0 < gs"  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
205  | 
and gs_lt: "\<forall>x. x \<noteq> a & norm (x-a) < gs --> norm (g x) < r"  | 
| 14477 | 206  | 
by auto  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
207  | 
show "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> a \<and> norm (x-a) < s \<longrightarrow> norm (f x * g x) < r)"  | 
| 14477 | 208  | 
proof (intro exI conjI strip)  | 
209  | 
show "0 < min fs gs" by (simp add: fs gs)  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
210  | 
fix x :: 'a  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
211  | 
assume "x \<noteq> a \<and> norm (x-a) < min fs gs"  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
212  | 
hence "x \<noteq> a \<and> norm (x-a) < fs \<and> norm (x-a) < gs" by simp  | 
| 14477 | 213  | 
with fs_lt gs_lt  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
214  | 
have "norm (f x) < 1" and "norm (g x) < r" by blast+  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
215  | 
hence "norm (f x) * norm (g x) < 1*r"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
216  | 
by (rule mult_strict_mono' [OF _ _ norm_ge_zero norm_ge_zero])  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
217  | 
thus "norm (f x * g x) < r"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
218  | 
by (simp add: order_le_less_trans [OF norm_mult_ineq])  | 
| 14477 | 219  | 
qed  | 
220  | 
qed  | 
|
221  | 
||
222  | 
lemma LIM_self: "(%x. x) -- a --> a"  | 
|
223  | 
by (auto simp add: LIM_def)  | 
|
224  | 
||
225  | 
text{*Limits are equal for functions equal except at limit point*}
 | 
|
226  | 
lemma LIM_equal:  | 
|
227  | 
"[| \<forall>x. x \<noteq> a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)"  | 
|
228  | 
by (simp add: LIM_def)  | 
|
229  | 
||
| 20796 | 230  | 
lemma LIM_cong:  | 
231  | 
"\<lbrakk>a = b; \<And>x. x \<noteq> b \<Longrightarrow> f x = g x; l = m\<rbrakk>  | 
|
| 21399 | 232  | 
\<Longrightarrow> ((\<lambda>x. f x) -- a --> l) = ((\<lambda>x. g x) -- b --> m)"  | 
| 20796 | 233  | 
by (simp add: LIM_def)  | 
234  | 
||
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
235  | 
lemma LIM_equal2:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
236  | 
assumes 1: "0 < R"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
237  | 
assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
238  | 
shows "g -- a --> l \<Longrightarrow> f -- a --> l"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
239  | 
apply (unfold LIM_def, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
240  | 
apply (drule_tac x="r" in spec, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
241  | 
apply (rule_tac x="min s R" in exI, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
242  | 
apply (simp add: 1)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
243  | 
apply (simp add: 2)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
244  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
245  | 
|
| 14477 | 246  | 
text{*Two uses in Hyperreal/Transcendental.ML*}
 | 
247  | 
lemma LIM_trans:  | 
|
248  | 
"[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l"  | 
|
249  | 
apply (drule LIM_add, assumption)  | 
|
250  | 
apply (auto simp add: add_assoc)  | 
|
251  | 
done  | 
|
252  | 
||
| 21239 | 253  | 
lemma LIM_compose:  | 
254  | 
assumes g: "g -- l --> g l"  | 
|
255  | 
assumes f: "f -- a --> l"  | 
|
256  | 
shows "(\<lambda>x. g (f x)) -- a --> g l"  | 
|
257  | 
proof (rule LIM_I)  | 
|
258  | 
fix r::real assume r: "0 < r"  | 
|
259  | 
obtain s where s: "0 < s"  | 
|
260  | 
and less_r: "\<And>y. \<lbrakk>y \<noteq> l; norm (y - l) < s\<rbrakk> \<Longrightarrow> norm (g y - g l) < r"  | 
|
261  | 
using LIM_D [OF g r] by fast  | 
|
262  | 
obtain t where t: "0 < t"  | 
|
263  | 
and less_s: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < t\<rbrakk> \<Longrightarrow> norm (f x - l) < s"  | 
|
264  | 
using LIM_D [OF f s] by fast  | 
|
265  | 
||
266  | 
show "\<exists>t>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < t \<longrightarrow> norm (g (f x) - g l) < r"  | 
|
267  | 
proof (rule exI, safe)  | 
|
268  | 
show "0 < t" using t .  | 
|
269  | 
next  | 
|
270  | 
fix x assume "x \<noteq> a" and "norm (x - a) < t"  | 
|
271  | 
hence "norm (f x - l) < s" by (rule less_s)  | 
|
272  | 
thus "norm (g (f x) - g l) < r"  | 
|
273  | 
using r less_r by (case_tac "f x = l", simp_all)  | 
|
274  | 
qed  | 
|
275  | 
qed  | 
|
276  | 
||
277  | 
lemma LIM_o: "\<lbrakk>g -- l --> g l; f -- a --> l\<rbrakk> \<Longrightarrow> (g \<circ> f) -- a --> g l"  | 
|
278  | 
unfolding o_def by (rule LIM_compose)  | 
|
279  | 
||
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
280  | 
lemma real_LIM_sandwich_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
281  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> real"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
282  | 
assumes f: "f -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
283  | 
assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
284  | 
assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
285  | 
shows "g -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
286  | 
proof (rule LIM_imp_LIM [OF f])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
287  | 
fix x assume x: "x \<noteq> a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
288  | 
have "norm (g x - 0) = g x" by (simp add: 1 x)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
289  | 
also have "g x \<le> f x" by (rule 2 [OF x])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
290  | 
also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
291  | 
also have "\<bar>f x\<bar> = norm (f x - 0)" by simp  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
292  | 
finally show "norm (g x - 0) \<le> norm (f x - 0)" .  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
293  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
294  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
295  | 
subsubsection {* Bounded Linear Operators *}
 | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
296  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
297  | 
locale bounded_linear = additive +  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
298  | 
constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
299  | 
assumes scaleR: "f (scaleR r x) = scaleR r (f x)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
300  | 
assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
301  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
302  | 
lemma (in bounded_linear) pos_bounded:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
303  | 
"\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
304  | 
apply (cut_tac bounded, erule exE)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
305  | 
apply (rule_tac x="max 1 K" in exI, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
306  | 
apply (rule order_less_le_trans [OF zero_less_one le_maxI1])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
307  | 
apply (drule spec, erule order_trans)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
308  | 
apply (rule mult_left_mono [OF le_maxI2 norm_ge_zero])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
309  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
310  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
311  | 
lemma (in bounded_linear) pos_boundedE:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
312  | 
obtains K where "0 < K" and "\<forall>x. norm (f x) \<le> norm x * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
313  | 
using pos_bounded by fast  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
314  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
315  | 
lemma (in bounded_linear) cont: "f -- a --> f a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
316  | 
proof (rule LIM_I)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
317  | 
fix r::real assume r: "0 < r"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
318  | 
obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
319  | 
using pos_bounded by fast  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
320  | 
show "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - f a) < r"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
321  | 
proof (rule exI, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
322  | 
from r K show "0 < r / K" by (rule divide_pos_pos)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
323  | 
next  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
324  | 
fix x assume x: "norm (x - a) < r / K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
325  | 
have "norm (f x - f a) = norm (f (x - a))" by (simp only: diff)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
326  | 
also have "\<dots> \<le> norm (x - a) * K" by (rule norm_le)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
327  | 
also from K x have "\<dots> < r" by (simp only: pos_less_divide_eq)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
328  | 
finally show "norm (f x - f a) < r" .  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
329  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
330  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
331  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
332  | 
lemma (in bounded_linear) LIM:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
333  | 
"g -- a --> l \<Longrightarrow> (\<lambda>x. f (g x)) -- a --> f l"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
334  | 
by (rule LIM_compose [OF cont])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
335  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
336  | 
lemma (in bounded_linear) LIM_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
337  | 
"g -- a --> 0 \<Longrightarrow> (\<lambda>x. f (g x)) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
338  | 
by (drule LIM, simp only: zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
339  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
340  | 
subsubsection {* Bounded Bilinear Operators *}
 | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
341  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
342  | 
locale bounded_bilinear =  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
343  | 
fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
344  | 
\<Rightarrow> 'c::real_normed_vector"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
345  | 
(infixl "**" 70)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
346  | 
assumes add_left: "prod (a + a') b = prod a b + prod a' b"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
347  | 
assumes add_right: "prod a (b + b') = prod a b + prod a b'"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
348  | 
assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
349  | 
assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
350  | 
assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
351  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
352  | 
lemma (in bounded_bilinear) pos_bounded:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
353  | 
"\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
354  | 
apply (cut_tac bounded, erule exE)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
355  | 
apply (rule_tac x="max 1 K" in exI, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
356  | 
apply (rule order_less_le_trans [OF zero_less_one le_maxI1])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
357  | 
apply (drule spec, drule spec, erule order_trans)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
358  | 
apply (rule mult_left_mono [OF le_maxI2])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
359  | 
apply (intro mult_nonneg_nonneg norm_ge_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
360  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
361  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
362  | 
lemma (in bounded_bilinear) additive_right: "additive (\<lambda>b. prod a b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
363  | 
by (rule additive.intro, rule add_right)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
364  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
365  | 
lemma (in bounded_bilinear) additive_left: "additive (\<lambda>a. prod a b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
366  | 
by (rule additive.intro, rule add_left)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
367  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
368  | 
lemma (in bounded_bilinear) zero_left: "prod 0 b = 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
369  | 
by (rule additive.zero [OF additive_left])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
370  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
371  | 
lemma (in bounded_bilinear) zero_right: "prod a 0 = 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
372  | 
by (rule additive.zero [OF additive_right])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
373  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
374  | 
lemma (in bounded_bilinear) minus_left: "prod (- a) b = - prod a b"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
375  | 
by (rule additive.minus [OF additive_left])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
376  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
377  | 
lemma (in bounded_bilinear) minus_right: "prod a (- b) = - prod a b"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
378  | 
by (rule additive.minus [OF additive_right])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
379  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
380  | 
lemma (in bounded_bilinear) diff_left:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
381  | 
"prod (a - a') b = prod a b - prod a' b"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
382  | 
by (rule additive.diff [OF additive_left])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
383  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
384  | 
lemma (in bounded_bilinear) diff_right:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
385  | 
"prod a (b - b') = prod a b - prod a b'"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
386  | 
by (rule additive.diff [OF additive_right])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
387  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
388  | 
lemma (in bounded_bilinear) LIM_prod_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
389  | 
assumes f: "f -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
390  | 
assumes g: "g -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
391  | 
shows "(\<lambda>x. f x ** g x) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
392  | 
proof (rule LIM_I)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
393  | 
fix r::real assume r: "0 < r"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
394  | 
obtain K where K: "0 < K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
395  | 
and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
396  | 
using pos_bounded by fast  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
397  | 
from K have K': "0 < inverse K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
398  | 
by (rule positive_imp_inverse_positive)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
399  | 
obtain s where s: "0 < s"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
400  | 
and norm_f: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < s\<rbrakk> \<Longrightarrow> norm (f x) < r"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
401  | 
using LIM_D [OF f r] by auto  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
402  | 
obtain t where t: "0 < t"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
403  | 
and norm_g: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < t\<rbrakk> \<Longrightarrow> norm (g x) < inverse K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
404  | 
using LIM_D [OF g K'] by auto  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
405  | 
show "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x ** g x - 0) < r"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
406  | 
proof (rule exI, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
407  | 
from s t show "0 < min s t" by simp  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
408  | 
next  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
409  | 
fix x assume x: "x \<noteq> a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
410  | 
assume "norm (x - a) < min s t"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
411  | 
hence xs: "norm (x - a) < s" and xt: "norm (x - a) < t" by simp_all  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
412  | 
from x xs have 1: "norm (f x) < r" by (rule norm_f)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
413  | 
from x xt have 2: "norm (g x) < inverse K" by (rule norm_g)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
414  | 
have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" by (rule norm_le)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
415  | 
also from 1 2 K have "\<dots> < r * inverse K * K"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
416  | 
by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
417  | 
also from K have "r * inverse K * K = r" by simp  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
418  | 
finally show "norm (f x ** g x - 0) < r" by simp  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
419  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
420  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
421  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
422  | 
lemma (in bounded_bilinear) bounded_linear_left:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
423  | 
"bounded_linear (\<lambda>a. a ** b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
424  | 
apply (unfold_locales)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
425  | 
apply (rule add_left)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
426  | 
apply (rule scaleR_left)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
427  | 
apply (cut_tac bounded, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
428  | 
apply (rule_tac x="norm b * K" in exI)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
429  | 
apply (simp add: mult_ac)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
430  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
431  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
432  | 
lemma (in bounded_bilinear) bounded_linear_right:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
433  | 
"bounded_linear (\<lambda>b. a ** b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
434  | 
apply (unfold_locales)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
435  | 
apply (rule add_right)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
436  | 
apply (rule scaleR_right)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
437  | 
apply (cut_tac bounded, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
438  | 
apply (rule_tac x="norm a * K" in exI)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
439  | 
apply (simp add: mult_ac)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
440  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
441  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
442  | 
lemma (in bounded_bilinear) LIM_left_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
443  | 
"f -- a --> 0 \<Longrightarrow> (\<lambda>x. f x ** c) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
444  | 
by (rule bounded_linear.LIM_zero [OF bounded_linear_left])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
445  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
446  | 
lemma (in bounded_bilinear) LIM_right_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
447  | 
"f -- a --> 0 \<Longrightarrow> (\<lambda>x. c ** f x) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
448  | 
by (rule bounded_linear.LIM_zero [OF bounded_linear_right])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
449  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
450  | 
lemma (in bounded_bilinear) prod_diff_prod:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
451  | 
"(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
452  | 
by (simp add: diff_left diff_right)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
453  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
454  | 
lemma (in bounded_bilinear) LIM:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
455  | 
"\<lbrakk>f -- a --> L; g -- a --> M\<rbrakk> \<Longrightarrow> (\<lambda>x. f x ** g x) -- a --> L ** M"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
456  | 
apply (drule LIM_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
457  | 
apply (drule LIM_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
458  | 
apply (rule LIM_zero_cancel)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
459  | 
apply (subst prod_diff_prod)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
460  | 
apply (rule LIM_add_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
461  | 
apply (rule LIM_add_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
462  | 
apply (erule (1) LIM_prod_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
463  | 
apply (erule LIM_left_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
464  | 
apply (erule LIM_right_zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
465  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
466  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
467  | 
interpretation bounded_bilinear_mult:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
468  | 
bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
469  | 
apply (rule bounded_bilinear.intro)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
470  | 
apply (rule left_distrib)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
471  | 
apply (rule right_distrib)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
472  | 
apply (rule mult_scaleR_left)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
473  | 
apply (rule mult_scaleR_right)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
474  | 
apply (rule_tac x="1" in exI)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
475  | 
apply (simp add: norm_mult_ineq)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
476  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
477  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
478  | 
interpretation bounded_bilinear_scaleR:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
479  | 
bounded_bilinear ["scaleR"]  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
480  | 
apply (rule bounded_bilinear.intro)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
481  | 
apply (rule scaleR_left_distrib)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
482  | 
apply (rule scaleR_right_distrib)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
483  | 
apply (simp add: real_scaleR_def)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
484  | 
apply (rule scaleR_left_commute)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
485  | 
apply (rule_tac x="1" in exI)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
486  | 
apply (simp add: norm_scaleR)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
487  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
488  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
489  | 
lemmas LIM_mult = bounded_bilinear_mult.LIM  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
490  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
491  | 
lemmas LIM_mult_zero = bounded_bilinear_mult.LIM_prod_zero  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
492  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
493  | 
lemmas LIM_mult_left_zero = bounded_bilinear_mult.LIM_left_zero  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
494  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
495  | 
lemmas LIM_mult_right_zero = bounded_bilinear_mult.LIM_right_zero  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
496  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
497  | 
lemmas LIM_scaleR = bounded_bilinear_scaleR.LIM  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
498  | 
|
| 20755 | 499  | 
subsubsection {* Purely nonstandard proofs *}
 | 
| 14477 | 500  | 
|
| 
20754
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
501  | 
lemma NSLIM_I:  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
502  | 
"(\<And>x. \<lbrakk>x \<noteq> star_of a; x \<approx> star_of a\<rbrakk> \<Longrightarrow> starfun f x \<approx> star_of L)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
503  | 
\<Longrightarrow> f -- a --NS> L"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
504  | 
by (simp add: NSLIM_def)  | 
| 14477 | 505  | 
|
| 
20754
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
506  | 
lemma NSLIM_D:  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
507  | 
"\<lbrakk>f -- a --NS> L; x \<noteq> star_of a; x \<approx> star_of a\<rbrakk>  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
508  | 
\<Longrightarrow> starfun f x \<approx> star_of L"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
509  | 
by (simp add: NSLIM_def)  | 
| 14477 | 510  | 
|
| 20755 | 511  | 
text{*Proving properties of limits using nonstandard definition.
 | 
512  | 
The properties hold for standard limits as well!*}  | 
|
513  | 
||
514  | 
lemma NSLIM_mult:  | 
|
515  | 
fixes l m :: "'a::real_normed_algebra"  | 
|
516  | 
shows "[| f -- x --NS> l; g -- x --NS> m |]  | 
|
517  | 
==> (%x. f(x) * g(x)) -- x --NS> (l * m)"  | 
|
518  | 
by (auto simp add: NSLIM_def intro!: approx_mult_HFinite)  | 
|
519  | 
||
| 20794 | 520  | 
lemma starfun_scaleR [simp]:  | 
521  | 
"starfun (\<lambda>x. f x *# g x) = (\<lambda>x. scaleHR (starfun f x) (starfun g x))"  | 
|
522  | 
by transfer (rule refl)  | 
|
523  | 
||
524  | 
lemma NSLIM_scaleR:  | 
|
525  | 
"[| f -- x --NS> l; g -- x --NS> m |]  | 
|
526  | 
==> (%x. f(x) *# g(x)) -- x --NS> (l *# m)"  | 
|
527  | 
by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite)  | 
|
528  | 
||
| 20755 | 529  | 
lemma NSLIM_add:  | 
530  | 
"[| f -- x --NS> l; g -- x --NS> m |]  | 
|
531  | 
==> (%x. f(x) + g(x)) -- x --NS> (l + m)"  | 
|
532  | 
by (auto simp add: NSLIM_def intro!: approx_add)  | 
|
533  | 
||
534  | 
lemma NSLIM_const [simp]: "(%x. k) -- x --NS> k"  | 
|
535  | 
by (simp add: NSLIM_def)  | 
|
536  | 
||
537  | 
lemma NSLIM_minus: "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L"  | 
|
538  | 
by (simp add: NSLIM_def)  | 
|
539  | 
||
| 21786 | 540  | 
lemma NSLIM_diff:  | 
541  | 
"\<lbrakk>f -- x --NS> l; g -- x --NS> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --NS> (l - m)"  | 
|
542  | 
by (simp only: diff_def NSLIM_add NSLIM_minus)  | 
|
543  | 
||
| 20755 | 544  | 
lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)"  | 
545  | 
by (simp only: NSLIM_add NSLIM_minus)  | 
|
546  | 
||
547  | 
lemma NSLIM_inverse:  | 
|
548  | 
fixes L :: "'a::real_normed_div_algebra"  | 
|
549  | 
shows "[| f -- a --NS> L; L \<noteq> 0 |]  | 
|
550  | 
==> (%x. inverse(f(x))) -- a --NS> (inverse L)"  | 
|
551  | 
apply (simp add: NSLIM_def, clarify)  | 
|
552  | 
apply (drule spec)  | 
|
553  | 
apply (auto simp add: star_of_approx_inverse)  | 
|
554  | 
done  | 
|
555  | 
||
556  | 
lemma NSLIM_zero:  | 
|
| 21786 | 557  | 
assumes f: "f -- a --NS> l" shows "(%x. f(x) - l) -- a --NS> 0"  | 
| 20755 | 558  | 
proof -  | 
| 21786 | 559  | 
have "(\<lambda>x. f x - l) -- a --NS> l - l"  | 
560  | 
by (rule NSLIM_diff [OF f NSLIM_const])  | 
|
| 20755 | 561  | 
thus ?thesis by simp  | 
562  | 
qed  | 
|
563  | 
||
564  | 
lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l"  | 
|
565  | 
apply (drule_tac g = "%x. l" and m = l in NSLIM_add)  | 
|
566  | 
apply (auto simp add: diff_minus add_assoc)  | 
|
567  | 
done  | 
|
568  | 
||
569  | 
lemma NSLIM_const_not_eq:  | 
|
570  | 
fixes a :: real (* TODO: generalize to real_normed_div_algebra *)  | 
|
571  | 
shows "k \<noteq> L ==> ~ ((%x. k) -- a --NS> L)"  | 
|
572  | 
apply (simp add: NSLIM_def)  | 
|
573  | 
apply (rule_tac x="star_of a + epsilon" in exI)  | 
|
574  | 
apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym]  | 
|
575  | 
simp add: hypreal_epsilon_not_zero)  | 
|
576  | 
done  | 
|
577  | 
||
578  | 
lemma NSLIM_not_zero:  | 
|
579  | 
fixes a :: real  | 
|
580  | 
shows "k \<noteq> 0 ==> ~ ((%x. k) -- a --NS> 0)"  | 
|
581  | 
by (rule NSLIM_const_not_eq)  | 
|
582  | 
||
583  | 
lemma NSLIM_const_eq:  | 
|
584  | 
fixes a :: real  | 
|
585  | 
shows "(%x. k) -- a --NS> L ==> k = L"  | 
|
586  | 
apply (rule ccontr)  | 
|
587  | 
apply (blast dest: NSLIM_const_not_eq)  | 
|
588  | 
done  | 
|
589  | 
||
590  | 
text{* can actually be proved more easily by unfolding the definition!*}
 | 
|
591  | 
lemma NSLIM_unique:  | 
|
592  | 
fixes a :: real  | 
|
593  | 
shows "[| f -- a --NS> L; f -- a --NS> M |] ==> L = M"  | 
|
594  | 
apply (drule NSLIM_minus)  | 
|
595  | 
apply (drule NSLIM_add, assumption)  | 
|
596  | 
apply (auto dest!: NSLIM_const_eq [symmetric])  | 
|
597  | 
apply (simp add: diff_def [symmetric])  | 
|
598  | 
done  | 
|
599  | 
||
600  | 
lemma NSLIM_mult_zero:  | 
|
601  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"  | 
|
602  | 
shows "[| f -- x --NS> 0; g -- x --NS> 0 |] ==> (%x. f(x)*g(x)) -- x --NS> 0"  | 
|
603  | 
by (drule NSLIM_mult, auto)  | 
|
604  | 
||
605  | 
lemma NSLIM_self: "(%x. x) -- a --NS> a"  | 
|
606  | 
by (simp add: NSLIM_def)  | 
|
607  | 
||
608  | 
subsubsection {* Equivalence of @{term LIM} and @{term NSLIM} *}
 | 
|
609  | 
||
| 
20754
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
610  | 
lemma LIM_NSLIM:  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
611  | 
assumes f: "f -- a --> L" shows "f -- a --NS> L"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
612  | 
proof (rule NSLIM_I)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
613  | 
fix x  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
614  | 
assume neq: "x \<noteq> star_of a"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
615  | 
assume approx: "x \<approx> star_of a"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
616  | 
have "starfun f x - star_of L \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
617  | 
proof (rule InfinitesimalI2)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
618  | 
fix r::real assume r: "0 < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
619  | 
from LIM_D [OF f r]  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
620  | 
obtain s where s: "0 < s" and  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
621  | 
less_r: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < s\<rbrakk> \<Longrightarrow> norm (f x - L) < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
622  | 
by fast  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
623  | 
from less_r have less_r':  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
624  | 
"\<And>x. \<lbrakk>x \<noteq> star_of a; hnorm (x - star_of a) < star_of s\<rbrakk>  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
625  | 
\<Longrightarrow> hnorm (starfun f x - star_of L) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
626  | 
by transfer  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
627  | 
from approx have "x - star_of a \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
628  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
629  | 
hence "hnorm (x - star_of a) < star_of s"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
630  | 
using s by (rule InfinitesimalD2)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
631  | 
with neq show "hnorm (starfun f x - star_of L) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
632  | 
by (rule less_r')  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
633  | 
qed  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
634  | 
thus "starfun f x \<approx> star_of L"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
635  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
636  | 
qed  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
637  | 
|
| 
20754
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
638  | 
lemma NSLIM_LIM:  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
639  | 
assumes f: "f -- a --NS> L" shows "f -- a --> L"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
640  | 
proof (rule LIM_I)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
641  | 
fix r::real assume r: "0 < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
642  | 
have "\<exists>s>0. \<forall>x. x \<noteq> star_of a \<and> hnorm (x - star_of a) < s  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
643  | 
\<longrightarrow> hnorm (starfun f x - star_of L) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
644  | 
proof (rule exI, safe)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
645  | 
show "0 < epsilon" by (rule hypreal_epsilon_gt_zero)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
646  | 
next  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
647  | 
fix x assume neq: "x \<noteq> star_of a"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
648  | 
assume "hnorm (x - star_of a) < epsilon"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
649  | 
with Infinitesimal_epsilon  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
650  | 
have "x - star_of a \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
651  | 
by (rule hnorm_less_Infinitesimal)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
652  | 
hence "x \<approx> star_of a"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
653  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
654  | 
with f neq have "starfun f x \<approx> star_of L"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
655  | 
by (rule NSLIM_D)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
656  | 
hence "starfun f x - star_of L \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
657  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
658  | 
thus "hnorm (starfun f x - star_of L) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
659  | 
using r by (rule InfinitesimalD2)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
660  | 
qed  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
661  | 
thus "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
662  | 
by transfer  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
663  | 
qed  | 
| 14477 | 664  | 
|
| 15228 | 665  | 
theorem LIM_NSLIM_iff: "(f -- x --> L) = (f -- x --NS> L)"  | 
| 14477 | 666  | 
by (blast intro: LIM_NSLIM NSLIM_LIM)  | 
667  | 
||
| 20755 | 668  | 
subsubsection {* Derived theorems about @{term LIM} *}
 | 
| 14477 | 669  | 
|
| 15228 | 670  | 
lemma LIM_mult2:  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
671  | 
fixes l m :: "'a::real_normed_algebra"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
672  | 
shows "[| f -- x --> l; g -- x --> m |]  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
673  | 
==> (%x. f(x) * g(x)) -- x --> (l * m)"  | 
| 14477 | 674  | 
by (simp add: LIM_NSLIM_iff NSLIM_mult)  | 
675  | 
||
| 20794 | 676  | 
lemma LIM_scaleR:  | 
677  | 
"[| f -- x --> l; g -- x --> m |]  | 
|
678  | 
==> (%x. f(x) *# g(x)) -- x --> (l *# m)"  | 
|
679  | 
by (simp add: LIM_NSLIM_iff NSLIM_scaleR)  | 
|
680  | 
||
| 15228 | 681  | 
lemma LIM_add2:  | 
682  | 
"[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + g(x)) -- x --> (l + m)"  | 
|
| 14477 | 683  | 
by (simp add: LIM_NSLIM_iff NSLIM_add)  | 
684  | 
||
685  | 
lemma LIM_const2: "(%x. k) -- x --> k"  | 
|
686  | 
by (simp add: LIM_NSLIM_iff)  | 
|
687  | 
||
688  | 
lemma LIM_minus2: "f -- a --> L ==> (%x. -f(x)) -- a --> -L"  | 
|
689  | 
by (simp add: LIM_NSLIM_iff NSLIM_minus)  | 
|
690  | 
||
691  | 
lemma LIM_add_minus2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)"  | 
|
692  | 
by (simp add: LIM_NSLIM_iff NSLIM_add_minus)  | 
|
693  | 
||
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
694  | 
lemma LIM_inverse:  | 
| 
20653
 
24cda2c5fd40
removed division_by_zero class requirements from several lemmas
 
huffman 
parents: 
20635 
diff
changeset
 | 
695  | 
fixes L :: "'a::real_normed_div_algebra"  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
696  | 
shows "[| f -- a --> L; L \<noteq> 0 |]  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
697  | 
==> (%x. inverse(f(x))) -- a --> (inverse L)"  | 
| 14477 | 698  | 
by (simp add: LIM_NSLIM_iff NSLIM_inverse)  | 
699  | 
||
| 21786 | 700  | 
lemma LIM_zero2: "f -- a --> l ==> (%x. f(x) - l) -- a --> 0"  | 
| 14477 | 701  | 
by (simp add: LIM_NSLIM_iff NSLIM_zero)  | 
702  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
703  | 
lemma LIM_unique2:  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
704  | 
fixes a :: real  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
705  | 
shows "[| f -- a --> L; f -- a --> M |] ==> L = M"  | 
| 14477 | 706  | 
by (simp add: LIM_NSLIM_iff NSLIM_unique)  | 
707  | 
||
708  | 
(* we can use the corresponding thm LIM_mult2 *)  | 
|
709  | 
(* for standard definition of limit *)  | 
|
710  | 
||
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
711  | 
lemma LIM_mult_zero2:  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
712  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
713  | 
shows "[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0"  | 
| 14477 | 714  | 
by (drule LIM_mult2, auto)  | 
715  | 
||
716  | 
||
| 20755 | 717  | 
subsection {* Continuity *}
 | 
| 14477 | 718  | 
|
| 21239 | 719  | 
subsubsection {* Purely standard proofs *}
 | 
720  | 
||
721  | 
lemma LIM_isCont_iff: "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"  | 
|
722  | 
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])  | 
|
723  | 
||
724  | 
lemma isCont_iff: "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"  | 
|
725  | 
by (simp add: isCont_def LIM_isCont_iff)  | 
|
726  | 
||
727  | 
lemma isCont_Id: "isCont (\<lambda>x. x) a"  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
728  | 
unfolding isCont_def by (rule LIM_self)  | 
| 21239 | 729  | 
|
| 21786 | 730  | 
lemma isCont_const [simp]: "isCont (\<lambda>x. k) a"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
731  | 
unfolding isCont_def by (rule LIM_const)  | 
| 21239 | 732  | 
|
| 21786 | 733  | 
lemma isCont_norm: "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"  | 
734  | 
unfolding isCont_def by (rule LIM_norm)  | 
|
735  | 
||
| 21239 | 736  | 
lemma isCont_add: "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
737  | 
unfolding isCont_def by (rule LIM_add)  | 
| 21239 | 738  | 
|
739  | 
lemma isCont_minus: "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
740  | 
unfolding isCont_def by (rule LIM_minus)  | 
| 21239 | 741  | 
|
742  | 
lemma isCont_diff: "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
743  | 
unfolding isCont_def by (rule LIM_diff)  | 
| 21239 | 744  | 
|
745  | 
lemma isCont_mult:  | 
|
746  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"  | 
|
| 21786 | 747  | 
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
748  | 
unfolding isCont_def by (rule LIM_mult)  | 
| 21239 | 749  | 
|
750  | 
lemma isCont_inverse:  | 
|
751  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_div_algebra"  | 
|
| 21786 | 752  | 
shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. inverse (f x)) a"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
753  | 
unfolding isCont_def by (rule LIM_inverse)  | 
| 21239 | 754  | 
|
755  | 
lemma isCont_LIM_compose:  | 
|
756  | 
"\<lbrakk>isCont g l; f -- a --> l\<rbrakk> \<Longrightarrow> (\<lambda>x. g (f x)) -- a --> g l"  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
757  | 
unfolding isCont_def by (rule LIM_compose)  | 
| 21239 | 758  | 
|
759  | 
lemma isCont_o2: "\<lbrakk>isCont f a; isCont g (f a)\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. g (f x)) a"  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
760  | 
unfolding isCont_def by (rule LIM_compose)  | 
| 21239 | 761  | 
|
762  | 
lemma isCont_o: "\<lbrakk>isCont f a; isCont g (f a)\<rbrakk> \<Longrightarrow> isCont (g o f) a"  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
763  | 
unfolding o_def by (rule isCont_o2)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
764  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
765  | 
lemma (in bounded_linear) isCont: "isCont f a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
766  | 
unfolding isCont_def by (rule cont)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
767  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
768  | 
lemma (in bounded_bilinear) isCont:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
769  | 
"\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
770  | 
unfolding isCont_def by (rule LIM)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
771  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
772  | 
lemmas isCont_scaleR = bounded_bilinear_scaleR.isCont  | 
| 21239 | 773  | 
|
774  | 
subsubsection {* Nonstandard proofs *}
 | 
|
775  | 
||
| 21786 | 776  | 
lemma isNSContD:  | 
777  | 
"\<lbrakk>isNSCont f a; y \<approx> star_of a\<rbrakk> \<Longrightarrow> ( *f* f) y \<approx> star_of (f a)"  | 
|
| 14477 | 778  | 
by (simp add: isNSCont_def)  | 
779  | 
||
780  | 
lemma isNSCont_NSLIM: "isNSCont f a ==> f -- a --NS> (f a) "  | 
|
781  | 
by (simp add: isNSCont_def NSLIM_def)  | 
|
782  | 
||
783  | 
lemma NSLIM_isNSCont: "f -- a --NS> (f a) ==> isNSCont f a"  | 
|
784  | 
apply (simp add: isNSCont_def NSLIM_def, auto)  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
785  | 
apply (case_tac "y = star_of a", auto)  | 
| 14477 | 786  | 
done  | 
787  | 
||
| 15228 | 788  | 
text{*NS continuity can be defined using NS Limit in
 | 
789  | 
similar fashion to standard def of continuity*}  | 
|
| 14477 | 790  | 
lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f -- a --NS> (f a))"  | 
791  | 
by (blast intro: isNSCont_NSLIM NSLIM_isNSCont)  | 
|
792  | 
||
| 15228 | 793  | 
text{*Hence, NS continuity can be given
 | 
794  | 
in terms of standard limit*}  | 
|
| 14477 | 795  | 
lemma isNSCont_LIM_iff: "(isNSCont f a) = (f -- a --> (f a))"  | 
796  | 
by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff)  | 
|
797  | 
||
| 15228 | 798  | 
text{*Moreover, it's trivial now that NS continuity
 | 
799  | 
is equivalent to standard continuity*}  | 
|
| 14477 | 800  | 
lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)"  | 
801  | 
apply (simp add: isCont_def)  | 
|
802  | 
apply (rule isNSCont_LIM_iff)  | 
|
803  | 
done  | 
|
804  | 
||
| 15228 | 805  | 
text{*Standard continuity ==> NS continuity*}
 | 
| 14477 | 806  | 
lemma isCont_isNSCont: "isCont f a ==> isNSCont f a"  | 
807  | 
by (erule isNSCont_isCont_iff [THEN iffD2])  | 
|
808  | 
||
| 15228 | 809  | 
text{*NS continuity ==> Standard continuity*}
 | 
| 14477 | 810  | 
lemma isNSCont_isCont: "isNSCont f a ==> isCont f a"  | 
811  | 
by (erule isNSCont_isCont_iff [THEN iffD1])  | 
|
812  | 
||
813  | 
text{*Alternative definition of continuity*}
 | 
|
814  | 
(* Prove equivalence between NS limits - *)  | 
|
815  | 
(* seems easier than using standard def *)  | 
|
816  | 
lemma NSLIM_h_iff: "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> L)"  | 
|
817  | 
apply (simp add: NSLIM_def, auto)  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
818  | 
apply (drule_tac x = "star_of a + x" in spec)  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
819  | 
apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp)  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
820  | 
apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]])  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
821  | 
apply (erule_tac [3] approx_minus_iff2 [THEN iffD1])  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
822  | 
prefer 2 apply (simp add: add_commute diff_def [symmetric])  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
823  | 
apply (rule_tac x = x in star_cases)  | 
| 
17318
 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 
huffman 
parents: 
17298 
diff
changeset
 | 
824  | 
apply (rule_tac [2] x = x in star_cases)  | 
| 
 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 
huffman 
parents: 
17298 
diff
changeset
 | 
825  | 
apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc approx_refl star_n_zero_num)  | 
| 14477 | 826  | 
done  | 
827  | 
||
828  | 
lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"  | 
|
829  | 
by (rule NSLIM_h_iff)  | 
|
830  | 
||
831  | 
lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a"  | 
|
832  | 
by (simp add: isNSCont_def)  | 
|
833  | 
||
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
834  | 
lemma isNSCont_inverse:  | 
| 
20653
 
24cda2c5fd40
removed division_by_zero class requirements from several lemmas
 
huffman 
parents: 
20635 
diff
changeset
 | 
835  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_div_algebra"  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
836  | 
shows "[| isNSCont f x; f x \<noteq> 0 |] ==> isNSCont (%x. inverse (f x)) x"  | 
| 14477 | 837  | 
by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff)  | 
838  | 
||
| 15228 | 839  | 
lemma isNSCont_const [simp]: "isNSCont (%x. k) a"  | 
| 14477 | 840  | 
by (simp add: isNSCont_def)  | 
841  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
842  | 
lemma isNSCont_abs [simp]: "isNSCont abs (a::real)"  | 
| 14477 | 843  | 
apply (simp add: isNSCont_def)  | 
| 
21810
 
b2d23672b003
generalized some lemmas; removed redundant lemmas; cleaned up some proofs
 
huffman 
parents: 
21786 
diff
changeset
 | 
844  | 
apply (auto intro: approx_hrabs simp add: starfun_rabs_hrabs)  | 
| 14477 | 845  | 
done  | 
846  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
847  | 
lemma isCont_abs [simp]: "isCont abs (a::real)"  | 
| 14477 | 848  | 
by (auto simp add: isNSCont_isCont_iff [symmetric])  | 
| 15228 | 849  | 
|
| 14477 | 850  | 
|
851  | 
(****************************************************************  | 
|
852  | 
(%* Leave as commented until I add topology theory or remove? *%)  | 
|
853  | 
(%*------------------------------------------------------------  | 
|
854  | 
Elementary topology proof for a characterisation of  | 
|
855  | 
continuity now: a function f is continuous if and only  | 
|
856  | 
  if the inverse image, {x. f(x) \<in> A}, of any open set A
 | 
|
857  | 
is always an open set  | 
|
858  | 
------------------------------------------------------------*%)  | 
|
859  | 
Goal "[| isNSopen A; \<forall>x. isNSCont f x |]  | 
|
860  | 
               ==> isNSopen {x. f x \<in> A}"
 | 
|
861  | 
by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1]));  | 
|
862  | 
by (dtac (mem_monad_approx RS approx_sym);  | 
|
863  | 
by (dres_inst_tac [("x","a")] spec 1);
 | 
|
864  | 
by (dtac isNSContD 1 THEN assume_tac 1)  | 
|
865  | 
by (dtac bspec 1 THEN assume_tac 1)  | 
|
866  | 
by (dres_inst_tac [("x","( *f* f) x")] approx_mem_monad2 1);
 | 
|
867  | 
by (blast_tac (claset() addIs [starfun_mem_starset]);  | 
|
868  | 
qed "isNSCont_isNSopen";  | 
|
869  | 
||
870  | 
Goalw [isNSCont_def]  | 
|
871  | 
          "\<forall>A. isNSopen A --> isNSopen {x. f x \<in> A} \
 | 
|
872  | 
\ ==> isNSCont f x";  | 
|
873  | 
by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS  | 
|
874  | 
(approx_minus_iff RS iffD2)],simpset() addsimps  | 
|
875  | 
[Infinitesimal_def,SReal_iff]));  | 
|
876  | 
by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1);
 | 
|
877  | 
by (etac (isNSopen_open_interval RSN (2,impE));  | 
|
878  | 
by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def]));  | 
|
879  | 
by (dres_inst_tac [("x","x")] spec 1);
 | 
|
880  | 
by (auto_tac (claset() addDs [approx_sym RS approx_mem_monad],  | 
|
881  | 
simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus]));  | 
|
882  | 
qed "isNSopen_isNSCont";  | 
|
883  | 
||
884  | 
Goal "(\<forall>x. isNSCont f x) = \  | 
|
885  | 
\     (\<forall>A. isNSopen A --> isNSopen {x. f(x) \<in> A})";
 | 
|
886  | 
by (blast_tac (claset() addIs [isNSCont_isNSopen,  | 
|
887  | 
isNSopen_isNSCont]);  | 
|
888  | 
qed "isNSCont_isNSopen_iff";  | 
|
889  | 
||
890  | 
(%*------- Standard version of same theorem --------*%)  | 
|
891  | 
Goal "(\<forall>x. isCont f x) = \  | 
|
892  | 
\         (\<forall>A. isopen A --> isopen {x. f(x) \<in> A})";
 | 
|
893  | 
by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff],  | 
|
894  | 
simpset() addsimps [isNSopen_isopen_iff RS sym,  | 
|
895  | 
isNSCont_isCont_iff RS sym]));  | 
|
896  | 
qed "isCont_isopen_iff";  | 
|
897  | 
*******************************************************************)  | 
|
898  | 
||
| 20755 | 899  | 
subsection {* Uniform Continuity *}
 | 
900  | 
||
| 14477 | 901  | 
lemma isNSUContD: "[| isNSUCont f; x \<approx> y|] ==> ( *f* f) x \<approx> ( *f* f) y"  | 
902  | 
by (simp add: isNSUCont_def)  | 
|
903  | 
||
904  | 
lemma isUCont_isCont: "isUCont f ==> isCont f x"  | 
|
905  | 
by (simp add: isUCont_def isCont_def LIM_def, meson)  | 
|
906  | 
||
| 
20754
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
907  | 
lemma isUCont_isNSUCont:  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
908  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
909  | 
assumes f: "isUCont f" shows "isNSUCont f"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
910  | 
proof (unfold isNSUCont_def, safe)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
911  | 
fix x y :: "'a star"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
912  | 
assume approx: "x \<approx> y"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
913  | 
have "starfun f x - starfun f y \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
914  | 
proof (rule InfinitesimalI2)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
915  | 
fix r::real assume r: "0 < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
916  | 
with f obtain s where s: "0 < s" and  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
917  | 
less_r: "\<And>x y. norm (x - y) < s \<Longrightarrow> norm (f x - f y) < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
918  | 
by (auto simp add: isUCont_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
919  | 
from less_r have less_r':  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
920  | 
"\<And>x y. hnorm (x - y) < star_of s  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
921  | 
\<Longrightarrow> hnorm (starfun f x - starfun f y) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
922  | 
by transfer  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
923  | 
from approx have "x - y \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
924  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
925  | 
hence "hnorm (x - y) < star_of s"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
926  | 
using s by (rule InfinitesimalD2)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
927  | 
thus "hnorm (starfun f x - starfun f y) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
928  | 
by (rule less_r')  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
929  | 
qed  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
930  | 
thus "starfun f x \<approx> starfun f y"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
931  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
932  | 
qed  | 
| 14477 | 933  | 
|
934  | 
lemma isNSUCont_isUCont:  | 
|
| 
20754
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
935  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
936  | 
assumes f: "isNSUCont f" shows "isUCont f"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
937  | 
proof (unfold isUCont_def, safe)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
938  | 
fix r::real assume r: "0 < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
939  | 
have "\<exists>s>0. \<forall>x y. hnorm (x - y) < s  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
940  | 
\<longrightarrow> hnorm (starfun f x - starfun f y) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
941  | 
proof (rule exI, safe)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
942  | 
show "0 < epsilon" by (rule hypreal_epsilon_gt_zero)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
943  | 
next  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
944  | 
fix x y :: "'a star"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
945  | 
assume "hnorm (x - y) < epsilon"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
946  | 
with Infinitesimal_epsilon  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
947  | 
have "x - y \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
948  | 
by (rule hnorm_less_Infinitesimal)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
949  | 
hence "x \<approx> y"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
950  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
951  | 
with f have "starfun f x \<approx> starfun f y"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
952  | 
by (simp add: isNSUCont_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
953  | 
hence "starfun f x - starfun f y \<in> Infinitesimal"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
954  | 
by (unfold approx_def)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
955  | 
thus "hnorm (starfun f x - starfun f y) < star_of r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
956  | 
using r by (rule InfinitesimalD2)  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
957  | 
qed  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
958  | 
thus "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
959  | 
by transfer  | 
| 
 
9c053a494dc6
add intro/dest rules for NSLIM; rewrite equivalence proofs using transfer
 
huffman 
parents: 
20752 
diff
changeset
 | 
960  | 
qed  | 
| 14477 | 961  | 
|
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
962  | 
subsection {* Relation of LIM and LIMSEQ *}
 | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
963  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
964  | 
lemma LIMSEQ_SEQ_conv1:  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
965  | 
fixes a :: "'a::real_normed_vector"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
966  | 
assumes X: "X -- a --> L"  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
967  | 
shows "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
968  | 
proof (safe intro!: LIMSEQ_I)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
969  | 
fix S :: "nat \<Rightarrow> 'a"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
970  | 
fix r :: real  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
971  | 
assume rgz: "0 < r"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
972  | 
assume as: "\<forall>n. S n \<noteq> a"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
973  | 
assume S: "S ----> a"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
974  | 
from LIM_D [OF X rgz] obtain s  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
975  | 
where sgz: "0 < s"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
976  | 
and aux: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < s\<rbrakk> \<Longrightarrow> norm (X x - L) < r"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
977  | 
by fast  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
978  | 
from LIMSEQ_D [OF S sgz]  | 
| 21733 | 979  | 
obtain no where "\<forall>n\<ge>no. norm (S n - a) < s" by blast  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
980  | 
hence "\<forall>n\<ge>no. norm (X (S n) - L) < r" by (simp add: aux as)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
981  | 
thus "\<exists>no. \<forall>n\<ge>no. norm (X (S n) - L) < r" ..  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
982  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
983  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
984  | 
lemma LIMSEQ_SEQ_conv2:  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
985  | 
fixes a :: real  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
986  | 
assumes "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
987  | 
shows "X -- a --> L"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
988  | 
proof (rule ccontr)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
989  | 
assume "\<not> (X -- a --> L)"  | 
| 20563 | 990  | 
hence "\<not> (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & norm (x - a) < s --> norm (X x - L) < r)" by (unfold LIM_def)  | 
991  | 
hence "\<exists>r > 0. \<forall>s > 0. \<exists>x. \<not>(x \<noteq> a \<and> \<bar>x - a\<bar> < s --> norm (X x - L) < r)" by simp  | 
|
992  | 
hence "\<exists>r > 0. \<forall>s > 0. \<exists>x. (x \<noteq> a \<and> \<bar>x - a\<bar> < s \<and> norm (X x - L) \<ge> r)" by (simp add: linorder_not_less)  | 
|
993  | 
then obtain r where rdef: "r > 0 \<and> (\<forall>s > 0. \<exists>x. (x \<noteq> a \<and> \<bar>x - a\<bar> < s \<and> norm (X x - L) \<ge> r))" by auto  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
994  | 
|
| 20563 | 995  | 
let ?F = "\<lambda>n::nat. SOME x. x\<noteq>a \<and> \<bar>x - a\<bar> < inverse (real (Suc n)) \<and> norm (X x - L) \<ge> r"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
996  | 
have "\<And>n. \<exists>x. x\<noteq>a \<and> \<bar>x - a\<bar> < inverse (real (Suc n)) \<and> norm (X x - L) \<ge> r"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
997  | 
using rdef by simp  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
998  | 
hence F: "\<And>n. ?F n \<noteq> a \<and> \<bar>?F n - a\<bar> < inverse (real (Suc n)) \<and> norm (X (?F n) - L) \<ge> r"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
999  | 
by (rule someI_ex)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1000  | 
hence F1: "\<And>n. ?F n \<noteq> a"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1001  | 
and F2: "\<And>n. \<bar>?F n - a\<bar> < inverse (real (Suc n))"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1002  | 
and F3: "\<And>n. norm (X (?F n) - L) \<ge> r"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1003  | 
by fast+  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1004  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1005  | 
have "?F ----> a"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1006  | 
proof (rule LIMSEQ_I, unfold real_norm_def)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1007  | 
fix e::real  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1008  | 
assume "0 < e"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1009  | 
(* choose no such that inverse (real (Suc n)) < e *)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1010  | 
have "\<exists>no. inverse (real (Suc no)) < e" by (rule reals_Archimedean)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1011  | 
then obtain m where nodef: "inverse (real (Suc m)) < e" by auto  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1012  | 
show "\<exists>no. \<forall>n. no \<le> n --> \<bar>?F n - a\<bar> < e"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1013  | 
proof (intro exI allI impI)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1014  | 
fix n  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1015  | 
assume mlen: "m \<le> n"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1016  | 
have "\<bar>?F n - a\<bar> < inverse (real (Suc n))"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1017  | 
by (rule F2)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1018  | 
also have "inverse (real (Suc n)) \<le> inverse (real (Suc m))"  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1019  | 
by auto  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1020  | 
also from nodef have  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1021  | 
"inverse (real (Suc m)) < e" .  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1022  | 
finally show "\<bar>?F n - a\<bar> < e" .  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1023  | 
qed  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1024  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1025  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1026  | 
moreover have "\<forall>n. ?F n \<noteq> a"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1027  | 
by (rule allI) (rule F1)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1028  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1029  | 
moreover from prems have "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L" by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1030  | 
ultimately have "(\<lambda>n. X (?F n)) ----> L" by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1031  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1032  | 
moreover have "\<not> ((\<lambda>n. X (?F n)) ----> L)"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1033  | 
proof -  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1034  | 
    {
 | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1035  | 
fix no::nat  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1036  | 
obtain n where "n = no + 1" by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1037  | 
then have nolen: "no \<le> n" by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1038  | 
(* We prove this by showing that for any m there is an n\<ge>m such that |X (?F n) - L| \<ge> r *)  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1039  | 
have "norm (X (?F n) - L) \<ge> r"  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1040  | 
by (rule F3)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
1041  | 
with nolen have "\<exists>n. no \<le> n \<and> norm (X (?F n) - L) \<ge> r" by fast  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1042  | 
}  | 
| 20563 | 1043  | 
then have "(\<forall>no. \<exists>n. no \<le> n \<and> norm (X (?F n) - L) \<ge> r)" by simp  | 
1044  | 
with rdef have "\<exists>e>0. (\<forall>no. \<exists>n. no \<le> n \<and> norm (X (?F n) - L) \<ge> e)" by auto  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1045  | 
thus ?thesis by (unfold LIMSEQ_def, auto simp add: linorder_not_less)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1046  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1047  | 
ultimately show False by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1048  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1049  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1050  | 
lemma LIMSEQ_SEQ_conv:  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
1051  | 
"(\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> (a::real) \<longrightarrow> (\<lambda>n. X (S n)) ----> L) =  | 
| 
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
1052  | 
(X -- a --> L)"  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1053  | 
proof  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1054  | 
assume "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1055  | 
show "X -- a --> L" by (rule LIMSEQ_SEQ_conv2)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1056  | 
next  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1057  | 
assume "(X -- a --> L)"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1058  | 
show "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L" by (rule LIMSEQ_SEQ_conv1)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1059  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
1060  | 
|
| 10751 | 1061  | 
end  |