| author | huffman | 
| Wed, 24 Aug 2011 11:56:57 -0700 | |
| changeset 44514 | d02b01e5ab8f | 
| parent 44314 | dbad46932536 | 
| child 44532 | a2e9b39df938 | 
| 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: 
21141diff
changeset | 7 | header{* Limits and Continuity *}
 | 
| 10751 | 8 | |
| 15131 | 9 | theory Lim | 
| 29197 
6d4cb27ed19c
adapted HOL source structure to distribution layout
 haftmann parents: 
28952diff
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: 
22637diff
changeset | 13 | text{*Standard Definitions*}
 | 
| 10751 | 14 | |
| 36661 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
changeset | 15 | abbreviation | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
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: 
21399diff
changeset | 17 |         ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where
 | 
| 36661 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
changeset | 18 | "f -- a --> L \<equiv> (f ---> L) (at a)" | 
| 10751 | 19 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21399diff
changeset | 20 | definition | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
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: 
21399diff
changeset | 24 | definition | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31338diff
changeset | 29 | |
| 36661 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
changeset | 30 | lemma LIM_def: "f -- a --> L = | 
| 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
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: 
32650diff
changeset | 32 | --> dist (f x) L < r)" | 
| 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
changeset | 33 | unfolding tendsto_iff eventually_at .. | 
| 10751 | 34 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 35 | lemma metric_LIM_I: | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
changeset | 37 | \<Longrightarrow> f -- a --> L" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 38 | by (simp add: LIM_def) | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 39 | |
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 40 | lemma metric_LIM_D: | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 41 | "\<lbrakk>f -- a --> L; 0 < r\<rbrakk> | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
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: 
31336diff
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: 
31336diff
changeset | 47 | shows "f -- a --> L = | 
| 20561 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 huffman parents: 
20552diff
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: 
31336diff
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: 
20432diff
changeset | 51 | lemma LIM_I: | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
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: 
20432diff
changeset | 54 | ==> f -- a --> L" | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
20432diff
changeset | 55 | by (simp add: LIM_eq) | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
20432diff
changeset | 56 | |
| 14477 | 57 | lemma LIM_D: | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
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: 
20552diff
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: 
31336diff
changeset | 63 | lemma LIM_offset: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 64 | fixes a :: "'a::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
36661diff
changeset | 66 | apply (rule topological_tendstoI) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 67 | apply (drule (2) topological_tendstoD) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 68 | apply (simp only: eventually_at dist_norm) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
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: 
31336diff
changeset | 74 | lemma LIM_offset_zero: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 75 | fixes a :: "'a::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
changeset | 79 | lemma LIM_offset_zero_cancel: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 80 | fixes a :: "'a::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
changeset | 86 | lemma LIM_zero: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 87 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 88 | shows "f -- a --> l \<Longrightarrow> (\<lambda>x. f x - l) -- a --> 0" | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 89 | unfolding tendsto_iff dist_norm by simp | 
| 21239 | 90 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 91 | lemma LIM_zero_cancel: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 92 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 93 | shows "(\<lambda>x. f x - l) -- a --> 0 \<Longrightarrow> f -- a --> l" | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 94 | unfolding tendsto_iff dist_norm by simp | 
| 21239 | 95 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 96 | lemma LIM_zero_iff: | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 97 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 98 | shows "(\<lambda>x. f x - l) -- a --> 0 = f -- a --> l" | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 99 | unfolding tendsto_iff dist_norm by simp | 
| 21399 | 100 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
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: 
31336diff
changeset | 108 | lemma LIM_imp_LIM: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 109 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
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: 
31336diff
changeset | 111 | assumes f: "f -- a --> l" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
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: 
31336diff
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: 
31336diff
changeset | 116 | |
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 117 | lemma trivial_limit_at: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 118 | fixes a :: "'a::real_normed_algebra_1" | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 119 |   shows "\<not> trivial_limit (at a)"  -- {* TODO: find a more appropriate class *}
 | 
| 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 120 | unfolding trivial_limit_def | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 121 | unfolding eventually_at dist_norm | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 122 | by (clarsimp, rule_tac x="a + of_real (d/2)" in exI, simp) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 123 | |
| 20561 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 huffman parents: 
20552diff
changeset | 124 | lemma LIM_const_not_eq: | 
| 23076 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 huffman parents: 
23069diff
changeset | 125 | fixes a :: "'a::real_normed_algebra_1" | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 126 | fixes k L :: "'b::t2_space" | 
| 23076 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 huffman parents: 
23069diff
changeset | 127 | shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) -- a --> L" | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 128 | by (simp add: tendsto_const_iff trivial_limit_at) | 
| 14477 | 129 | |
| 21786 | 130 | lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] | 
| 131 | ||
| 20561 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 huffman parents: 
20552diff
changeset | 132 | lemma LIM_const_eq: | 
| 23076 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 huffman parents: 
23069diff
changeset | 133 | fixes a :: "'a::real_normed_algebra_1" | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 134 | fixes k L :: "'b::t2_space" | 
| 23076 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 huffman parents: 
23069diff
changeset | 135 | shows "(\<lambda>x. k) -- a --> L \<Longrightarrow> k = L" | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 136 | by (simp add: tendsto_const_iff trivial_limit_at) | 
| 14477 | 137 | |
| 20561 
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
 huffman parents: 
20552diff
changeset | 138 | lemma LIM_unique: | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 139 |   fixes a :: "'a::real_normed_algebra_1" -- {* TODO: find a more appropriate class *}
 | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 140 | fixes L M :: "'b::t2_space" | 
| 23076 
1b2acb3ccb29
generalize uniqueness of limits to class real_normed_algebra_1
 huffman parents: 
23069diff
changeset | 141 | shows "\<lbrakk>f -- a --> L; f -- a --> M\<rbrakk> \<Longrightarrow> L = M" | 
| 44205 
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
 huffman parents: 
44194diff
changeset | 142 | using trivial_limit_at by (rule tendsto_unique) | 
| 14477 | 143 | |
| 144 | text{*Limits are equal for functions equal except at limit point*}
 | |
| 145 | lemma LIM_equal: | |
| 146 | "[| \<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: 
36661diff
changeset | 147 | unfolding tendsto_def eventually_at_topological by simp | 
| 14477 | 148 | |
| 20796 | 149 | lemma LIM_cong: | 
| 150 | "\<lbrakk>a = b; \<And>x. x \<noteq> b \<Longrightarrow> f x = g x; l = m\<rbrakk> | |
| 21399 | 151 | \<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: 
36661diff
changeset | 152 | by (simp add: LIM_equal) | 
| 20796 | 153 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 154 | lemma metric_LIM_equal2: | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 155 | assumes 1: "0 < R" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 156 | 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: 
21257diff
changeset | 157 | shows "g -- a --> l \<Longrightarrow> f -- a --> l" | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 158 | apply (rule topological_tendstoI) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 159 | apply (drule (2) topological_tendstoD) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 160 | apply (simp add: eventually_at, safe) | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 161 | apply (rule_tac x="min d R" in exI, safe) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 162 | apply (simp add: 1) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 163 | apply (simp add: 2) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 164 | done | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 165 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 166 | lemma LIM_equal2: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 167 | fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 168 | assumes 1: "0 < R" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 169 | 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: 
31336diff
changeset | 170 | shows "g -- a --> l \<Longrightarrow> f -- a --> l" | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 171 | 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: 
31336diff
changeset | 172 | |
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 173 | lemma LIM_compose_eventually: | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 174 | assumes f: "f -- a --> b" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 175 | assumes g: "g -- b --> c" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 176 | assumes inj: "eventually (\<lambda>x. f x \<noteq> b) (at a)" | 
| 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 177 | shows "(\<lambda>x. g (f x)) -- a --> c" | 
| 44253 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 178 | using g f inj by (rule tendsto_compose_eventually) | 
| 21239 | 179 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 180 | lemma metric_LIM_compose2: | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 181 | assumes f: "f -- a --> b" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 182 | assumes g: "g -- b --> c" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 183 | 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: 
31336diff
changeset | 184 | shows "(\<lambda>x. g (f x)) -- a --> c" | 
| 44314 | 185 | using g f inj [folded eventually_at] | 
| 186 | by (rule tendsto_compose_eventually) | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 187 | |
| 23040 | 188 | lemma LIM_compose2: | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 189 | fixes a :: "'a::real_normed_vector" | 
| 23040 | 190 | assumes f: "f -- a --> b" | 
| 191 | assumes g: "g -- b --> c" | |
| 192 | assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b" | |
| 193 | shows "(\<lambda>x. g (f x)) -- a --> c" | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 194 | by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) | 
| 23040 | 195 | |
| 21239 | 196 | lemma LIM_o: "\<lbrakk>g -- l --> g l; f -- a --> l\<rbrakk> \<Longrightarrow> (g \<circ> f) -- a --> g l" | 
| 44314 | 197 | unfolding o_def by (rule tendsto_compose) | 
| 21239 | 198 | |
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 199 | lemma real_LIM_sandwich_zero: | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36661diff
changeset | 200 | fixes f g :: "'a::topological_space \<Rightarrow> real" | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 201 | assumes f: "f -- a --> 0" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 202 | assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 203 | assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 204 | shows "g -- a --> 0" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 205 | proof (rule LIM_imp_LIM [OF f]) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 206 | fix x assume x: "x \<noteq> a" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 207 | have "norm (g x - 0) = g x" by (simp add: 1 x) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 208 | also have "g x \<le> f x" by (rule 2 [OF x]) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 209 | also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 210 | also have "\<bar>f x\<bar> = norm (f x - 0)" by simp | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 211 | finally show "norm (g x - 0) \<le> norm (f x - 0)" . | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 212 | qed | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 213 | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21810diff
changeset | 214 | text {* Bounded Linear Operators *}
 | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 215 | |
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 216 | lemma (in bounded_linear) LIM: | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 217 | "g -- a --> l \<Longrightarrow> (\<lambda>x. f (g x)) -- a --> f l" | 
| 36661 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
changeset | 218 | by (rule tendsto) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: 
31338diff
changeset | 219 | |
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 220 | lemma (in bounded_linear) LIM_zero: | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 221 | "g -- a --> 0 \<Longrightarrow> (\<lambda>x. f (g x)) -- a --> 0" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
41550diff
changeset | 222 | by (rule tendsto_zero) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 223 | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21810diff
changeset | 224 | text {* Bounded Bilinear Operators *}
 | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 225 | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: 
31338diff
changeset | 226 | lemma (in bounded_bilinear) LIM: | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: 
31338diff
changeset | 227 | "\<lbrakk>f -- a --> L; g -- a --> M\<rbrakk> \<Longrightarrow> (\<lambda>x. f x ** g x) -- a --> L ** M" | 
| 36661 
0a5b7b818d65
make (f -- a --> x) an abbreviation for (f ---> x) (at a)
 huffman parents: 
32650diff
changeset | 228 | by (rule tendsto) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: 
31338diff
changeset | 229 | |
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 230 | lemma (in bounded_bilinear) LIM_prod_zero: | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 231 | fixes a :: "'d::metric_space" | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 232 | assumes f: "f -- a --> 0" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 233 | assumes g: "g -- a --> 0" | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 234 | shows "(\<lambda>x. f x ** g x) -- a --> 0" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
41550diff
changeset | 235 | using f g by (rule tendsto_zero) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 236 | |
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 237 | lemma (in bounded_bilinear) LIM_left_zero: | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 238 | "f -- a --> 0 \<Longrightarrow> (\<lambda>x. f x ** c) -- a --> 0" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
41550diff
changeset | 239 | by (rule tendsto_left_zero) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 240 | |
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 241 | lemma (in bounded_bilinear) LIM_right_zero: | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 242 | "f -- a --> 0 \<Longrightarrow> (\<lambda>x. c ** f x) -- a --> 0" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
41550diff
changeset | 243 | by (rule tendsto_right_zero) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 244 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 245 | lemmas LIM_mult_zero = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 246 | bounded_bilinear.LIM_prod_zero [OF bounded_bilinear_mult] | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 247 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 248 | lemmas LIM_mult_left_zero = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 249 | bounded_bilinear.LIM_left_zero [OF bounded_bilinear_mult] | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 250 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 251 | lemmas LIM_mult_right_zero = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 252 | bounded_bilinear.LIM_right_zero [OF bounded_bilinear_mult] | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 253 | |
| 22637 | 254 | lemma LIM_inverse_fun: | 
| 255 | assumes a: "a \<noteq> (0::'a::real_normed_div_algebra)" | |
| 256 | shows "inverse -- a --> inverse a" | |
| 44314 | 257 | by (rule tendsto_inverse [OF tendsto_ident_at a]) | 
| 29885 | 258 | |
| 14477 | 259 | |
| 20755 | 260 | subsection {* Continuity *}
 | 
| 14477 | 261 | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 262 | lemma LIM_isCont_iff: | 
| 36665 | 263 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 264 | shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)" | 
| 21239 | 265 | by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) | 
| 266 | ||
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 267 | lemma isCont_iff: | 
| 36665 | 268 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 269 | shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x" | 
| 21239 | 270 | by (simp add: isCont_def LIM_isCont_iff) | 
| 271 | ||
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
23040diff
changeset | 272 | lemma isCont_ident [simp]: "isCont (\<lambda>x. x) a" | 
| 44314 | 273 | unfolding isCont_def by (rule tendsto_ident_at) | 
| 21239 | 274 | |
| 21786 | 275 | lemma isCont_const [simp]: "isCont (\<lambda>x. k) a" | 
| 44314 | 276 | unfolding isCont_def by (rule tendsto_const) | 
| 21239 | 277 | |
| 44233 | 278 | lemma isCont_norm [simp]: | 
| 36665 | 279 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 280 | shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a" | 
| 44314 | 281 | unfolding isCont_def by (rule tendsto_norm) | 
| 21786 | 282 | |
| 44233 | 283 | lemma isCont_rabs [simp]: | 
| 284 | fixes f :: "'a::topological_space \<Rightarrow> real" | |
| 285 | shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a" | |
| 44314 | 286 | unfolding isCont_def by (rule tendsto_rabs) | 
| 22627 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 huffman parents: 
22613diff
changeset | 287 | |
| 44233 | 288 | lemma isCont_add [simp]: | 
| 36665 | 289 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 290 | shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a" | 
| 44314 | 291 | unfolding isCont_def by (rule tendsto_add) | 
| 21239 | 292 | |
| 44233 | 293 | lemma isCont_minus [simp]: | 
| 36665 | 294 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 295 | shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a" | 
| 44314 | 296 | unfolding isCont_def by (rule tendsto_minus) | 
| 21239 | 297 | |
| 44233 | 298 | lemma isCont_diff [simp]: | 
| 36665 | 299 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 300 | shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a" | 
| 44314 | 301 | unfolding isCont_def by (rule tendsto_diff) | 
| 21239 | 302 | |
| 44233 | 303 | lemma isCont_mult [simp]: | 
| 36665 | 304 | fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra" | 
| 21786 | 305 | shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a" | 
| 44314 | 306 | unfolding isCont_def by (rule tendsto_mult) | 
| 21239 | 307 | |
| 44233 | 308 | lemma isCont_inverse [simp]: | 
| 36665 | 309 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra" | 
| 21786 | 310 | shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. inverse (f x)) a" | 
| 44314 | 311 | unfolding isCont_def by (rule tendsto_inverse) | 
| 21239 | 312 | |
| 44233 | 313 | lemma isCont_divide [simp]: | 
| 314 | fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_field" | |
| 315 | shows "\<lbrakk>isCont f a; isCont g a; g a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x / g x) a" | |
| 316 | unfolding isCont_def by (rule tendsto_divide) | |
| 317 | ||
| 44310 | 318 | lemma isCont_tendsto_compose: | 
| 319 | "\<lbrakk>isCont g l; (f ---> l) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. g (f x)) ---> g l) F" | |
| 320 | unfolding isCont_def by (rule tendsto_compose) | |
| 321 | ||
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 322 | lemma metric_isCont_LIM_compose2: | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 323 | assumes f [unfolded isCont_def]: "isCont f a" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 324 | assumes g: "g -- f a --> l" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 325 | 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: 
31336diff
changeset | 326 | shows "(\<lambda>x. g (f x)) -- a --> l" | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 327 | by (rule metric_LIM_compose2 [OF f g inj]) | 
| 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 328 | |
| 23040 | 329 | lemma isCont_LIM_compose2: | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 330 | fixes a :: "'a::real_normed_vector" | 
| 23040 | 331 | assumes f [unfolded isCont_def]: "isCont f a" | 
| 332 | assumes g: "g -- f a --> l" | |
| 333 | assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a" | |
| 334 | shows "(\<lambda>x. g (f x)) -- a --> l" | |
| 335 | by (rule LIM_compose2 [OF f g inj]) | |
| 336 | ||
| 21239 | 337 | lemma isCont_o2: "\<lbrakk>isCont f a; isCont g (f a)\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. g (f x)) a" | 
| 44314 | 338 | unfolding isCont_def by (rule tendsto_compose) | 
| 21239 | 339 | |
| 340 | 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: 
21257diff
changeset | 341 | unfolding o_def by (rule isCont_o2) | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 342 | |
| 44233 | 343 | lemma (in bounded_linear) isCont: | 
| 344 | "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a" | |
| 44314 | 345 | unfolding isCont_def by (rule tendsto) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 346 | |
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 347 | lemma (in bounded_bilinear) isCont: | 
| 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 348 | "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a" | 
| 44314 | 349 | unfolding isCont_def by (rule tendsto) | 
| 21282 
dd647b4d7952
added bounded_linear and bounded_bilinear locales
 huffman parents: 
21257diff
changeset | 350 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 351 | lemmas isCont_scaleR [simp] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 352 | bounded_bilinear.isCont [OF bounded_bilinear_scaleR] | 
| 21239 | 353 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 354 | lemmas isCont_of_real [simp] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44254diff
changeset | 355 | bounded_linear.isCont [OF bounded_linear_of_real] | 
| 22627 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 huffman parents: 
22613diff
changeset | 356 | |
| 44233 | 357 | lemma isCont_power [simp]: | 
| 36665 | 358 |   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: 
22613diff
changeset | 359 | shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a" | 
| 44314 | 360 | unfolding isCont_def by (rule tendsto_power) | 
| 22627 
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
 huffman parents: 
22613diff
changeset | 361 | |
| 44233 | 362 | lemma isCont_sgn [simp]: | 
| 36665 | 363 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 364 | shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. sgn (f x)) a" | 
| 44314 | 365 | unfolding isCont_def by (rule tendsto_sgn) | 
| 29885 | 366 | |
| 44233 | 367 | lemma isCont_setsum [simp]: | 
| 368 | fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::real_normed_vector" | |
| 369 | fixes A :: "'a set" | |
| 370 | shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a" | |
| 371 | unfolding isCont_def by (simp add: tendsto_setsum) | |
| 15228 | 372 | |
| 44233 | 373 | lemmas isCont_intros = | 
| 374 | isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus | |
| 375 | isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR | |
| 376 | 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: 
29667diff
changeset | 377 | |
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 378 | 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: 
29667diff
changeset | 379 |   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: 
29667diff
changeset | 380 | shows "0 \<le> f x" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 381 | proof (rule ccontr) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 382 | 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: 
29667diff
changeset | 383 | 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: 
29667diff
changeset | 384 | 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: 
29667diff
changeset | 385 | 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: 
29667diff
changeset | 386 | |
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 387 | 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: 
29667diff
changeset | 388 | 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: 
29667diff
changeset | 389 | 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: 
29667diff
changeset | 390 | have "b < ?x" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 391 | proof (cases "s < x - b") | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 392 | 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: 
29667diff
changeset | 393 | next | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 394 | 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: 
32436diff
changeset | 395 | 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: 
29667diff
changeset | 396 | 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: 
29667diff
changeset | 397 | qed | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 398 | 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: 
29667diff
changeset | 399 | 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: 
29667diff
changeset | 400 | 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: 
29667diff
changeset | 401 | 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: 
29667diff
changeset | 402 | 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: 
29667diff
changeset | 403 | 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: 
29667diff
changeset | 404 | 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: 
29667diff
changeset | 405 | qed | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 406 | |
| 14477 | 407 | |
| 20755 | 408 | subsection {* Uniform Continuity *}
 | 
| 409 | ||
| 14477 | 410 | lemma isUCont_isCont: "isUCont f ==> isCont f x" | 
| 23012 | 411 | by (simp add: isUCont_def isCont_def LIM_def, force) | 
| 14477 | 412 | |
| 23118 | 413 | lemma isUCont_Cauchy: | 
| 414 | "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" | |
| 415 | unfolding isUCont_def | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 416 | apply (rule metric_CauchyI) | 
| 23118 | 417 | apply (drule_tac x=e in spec, safe) | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 418 | apply (drule_tac e=s in metric_CauchyD, safe) | 
| 23118 | 419 | apply (rule_tac x=M in exI, simp) | 
| 420 | done | |
| 421 | ||
| 422 | lemma (in bounded_linear) isUCont: "isUCont f" | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 423 | unfolding isUCont_def dist_norm | 
| 23118 | 424 | proof (intro allI impI) | 
| 425 | fix r::real assume r: "0 < r" | |
| 426 | obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K" | |
| 427 | using pos_bounded by fast | |
| 428 | show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r" | |
| 429 | proof (rule exI, safe) | |
| 430 | from r K show "0 < r / K" by (rule divide_pos_pos) | |
| 431 | next | |
| 432 | fix x y :: 'a | |
| 433 | assume xy: "norm (x - y) < r / K" | |
| 434 | have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) | |
| 435 | also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le) | |
| 436 | also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq) | |
| 437 | finally show "norm (f x - f y) < r" . | |
| 438 | qed | |
| 439 | qed | |
| 440 | ||
| 441 | lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" | |
| 442 | by (rule isUCont [THEN isUCont_Cauchy]) | |
| 443 | ||
| 14477 | 444 | |
| 21165 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 huffman parents: 
21141diff
changeset | 445 | subsection {* Relation of LIM and LIMSEQ *}
 | 
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 446 | |
| 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 447 | lemma LIMSEQ_SEQ_conv1: | 
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 448 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 449 | assumes f: "f -- a --> l" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 450 | 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: 
44253diff
changeset | 451 | using tendsto_compose_eventually [OF f, where F=sequentially] by simp | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 452 | |
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 453 | lemma LIMSEQ_SEQ_conv2_lemma: | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 454 | fixes a :: "'a::metric_space" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 455 | assumes "\<forall>S. (\<forall>n. S n \<noteq> a) \<and> S ----> a \<longrightarrow> eventually (\<lambda>x. P (S x)) sequentially" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 456 | shows "eventually P (at a)" | 
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 457 | proof (rule ccontr) | 
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 458 | let ?I = "\<lambda>n. inverse (real (Suc n))" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 459 | let ?F = "\<lambda>n::nat. SOME x. x \<noteq> a \<and> dist x a < ?I n \<and> \<not> P x" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 460 | assume "\<not> eventually P (at a)" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 461 | hence P: "\<forall>d>0. \<exists>x. x \<noteq> a \<and> dist x a < d \<and> \<not> P x" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 462 | unfolding eventually_at by simp | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 463 | hence "\<And>n. \<exists>x. x \<noteq> a \<and> dist x a < ?I n \<and> \<not> P x" by simp | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 464 | hence F: "\<And>n. ?F n \<noteq> a \<and> dist (?F n) a < ?I n \<and> \<not> P (?F n)" | 
| 21165 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 huffman parents: 
21141diff
changeset | 465 | by (rule someI_ex) | 
| 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 huffman parents: 
21141diff
changeset | 466 | hence F1: "\<And>n. ?F n \<noteq> a" | 
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 467 | and F2: "\<And>n. dist (?F n) a < ?I n" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 468 | and F3: "\<And>n. \<not> P (?F n)" | 
| 21165 
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
 huffman parents: 
21141diff
changeset | 469 | by fast+ | 
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 470 | have "?F ----> a" | 
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 471 | using LIMSEQ_inverse_real_of_nat | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 472 | by (rule metric_tendsto_imp_tendsto, | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 473 | simp add: dist_norm F2 [THEN less_imp_le]) | 
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 474 | 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: 
21141diff
changeset | 475 | by (rule allI) (rule F1) | 
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 476 | ultimately have "eventually (\<lambda>n. P (?F n)) sequentially" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 477 | using assms by simp | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 478 | thus "False" by (simp add: F3) | 
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 479 | qed | 
| 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 480 | |
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 481 | lemma LIMSEQ_SEQ_conv2: | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 482 | fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 483 | 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: 
44253diff
changeset | 484 | shows "f -- a --> l" | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 485 | using assms unfolding tendsto_def [where l=l] | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 486 | by (simp add: LIMSEQ_SEQ_conv2_lemma) | 
| 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 487 | |
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 488 | lemma LIMSEQ_SEQ_conv: | 
| 44254 
336dd390e4a4
Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems
 huffman parents: 
44253diff
changeset | 489 | "(\<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: 
44253diff
changeset | 490 | (X -- a --> (L::'b::topological_space))" | 
| 44253 
c073a0bd8458
add lemma tendsto_compose_eventually; use it to shorten some proofs
 huffman parents: 
44251diff
changeset | 491 | using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 .. | 
| 19023 
5652a536b7e8
* include generalised MVT in HyperReal (contributed by Benjamin Porter)
 kleing parents: 
17318diff
changeset | 492 | |
| 44314 | 493 | subsection {* Legacy theorem names *}
 | 
| 494 | ||
| 495 | lemmas LIM_ident [simp] = tendsto_ident_at | |
| 496 | lemmas LIM_const [simp] = tendsto_const [where F="at x", standard] | |
| 497 | lemmas LIM_add = tendsto_add [where F="at x", standard] | |
| 498 | lemmas LIM_add_zero = tendsto_add_zero [where F="at x", standard] | |
| 499 | lemmas LIM_minus = tendsto_minus [where F="at x", standard] | |
| 500 | lemmas LIM_diff = tendsto_diff [where F="at x", standard] | |
| 501 | lemmas LIM_norm = tendsto_norm [where F="at x", standard] | |
| 502 | lemmas LIM_norm_zero = tendsto_norm_zero [where F="at x", standard] | |
| 503 | lemmas LIM_norm_zero_cancel = tendsto_norm_zero_cancel [where F="at x", standard] | |
| 504 | lemmas LIM_norm_zero_iff = tendsto_norm_zero_iff [where F="at x", standard] | |
| 505 | lemmas LIM_rabs = tendsto_rabs [where F="at x", standard] | |
| 506 | lemmas LIM_rabs_zero = tendsto_rabs_zero [where F="at x", standard] | |
| 507 | lemmas LIM_rabs_zero_cancel = tendsto_rabs_zero_cancel [where F="at x", standard] | |
| 508 | lemmas LIM_rabs_zero_iff = tendsto_rabs_zero_iff [where F="at x", standard] | |
| 509 | lemmas LIM_compose = tendsto_compose [where F="at x", standard] | |
| 510 | lemmas LIM_mult = tendsto_mult [where F="at x", standard] | |
| 511 | lemmas LIM_scaleR = tendsto_scaleR [where F="at x", standard] | |
| 512 | lemmas LIM_of_real = tendsto_of_real [where F="at x", standard] | |
| 513 | lemmas LIM_power = tendsto_power [where F="at x", standard] | |
| 514 | lemmas LIM_inverse = tendsto_inverse [where F="at x", standard] | |
| 515 | lemmas LIM_sgn = tendsto_sgn [where F="at x", standard] | |
| 516 | lemmas isCont_LIM_compose = isCont_tendsto_compose [where F="at x", standard] | |
| 517 | ||
| 10751 | 518 | end |