| author | huffman | 
| Sat, 06 Jun 2009 09:11:12 -0700 | |
| changeset 31488 | 5691ccb8d6b5 | 
| parent 31487 | 93938cafc0e6 | 
| child 32436 | 10cd49e0c067 | 
| permissions | -rw-r--r-- | 
| 10751 | 1  | 
(* Title : Lim.thy  | 
2  | 
Author : Jacques D. Fleuriot  | 
|
3  | 
Copyright : 1998 University of Cambridge  | 
|
| 14477 | 4  | 
Conversion to Isar and new proofs by Lawrence C Paulson, 2004  | 
| 10751 | 5  | 
*)  | 
6  | 
||
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
7  | 
header{* Limits and Continuity *}
 | 
| 10751 | 8  | 
|
| 15131 | 9  | 
theory Lim  | 
| 
29197
 
6d4cb27ed19c
adapted HOL source structure to distribution layout
 
haftmann 
parents: 
28952 
diff
changeset
 | 
10  | 
imports SEQ  | 
| 15131 | 11  | 
begin  | 
| 14477 | 12  | 
|
| 
22641
 
a5dc96fad632
moved nonstandard limit stuff from Lim.thy into new theory HLim.thy
 
huffman 
parents: 
22637 
diff
changeset
 | 
13  | 
text{*Standard Definitions*}
 | 
| 10751 | 14  | 
|
| 19765 | 15  | 
definition  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
16  | 
LIM :: "['a::metric_space \<Rightarrow> 'b::metric_space, 'a, 'b] \<Rightarrow> bool"  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
17  | 
        ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where
 | 
| 28562 | 18  | 
[code del]: "f -- a --> L =  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
19  | 
(\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
20  | 
--> dist (f x) L < r)"  | 
| 10751 | 21  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
22  | 
definition  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
23  | 
isCont :: "['a::metric_space \<Rightarrow> 'b::metric_space, 'a] \<Rightarrow> bool" where  | 
| 19765 | 24  | 
"isCont f a = (f -- a --> (f a))"  | 
| 10751 | 25  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
26  | 
definition  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
27  | 
isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
28  | 
[code del]: "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"  | 
| 10751 | 29  | 
|
| 31392 | 30  | 
subsection {* Limits of Functions *}
 | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
31  | 
|
| 
31487
 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 
huffman 
parents: 
31392 
diff
changeset
 | 
32  | 
lemma LIM_conv_tendsto: "(f -- a --> L) \<longleftrightarrow> (f ---> L) (at a)"  | 
| 31488 | 33  | 
unfolding LIM_def tendsto_iff eventually_at ..  | 
| 10751 | 34  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
35  | 
lemma metric_LIM_I:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
36  | 
"(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
37  | 
\<Longrightarrow> f -- a --> L"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
38  | 
by (simp add: LIM_def)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
39  | 
|
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
40  | 
lemma metric_LIM_D:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
41  | 
"\<lbrakk>f -- a --> L; 0 < r\<rbrakk>  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
42  | 
\<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
43  | 
by (simp add: LIM_def)  | 
| 14477 | 44  | 
|
45  | 
lemma LIM_eq:  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
46  | 
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
47  | 
shows "f -- a --> L =  | 
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
48  | 
(\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
49  | 
by (simp add: LIM_def dist_norm)  | 
| 14477 | 50  | 
|
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
51  | 
lemma LIM_I:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
52  | 
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
53  | 
shows "(!!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
 | 
54  | 
==> f -- a --> L"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
55  | 
by (simp add: LIM_eq)  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
56  | 
|
| 14477 | 57  | 
lemma LIM_D:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
58  | 
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
59  | 
shows "[| 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
 | 
60  | 
==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"  | 
| 14477 | 61  | 
by (simp add: LIM_eq)  | 
62  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
63  | 
lemma LIM_offset:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
64  | 
fixes a :: "'a::real_normed_vector" and L :: "'b::metric_space"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
65  | 
shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
66  | 
unfolding LIM_def dist_norm  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
67  | 
apply clarify  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
68  | 
apply (drule_tac x="r" in spec, safe)  | 
| 20756 | 69  | 
apply (rule_tac x="s" in exI, safe)  | 
70  | 
apply (drule_tac x="x + k" in spec)  | 
|
| 29667 | 71  | 
apply (simp add: algebra_simps)  | 
| 20756 | 72  | 
done  | 
73  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
74  | 
lemma LIM_offset_zero:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
75  | 
fixes a :: "'a::real_normed_vector" and L :: "'b::metric_space"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
76  | 
shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"  | 
| 21239 | 77  | 
by (drule_tac k="a" in LIM_offset, simp add: add_commute)  | 
78  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
79  | 
lemma LIM_offset_zero_cancel:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
80  | 
fixes a :: "'a::real_normed_vector" and L :: "'b::metric_space"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
81  | 
shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"  | 
| 21239 | 82  | 
by (drule_tac k="- a" in LIM_offset, simp)  | 
83  | 
||
| 15228 | 84  | 
lemma LIM_const [simp]: "(%x. k) -- x --> k"  | 
| 14477 | 85  | 
by (simp add: LIM_def)  | 
86  | 
||
87  | 
lemma LIM_add:  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
88  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 14477 | 89  | 
assumes f: "f -- a --> L" and g: "g -- a --> M"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
90  | 
shows "(\<lambda>x. f x + g x) -- a --> (L + M)"  | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
91  | 
using assms unfolding LIM_conv_tendsto by (rule tendsto_add)  | 
| 14477 | 92  | 
|
| 21257 | 93  | 
lemma LIM_add_zero:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
94  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
95  | 
shows "\<lbrakk>f -- a --> 0; g -- a --> 0\<rbrakk> \<Longrightarrow> (\<lambda>x. f x + g x) -- a --> 0"  | 
| 21257 | 96  | 
by (drule (1) LIM_add, simp)  | 
97  | 
||
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
98  | 
lemma minus_diff_minus:  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
99  | 
fixes a b :: "'a::ab_group_add"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
100  | 
shows "(- a) - (- b) = - (a - b)"  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
101  | 
by simp  | 
| 
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
102  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
103  | 
lemma LIM_minus:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
104  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
105  | 
shows "f -- a --> L \<Longrightarrow> (\<lambda>x. - f x) -- a --> - L"  | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
106  | 
unfolding LIM_conv_tendsto by (rule tendsto_minus)  | 
| 14477 | 107  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
108  | 
(* TODO: delete *)  | 
| 14477 | 109  | 
lemma LIM_add_minus:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
110  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
111  | 
shows "[| 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
 | 
112  | 
by (intro LIM_add LIM_minus)  | 
| 14477 | 113  | 
|
114  | 
lemma LIM_diff:  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
115  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
116  | 
shows "\<lbrakk>f -- x --> l; g -- x --> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --> l - m"  | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
117  | 
unfolding LIM_conv_tendsto by (rule tendsto_diff)  | 
| 14477 | 118  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
119  | 
lemma LIM_zero:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
120  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
121  | 
shows "f -- a --> l \<Longrightarrow> (\<lambda>x. f x - l) -- a --> 0"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
122  | 
by (simp add: LIM_def dist_norm)  | 
| 21239 | 123  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
124  | 
lemma LIM_zero_cancel:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
125  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
126  | 
shows "(\<lambda>x. f x - l) -- a --> 0 \<Longrightarrow> f -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
127  | 
by (simp add: LIM_def dist_norm)  | 
| 21239 | 128  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
129  | 
lemma LIM_zero_iff:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
130  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
131  | 
shows "(\<lambda>x. f x - l) -- a --> 0 = f -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
132  | 
by (simp add: LIM_def dist_norm)  | 
| 21399 | 133  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
134  | 
lemma metric_LIM_imp_LIM:  | 
| 21257 | 135  | 
assumes f: "f -- a --> l"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
136  | 
assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"  | 
| 21257 | 137  | 
shows "g -- a --> m"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
138  | 
apply (rule metric_LIM_I, drule metric_LIM_D [OF f], safe)  | 
| 21257 | 139  | 
apply (rule_tac x="s" in exI, safe)  | 
140  | 
apply (drule_tac x="x" in spec, safe)  | 
|
141  | 
apply (erule (1) order_le_less_trans [OF le])  | 
|
142  | 
done  | 
|
143  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
144  | 
lemma LIM_imp_LIM:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
145  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
146  | 
fixes g :: "'a::metric_space \<Rightarrow> 'c::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
147  | 
assumes f: "f -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
148  | 
assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
149  | 
shows "g -- a --> m"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
150  | 
apply (rule metric_LIM_imp_LIM [OF f])  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
151  | 
apply (simp add: dist_norm le)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
152  | 
done  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
153  | 
|
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
154  | 
lemma LIM_norm:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
155  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
156  | 
shows "f -- a --> l \<Longrightarrow> (\<lambda>x. norm (f x)) -- a --> norm l"  | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
157  | 
unfolding LIM_conv_tendsto by (rule tendsto_norm)  | 
| 21257 | 158  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
159  | 
lemma LIM_norm_zero:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
160  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
161  | 
shows "f -- a --> 0 \<Longrightarrow> (\<lambda>x. norm (f x)) -- a --> 0"  | 
| 21257 | 162  | 
by (drule LIM_norm, simp)  | 
163  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
164  | 
lemma LIM_norm_zero_cancel:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
165  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
166  | 
shows "(\<lambda>x. norm (f x)) -- a --> 0 \<Longrightarrow> f -- a --> 0"  | 
| 21257 | 167  | 
by (erule LIM_imp_LIM, simp)  | 
168  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
169  | 
lemma LIM_norm_zero_iff:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
170  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
171  | 
shows "(\<lambda>x. norm (f x)) -- a --> 0 = f -- a --> 0"  | 
| 21399 | 172  | 
by (rule iffI [OF LIM_norm_zero_cancel LIM_norm_zero])  | 
173  | 
||
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
174  | 
lemma LIM_rabs: "f -- a --> (l::real) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) -- a --> \<bar>l\<bar>"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
175  | 
by (fold real_norm_def, rule LIM_norm)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
176  | 
|
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
177  | 
lemma LIM_rabs_zero: "f -- a --> (0::real) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) -- a --> 0"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
178  | 
by (fold real_norm_def, rule LIM_norm_zero)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
179  | 
|
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
180  | 
lemma LIM_rabs_zero_cancel: "(\<lambda>x. \<bar>f x\<bar>) -- a --> (0::real) \<Longrightarrow> f -- a --> 0"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
181  | 
by (fold real_norm_def, rule LIM_norm_zero_cancel)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
182  | 
|
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
183  | 
lemma LIM_rabs_zero_iff: "(\<lambda>x. \<bar>f x\<bar>) -- a --> (0::real) = f -- a --> 0"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
184  | 
by (fold real_norm_def, rule LIM_norm_zero_iff)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
185  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
186  | 
lemma LIM_const_not_eq:  | 
| 
23076
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
187  | 
fixes a :: "'a::real_normed_algebra_1"  | 
| 
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
188  | 
shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) -- a --> L"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
189  | 
apply (simp add: LIM_def)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
190  | 
apply (rule_tac x="dist k L" in exI, simp add: zero_less_dist_iff, safe)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
191  | 
apply (rule_tac x="a + of_real (s/2)" in exI, simp add: dist_norm)  | 
| 
20552
 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 
huffman 
parents: 
20432 
diff
changeset
 | 
192  | 
done  | 
| 14477 | 193  | 
|
| 21786 | 194  | 
lemmas LIM_not_zero = LIM_const_not_eq [where L = 0]  | 
195  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
196  | 
lemma LIM_const_eq:  | 
| 
23076
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
197  | 
fixes a :: "'a::real_normed_algebra_1"  | 
| 
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
198  | 
shows "(\<lambda>x. k) -- a --> L \<Longrightarrow> k = L"  | 
| 14477 | 199  | 
apply (rule ccontr)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
200  | 
apply (blast dest: LIM_const_not_eq)  | 
| 14477 | 201  | 
done  | 
202  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
203  | 
lemma LIM_unique:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
204  | 
  fixes a :: "'a::real_normed_algebra_1" -- {* TODO: find a more appropriate class *}
 | 
| 
23076
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
205  | 
shows "\<lbrakk>f -- a --> L; f -- a --> M\<rbrakk> \<Longrightarrow> L = M"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
206  | 
apply (rule ccontr)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
207  | 
apply (drule_tac r="dist L M / 2" in metric_LIM_D, simp add: zero_less_dist_iff)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
208  | 
apply (drule_tac r="dist L M / 2" in metric_LIM_D, simp add: zero_less_dist_iff)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
209  | 
apply (clarify, rename_tac r s)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
210  | 
apply (subgoal_tac "min r s \<noteq> 0")  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
211  | 
apply (subgoal_tac "dist L M < dist L M / 2 + dist L M / 2", simp)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
212  | 
apply (subgoal_tac "dist L M \<le> dist (f (a + of_real (min r s / 2))) L +  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
213  | 
dist (f (a + of_real (min r s / 2))) M")  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
214  | 
apply (erule le_less_trans, rule add_strict_mono)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
215  | 
apply (drule spec, erule mp, simp add: dist_norm)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
216  | 
apply (drule spec, erule mp, simp add: dist_norm)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
217  | 
apply (subst dist_commute, rule dist_triangle)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
218  | 
apply simp  | 
| 14477 | 219  | 
done  | 
220  | 
||
| 
23069
 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 
huffman 
parents: 
23040 
diff
changeset
 | 
221  | 
lemma LIM_ident [simp]: "(\<lambda>x. x) -- a --> a"  | 
| 14477 | 222  | 
by (auto simp add: LIM_def)  | 
223  | 
||
224  | 
text{*Limits are equal for functions equal except at limit point*}
 | 
|
225  | 
lemma LIM_equal:  | 
|
226  | 
"[| \<forall>x. x \<noteq> a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)"  | 
|
227  | 
by (simp add: LIM_def)  | 
|
228  | 
||
| 20796 | 229  | 
lemma LIM_cong:  | 
230  | 
"\<lbrakk>a = b; \<And>x. x \<noteq> b \<Longrightarrow> f x = g x; l = m\<rbrakk>  | 
|
| 21399 | 231  | 
\<Longrightarrow> ((\<lambda>x. f x) -- a --> l) = ((\<lambda>x. g x) -- b --> m)"  | 
| 20796 | 232  | 
by (simp add: LIM_def)  | 
233  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
234  | 
lemma metric_LIM_equal2:  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
235  | 
assumes 1: "0 < R"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
236  | 
assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < R\<rbrakk> \<Longrightarrow> f x = g x"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
237  | 
shows "g -- a --> l \<Longrightarrow> f -- a --> l"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
238  | 
apply (unfold LIM_def, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
239  | 
apply (drule_tac x="r" in spec, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
240  | 
apply (rule_tac x="min s R" in exI, safe)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
241  | 
apply (simp add: 1)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
242  | 
apply (simp add: 2)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
243  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
244  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
245  | 
lemma LIM_equal2:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
246  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::metric_space"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
247  | 
assumes 1: "0 < R"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
248  | 
assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
249  | 
shows "g -- a --> l \<Longrightarrow> f -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
250  | 
apply (unfold LIM_def dist_norm, safe)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
251  | 
apply (drule_tac x="r" in spec, safe)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
252  | 
apply (rule_tac x="min s R" in exI, safe)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
253  | 
apply (simp add: 1)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
254  | 
apply (simp add: 2)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
255  | 
done  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
256  | 
|
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
257  | 
text{*Two uses in Transcendental.ML*}
 | 
| 14477 | 258  | 
lemma LIM_trans:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
259  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
260  | 
shows "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l"  | 
| 14477 | 261  | 
apply (drule LIM_add, assumption)  | 
262  | 
apply (auto simp add: add_assoc)  | 
|
263  | 
done  | 
|
264  | 
||
| 21239 | 265  | 
lemma LIM_compose:  | 
266  | 
assumes g: "g -- l --> g l"  | 
|
267  | 
assumes f: "f -- a --> l"  | 
|
268  | 
shows "(\<lambda>x. g (f x)) -- a --> g l"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
269  | 
proof (rule metric_LIM_I)  | 
| 21239 | 270  | 
fix r::real assume r: "0 < r"  | 
271  | 
obtain s where s: "0 < s"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
272  | 
and less_r: "\<And>y. \<lbrakk>y \<noteq> l; dist y l < s\<rbrakk> \<Longrightarrow> dist (g y) (g l) < r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
273  | 
using metric_LIM_D [OF g r] by fast  | 
| 21239 | 274  | 
obtain t where t: "0 < t"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
275  | 
and less_s: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < t\<rbrakk> \<Longrightarrow> dist (f x) l < s"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
276  | 
using metric_LIM_D [OF f s] by fast  | 
| 21239 | 277  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
278  | 
show "\<exists>t>0. \<forall>x. x \<noteq> a \<and> dist x a < t \<longrightarrow> dist (g (f x)) (g l) < r"  | 
| 21239 | 279  | 
proof (rule exI, safe)  | 
280  | 
show "0 < t" using t .  | 
|
281  | 
next  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
282  | 
fix x assume "x \<noteq> a" and "dist x a < t"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
283  | 
hence "dist (f x) l < s" by (rule less_s)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
284  | 
thus "dist (g (f x)) (g l) < r"  | 
| 21239 | 285  | 
using r less_r by (case_tac "f x = l", simp_all)  | 
286  | 
qed  | 
|
287  | 
qed  | 
|
288  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
289  | 
lemma metric_LIM_compose2:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
290  | 
assumes f: "f -- a --> b"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
291  | 
assumes g: "g -- b --> c"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
292  | 
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
293  | 
shows "(\<lambda>x. g (f x)) -- a --> c"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
294  | 
proof (rule metric_LIM_I)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
295  | 
fix r :: real  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
296  | 
assume r: "0 < r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
297  | 
obtain s where s: "0 < s"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
298  | 
and less_r: "\<And>y. \<lbrakk>y \<noteq> b; dist y b < s\<rbrakk> \<Longrightarrow> dist (g y) c < r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
299  | 
using metric_LIM_D [OF g r] by fast  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
300  | 
obtain t where t: "0 < t"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
301  | 
and less_s: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < t\<rbrakk> \<Longrightarrow> dist (f x) b < s"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
302  | 
using metric_LIM_D [OF f s] by fast  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
303  | 
obtain d where d: "0 < d"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
304  | 
and neq_b: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < d\<rbrakk> \<Longrightarrow> f x \<noteq> b"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
305  | 
using inj by fast  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
306  | 
|
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
307  | 
show "\<exists>t>0. \<forall>x. x \<noteq> a \<and> dist x a < t \<longrightarrow> dist (g (f x)) c < r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
308  | 
proof (safe intro!: exI)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
309  | 
show "0 < min d t" using d t by simp  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
310  | 
next  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
311  | 
fix x  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
312  | 
assume "x \<noteq> a" and "dist x a < min d t"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
313  | 
hence "f x \<noteq> b" and "dist (f x) b < s"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
314  | 
using neq_b less_s by simp_all  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
315  | 
thus "dist (g (f x)) c < r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
316  | 
by (rule less_r)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
317  | 
qed  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
318  | 
qed  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
319  | 
|
| 23040 | 320  | 
lemma LIM_compose2:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
321  | 
fixes a :: "'a::real_normed_vector"  | 
| 23040 | 322  | 
assumes f: "f -- a --> b"  | 
323  | 
assumes g: "g -- b --> c"  | 
|
324  | 
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"  | 
|
325  | 
shows "(\<lambda>x. g (f x)) -- a --> c"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
326  | 
by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])  | 
| 23040 | 327  | 
|
| 21239 | 328  | 
lemma LIM_o: "\<lbrakk>g -- l --> g l; f -- a --> l\<rbrakk> \<Longrightarrow> (g \<circ> f) -- a --> g l"  | 
329  | 
unfolding o_def by (rule LIM_compose)  | 
|
330  | 
||
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
331  | 
lemma real_LIM_sandwich_zero:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
332  | 
fixes f g :: "'a::metric_space \<Rightarrow> real"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
333  | 
assumes f: "f -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
334  | 
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
 | 
335  | 
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
 | 
336  | 
shows "g -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
337  | 
proof (rule LIM_imp_LIM [OF f])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
338  | 
fix x assume x: "x \<noteq> a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
339  | 
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
 | 
340  | 
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
 | 
341  | 
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
 | 
342  | 
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
 | 
343  | 
finally show "norm (g x - 0) \<le> norm (f x - 0)" .  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
344  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
345  | 
|
| 
22442
 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 
huffman 
parents: 
21810 
diff
changeset
 | 
346  | 
text {* Bounded Linear Operators *}
 | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
347  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
348  | 
lemma (in bounded_linear) LIM:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
349  | 
"g -- a --> l \<Longrightarrow> (\<lambda>x. f (g x)) -- a --> f l"  | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
350  | 
unfolding LIM_conv_tendsto by (rule tendsto)  | 
| 
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
351  | 
|
| 
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
352  | 
lemma (in bounded_linear) cont: "f -- a --> f a"  | 
| 
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
353  | 
by (rule LIM [OF LIM_ident])  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
354  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
355  | 
lemma (in bounded_linear) LIM_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
356  | 
"g -- a --> 0 \<Longrightarrow> (\<lambda>x. f (g x)) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
357  | 
by (drule LIM, simp only: zero)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
358  | 
|
| 
22442
 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 
huffman 
parents: 
21810 
diff
changeset
 | 
359  | 
text {* Bounded Bilinear Operators *}
 | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
360  | 
|
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
361  | 
lemma (in bounded_bilinear) LIM:  | 
| 
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
362  | 
"\<lbrakk>f -- a --> L; g -- a --> M\<rbrakk> \<Longrightarrow> (\<lambda>x. f x ** g x) -- a --> L ** M"  | 
| 
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
363  | 
unfolding LIM_conv_tendsto by (rule tendsto)  | 
| 
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
364  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
365  | 
lemma (in bounded_bilinear) LIM_prod_zero:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
366  | 
fixes a :: "'d::metric_space"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
367  | 
assumes f: "f -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
368  | 
assumes g: "g -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
369  | 
shows "(\<lambda>x. f x ** g x) -- a --> 0"  | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
370  | 
using LIM [OF f g] by (simp add: zero_left)  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
371  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
372  | 
lemma (in bounded_bilinear) LIM_left_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
373  | 
"f -- a --> 0 \<Longrightarrow> (\<lambda>x. f x ** c) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
374  | 
by (rule bounded_linear.LIM_zero [OF bounded_linear_left])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
375  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
376  | 
lemma (in bounded_bilinear) LIM_right_zero:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
377  | 
"f -- a --> 0 \<Longrightarrow> (\<lambda>x. c ** f x) -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
378  | 
by (rule bounded_linear.LIM_zero [OF bounded_linear_right])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
379  | 
|
| 23127 | 380  | 
lemmas LIM_mult = mult.LIM  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
381  | 
|
| 23127 | 382  | 
lemmas LIM_mult_zero = mult.LIM_prod_zero  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
383  | 
|
| 23127 | 384  | 
lemmas LIM_mult_left_zero = mult.LIM_left_zero  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
385  | 
|
| 23127 | 386  | 
lemmas LIM_mult_right_zero = mult.LIM_right_zero  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
387  | 
|
| 23127 | 388  | 
lemmas LIM_scaleR = scaleR.LIM  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
389  | 
|
| 23127 | 390  | 
lemmas LIM_of_real = of_real.LIM  | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
391  | 
|
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
392  | 
lemma LIM_power:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
393  | 
  fixes f :: "'a::metric_space \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
394  | 
assumes f: "f -- a --> l"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
395  | 
shows "(\<lambda>x. f x ^ n) -- a --> l ^ n"  | 
| 
30273
 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 
huffman 
parents: 
29885 
diff
changeset
 | 
396  | 
by (induct n, simp, simp add: LIM_mult f)  | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
397  | 
|
| 
22641
 
a5dc96fad632
moved nonstandard limit stuff from Lim.thy into new theory HLim.thy
 
huffman 
parents: 
22637 
diff
changeset
 | 
398  | 
subsubsection {* Derived theorems about @{term LIM} *}
 | 
| 
 
a5dc96fad632
moved nonstandard limit stuff from Lim.thy into new theory HLim.thy
 
huffman 
parents: 
22637 
diff
changeset
 | 
399  | 
|
| 31355 | 400  | 
lemma LIM_inverse:  | 
401  | 
fixes L :: "'a::real_normed_div_algebra"  | 
|
402  | 
shows "\<lbrakk>f -- a --> L; L \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>x. inverse (f x)) -- a --> inverse L"  | 
|
403  | 
unfolding LIM_conv_tendsto  | 
|
404  | 
by (rule tendsto_inverse)  | 
|
| 22637 | 405  | 
|
406  | 
lemma LIM_inverse_fun:  | 
|
407  | 
assumes a: "a \<noteq> (0::'a::real_normed_div_algebra)"  | 
|
408  | 
shows "inverse -- a --> inverse a"  | 
|
| 31355 | 409  | 
by (rule LIM_inverse [OF LIM_ident a])  | 
| 22637 | 410  | 
|
| 29885 | 411  | 
lemma LIM_sgn:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
412  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
413  | 
shows "\<lbrakk>f -- a --> l; l \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>x. sgn (f x)) -- a --> sgn l"  | 
| 29885 | 414  | 
unfolding sgn_div_norm  | 
415  | 
by (simp add: LIM_scaleR LIM_inverse LIM_norm)  | 
|
416  | 
||
| 14477 | 417  | 
|
| 20755 | 418  | 
subsection {* Continuity *}
 | 
| 14477 | 419  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
420  | 
lemma LIM_isCont_iff:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
421  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::metric_space"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
422  | 
shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"  | 
| 21239 | 423  | 
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])  | 
424  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
425  | 
lemma isCont_iff:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
426  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::metric_space"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
427  | 
shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"  | 
| 21239 | 428  | 
by (simp add: isCont_def LIM_isCont_iff)  | 
429  | 
||
| 
23069
 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 
huffman 
parents: 
23040 
diff
changeset
 | 
430  | 
lemma isCont_ident [simp]: "isCont (\<lambda>x. x) a"  | 
| 
 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 
huffman 
parents: 
23040 
diff
changeset
 | 
431  | 
unfolding isCont_def by (rule LIM_ident)  | 
| 21239 | 432  | 
|
| 21786 | 433  | 
lemma isCont_const [simp]: "isCont (\<lambda>x. k) a"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
434  | 
unfolding isCont_def by (rule LIM_const)  | 
| 21239 | 435  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
436  | 
lemma isCont_norm:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
437  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
438  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"  | 
| 21786 | 439  | 
unfolding isCont_def by (rule LIM_norm)  | 
440  | 
||
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
441  | 
lemma isCont_rabs: "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x :: real\<bar>) a"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
442  | 
unfolding isCont_def by (rule LIM_rabs)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
443  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
444  | 
lemma isCont_add:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
445  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
446  | 
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
 | 
447  | 
unfolding isCont_def by (rule LIM_add)  | 
| 21239 | 448  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
449  | 
lemma isCont_minus:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
450  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
451  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
452  | 
unfolding isCont_def by (rule LIM_minus)  | 
| 21239 | 453  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
454  | 
lemma isCont_diff:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
455  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
456  | 
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
 | 
457  | 
unfolding isCont_def by (rule LIM_diff)  | 
| 21239 | 458  | 
|
459  | 
lemma isCont_mult:  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
460  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_algebra"  | 
| 21786 | 461  | 
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
 | 
462  | 
unfolding isCont_def by (rule LIM_mult)  | 
| 21239 | 463  | 
|
464  | 
lemma isCont_inverse:  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
465  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_div_algebra"  | 
| 21786 | 466  | 
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
 | 
467  | 
unfolding isCont_def by (rule LIM_inverse)  | 
| 21239 | 468  | 
|
469  | 
lemma isCont_LIM_compose:  | 
|
470  | 
"\<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
 | 
471  | 
unfolding isCont_def by (rule LIM_compose)  | 
| 21239 | 472  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
473  | 
lemma metric_isCont_LIM_compose2:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
474  | 
assumes f [unfolded isCont_def]: "isCont f a"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
475  | 
assumes g: "g -- f a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
476  | 
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
477  | 
shows "(\<lambda>x. g (f x)) -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
478  | 
by (rule metric_LIM_compose2 [OF f g inj])  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
479  | 
|
| 23040 | 480  | 
lemma isCont_LIM_compose2:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
481  | 
fixes a :: "'a::real_normed_vector"  | 
| 23040 | 482  | 
assumes f [unfolded isCont_def]: "isCont f a"  | 
483  | 
assumes g: "g -- f a --> l"  | 
|
484  | 
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"  | 
|
485  | 
shows "(\<lambda>x. g (f x)) -- a --> l"  | 
|
486  | 
by (rule LIM_compose2 [OF f g inj])  | 
|
487  | 
||
| 21239 | 488  | 
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
 | 
489  | 
unfolding isCont_def by (rule LIM_compose)  | 
| 21239 | 490  | 
|
491  | 
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
 | 
492  | 
unfolding o_def by (rule isCont_o2)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
493  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
494  | 
lemma (in bounded_linear) isCont: "isCont f a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
495  | 
unfolding isCont_def by (rule cont)  | 
| 
 
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  | 
lemma (in bounded_bilinear) isCont:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
498  | 
"\<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
 | 
499  | 
unfolding isCont_def by (rule LIM)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
500  | 
|
| 23127 | 501  | 
lemmas isCont_scaleR = scaleR.isCont  | 
| 21239 | 502  | 
|
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
503  | 
lemma isCont_of_real:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
504  | 
"isCont f a \<Longrightarrow> isCont (\<lambda>x. of_real (f x)::'b::real_normed_algebra_1) a"  | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
505  | 
unfolding isCont_def by (rule LIM_of_real)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
506  | 
|
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
507  | 
lemma isCont_power:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
508  | 
  fixes f :: "'a::metric_space \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
509  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
510  | 
unfolding isCont_def by (rule LIM_power)  | 
| 
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
511  | 
|
| 29885 | 512  | 
lemma isCont_sgn:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
513  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
514  | 
shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. sgn (f x)) a"  | 
| 29885 | 515  | 
unfolding isCont_def by (rule LIM_sgn)  | 
516  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
517  | 
lemma isCont_abs [simp]: "isCont abs (a::real)"  | 
| 
23069
 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 
huffman 
parents: 
23040 
diff
changeset
 | 
518  | 
by (rule isCont_rabs [OF isCont_ident])  | 
| 15228 | 519  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
520  | 
lemma isCont_setsum:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
521  | 
fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::real_normed_vector"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
522  | 
fixes A :: "'a set" assumes "finite A"  | 
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
523  | 
shows "\<forall> i \<in> A. isCont (f i) x \<Longrightarrow> isCont (\<lambda> x. \<Sum> i \<in> A. f i x) x"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
524  | 
using `finite A`  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
525  | 
proof induct  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
526  | 
case (insert a F) show "isCont (\<lambda> x. \<Sum> i \<in> (insert a F). f i x) x"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
527  | 
unfolding setsum_insert[OF `finite F` `a \<notin> F`] by (rule isCont_add, auto simp add: insert)  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
528  | 
qed auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
529  | 
|
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
530  | 
lemma LIM_less_bound: fixes f :: "real \<Rightarrow> real" assumes "b < x"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
531  | 
  and all_le: "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and isCont: "isCont f x"
 | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
532  | 
shows "0 \<le> f x"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
533  | 
proof (rule ccontr)  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
534  | 
assume "\<not> 0 \<le> f x" hence "f x < 0" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
535  | 
hence "0 < - f x / 2" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
536  | 
from isCont[unfolded isCont_def, THEN LIM_D, OF this]  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
537  | 
obtain s where "s > 0" and s_D: "\<And>x'. \<lbrakk> x' \<noteq> x ; \<bar> x' - x \<bar> < s \<rbrakk> \<Longrightarrow> \<bar> f x' - f x \<bar> < - f x / 2" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
538  | 
|
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
539  | 
let ?x = "x - min (s / 2) ((x - b) / 2)"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
540  | 
have "?x < x" and "\<bar> ?x - x \<bar> < s"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
541  | 
using `b < x` and `0 < s` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
542  | 
have "b < ?x"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
543  | 
proof (cases "s < x - b")  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
544  | 
case True thus ?thesis using `0 < s` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
545  | 
next  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
546  | 
case False hence "s / 2 \<ge> (x - b) / 2" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
547  | 
from inf_absorb2[OF this, unfolded inf_real_def]  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
548  | 
have "?x = (x + b) / 2" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
549  | 
thus ?thesis using `b < x` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
550  | 
qed  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
551  | 
hence "0 \<le> f ?x" using all_le `?x < x` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
552  | 
moreover have "\<bar>f ?x - f x\<bar> < - f x / 2"  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
553  | 
using s_D[OF _ `\<bar> ?x - x \<bar> < s`] `?x < x` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
554  | 
hence "f ?x - f x < - f x / 2" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
555  | 
hence "f ?x < f x / 2" by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
556  | 
hence "f ?x < 0" using `f x < 0` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
557  | 
thus False using `0 \<le> f ?x` by auto  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
558  | 
qed  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
559  | 
|
| 14477 | 560  | 
|
| 20755 | 561  | 
subsection {* Uniform Continuity *}
 | 
562  | 
||
| 14477 | 563  | 
lemma isUCont_isCont: "isUCont f ==> isCont f x"  | 
| 23012 | 564  | 
by (simp add: isUCont_def isCont_def LIM_def, force)  | 
| 14477 | 565  | 
|
| 23118 | 566  | 
lemma isUCont_Cauchy:  | 
567  | 
"\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"  | 
|
568  | 
unfolding isUCont_def  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
569  | 
apply (rule metric_CauchyI)  | 
| 23118 | 570  | 
apply (drule_tac x=e in spec, safe)  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
571  | 
apply (drule_tac e=s in metric_CauchyD, safe)  | 
| 23118 | 572  | 
apply (rule_tac x=M in exI, simp)  | 
573  | 
done  | 
|
574  | 
||
575  | 
lemma (in bounded_linear) isUCont: "isUCont f"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
576  | 
unfolding isUCont_def dist_norm  | 
| 23118 | 577  | 
proof (intro allI impI)  | 
578  | 
fix r::real assume r: "0 < r"  | 
|
579  | 
obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"  | 
|
580  | 
using pos_bounded by fast  | 
|
581  | 
show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"  | 
|
582  | 
proof (rule exI, safe)  | 
|
583  | 
from r K show "0 < r / K" by (rule divide_pos_pos)  | 
|
584  | 
next  | 
|
585  | 
fix x y :: 'a  | 
|
586  | 
assume xy: "norm (x - y) < r / K"  | 
|
587  | 
have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)  | 
|
588  | 
also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)  | 
|
589  | 
also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)  | 
|
590  | 
finally show "norm (f x - f y) < r" .  | 
|
591  | 
qed  | 
|
592  | 
qed  | 
|
593  | 
||
594  | 
lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"  | 
|
595  | 
by (rule isUCont [THEN isUCont_Cauchy])  | 
|
596  | 
||
| 14477 | 597  | 
|
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
598  | 
subsection {* Relation of LIM and LIMSEQ *}
 | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
599  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
600  | 
lemma LIMSEQ_SEQ_conv1:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
601  | 
fixes a :: "'a::metric_space"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
602  | 
assumes X: "X -- a --> L"  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
603  | 
shows "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
604  | 
proof (safe intro!: metric_LIMSEQ_I)  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
605  | 
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
 | 
606  | 
fix r :: real  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
607  | 
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
 | 
608  | 
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
 | 
609  | 
assume S: "S ----> a"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
610  | 
from metric_LIM_D [OF X rgz] obtain s  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
611  | 
where sgz: "0 < s"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
612  | 
and aux: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < s\<rbrakk> \<Longrightarrow> dist (X x) L < r"  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
613  | 
by fast  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
614  | 
from metric_LIMSEQ_D [OF S sgz]  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
615  | 
obtain no where "\<forall>n\<ge>no. dist (S n) a < s" by blast  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
616  | 
hence "\<forall>n\<ge>no. dist (X (S n)) L < r" by (simp add: aux as)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
617  | 
thus "\<exists>no. \<forall>n\<ge>no. dist (X (S n)) L < r" ..  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
618  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
619  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
620  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
621  | 
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
 | 
622  | 
fixes a :: real  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
623  | 
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
 | 
624  | 
shows "X -- a --> L"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
625  | 
proof (rule ccontr)  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
626  | 
assume "\<not> (X -- a --> L)"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
627  | 
hence "\<not> (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & norm (x - a) < s --> dist (X x) L < r)"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
628  | 
unfolding LIM_def dist_norm .  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
629  | 
hence "\<exists>r > 0. \<forall>s > 0. \<exists>x. \<not>(x \<noteq> a \<and> \<bar>x - a\<bar> < s --> dist (X x) L < r)" by simp  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
630  | 
hence "\<exists>r > 0. \<forall>s > 0. \<exists>x. (x \<noteq> a \<and> \<bar>x - a\<bar> < s \<and> dist (X x) L \<ge> r)" by (simp add: not_less)  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
631  | 
then obtain r where rdef: "r > 0 \<and> (\<forall>s > 0. \<exists>x. (x \<noteq> a \<and> \<bar>x - a\<bar> < s \<and> dist (X x) L \<ge> r))" by auto  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
632  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
633  | 
let ?F = "\<lambda>n::nat. SOME x. x\<noteq>a \<and> \<bar>x - a\<bar> < inverse (real (Suc n)) \<and> dist (X x) L \<ge> r"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
634  | 
have "\<And>n. \<exists>x. x\<noteq>a \<and> \<bar>x - a\<bar> < inverse (real (Suc n)) \<and> dist (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
 | 
635  | 
using rdef by simp  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
636  | 
hence F: "\<And>n. ?F n \<noteq> a \<and> \<bar>?F n - a\<bar> < inverse (real (Suc n)) \<and> dist (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
 | 
637  | 
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
 | 
638  | 
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
 | 
639  | 
and F2: "\<And>n. \<bar>?F n - a\<bar> < inverse (real (Suc n))"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
640  | 
and F3: "\<And>n. dist (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
 | 
641  | 
by fast+  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
642  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
643  | 
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
 | 
644  | 
proof (rule LIMSEQ_I, unfold real_norm_def)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
645  | 
fix e::real  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
646  | 
assume "0 < e"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
647  | 
(* choose no such that inverse (real (Suc n)) < e *)  | 
| 23441 | 648  | 
then have "\<exists>no. inverse (real (Suc no)) < e" by (rule reals_Archimedean)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
649  | 
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
 | 
650  | 
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
 | 
651  | 
proof (intro exI allI impI)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
652  | 
fix n  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
653  | 
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
 | 
654  | 
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
 | 
655  | 
by (rule F2)  | 
| 
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
656  | 
also have "inverse (real (Suc n)) \<le> inverse (real (Suc m))"  | 
| 23441 | 657  | 
using mlen by auto  | 
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
658  | 
also from nodef have  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
659  | 
"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
 | 
660  | 
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
 | 
661  | 
qed  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
662  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
663  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
664  | 
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
 | 
665  | 
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
 | 
666  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
667  | 
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
 | 
668  | 
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
 | 
669  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
670  | 
moreover have "\<not> ((\<lambda>n. X (?F n)) ----> L)"  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
671  | 
proof -  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
672  | 
    {
 | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
673  | 
fix no::nat  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
674  | 
obtain n where "n = no + 1" by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
675  | 
then have nolen: "no \<le> n" by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
676  | 
(* We prove this by showing that for any m there is an n\<ge>m such that |X (?F n) - L| \<ge> r *)  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
677  | 
have "dist (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
 | 
678  | 
by (rule F3)  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
679  | 
with nolen have "\<exists>n. no \<le> n \<and> dist (X (?F n)) L \<ge> r" by fast  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
680  | 
}  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
681  | 
then have "(\<forall>no. \<exists>n. no \<le> n \<and> dist (X (?F n)) L \<ge> r)" by simp  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
682  | 
with rdef have "\<exists>e>0. (\<forall>no. \<exists>n. no \<le> n \<and> dist (X (?F n)) L \<ge> e)" by auto  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
683  | 
thus ?thesis by (unfold LIMSEQ_def, auto simp add: not_less)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
684  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
685  | 
ultimately show False by simp  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
686  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
687  | 
|
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
688  | 
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
 | 
689  | 
"(\<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
 | 
690  | 
(X -- a --> L)"  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
691  | 
proof  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
692  | 
assume "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"  | 
| 23441 | 693  | 
thus "X -- a --> L" by (rule LIMSEQ_SEQ_conv2)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
694  | 
next  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
695  | 
assume "(X -- a --> L)"  | 
| 23441 | 696  | 
thus "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L" by (rule LIMSEQ_SEQ_conv1)  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
697  | 
qed  | 
| 
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
698  | 
|
| 10751 | 699  | 
end  |