| author | haftmann |
| Sat, 23 Mar 2013 17:11:06 +0100 | |
| changeset 51487 | f4bfdee99304 |
| parent 51478 | 270b21f3ae0a |
| child 51524 | 7cb5ac44ca9e |
| permissions | -rw-r--r-- |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
1 |
(* Title : Limits.thy |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
2 |
Author : Brian Huffman |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
3 |
*) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
4 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
5 |
header {* Filters and Limits *}
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
6 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
7 |
theory Limits |
| 36822 | 8 |
imports RealVector |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
9 |
begin |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
10 |
|
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
11 |
definition at_infinity :: "'a::real_normed_vector filter" where |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
12 |
"at_infinity = Abs_filter (\<lambda>P. \<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x)" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
13 |
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
14 |
lemma eventually_at_infinity: |
| 50325 | 15 |
"eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
16 |
unfolding at_infinity_def |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
17 |
proof (rule eventually_Abs_filter, rule is_filter.intro) |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
18 |
fix P Q :: "'a \<Rightarrow> bool" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
19 |
assume "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x" and "\<exists>s. \<forall>x. s \<le> norm x \<longrightarrow> Q x" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
20 |
then obtain r s where |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
21 |
"\<forall>x. r \<le> norm x \<longrightarrow> P x" and "\<forall>x. s \<le> norm x \<longrightarrow> Q x" by auto |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
22 |
then have "\<forall>x. max r s \<le> norm x \<longrightarrow> P x \<and> Q x" by simp |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
23 |
then show "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x \<and> Q x" .. |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
24 |
qed auto |
| 31392 | 25 |
|
| 50325 | 26 |
lemma at_infinity_eq_at_top_bot: |
27 |
"(at_infinity \<Colon> real filter) = sup at_top at_bot" |
|
28 |
unfolding sup_filter_def at_infinity_def eventually_at_top_linorder eventually_at_bot_linorder |
|
29 |
proof (intro arg_cong[where f=Abs_filter] ext iffI) |
|
30 |
fix P :: "real \<Rightarrow> bool" assume "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x" |
|
31 |
then guess r .. |
|
32 |
then have "(\<forall>x\<ge>r. P x) \<and> (\<forall>x\<le>-r. P x)" by auto |
|
33 |
then show "(\<exists>r. \<forall>x\<ge>r. P x) \<and> (\<exists>r. \<forall>x\<le>r. P x)" by auto |
|
34 |
next |
|
35 |
fix P :: "real \<Rightarrow> bool" assume "(\<exists>r. \<forall>x\<ge>r. P x) \<and> (\<exists>r. \<forall>x\<le>r. P x)" |
|
36 |
then obtain p q where "\<forall>x\<ge>p. P x" "\<forall>x\<le>q. P x" by auto |
|
37 |
then show "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x" |
|
38 |
by (intro exI[of _ "max p (-q)"]) |
|
39 |
(auto simp: abs_real_def) |
|
40 |
qed |
|
41 |
||
42 |
lemma at_top_le_at_infinity: |
|
43 |
"at_top \<le> (at_infinity :: real filter)" |
|
44 |
unfolding at_infinity_eq_at_top_bot by simp |
|
45 |
||
46 |
lemma at_bot_le_at_infinity: |
|
47 |
"at_bot \<le> (at_infinity :: real filter)" |
|
48 |
unfolding at_infinity_eq_at_top_bot by simp |
|
49 |
||
| 31355 | 50 |
subsection {* Boundedness *}
|
51 |
||
|
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
52 |
lemma Bfun_def: |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
53 |
"Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)" |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
54 |
unfolding Bfun_metric_def norm_conv_dist |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
55 |
proof safe |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
56 |
fix y K assume "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F" |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
57 |
moreover have "eventually (\<lambda>x. dist (f x) 0 \<le> dist (f x) y + dist 0 y) F" |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
58 |
by (intro always_eventually) (metis dist_commute dist_triangle) |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
59 |
with * have "eventually (\<lambda>x. dist (f x) 0 \<le> K + dist 0 y) F" |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
60 |
by eventually_elim auto |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
61 |
with `0 < K` show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F" |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
62 |
by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto |
|
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
63 |
qed auto |
| 31355 | 64 |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
65 |
lemma BfunI: |
| 44195 | 66 |
assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F" |
| 31355 | 67 |
unfolding Bfun_def |
68 |
proof (intro exI conjI allI) |
|
69 |
show "0 < max K 1" by simp |
|
70 |
next |
|
| 44195 | 71 |
show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F" |
| 31355 | 72 |
using K by (rule eventually_elim1, simp) |
73 |
qed |
|
74 |
||
75 |
lemma BfunE: |
|
| 44195 | 76 |
assumes "Bfun f F" |
77 |
obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F" |
|
| 31355 | 78 |
using assms unfolding Bfun_def by fast |
79 |
||
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
80 |
subsection {* Convergence to Zero *}
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
81 |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
82 |
definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
|
| 44195 | 83 |
where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
84 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
85 |
lemma ZfunI: |
| 44195 | 86 |
"(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
87 |
unfolding Zfun_def by simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
88 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
89 |
lemma ZfunD: |
| 44195 | 90 |
"\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
91 |
unfolding Zfun_def by simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
92 |
|
| 31355 | 93 |
lemma Zfun_ssubst: |
| 44195 | 94 |
"eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
95 |
unfolding Zfun_def by (auto elim!: eventually_rev_mp) |
| 31355 | 96 |
|
| 44195 | 97 |
lemma Zfun_zero: "Zfun (\<lambda>x. 0) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
98 |
unfolding Zfun_def by simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
99 |
|
| 44195 | 100 |
lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) F = Zfun (\<lambda>x. f x) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
101 |
unfolding Zfun_def by simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
102 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
103 |
lemma Zfun_imp_Zfun: |
| 44195 | 104 |
assumes f: "Zfun f F" |
105 |
assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F" |
|
106 |
shows "Zfun (\<lambda>x. g x) F" |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
107 |
proof (cases) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
108 |
assume K: "0 < K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
109 |
show ?thesis |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
110 |
proof (rule ZfunI) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
111 |
fix r::real assume "0 < r" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
112 |
hence "0 < r / K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
113 |
using K by (rule divide_pos_pos) |
| 44195 | 114 |
then have "eventually (\<lambda>x. norm (f x) < r / K) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
115 |
using ZfunD [OF f] by fast |
| 44195 | 116 |
with g show "eventually (\<lambda>x. norm (g x) < r) F" |
| 46887 | 117 |
proof eventually_elim |
118 |
case (elim x) |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
119 |
hence "norm (f x) * K < r" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
120 |
by (simp add: pos_less_divide_eq K) |
| 46887 | 121 |
thus ?case |
122 |
by (simp add: order_le_less_trans [OF elim(1)]) |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
123 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
124 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
125 |
next |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
126 |
assume "\<not> 0 < K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
127 |
hence K: "K \<le> 0" by (simp only: not_less) |
| 31355 | 128 |
show ?thesis |
129 |
proof (rule ZfunI) |
|
130 |
fix r :: real |
|
131 |
assume "0 < r" |
|
| 44195 | 132 |
from g show "eventually (\<lambda>x. norm (g x) < r) F" |
| 46887 | 133 |
proof eventually_elim |
134 |
case (elim x) |
|
135 |
also have "norm (f x) * K \<le> norm (f x) * 0" |
|
| 31355 | 136 |
using K norm_ge_zero by (rule mult_left_mono) |
| 46887 | 137 |
finally show ?case |
| 31355 | 138 |
using `0 < r` by simp |
139 |
qed |
|
140 |
qed |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
141 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
142 |
|
| 44195 | 143 |
lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
144 |
by (erule_tac K="1" in Zfun_imp_Zfun, simp) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
145 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
146 |
lemma Zfun_add: |
| 44195 | 147 |
assumes f: "Zfun f F" and g: "Zfun g F" |
148 |
shows "Zfun (\<lambda>x. f x + g x) F" |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
149 |
proof (rule ZfunI) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
150 |
fix r::real assume "0 < r" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
151 |
hence r: "0 < r / 2" by simp |
| 44195 | 152 |
have "eventually (\<lambda>x. norm (f x) < r/2) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
153 |
using f r by (rule ZfunD) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
154 |
moreover |
| 44195 | 155 |
have "eventually (\<lambda>x. norm (g x) < r/2) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
156 |
using g r by (rule ZfunD) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
157 |
ultimately |
| 44195 | 158 |
show "eventually (\<lambda>x. norm (f x + g x) < r) F" |
| 46887 | 159 |
proof eventually_elim |
160 |
case (elim x) |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
161 |
have "norm (f x + g x) \<le> norm (f x) + norm (g x)" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
162 |
by (rule norm_triangle_ineq) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
163 |
also have "\<dots> < r/2 + r/2" |
| 46887 | 164 |
using elim by (rule add_strict_mono) |
165 |
finally show ?case |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
166 |
by simp |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
167 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
168 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
169 |
|
| 44195 | 170 |
lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
171 |
unfolding Zfun_def by simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
172 |
|
| 44195 | 173 |
lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
174 |
by (simp only: diff_minus Zfun_add Zfun_minus) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
175 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
176 |
lemma (in bounded_linear) Zfun: |
| 44195 | 177 |
assumes g: "Zfun g F" |
178 |
shows "Zfun (\<lambda>x. f (g x)) F" |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
179 |
proof - |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
180 |
obtain K where "\<And>x. norm (f x) \<le> norm x * K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
181 |
using bounded by fast |
| 44195 | 182 |
then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F" |
| 31355 | 183 |
by simp |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
184 |
with g show ?thesis |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
185 |
by (rule Zfun_imp_Zfun) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
186 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
187 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
188 |
lemma (in bounded_bilinear) Zfun: |
| 44195 | 189 |
assumes f: "Zfun f F" |
190 |
assumes g: "Zfun g F" |
|
191 |
shows "Zfun (\<lambda>x. f x ** g x) F" |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
192 |
proof (rule ZfunI) |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
193 |
fix r::real assume r: "0 < r" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
194 |
obtain K where K: "0 < K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
195 |
and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
196 |
using pos_bounded by fast |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
197 |
from K have K': "0 < inverse K" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
198 |
by (rule positive_imp_inverse_positive) |
| 44195 | 199 |
have "eventually (\<lambda>x. norm (f x) < r) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
200 |
using f r by (rule ZfunD) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
201 |
moreover |
| 44195 | 202 |
have "eventually (\<lambda>x. norm (g x) < inverse K) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
203 |
using g K' by (rule ZfunD) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
204 |
ultimately |
| 44195 | 205 |
show "eventually (\<lambda>x. norm (f x ** g x) < r) F" |
| 46887 | 206 |
proof eventually_elim |
207 |
case (elim x) |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
208 |
have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
209 |
by (rule norm_le) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
210 |
also have "norm (f x) * norm (g x) * K < r * inverse K * K" |
| 46887 | 211 |
by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
212 |
also from K have "r * inverse K * K = r" |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
213 |
by simp |
| 46887 | 214 |
finally show ?case . |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
215 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
216 |
qed |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
217 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
218 |
lemma (in bounded_bilinear) Zfun_left: |
| 44195 | 219 |
"Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
220 |
by (rule bounded_linear_left [THEN bounded_linear.Zfun]) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
221 |
|
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
222 |
lemma (in bounded_bilinear) Zfun_right: |
| 44195 | 223 |
"Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
224 |
by (rule bounded_linear_right [THEN bounded_linear.Zfun]) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
225 |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
226 |
lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
227 |
lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
228 |
lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
229 |
|
| 44195 | 230 |
lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\<lambda>x. f x - a) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
231 |
by (simp only: tendsto_iff Zfun_def dist_norm) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
232 |
|
|
44205
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents:
44195
diff
changeset
|
233 |
subsubsection {* Distance and norms *}
|
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
234 |
|
| 31565 | 235 |
lemma tendsto_norm [tendsto_intros]: |
| 44195 | 236 |
"(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
237 |
unfolding norm_conv_dist by (intro tendsto_intros) |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
238 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
239 |
lemma continuous_norm [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
240 |
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
241 |
unfolding continuous_def by (rule tendsto_norm) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
242 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
243 |
lemma continuous_on_norm [continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
244 |
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. norm (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
245 |
unfolding continuous_on_def by (auto intro: tendsto_norm) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
246 |
|
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
247 |
lemma tendsto_norm_zero: |
| 44195 | 248 |
"(f ---> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
249 |
by (drule tendsto_norm, simp) |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
250 |
|
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
251 |
lemma tendsto_norm_zero_cancel: |
| 44195 | 252 |
"((\<lambda>x. norm (f x)) ---> 0) F \<Longrightarrow> (f ---> 0) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
253 |
unfolding tendsto_iff dist_norm by simp |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
254 |
|
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
255 |
lemma tendsto_norm_zero_iff: |
| 44195 | 256 |
"((\<lambda>x. norm (f x)) ---> 0) F \<longleftrightarrow> (f ---> 0) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
257 |
unfolding tendsto_iff dist_norm by simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
258 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
259 |
lemma tendsto_rabs [tendsto_intros]: |
| 44195 | 260 |
"(f ---> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
261 |
by (fold real_norm_def, rule tendsto_norm) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
262 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
263 |
lemma continuous_rabs [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
264 |
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. \<bar>f x :: real\<bar>)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
265 |
unfolding real_norm_def[symmetric] by (rule continuous_norm) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
266 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
267 |
lemma continuous_on_rabs [continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
268 |
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. \<bar>f x :: real\<bar>)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
269 |
unfolding real_norm_def[symmetric] by (rule continuous_on_norm) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
270 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
271 |
lemma tendsto_rabs_zero: |
| 44195 | 272 |
"(f ---> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
273 |
by (fold real_norm_def, rule tendsto_norm_zero) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
274 |
|
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
275 |
lemma tendsto_rabs_zero_cancel: |
| 44195 | 276 |
"((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<Longrightarrow> (f ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
277 |
by (fold real_norm_def, rule tendsto_norm_zero_cancel) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
278 |
|
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
279 |
lemma tendsto_rabs_zero_iff: |
| 44195 | 280 |
"((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<longleftrightarrow> (f ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
281 |
by (fold real_norm_def, rule tendsto_norm_zero_iff) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
282 |
|
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
283 |
subsubsection {* Addition and subtraction *}
|
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
284 |
|
| 31565 | 285 |
lemma tendsto_add [tendsto_intros]: |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
286 |
fixes a b :: "'a::real_normed_vector" |
| 44195 | 287 |
shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
288 |
by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
289 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
290 |
lemma continuous_add [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
291 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
292 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x + g x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
293 |
unfolding continuous_def by (rule tendsto_add) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
294 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
295 |
lemma continuous_on_add [continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
296 |
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
297 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
298 |
unfolding continuous_on_def by (auto intro: tendsto_add) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
299 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
300 |
lemma tendsto_add_zero: |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
301 |
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" |
| 44195 | 302 |
shows "\<lbrakk>(f ---> 0) F; (g ---> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
303 |
by (drule (1) tendsto_add, simp) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
304 |
|
| 31565 | 305 |
lemma tendsto_minus [tendsto_intros]: |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
306 |
fixes a :: "'a::real_normed_vector" |
| 44195 | 307 |
shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
308 |
by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
309 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
310 |
lemma continuous_minus [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
311 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
312 |
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
313 |
unfolding continuous_def by (rule tendsto_minus) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
314 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
315 |
lemma continuous_on_minus [continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
316 |
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
317 |
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
318 |
unfolding continuous_on_def by (auto intro: tendsto_minus) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
319 |
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
320 |
lemma tendsto_minus_cancel: |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
321 |
fixes a :: "'a::real_normed_vector" |
| 44195 | 322 |
shows "((\<lambda>x. - f x) ---> - a) F \<Longrightarrow> (f ---> a) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
323 |
by (drule tendsto_minus, simp) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
324 |
|
|
50330
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
325 |
lemma tendsto_minus_cancel_left: |
|
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
326 |
"(f ---> - (y::_::real_normed_vector)) F \<longleftrightarrow> ((\<lambda>x. - f x) ---> y) F" |
|
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
327 |
using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] |
|
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
328 |
by auto |
|
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
329 |
|
| 31565 | 330 |
lemma tendsto_diff [tendsto_intros]: |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
331 |
fixes a b :: "'a::real_normed_vector" |
| 44195 | 332 |
shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
333 |
by (simp add: diff_minus tendsto_add tendsto_minus) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
334 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
335 |
lemma continuous_diff [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
336 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
337 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x - g x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
338 |
unfolding continuous_def by (rule tendsto_diff) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
339 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
340 |
lemma continuous_on_diff [continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
341 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
342 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
343 |
unfolding continuous_on_def by (auto intro: tendsto_diff) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
344 |
|
| 31588 | 345 |
lemma tendsto_setsum [tendsto_intros]: |
346 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector" |
|
| 44195 | 347 |
assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) F" |
348 |
shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) F" |
|
| 31588 | 349 |
proof (cases "finite S") |
350 |
assume "finite S" thus ?thesis using assms |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
351 |
by (induct, simp add: tendsto_const, simp add: tendsto_add) |
| 31588 | 352 |
next |
353 |
assume "\<not> finite S" thus ?thesis |
|
354 |
by (simp add: tendsto_const) |
|
355 |
qed |
|
356 |
||
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
357 |
lemma continuous_setsum [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
358 |
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
359 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>S. f i x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
360 |
unfolding continuous_def by (rule tendsto_setsum) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
361 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
362 |
lemma continuous_on_setsum [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
363 |
fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
364 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Sum>i\<in>S. f i x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
365 |
unfolding continuous_on_def by (auto intro: tendsto_setsum) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
366 |
|
| 50999 | 367 |
lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real] |
368 |
||
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
369 |
subsubsection {* Linear operators and multiplication *}
|
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
370 |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
371 |
lemma (in bounded_linear) tendsto: |
| 44195 | 372 |
"(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
373 |
by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
374 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
375 |
lemma (in bounded_linear) continuous: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
376 |
"continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
377 |
using tendsto[of g _ F] by (auto simp: continuous_def) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
378 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
379 |
lemma (in bounded_linear) continuous_on: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
380 |
"continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
381 |
using tendsto[of g] by (auto simp: continuous_on_def) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
382 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
383 |
lemma (in bounded_linear) tendsto_zero: |
| 44195 | 384 |
"(g ---> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
385 |
by (drule tendsto, simp only: zero) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
386 |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
387 |
lemma (in bounded_bilinear) tendsto: |
| 44195 | 388 |
"\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
389 |
by (simp only: tendsto_Zfun_iff prod_diff_prod |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
390 |
Zfun_add Zfun Zfun_left Zfun_right) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
391 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
392 |
lemma (in bounded_bilinear) continuous: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
393 |
"continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ** g x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
394 |
using tendsto[of f _ F g] by (auto simp: continuous_def) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
395 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
396 |
lemma (in bounded_bilinear) continuous_on: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
397 |
"continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x ** g x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
398 |
using tendsto[of f _ _ g] by (auto simp: continuous_on_def) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
399 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
400 |
lemma (in bounded_bilinear) tendsto_zero: |
| 44195 | 401 |
assumes f: "(f ---> 0) F" |
402 |
assumes g: "(g ---> 0) F" |
|
403 |
shows "((\<lambda>x. f x ** g x) ---> 0) F" |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
404 |
using tendsto [OF f g] by (simp add: zero_left) |
| 31355 | 405 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
406 |
lemma (in bounded_bilinear) tendsto_left_zero: |
| 44195 | 407 |
"(f ---> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
408 |
by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
409 |
|
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
410 |
lemma (in bounded_bilinear) tendsto_right_zero: |
| 44195 | 411 |
"(f ---> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
412 |
by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
413 |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
414 |
lemmas tendsto_of_real [tendsto_intros] = |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
415 |
bounded_linear.tendsto [OF bounded_linear_of_real] |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
416 |
|
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
417 |
lemmas tendsto_scaleR [tendsto_intros] = |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
418 |
bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
419 |
|
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
420 |
lemmas tendsto_mult [tendsto_intros] = |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
421 |
bounded_bilinear.tendsto [OF bounded_bilinear_mult] |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
422 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
423 |
lemmas continuous_of_real [continuous_intros] = |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
424 |
bounded_linear.continuous [OF bounded_linear_of_real] |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
425 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
426 |
lemmas continuous_scaleR [continuous_intros] = |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
427 |
bounded_bilinear.continuous [OF bounded_bilinear_scaleR] |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
428 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
429 |
lemmas continuous_mult [continuous_intros] = |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
430 |
bounded_bilinear.continuous [OF bounded_bilinear_mult] |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
431 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
432 |
lemmas continuous_on_of_real [continuous_on_intros] = |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
433 |
bounded_linear.continuous_on [OF bounded_linear_of_real] |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
434 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
435 |
lemmas continuous_on_scaleR [continuous_on_intros] = |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
436 |
bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
437 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
438 |
lemmas continuous_on_mult [continuous_on_intros] = |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
439 |
bounded_bilinear.continuous_on [OF bounded_bilinear_mult] |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
440 |
|
|
44568
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
441 |
lemmas tendsto_mult_zero = |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
442 |
bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
443 |
|
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
444 |
lemmas tendsto_mult_left_zero = |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
445 |
bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
446 |
|
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
447 |
lemmas tendsto_mult_right_zero = |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
448 |
bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
449 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
450 |
lemma tendsto_power [tendsto_intros]: |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
451 |
fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
|
| 44195 | 452 |
shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
453 |
by (induct n) (simp_all add: tendsto_const tendsto_mult) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
454 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
455 |
lemma continuous_power [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
456 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
457 |
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
458 |
unfolding continuous_def by (rule tendsto_power) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
459 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
460 |
lemma continuous_on_power [continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
461 |
fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
462 |
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. (f x)^n)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
463 |
unfolding continuous_on_def by (auto intro: tendsto_power) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
464 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
465 |
lemma tendsto_setprod [tendsto_intros]: |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
466 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
|
| 44195 | 467 |
assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) F" |
468 |
shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) F" |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
469 |
proof (cases "finite S") |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
470 |
assume "finite S" thus ?thesis using assms |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
471 |
by (induct, simp add: tendsto_const, simp add: tendsto_mult) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
472 |
next |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
473 |
assume "\<not> finite S" thus ?thesis |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
474 |
by (simp add: tendsto_const) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
475 |
qed |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
476 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
477 |
lemma continuous_setprod [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
478 |
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
479 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>S. f i x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
480 |
unfolding continuous_def by (rule tendsto_setprod) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
481 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
482 |
lemma continuous_on_setprod [continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
483 |
fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
484 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Prod>i\<in>S. f i x)" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
485 |
unfolding continuous_on_def by (auto intro: tendsto_setprod) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
486 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
487 |
subsubsection {* Inverse and division *}
|
| 31355 | 488 |
|
489 |
lemma (in bounded_bilinear) Zfun_prod_Bfun: |
|
| 44195 | 490 |
assumes f: "Zfun f F" |
491 |
assumes g: "Bfun g F" |
|
492 |
shows "Zfun (\<lambda>x. f x ** g x) F" |
|
| 31355 | 493 |
proof - |
494 |
obtain K where K: "0 \<le> K" |
|
495 |
and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" |
|
496 |
using nonneg_bounded by fast |
|
497 |
obtain B where B: "0 < B" |
|
| 44195 | 498 |
and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
499 |
using g by (rule BfunE) |
| 44195 | 500 |
have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) F" |
| 46887 | 501 |
using norm_g proof eventually_elim |
502 |
case (elim x) |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
503 |
have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" |
| 31355 | 504 |
by (rule norm_le) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
505 |
also have "\<dots> \<le> norm (f x) * B * K" |
|
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
506 |
by (intro mult_mono' order_refl norm_g norm_ge_zero |
| 46887 | 507 |
mult_nonneg_nonneg K elim) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
508 |
also have "\<dots> = norm (f x) * (B * K)" |
| 31355 | 509 |
by (rule mult_assoc) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
510 |
finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" . |
| 31355 | 511 |
qed |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
512 |
with f show ?thesis |
|
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
513 |
by (rule Zfun_imp_Zfun) |
| 31355 | 514 |
qed |
515 |
||
516 |
lemma (in bounded_bilinear) flip: |
|
517 |
"bounded_bilinear (\<lambda>x y. y ** x)" |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
518 |
apply default |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
519 |
apply (rule add_right) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
520 |
apply (rule add_left) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
521 |
apply (rule scaleR_right) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
522 |
apply (rule scaleR_left) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
523 |
apply (subst mult_commute) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
524 |
using bounded by fast |
| 31355 | 525 |
|
526 |
lemma (in bounded_bilinear) Bfun_prod_Zfun: |
|
| 44195 | 527 |
assumes f: "Bfun f F" |
528 |
assumes g: "Zfun g F" |
|
529 |
shows "Zfun (\<lambda>x. f x ** g x) F" |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
530 |
using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) |
| 31355 | 531 |
|
532 |
lemma Bfun_inverse_lemma: |
|
533 |
fixes x :: "'a::real_normed_div_algebra" |
|
534 |
shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r" |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
535 |
apply (subst nonzero_norm_inverse, clarsimp) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
536 |
apply (erule (1) le_imp_inverse_le) |
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
537 |
done |
| 31355 | 538 |
|
539 |
lemma Bfun_inverse: |
|
540 |
fixes a :: "'a::real_normed_div_algebra" |
|
| 44195 | 541 |
assumes f: "(f ---> a) F" |
| 31355 | 542 |
assumes a: "a \<noteq> 0" |
| 44195 | 543 |
shows "Bfun (\<lambda>x. inverse (f x)) F" |
| 31355 | 544 |
proof - |
545 |
from a have "0 < norm a" by simp |
|
546 |
hence "\<exists>r>0. r < norm a" by (rule dense) |
|
547 |
then obtain r where r1: "0 < r" and r2: "r < norm a" by fast |
|
| 44195 | 548 |
have "eventually (\<lambda>x. dist (f x) a < r) F" |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
549 |
using tendstoD [OF f r1] by fast |
| 44195 | 550 |
hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F" |
| 46887 | 551 |
proof eventually_elim |
552 |
case (elim x) |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
553 |
hence 1: "norm (f x - a) < r" |
| 31355 | 554 |
by (simp add: dist_norm) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
555 |
hence 2: "f x \<noteq> 0" using r2 by auto |
|
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
556 |
hence "norm (inverse (f x)) = inverse (norm (f x))" |
| 31355 | 557 |
by (rule nonzero_norm_inverse) |
558 |
also have "\<dots> \<le> inverse (norm a - r)" |
|
559 |
proof (rule le_imp_inverse_le) |
|
560 |
show "0 < norm a - r" using r2 by simp |
|
561 |
next |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
562 |
have "norm a - norm (f x) \<le> norm (a - f x)" |
| 31355 | 563 |
by (rule norm_triangle_ineq2) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
564 |
also have "\<dots> = norm (f x - a)" |
| 31355 | 565 |
by (rule norm_minus_commute) |
566 |
also have "\<dots> < r" using 1 . |
|
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
567 |
finally show "norm a - r \<le> norm (f x)" by simp |
| 31355 | 568 |
qed |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
569 |
finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" . |
| 31355 | 570 |
qed |
571 |
thus ?thesis by (rule BfunI) |
|
572 |
qed |
|
573 |
||
| 31565 | 574 |
lemma tendsto_inverse [tendsto_intros]: |
| 31355 | 575 |
fixes a :: "'a::real_normed_div_algebra" |
| 44195 | 576 |
assumes f: "(f ---> a) F" |
| 31355 | 577 |
assumes a: "a \<noteq> 0" |
| 44195 | 578 |
shows "((\<lambda>x. inverse (f x)) ---> inverse a) F" |
| 31355 | 579 |
proof - |
580 |
from a have "0 < norm a" by simp |
|
| 44195 | 581 |
with f have "eventually (\<lambda>x. dist (f x) a < norm a) F" |
| 31355 | 582 |
by (rule tendstoD) |
| 44195 | 583 |
then have "eventually (\<lambda>x. f x \<noteq> 0) F" |
| 31355 | 584 |
unfolding dist_norm by (auto elim!: eventually_elim1) |
| 44627 | 585 |
with a have "eventually (\<lambda>x. inverse (f x) - inverse a = |
586 |
- (inverse (f x) * (f x - a) * inverse a)) F" |
|
587 |
by (auto elim!: eventually_elim1 simp: inverse_diff_inverse) |
|
588 |
moreover have "Zfun (\<lambda>x. - (inverse (f x) * (f x - a) * inverse a)) F" |
|
589 |
by (intro Zfun_minus Zfun_mult_left |
|
590 |
bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] |
|
591 |
Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) |
|
592 |
ultimately show ?thesis |
|
593 |
unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) |
|
| 31355 | 594 |
qed |
595 |
||
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
596 |
lemma continuous_inverse: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
597 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
598 |
assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
599 |
shows "continuous F (\<lambda>x. inverse (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
600 |
using assms unfolding continuous_def by (rule tendsto_inverse) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
601 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
602 |
lemma continuous_at_within_inverse[continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
603 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
604 |
assumes "continuous (at a within s) f" and "f a \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
605 |
shows "continuous (at a within s) (\<lambda>x. inverse (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
606 |
using assms unfolding continuous_within by (rule tendsto_inverse) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
607 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
608 |
lemma isCont_inverse[continuous_intros, simp]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
609 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
610 |
assumes "isCont f a" and "f a \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
611 |
shows "isCont (\<lambda>x. inverse (f x)) a" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
612 |
using assms unfolding continuous_at by (rule tendsto_inverse) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
613 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
614 |
lemma continuous_on_inverse[continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
615 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
616 |
assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
617 |
shows "continuous_on s (\<lambda>x. inverse (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
618 |
using assms unfolding continuous_on_def by (fast intro: tendsto_inverse) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
619 |
|
| 31565 | 620 |
lemma tendsto_divide [tendsto_intros]: |
| 31355 | 621 |
fixes a b :: "'a::real_normed_field" |
| 44195 | 622 |
shows "\<lbrakk>(f ---> a) F; (g ---> b) F; b \<noteq> 0\<rbrakk> |
623 |
\<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) F" |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
624 |
by (simp add: tendsto_mult tendsto_inverse divide_inverse) |
| 31355 | 625 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
626 |
lemma continuous_divide: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
627 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
628 |
assumes "continuous F f" and "continuous F g" and "g (Lim F (\<lambda>x. x)) \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
629 |
shows "continuous F (\<lambda>x. (f x) / (g x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
630 |
using assms unfolding continuous_def by (rule tendsto_divide) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
631 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
632 |
lemma continuous_at_within_divide[continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
633 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
634 |
assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
635 |
shows "continuous (at a within s) (\<lambda>x. (f x) / (g x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
636 |
using assms unfolding continuous_within by (rule tendsto_divide) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
637 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
638 |
lemma isCont_divide[continuous_intros, simp]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
639 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
640 |
assumes "isCont f a" "isCont g a" "g a \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
641 |
shows "isCont (\<lambda>x. (f x) / g x) a" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
642 |
using assms unfolding continuous_at by (rule tendsto_divide) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
643 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
644 |
lemma continuous_on_divide[continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
645 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_field" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
646 |
assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. g x \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
647 |
shows "continuous_on s (\<lambda>x. (f x) / (g x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
648 |
using assms unfolding continuous_on_def by (fast intro: tendsto_divide) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
649 |
|
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
650 |
lemma tendsto_sgn [tendsto_intros]: |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
651 |
fixes l :: "'a::real_normed_vector" |
| 44195 | 652 |
shows "\<lbrakk>(f ---> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
653 |
unfolding sgn_div_norm by (simp add: tendsto_intros) |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
654 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
655 |
lemma continuous_sgn: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
656 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
657 |
assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
658 |
shows "continuous F (\<lambda>x. sgn (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
659 |
using assms unfolding continuous_def by (rule tendsto_sgn) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
660 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
661 |
lemma continuous_at_within_sgn[continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
662 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
663 |
assumes "continuous (at a within s) f" and "f a \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
664 |
shows "continuous (at a within s) (\<lambda>x. sgn (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
665 |
using assms unfolding continuous_within by (rule tendsto_sgn) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
666 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
667 |
lemma isCont_sgn[continuous_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
668 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
669 |
assumes "isCont f a" and "f a \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
670 |
shows "isCont (\<lambda>x. sgn (f x)) a" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
671 |
using assms unfolding continuous_at by (rule tendsto_sgn) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
672 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
673 |
lemma continuous_on_sgn[continuous_on_intros]: |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
674 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
675 |
assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
676 |
shows "continuous_on s (\<lambda>x. sgn (f x))" |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
677 |
using assms unfolding continuous_on_def by (fast intro: tendsto_sgn) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
678 |
|
| 50325 | 679 |
lemma filterlim_at_infinity: |
680 |
fixes f :: "_ \<Rightarrow> 'a\<Colon>real_normed_vector" |
|
681 |
assumes "0 \<le> c" |
|
682 |
shows "(LIM x F. f x :> at_infinity) \<longleftrightarrow> (\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F)" |
|
683 |
unfolding filterlim_iff eventually_at_infinity |
|
684 |
proof safe |
|
685 |
fix P :: "'a \<Rightarrow> bool" and b |
|
686 |
assume *: "\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F" |
|
687 |
and P: "\<forall>x. b \<le> norm x \<longrightarrow> P x" |
|
688 |
have "max b (c + 1) > c" by auto |
|
689 |
with * have "eventually (\<lambda>x. max b (c + 1) \<le> norm (f x)) F" |
|
690 |
by auto |
|
691 |
then show "eventually (\<lambda>x. P (f x)) F" |
|
692 |
proof eventually_elim |
|
693 |
fix x assume "max b (c + 1) \<le> norm (f x)" |
|
694 |
with P show "P (f x)" by auto |
|
695 |
qed |
|
696 |
qed force |
|
697 |
||
| 50347 | 698 |
subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *}
|
699 |
||
700 |
text {*
|
|
701 |
||
702 |
This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
|
|
703 |
@{term "at_right x"} and also @{term "at_right 0"}.
|
|
704 |
||
705 |
*} |
|
706 |
||
| 51471 | 707 |
lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] |
| 50323 | 708 |
|
| 50347 | 709 |
lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d::real)" |
710 |
unfolding filter_eq_iff eventually_filtermap eventually_nhds_metric |
|
711 |
by (intro allI ex_cong) (auto simp: dist_real_def field_simps) |
|
712 |
||
713 |
lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a::real)" |
|
714 |
unfolding filter_eq_iff eventually_filtermap eventually_nhds_metric |
|
715 |
apply (intro allI ex_cong) |
|
716 |
apply (auto simp: dist_real_def field_simps) |
|
717 |
apply (erule_tac x="-x" in allE) |
|
718 |
apply simp |
|
719 |
done |
|
720 |
||
721 |
lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d::real)" |
|
722 |
unfolding at_def filtermap_nhds_shift[symmetric] |
|
723 |
by (simp add: filter_eq_iff eventually_filtermap eventually_within) |
|
724 |
||
725 |
lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d::real)" |
|
726 |
unfolding filtermap_at_shift[symmetric] |
|
727 |
by (simp add: filter_eq_iff eventually_filtermap eventually_within) |
|
| 50323 | 728 |
|
| 50347 | 729 |
lemma at_right_to_0: "at_right (a::real) = filtermap (\<lambda>x. x + a) (at_right 0)" |
730 |
using filtermap_at_right_shift[of "-a" 0] by simp |
|
731 |
||
732 |
lemma filterlim_at_right_to_0: |
|
733 |
"filterlim f F (at_right (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)" |
|
734 |
unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. |
|
735 |
||
736 |
lemma eventually_at_right_to_0: |
|
737 |
"eventually P (at_right (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)" |
|
738 |
unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) |
|
739 |
||
740 |
lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a::real)" |
|
741 |
unfolding at_def filtermap_nhds_minus[symmetric] |
|
742 |
by (simp add: filter_eq_iff eventually_filtermap eventually_within) |
|
743 |
||
744 |
lemma at_left_minus: "at_left (a::real) = filtermap (\<lambda>x. - x) (at_right (- a))" |
|
745 |
by (simp add: filter_eq_iff eventually_filtermap eventually_within filtermap_at_minus[symmetric]) |
|
| 50323 | 746 |
|
| 50347 | 747 |
lemma at_right_minus: "at_right (a::real) = filtermap (\<lambda>x. - x) (at_left (- a))" |
748 |
by (simp add: filter_eq_iff eventually_filtermap eventually_within filtermap_at_minus[symmetric]) |
|
749 |
||
750 |
lemma filterlim_at_left_to_right: |
|
751 |
"filterlim f F (at_left (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))" |
|
752 |
unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. |
|
753 |
||
754 |
lemma eventually_at_left_to_right: |
|
755 |
"eventually P (at_left (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))" |
|
756 |
unfolding at_left_minus[of a] by (simp add: eventually_filtermap) |
|
757 |
||
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
758 |
lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: real filter)" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
759 |
unfolding filter_eq_iff eventually_filtermap eventually_at_top_linorder eventually_at_bot_linorder |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
760 |
by (metis le_minus_iff minus_minus) |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
761 |
|
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
762 |
lemma at_bot_mirror: "at_bot = filtermap uminus (at_top :: real filter)" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
763 |
unfolding at_top_mirror filtermap_filtermap by (simp add: filtermap_ident) |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
764 |
|
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
765 |
lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \<longleftrightarrow> (LIM x at_bot. f (-x::real) :> F)" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
766 |
unfolding filterlim_def at_top_mirror filtermap_filtermap .. |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
767 |
|
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
768 |
lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \<longleftrightarrow> (LIM x at_top. f (-x::real) :> F)" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
769 |
unfolding filterlim_def at_bot_mirror filtermap_filtermap .. |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
770 |
|
| 50323 | 771 |
lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" |
772 |
unfolding filterlim_at_top eventually_at_bot_dense |
|
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
773 |
by (metis leI minus_less_iff order_less_asym) |
| 50323 | 774 |
|
775 |
lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" |
|
776 |
unfolding filterlim_at_bot eventually_at_top_dense |
|
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
777 |
by (metis leI less_minus_iff order_less_asym) |
| 50323 | 778 |
|
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
779 |
lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_bot)" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
780 |
using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
781 |
using filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F] |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
782 |
by auto |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
783 |
|
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
784 |
lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_top)" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
785 |
unfolding filterlim_uminus_at_top by simp |
| 50323 | 786 |
|
| 50347 | 787 |
lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" |
788 |
unfolding filterlim_at_top_gt[where c=0] eventually_within at_def |
|
789 |
proof safe |
|
790 |
fix Z :: real assume [arith]: "0 < Z" |
|
791 |
then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)" |
|
792 |
by (auto simp add: eventually_nhds_metric dist_real_def intro!: exI[of _ "\<bar>inverse Z\<bar>"]) |
|
793 |
then show "eventually (\<lambda>x. x \<in> - {0} \<longrightarrow> x \<in> {0<..} \<longrightarrow> Z \<le> inverse x) (nhds 0)"
|
|
794 |
by (auto elim!: eventually_elim1 simp: inverse_eq_divide field_simps) |
|
795 |
qed |
|
796 |
||
797 |
lemma filterlim_inverse_at_top: |
|
798 |
"(f ---> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> LIM x F. inverse (f x) :> at_top" |
|
799 |
by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) |
|
800 |
(simp add: filterlim_def eventually_filtermap le_within_iff at_def eventually_elim1) |
|
801 |
||
802 |
lemma filterlim_inverse_at_bot_neg: |
|
803 |
"LIM x (at_left (0::real)). inverse x :> at_bot" |
|
804 |
by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) |
|
805 |
||
806 |
lemma filterlim_inverse_at_bot: |
|
807 |
"(f ---> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x < 0) F \<Longrightarrow> LIM x F. inverse (f x) :> at_bot" |
|
808 |
unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] |
|
809 |
by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) |
|
810 |
||
| 50325 | 811 |
lemma tendsto_inverse_0: |
812 |
fixes x :: "_ \<Rightarrow> 'a\<Colon>real_normed_div_algebra" |
|
813 |
shows "(inverse ---> (0::'a)) at_infinity" |
|
814 |
unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity |
|
815 |
proof safe |
|
816 |
fix r :: real assume "0 < r" |
|
817 |
show "\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> norm (inverse x :: 'a) < r" |
|
818 |
proof (intro exI[of _ "inverse (r / 2)"] allI impI) |
|
819 |
fix x :: 'a |
|
820 |
from `0 < r` have "0 < inverse (r / 2)" by simp |
|
821 |
also assume *: "inverse (r / 2) \<le> norm x" |
|
822 |
finally show "norm (inverse x) < r" |
|
823 |
using * `0 < r` by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) |
|
824 |
qed |
|
825 |
qed |
|
826 |
||
| 50347 | 827 |
lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" |
828 |
proof (rule antisym) |
|
829 |
have "(inverse ---> (0::real)) at_top" |
|
830 |
by (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) |
|
831 |
then show "filtermap inverse at_top \<le> at_right (0::real)" |
|
832 |
unfolding at_within_eq |
|
833 |
by (intro le_withinI) (simp_all add: eventually_filtermap eventually_gt_at_top filterlim_def) |
|
834 |
next |
|
835 |
have "filtermap inverse (filtermap inverse (at_right (0::real))) \<le> filtermap inverse at_top" |
|
836 |
using filterlim_inverse_at_top_right unfolding filterlim_def by (rule filtermap_mono) |
|
837 |
then show "at_right (0::real) \<le> filtermap inverse at_top" |
|
838 |
by (simp add: filtermap_ident filtermap_filtermap) |
|
839 |
qed |
|
840 |
||
841 |
lemma eventually_at_right_to_top: |
|
842 |
"eventually P (at_right (0::real)) \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) at_top" |
|
843 |
unfolding at_right_to_top eventually_filtermap .. |
|
844 |
||
845 |
lemma filterlim_at_right_to_top: |
|
846 |
"filterlim f F (at_right (0::real)) \<longleftrightarrow> (LIM x at_top. f (inverse x) :> F)" |
|
847 |
unfolding filterlim_def at_right_to_top filtermap_filtermap .. |
|
848 |
||
849 |
lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" |
|
850 |
unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. |
|
851 |
||
852 |
lemma eventually_at_top_to_right: |
|
853 |
"eventually P at_top \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) (at_right (0::real))" |
|
854 |
unfolding at_top_to_right eventually_filtermap .. |
|
855 |
||
856 |
lemma filterlim_at_top_to_right: |
|
857 |
"filterlim f F at_top \<longleftrightarrow> (LIM x (at_right (0::real)). f (inverse x) :> F)" |
|
858 |
unfolding filterlim_def at_top_to_right filtermap_filtermap .. |
|
859 |
||
| 50325 | 860 |
lemma filterlim_inverse_at_infinity: |
861 |
fixes x :: "_ \<Rightarrow> 'a\<Colon>{real_normed_div_algebra, division_ring_inverse_zero}"
|
|
862 |
shows "filterlim inverse at_infinity (at (0::'a))" |
|
863 |
unfolding filterlim_at_infinity[OF order_refl] |
|
864 |
proof safe |
|
865 |
fix r :: real assume "0 < r" |
|
866 |
then show "eventually (\<lambda>x::'a. r \<le> norm (inverse x)) (at 0)" |
|
867 |
unfolding eventually_at norm_inverse |
|
868 |
by (intro exI[of _ "inverse r"]) |
|
869 |
(auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) |
|
870 |
qed |
|
871 |
||
872 |
lemma filterlim_inverse_at_iff: |
|
873 |
fixes g :: "'a \<Rightarrow> 'b\<Colon>{real_normed_div_algebra, division_ring_inverse_zero}"
|
|
874 |
shows "(LIM x F. inverse (g x) :> at 0) \<longleftrightarrow> (LIM x F. g x :> at_infinity)" |
|
875 |
unfolding filterlim_def filtermap_filtermap[symmetric] |
|
876 |
proof |
|
877 |
assume "filtermap g F \<le> at_infinity" |
|
878 |
then have "filtermap inverse (filtermap g F) \<le> filtermap inverse at_infinity" |
|
879 |
by (rule filtermap_mono) |
|
880 |
also have "\<dots> \<le> at 0" |
|
881 |
using tendsto_inverse_0 |
|
882 |
by (auto intro!: le_withinI exI[of _ 1] |
|
883 |
simp: eventually_filtermap eventually_at_infinity filterlim_def at_def) |
|
884 |
finally show "filtermap inverse (filtermap g F) \<le> at 0" . |
|
885 |
next |
|
886 |
assume "filtermap inverse (filtermap g F) \<le> at 0" |
|
887 |
then have "filtermap inverse (filtermap inverse (filtermap g F)) \<le> filtermap inverse (at 0)" |
|
888 |
by (rule filtermap_mono) |
|
889 |
with filterlim_inverse_at_infinity show "filtermap g F \<le> at_infinity" |
|
890 |
by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) |
|
891 |
qed |
|
892 |
||
| 50419 | 893 |
lemma tendsto_inverse_0_at_top: |
894 |
"LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) ---> 0) F" |
|
895 |
by (metis at_top_le_at_infinity filterlim_at filterlim_inverse_at_iff filterlim_mono order_refl) |
|
896 |
||
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
897 |
text {*
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
898 |
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
899 |
We only show rules for multiplication and addition when the functions are either against a real |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
900 |
value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}.
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
901 |
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
902 |
*} |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
903 |
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
904 |
lemma filterlim_tendsto_pos_mult_at_top: |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
905 |
assumes f: "(f ---> c) F" and c: "0 < c" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
906 |
assumes g: "LIM x F. g x :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
907 |
shows "LIM x F. (f x * g x :: real) :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
908 |
unfolding filterlim_at_top_gt[where c=0] |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
909 |
proof safe |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
910 |
fix Z :: real assume "0 < Z" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
911 |
from f `0 < c` have "eventually (\<lambda>x. c / 2 < f x) F" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
912 |
by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_elim1 |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
913 |
simp: dist_real_def abs_real_def split: split_if_asm) |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
914 |
moreover from g have "eventually (\<lambda>x. (Z / c * 2) \<le> g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
915 |
unfolding filterlim_at_top by auto |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
916 |
ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
917 |
proof eventually_elim |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
918 |
fix x assume "c / 2 < f x" "Z / c * 2 \<le> g x" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
919 |
with `0 < Z` `0 < c` have "c / 2 * (Z / c * 2) \<le> f x * g x" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
920 |
by (intro mult_mono) (auto simp: zero_le_divide_iff) |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
921 |
with `0 < c` show "Z \<le> f x * g x" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
922 |
by simp |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
923 |
qed |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
924 |
qed |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
925 |
|
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
926 |
lemma filterlim_at_top_mult_at_top: |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
927 |
assumes f: "LIM x F. f x :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
928 |
assumes g: "LIM x F. g x :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
929 |
shows "LIM x F. (f x * g x :: real) :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
930 |
unfolding filterlim_at_top_gt[where c=0] |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
931 |
proof safe |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
932 |
fix Z :: real assume "0 < Z" |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
933 |
from f have "eventually (\<lambda>x. 1 \<le> f x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
934 |
unfolding filterlim_at_top by auto |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
935 |
moreover from g have "eventually (\<lambda>x. Z \<le> g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
936 |
unfolding filterlim_at_top by auto |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
937 |
ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
938 |
proof eventually_elim |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
939 |
fix x assume "1 \<le> f x" "Z \<le> g x" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
940 |
with `0 < Z` have "1 * Z \<le> f x * g x" |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
941 |
by (intro mult_mono) (auto simp: zero_le_divide_iff) |
|
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
942 |
then show "Z \<le> f x * g x" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
943 |
by simp |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
944 |
qed |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
945 |
qed |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
946 |
|
| 50419 | 947 |
lemma filterlim_tendsto_pos_mult_at_bot: |
948 |
assumes "(f ---> c) F" "0 < (c::real)" "filterlim g at_bot F" |
|
949 |
shows "LIM x F. f x * g x :> at_bot" |
|
950 |
using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\<lambda>x. - g x"] assms(3) |
|
951 |
unfolding filterlim_uminus_at_bot by simp |
|
952 |
||
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
953 |
lemma filterlim_tendsto_add_at_top: |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
954 |
assumes f: "(f ---> c) F" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
955 |
assumes g: "LIM x F. g x :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
956 |
shows "LIM x F. (f x + g x :: real) :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
957 |
unfolding filterlim_at_top_gt[where c=0] |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
958 |
proof safe |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
959 |
fix Z :: real assume "0 < Z" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
960 |
from f have "eventually (\<lambda>x. c - 1 < f x) F" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
961 |
by (auto dest!: tendstoD[where e=1] elim!: eventually_elim1 simp: dist_real_def) |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
962 |
moreover from g have "eventually (\<lambda>x. Z - (c - 1) \<le> g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
963 |
unfolding filterlim_at_top by auto |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
964 |
ultimately show "eventually (\<lambda>x. Z \<le> f x + g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
965 |
by eventually_elim simp |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
966 |
qed |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
967 |
|
| 50347 | 968 |
lemma LIM_at_top_divide: |
969 |
fixes f g :: "'a \<Rightarrow> real" |
|
970 |
assumes f: "(f ---> a) F" "0 < a" |
|
971 |
assumes g: "(g ---> 0) F" "eventually (\<lambda>x. 0 < g x) F" |
|
972 |
shows "LIM x F. f x / g x :> at_top" |
|
973 |
unfolding divide_inverse |
|
974 |
by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) |
|
975 |
||
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
976 |
lemma filterlim_at_top_add_at_top: |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
977 |
assumes f: "LIM x F. f x :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
978 |
assumes g: "LIM x F. g x :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
979 |
shows "LIM x F. (f x + g x :: real) :> at_top" |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
980 |
unfolding filterlim_at_top_gt[where c=0] |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
981 |
proof safe |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
982 |
fix Z :: real assume "0 < Z" |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
983 |
from f have "eventually (\<lambda>x. 0 \<le> f x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
984 |
unfolding filterlim_at_top by auto |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
985 |
moreover from g have "eventually (\<lambda>x. Z \<le> g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
986 |
unfolding filterlim_at_top by auto |
|
50346
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents:
50331
diff
changeset
|
987 |
ultimately show "eventually (\<lambda>x. Z \<le> f x + g x) F" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
988 |
by eventually_elim simp |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
989 |
qed |
|
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
990 |
|
| 50331 | 991 |
lemma tendsto_divide_0: |
992 |
fixes f :: "_ \<Rightarrow> 'a\<Colon>{real_normed_div_algebra, division_ring_inverse_zero}"
|
|
993 |
assumes f: "(f ---> c) F" |
|
994 |
assumes g: "LIM x F. g x :> at_infinity" |
|
995 |
shows "((\<lambda>x. f x / g x) ---> 0) F" |
|
996 |
using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse) |
|
997 |
||
998 |
lemma linear_plus_1_le_power: |
|
999 |
fixes x :: real |
|
1000 |
assumes x: "0 \<le> x" |
|
1001 |
shows "real n * x + 1 \<le> (x + 1) ^ n" |
|
1002 |
proof (induct n) |
|
1003 |
case (Suc n) |
|
1004 |
have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)" |
|
1005 |
by (simp add: field_simps real_of_nat_Suc mult_nonneg_nonneg x) |
|
1006 |
also have "\<dots> \<le> (x + 1)^Suc n" |
|
1007 |
using Suc x by (simp add: mult_left_mono) |
|
1008 |
finally show ?case . |
|
1009 |
qed simp |
|
1010 |
||
1011 |
lemma filterlim_realpow_sequentially_gt1: |
|
1012 |
fixes x :: "'a :: real_normed_div_algebra" |
|
1013 |
assumes x[arith]: "1 < norm x" |
|
1014 |
shows "LIM n sequentially. x ^ n :> at_infinity" |
|
1015 |
proof (intro filterlim_at_infinity[THEN iffD2] allI impI) |
|
1016 |
fix y :: real assume "0 < y" |
|
1017 |
have "0 < norm x - 1" by simp |
|
1018 |
then obtain N::nat where "y < real N * (norm x - 1)" by (blast dest: reals_Archimedean3) |
|
1019 |
also have "\<dots> \<le> real N * (norm x - 1) + 1" by simp |
|
1020 |
also have "\<dots> \<le> (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp |
|
1021 |
also have "\<dots> = norm x ^ N" by simp |
|
1022 |
finally have "\<forall>n\<ge>N. y \<le> norm x ^ n" |
|
1023 |
by (metis order_less_le_trans power_increasing order_less_imp_le x) |
|
1024 |
then show "eventually (\<lambda>n. y \<le> norm (x ^ n)) sequentially" |
|
1025 |
unfolding eventually_sequentially |
|
1026 |
by (auto simp: norm_power) |
|
1027 |
qed simp |
|
1028 |
||
| 51471 | 1029 |
|
1030 |
(* Unfortunately eventually_within was overwritten by Multivariate_Analysis. |
|
1031 |
Hence it was references as Limits.within, but now it is Basic_Topology.eventually_within *) |
|
1032 |
lemmas eventually_within = eventually_within |
|
1033 |
||
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
1034 |
end |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1035 |