| author | wenzelm | 
| Mon, 30 Jul 2012 17:03:24 +0200 | |
| changeset 48609 | 0090fab725e3 | 
| parent 45031 | 9583f2b56f85 | 
| child 50331 | 4b6dc5077e98 | 
| 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  | 
|
| 
36661
 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 
huffman 
parents: 
32650 
diff
changeset
 | 
15  | 
abbreviation  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
16  | 
LIM :: "['a::topological_space \<Rightarrow> 'b::topological_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
 | 
| 
36661
 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 
huffman 
parents: 
32650 
diff
changeset
 | 
18  | 
"f -- a --> L \<equiv> (f ---> L) (at a)"  | 
| 10751 | 19  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
20  | 
definition  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
21  | 
isCont :: "['a::topological_space \<Rightarrow> 'b::topological_space, 'a] \<Rightarrow> bool" where  | 
| 19765 | 22  | 
"isCont f a = (f -- a --> (f a))"  | 
| 10751 | 23  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21399 
diff
changeset
 | 
24  | 
definition  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
25  | 
isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where  | 
| 37767 | 26  | 
"isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"  | 
| 10751 | 27  | 
|
| 31392 | 28  | 
subsection {* Limits of Functions *}
 | 
| 
31349
 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 
huffman 
parents: 
31338 
diff
changeset
 | 
29  | 
|
| 
36661
 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 
huffman 
parents: 
32650 
diff
changeset
 | 
30  | 
lemma LIM_def: "f -- a --> L =  | 
| 
 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 
huffman 
parents: 
32650 
diff
changeset
 | 
31  | 
(\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s  | 
| 
 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 
huffman 
parents: 
32650 
diff
changeset
 | 
32  | 
--> dist (f x) L < r)"  | 
| 
 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 
huffman 
parents: 
32650 
diff
changeset
 | 
33  | 
unfolding 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:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
64  | 
fixes a :: "'a::real_normed_vector"  | 
| 
31338
 
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"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
66  | 
apply (rule topological_tendstoI)  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
67  | 
apply (drule (2) topological_tendstoD)  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
68  | 
apply (simp only: eventually_at dist_norm)  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
69  | 
apply (clarify, rule_tac x=d in exI, safe)  | 
| 20756 | 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:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
75  | 
fixes a :: "'a::real_normed_vector"  | 
| 
31338
 
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:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
80  | 
fixes a :: "'a::real_normed_vector"  | 
| 
31338
 
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  | 
||
| 32650 | 84  | 
lemma LIM_cong_limit: "\<lbrakk> f -- x --> L ; K = L \<rbrakk> \<Longrightarrow> f -- x --> K" by simp  | 
85  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
86  | 
lemma LIM_zero:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
87  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 44570 | 88  | 
shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
89  | 
unfolding tendsto_iff dist_norm by simp  | 
| 21239 | 90  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
91  | 
lemma LIM_zero_cancel:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
92  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 44570 | 93  | 
shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
94  | 
unfolding tendsto_iff dist_norm by simp  | 
| 21239 | 95  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
96  | 
lemma LIM_zero_iff:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
97  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 44570 | 98  | 
shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
99  | 
unfolding tendsto_iff dist_norm by simp  | 
| 21399 | 100  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
101  | 
lemma metric_LIM_imp_LIM:  | 
| 21257 | 102  | 
assumes f: "f -- a --> l"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
103  | 
assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"  | 
| 21257 | 104  | 
shows "g -- a --> m"  | 
| 44251 | 105  | 
by (rule metric_tendsto_imp_tendsto [OF f],  | 
106  | 
auto simp add: eventually_at_topological le)  | 
|
| 21257 | 107  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
108  | 
lemma LIM_imp_LIM:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
109  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
110  | 
fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
111  | 
assumes f: "f -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
112  | 
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
 | 
113  | 
shows "g -- a --> m"  | 
| 44251 | 114  | 
by (rule metric_LIM_imp_LIM [OF f],  | 
115  | 
simp add: dist_norm le)  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
116  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
117  | 
lemma LIM_const_not_eq:  | 
| 44571 | 118  | 
fixes a :: "'a::perfect_space"  | 
| 
44205
 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 
huffman 
parents: 
44194 
diff
changeset
 | 
119  | 
fixes k L :: "'b::t2_space"  | 
| 
23076
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
120  | 
shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) -- a --> L"  | 
| 44571 | 121  | 
by (simp add: tendsto_const_iff)  | 
| 14477 | 122  | 
|
| 21786 | 123  | 
lemmas LIM_not_zero = LIM_const_not_eq [where L = 0]  | 
124  | 
||
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
125  | 
lemma LIM_const_eq:  | 
| 44571 | 126  | 
fixes a :: "'a::perfect_space"  | 
| 
44205
 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 
huffman 
parents: 
44194 
diff
changeset
 | 
127  | 
fixes k L :: "'b::t2_space"  | 
| 
23076
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
128  | 
shows "(\<lambda>x. k) -- a --> L \<Longrightarrow> k = L"  | 
| 44571 | 129  | 
by (simp add: tendsto_const_iff)  | 
| 14477 | 130  | 
|
| 
20561
 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 
huffman 
parents: 
20552 
diff
changeset
 | 
131  | 
lemma LIM_unique:  | 
| 44571 | 132  | 
fixes a :: "'a::perfect_space"  | 
| 
44205
 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 
huffman 
parents: 
44194 
diff
changeset
 | 
133  | 
fixes L M :: "'b::t2_space"  | 
| 
23076
 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 
huffman 
parents: 
23069 
diff
changeset
 | 
134  | 
shows "\<lbrakk>f -- a --> L; f -- a --> M\<rbrakk> \<Longrightarrow> L = M"  | 
| 44571 | 135  | 
using at_neq_bot by (rule tendsto_unique)  | 
| 14477 | 136  | 
|
137  | 
text{*Limits are equal for functions equal except at limit point*}
 | 
|
138  | 
lemma LIM_equal:  | 
|
139  | 
"[| \<forall>x. x \<noteq> a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)"  | 
|
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
140  | 
unfolding tendsto_def eventually_at_topological by simp  | 
| 14477 | 141  | 
|
| 20796 | 142  | 
lemma LIM_cong:  | 
143  | 
"\<lbrakk>a = b; \<And>x. x \<noteq> b \<Longrightarrow> f x = g x; l = m\<rbrakk>  | 
|
| 21399 | 144  | 
\<Longrightarrow> ((\<lambda>x. f x) -- a --> l) = ((\<lambda>x. g x) -- b --> m)"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
145  | 
by (simp add: LIM_equal)  | 
| 20796 | 146  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
147  | 
lemma metric_LIM_equal2:  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
148  | 
assumes 1: "0 < R"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
149  | 
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
 | 
150  | 
shows "g -- a --> l \<Longrightarrow> f -- a --> l"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
151  | 
apply (rule topological_tendstoI)  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
152  | 
apply (drule (2) topological_tendstoD)  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
153  | 
apply (simp add: eventually_at, safe)  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
154  | 
apply (rule_tac x="min d R" in exI, safe)  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
155  | 
apply (simp add: 1)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
156  | 
apply (simp add: 2)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
157  | 
done  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
158  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
159  | 
lemma LIM_equal2:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
160  | 
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
161  | 
assumes 1: "0 < R"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
162  | 
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
 | 
163  | 
shows "g -- a --> l \<Longrightarrow> f -- a --> l"  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
164  | 
by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
165  | 
|
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
166  | 
lemma LIM_compose_eventually:  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
167  | 
assumes f: "f -- a --> b"  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
168  | 
assumes g: "g -- b --> c"  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
169  | 
assumes inj: "eventually (\<lambda>x. f x \<noteq> b) (at a)"  | 
| 
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
170  | 
shows "(\<lambda>x. g (f x)) -- a --> c"  | 
| 
44253
 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 
huffman 
parents: 
44251 
diff
changeset
 | 
171  | 
using g f inj by (rule tendsto_compose_eventually)  | 
| 21239 | 172  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
173  | 
lemma metric_LIM_compose2:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
174  | 
assumes f: "f -- a --> b"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
175  | 
assumes g: "g -- b --> c"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
176  | 
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
 | 
177  | 
shows "(\<lambda>x. g (f x)) -- a --> c"  | 
| 44314 | 178  | 
using g f inj [folded eventually_at]  | 
179  | 
by (rule tendsto_compose_eventually)  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
180  | 
|
| 23040 | 181  | 
lemma LIM_compose2:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
182  | 
fixes a :: "'a::real_normed_vector"  | 
| 23040 | 183  | 
assumes f: "f -- a --> b"  | 
184  | 
assumes g: "g -- b --> c"  | 
|
185  | 
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"  | 
|
186  | 
shows "(\<lambda>x. g (f x)) -- a --> c"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
187  | 
by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])  | 
| 23040 | 188  | 
|
| 21239 | 189  | 
lemma LIM_o: "\<lbrakk>g -- l --> g l; f -- a --> l\<rbrakk> \<Longrightarrow> (g \<circ> f) -- a --> g l"  | 
| 44314 | 190  | 
unfolding o_def by (rule tendsto_compose)  | 
| 21239 | 191  | 
|
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
192  | 
lemma real_LIM_sandwich_zero:  | 
| 
36662
 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 
huffman 
parents: 
36661 
diff
changeset
 | 
193  | 
fixes f g :: "'a::topological_space \<Rightarrow> real"  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
194  | 
assumes f: "f -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
195  | 
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
 | 
196  | 
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
 | 
197  | 
shows "g -- a --> 0"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
198  | 
proof (rule LIM_imp_LIM [OF f])  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
199  | 
fix x assume x: "x \<noteq> a"  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
200  | 
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
 | 
201  | 
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
 | 
202  | 
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
 | 
203  | 
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
 | 
204  | 
finally show "norm (g x - 0) \<le> norm (f x - 0)" .  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
205  | 
qed  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
206  | 
|
| 14477 | 207  | 
|
| 20755 | 208  | 
subsection {* Continuity *}
 | 
| 14477 | 209  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
210  | 
lemma LIM_isCont_iff:  | 
| 36665 | 211  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
212  | 
shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"  | 
| 21239 | 213  | 
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])  | 
214  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
215  | 
lemma isCont_iff:  | 
| 36665 | 216  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
217  | 
shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"  | 
| 21239 | 218  | 
by (simp add: isCont_def LIM_isCont_iff)  | 
219  | 
||
| 
23069
 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 
huffman 
parents: 
23040 
diff
changeset
 | 
220  | 
lemma isCont_ident [simp]: "isCont (\<lambda>x. x) a"  | 
| 44314 | 221  | 
unfolding isCont_def by (rule tendsto_ident_at)  | 
| 21239 | 222  | 
|
| 21786 | 223  | 
lemma isCont_const [simp]: "isCont (\<lambda>x. k) a"  | 
| 44314 | 224  | 
unfolding isCont_def by (rule tendsto_const)  | 
| 21239 | 225  | 
|
| 44233 | 226  | 
lemma isCont_norm [simp]:  | 
| 36665 | 227  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
228  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"  | 
| 44314 | 229  | 
unfolding isCont_def by (rule tendsto_norm)  | 
| 21786 | 230  | 
|
| 44233 | 231  | 
lemma isCont_rabs [simp]:  | 
232  | 
fixes f :: "'a::topological_space \<Rightarrow> real"  | 
|
233  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"  | 
|
| 44314 | 234  | 
unfolding isCont_def by (rule tendsto_rabs)  | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
235  | 
|
| 44233 | 236  | 
lemma isCont_add [simp]:  | 
| 36665 | 237  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
238  | 
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"  | 
| 44314 | 239  | 
unfolding isCont_def by (rule tendsto_add)  | 
| 21239 | 240  | 
|
| 44233 | 241  | 
lemma isCont_minus [simp]:  | 
| 36665 | 242  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
243  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"  | 
| 44314 | 244  | 
unfolding isCont_def by (rule tendsto_minus)  | 
| 21239 | 245  | 
|
| 44233 | 246  | 
lemma isCont_diff [simp]:  | 
| 36665 | 247  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
248  | 
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"  | 
| 44314 | 249  | 
unfolding isCont_def by (rule tendsto_diff)  | 
| 21239 | 250  | 
|
| 44233 | 251  | 
lemma isCont_mult [simp]:  | 
| 36665 | 252  | 
fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra"  | 
| 21786 | 253  | 
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"  | 
| 44314 | 254  | 
unfolding isCont_def by (rule tendsto_mult)  | 
| 21239 | 255  | 
|
| 44233 | 256  | 
lemma isCont_inverse [simp]:  | 
| 36665 | 257  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"  | 
| 21786 | 258  | 
shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. inverse (f x)) a"  | 
| 44314 | 259  | 
unfolding isCont_def by (rule tendsto_inverse)  | 
| 21239 | 260  | 
|
| 44233 | 261  | 
lemma isCont_divide [simp]:  | 
262  | 
fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_field"  | 
|
263  | 
shows "\<lbrakk>isCont f a; isCont g a; g a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x / g x) a"  | 
|
264  | 
unfolding isCont_def by (rule tendsto_divide)  | 
|
265  | 
||
| 44310 | 266  | 
lemma isCont_tendsto_compose:  | 
267  | 
"\<lbrakk>isCont g l; (f ---> l) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. g (f x)) ---> g l) F"  | 
|
268  | 
unfolding isCont_def by (rule tendsto_compose)  | 
|
269  | 
||
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
270  | 
lemma metric_isCont_LIM_compose2:  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
271  | 
assumes f [unfolded isCont_def]: "isCont f a"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
272  | 
assumes g: "g -- f a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
273  | 
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
 | 
274  | 
shows "(\<lambda>x. g (f x)) -- a --> l"  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
275  | 
by (rule metric_LIM_compose2 [OF f g inj])  | 
| 
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
276  | 
|
| 23040 | 277  | 
lemma isCont_LIM_compose2:  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
278  | 
fixes a :: "'a::real_normed_vector"  | 
| 23040 | 279  | 
assumes f [unfolded isCont_def]: "isCont f a"  | 
280  | 
assumes g: "g -- f a --> l"  | 
|
281  | 
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"  | 
|
282  | 
shows "(\<lambda>x. g (f x)) -- a --> l"  | 
|
283  | 
by (rule LIM_compose2 [OF f g inj])  | 
|
284  | 
||
| 21239 | 285  | 
lemma isCont_o2: "\<lbrakk>isCont f a; isCont g (f a)\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. g (f x)) a"  | 
| 44314 | 286  | 
unfolding isCont_def by (rule tendsto_compose)  | 
| 21239 | 287  | 
|
288  | 
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
 | 
289  | 
unfolding o_def by (rule isCont_o2)  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
290  | 
|
| 44233 | 291  | 
lemma (in bounded_linear) isCont:  | 
292  | 
"isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"  | 
|
| 44314 | 293  | 
unfolding isCont_def by (rule tendsto)  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
294  | 
|
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
295  | 
lemma (in bounded_bilinear) isCont:  | 
| 
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
296  | 
"\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"  | 
| 44314 | 297  | 
unfolding isCont_def by (rule tendsto)  | 
| 
21282
 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 
huffman 
parents: 
21257 
diff
changeset
 | 
298  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44254 
diff
changeset
 | 
299  | 
lemmas isCont_scaleR [simp] =  | 
| 
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44254 
diff
changeset
 | 
300  | 
bounded_bilinear.isCont [OF bounded_bilinear_scaleR]  | 
| 21239 | 301  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44254 
diff
changeset
 | 
302  | 
lemmas isCont_of_real [simp] =  | 
| 
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44254 
diff
changeset
 | 
303  | 
bounded_linear.isCont [OF bounded_linear_of_real]  | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
304  | 
|
| 44233 | 305  | 
lemma isCont_power [simp]:  | 
| 36665 | 306  | 
  fixes f :: "'a::topological_space \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
307  | 
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"  | 
| 44314 | 308  | 
unfolding isCont_def by (rule tendsto_power)  | 
| 
22627
 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 
huffman 
parents: 
22613 
diff
changeset
 | 
309  | 
|
| 44233 | 310  | 
lemma isCont_sgn [simp]:  | 
| 36665 | 311  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
312  | 
shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. sgn (f x)) a"  | 
| 44314 | 313  | 
unfolding isCont_def by (rule tendsto_sgn)  | 
| 29885 | 314  | 
|
| 44233 | 315  | 
lemma isCont_setsum [simp]:  | 
316  | 
fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::real_normed_vector"  | 
|
317  | 
fixes A :: "'a set"  | 
|
318  | 
shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"  | 
|
319  | 
unfolding isCont_def by (simp add: tendsto_setsum)  | 
|
| 15228 | 320  | 
|
| 44233 | 321  | 
lemmas isCont_intros =  | 
322  | 
isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus  | 
|
323  | 
isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR  | 
|
324  | 
isCont_of_real isCont_power isCont_sgn isCont_setsum  | 
|
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
325  | 
|
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
326  | 
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
 | 
327  | 
  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
 | 
328  | 
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
 | 
329  | 
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
 | 
330  | 
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
 | 
331  | 
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
 | 
332  | 
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
 | 
333  | 
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
 | 
334  | 
|
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
335  | 
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
 | 
336  | 
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
 | 
337  | 
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
 | 
338  | 
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
 | 
339  | 
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
 | 
340  | 
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
 | 
341  | 
next  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
342  | 
case False hence "s / 2 \<ge> (x - b) / 2" by auto  | 
| 
32642
 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 
haftmann 
parents: 
32436 
diff
changeset
 | 
343  | 
hence "?x = (x + b) / 2" by (simp add: field_simps min_max.inf_absorb2)  | 
| 
29803
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
344  | 
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
 | 
345  | 
qed  | 
| 
 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
346  | 
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
 | 
347  | 
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
 | 
348  | 
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
 | 
349  | 
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
 | 
350  | 
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
 | 
351  | 
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
 | 
352  | 
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
 | 
353  | 
qed  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
354  | 
|
| 14477 | 355  | 
|
| 20755 | 356  | 
subsection {* Uniform Continuity *}
 | 
357  | 
||
| 14477 | 358  | 
lemma isUCont_isCont: "isUCont f ==> isCont f x"  | 
| 23012 | 359  | 
by (simp add: isUCont_def isCont_def LIM_def, force)  | 
| 14477 | 360  | 
|
| 23118 | 361  | 
lemma isUCont_Cauchy:  | 
362  | 
"\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"  | 
|
363  | 
unfolding isUCont_def  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
364  | 
apply (rule metric_CauchyI)  | 
| 23118 | 365  | 
apply (drule_tac x=e in spec, safe)  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
366  | 
apply (drule_tac e=s in metric_CauchyD, safe)  | 
| 23118 | 367  | 
apply (rule_tac x=M in exI, simp)  | 
368  | 
done  | 
|
369  | 
||
370  | 
lemma (in bounded_linear) isUCont: "isUCont f"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
371  | 
unfolding isUCont_def dist_norm  | 
| 23118 | 372  | 
proof (intro allI impI)  | 
373  | 
fix r::real assume r: "0 < r"  | 
|
374  | 
obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"  | 
|
375  | 
using pos_bounded by fast  | 
|
376  | 
show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"  | 
|
377  | 
proof (rule exI, safe)  | 
|
378  | 
from r K show "0 < r / K" by (rule divide_pos_pos)  | 
|
379  | 
next  | 
|
380  | 
fix x y :: 'a  | 
|
381  | 
assume xy: "norm (x - y) < r / K"  | 
|
382  | 
have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)  | 
|
383  | 
also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)  | 
|
384  | 
also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)  | 
|
385  | 
finally show "norm (f x - f y) < r" .  | 
|
386  | 
qed  | 
|
387  | 
qed  | 
|
388  | 
||
389  | 
lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"  | 
|
390  | 
by (rule isUCont [THEN isUCont_Cauchy])  | 
|
391  | 
||
| 14477 | 392  | 
|
| 
21165
 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 
huffman 
parents: 
21141 
diff
changeset
 | 
393  | 
subsection {* Relation of LIM and LIMSEQ *}
 | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
394  | 
|
| 44532 | 395  | 
lemma sequentially_imp_eventually_within:  | 
396  | 
fixes a :: "'a::metric_space"  | 
|
397  | 
assumes "\<forall>f. (\<forall>n. f n \<in> s \<and> f n \<noteq> a) \<and> f ----> a \<longrightarrow>  | 
|
398  | 
eventually (\<lambda>n. P (f n)) sequentially"  | 
|
399  | 
shows "eventually P (at a within s)"  | 
|
400  | 
proof (rule ccontr)  | 
|
401  | 
let ?I = "\<lambda>n. inverse (real (Suc n))"  | 
|
402  | 
def F \<equiv> "\<lambda>n::nat. SOME x. x \<in> s \<and> x \<noteq> a \<and> dist x a < ?I n \<and> \<not> P x"  | 
|
403  | 
assume "\<not> eventually P (at a within s)"  | 
|
404  | 
hence P: "\<forall>d>0. \<exists>x. x \<in> s \<and> x \<noteq> a \<and> dist x a < d \<and> \<not> P x"  | 
|
405  | 
unfolding Limits.eventually_within Limits.eventually_at by fast  | 
|
406  | 
hence "\<And>n. \<exists>x. x \<in> s \<and> x \<noteq> a \<and> dist x a < ?I n \<and> \<not> P x" by simp  | 
|
407  | 
hence F: "\<And>n. F n \<in> s \<and> F n \<noteq> a \<and> dist (F n) a < ?I n \<and> \<not> P (F n)"  | 
|
408  | 
unfolding F_def by (rule someI_ex)  | 
|
409  | 
hence F0: "\<forall>n. F n \<in> s" and F1: "\<forall>n. F n \<noteq> a"  | 
|
410  | 
and F2: "\<forall>n. dist (F n) a < ?I n" and F3: "\<forall>n. \<not> P (F n)"  | 
|
411  | 
by fast+  | 
|
412  | 
from LIMSEQ_inverse_real_of_nat have "F ----> a"  | 
|
413  | 
by (rule metric_tendsto_imp_tendsto,  | 
|
414  | 
simp add: dist_norm F2 less_imp_le)  | 
|
415  | 
hence "eventually (\<lambda>n. P (F n)) sequentially"  | 
|
416  | 
using assms F0 F1 by simp  | 
|
417  | 
thus "False" by (simp add: F3)  | 
|
418  | 
qed  | 
|
419  | 
||
420  | 
lemma sequentially_imp_eventually_at:  | 
|
421  | 
fixes a :: "'a::metric_space"  | 
|
422  | 
assumes "\<forall>f. (\<forall>n. f n \<noteq> a) \<and> f ----> a \<longrightarrow>  | 
|
423  | 
eventually (\<lambda>n. P (f n)) sequentially"  | 
|
424  | 
shows "eventually P (at a)"  | 
|
| 45031 | 425  | 
using assms sequentially_imp_eventually_within [where s=UNIV] by simp  | 
| 44532 | 426  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
427  | 
lemma LIMSEQ_SEQ_conv1:  | 
| 
44254
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
428  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
429  | 
assumes f: "f -- a --> l"  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
430  | 
shows "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l"  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
431  | 
using tendsto_compose_eventually [OF f, where F=sequentially] by simp  | 
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
432  | 
|
| 
44254
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
433  | 
lemma LIMSEQ_SEQ_conv2:  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
434  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
435  | 
assumes "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> (\<lambda>n. f (S n)) ----> l"  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
436  | 
shows "f -- a --> l"  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
437  | 
using assms unfolding tendsto_def [where l=l]  | 
| 44532 | 438  | 
by (simp add: sequentially_imp_eventually_at)  | 
| 
44254
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
439  | 
|
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
440  | 
lemma LIMSEQ_SEQ_conv:  | 
| 
44254
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
441  | 
"(\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> (a::'a::metric_space) \<longrightarrow> (\<lambda>n. X (S n)) ----> L) =  | 
| 
 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 
huffman 
parents: 
44253 
diff
changeset
 | 
442  | 
(X -- a --> (L::'b::topological_space))"  | 
| 
44253
 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 
huffman 
parents: 
44251 
diff
changeset
 | 
443  | 
using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 ..  | 
| 
19023
 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 
kleing 
parents: 
17318 
diff
changeset
 | 
444  | 
|
| 10751 | 445  | 
end  |