author | paulson <lp15@cam.ac.uk> |
Thu, 26 Jan 2023 13:59:51 +0000 | |
changeset 77102 | 780161d4b55c |
parent 76724 | 7ff71bdcf731 |
child 77221 | 0cdb384bf56a |
permissions | -rw-r--r-- |
52265 | 1 |
(* Title: HOL/Limits.thy |
51526 | 2 |
Author: Brian Huffman |
3 |
Author: Jacques D. Fleuriot, University of Cambridge |
|
4 |
Author: Lawrence C Paulson |
|
5 |
Author: Jeremy Avigad |
|
31349
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 |
|
60758 | 8 |
section \<open>Limits on Real Vector Spaces\<close> |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
9 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
10 |
theory Limits |
63546 | 11 |
imports Real_Vector_Spaces |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
12 |
begin |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
13 |
|
73795 | 14 |
lemma range_mult [simp]: |
15 |
fixes a::"real" shows "range ((*) a) = (if a=0 then {0} else UNIV)" |
|
16 |
by (simp add: surj_def) (meson dvdE dvd_field_iff) |
|
17 |
||
18 |
||
60758 | 19 |
subsection \<open>Filter going to infinity norm\<close> |
51526 | 20 |
|
63546 | 21 |
definition at_infinity :: "'a::real_normed_vector filter" |
22 |
where "at_infinity = (INF r. principal {x. r \<le> norm x})" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
23 |
|
57276 | 24 |
lemma eventually_at_infinity: "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)" |
25 |
unfolding at_infinity_def |
|
26 |
by (subst eventually_INF_base) |
|
27 |
(auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) |
|
31392 | 28 |
|
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62369
diff
changeset
|
29 |
corollary eventually_at_infinity_pos: |
63546 | 30 |
"eventually p at_infinity \<longleftrightarrow> (\<exists>b. 0 < b \<and> (\<forall>x. norm x \<ge> b \<longrightarrow> p x))" |
68614 | 31 |
unfolding eventually_at_infinity |
32 |
by (meson le_less_trans norm_ge_zero not_le zero_less_one) |
|
63546 | 33 |
|
34 |
lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" |
|
68614 | 35 |
proof - |
36 |
have 1: "\<lbrakk>\<forall>n\<ge>u. A n; \<forall>n\<le>v. A n\<rbrakk> |
|
37 |
\<Longrightarrow> \<exists>b. \<forall>x. b \<le> \<bar>x\<bar> \<longrightarrow> A x" for A and u v::real |
|
38 |
by (rule_tac x="max (- v) u" in exI) (auto simp: abs_real_def) |
|
39 |
have 2: "\<forall>x. u \<le> \<bar>x\<bar> \<longrightarrow> A x \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. A n" for A and u::real |
|
40 |
by (meson abs_less_iff le_cases less_le_not_le) |
|
41 |
have 3: "\<forall>x. u \<le> \<bar>x\<bar> \<longrightarrow> A x \<Longrightarrow> \<exists>N. \<forall>n\<le>N. A n" for A and u::real |
|
42 |
by (metis (full_types) abs_ge_self abs_minus_cancel le_minus_iff order_trans) |
|
43 |
show ?thesis |
|
68615 | 44 |
by (auto simp: filter_eq_iff eventually_sup eventually_at_infinity |
68614 | 45 |
eventually_at_top_linorder eventually_at_bot_linorder intro: 1 2 3) |
46 |
qed |
|
50325 | 47 |
|
57276 | 48 |
lemma at_top_le_at_infinity: "at_top \<le> (at_infinity :: real filter)" |
50325 | 49 |
unfolding at_infinity_eq_at_top_bot by simp |
50 |
||
57276 | 51 |
lemma at_bot_le_at_infinity: "at_bot \<le> (at_infinity :: real filter)" |
50325 | 52 |
unfolding at_infinity_eq_at_top_bot by simp |
53 |
||
63546 | 54 |
lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F" |
55 |
for f :: "_ \<Rightarrow> real" |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
56 |
by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
57 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
58 |
lemma filterlim_real_at_infinity_sequentially: "filterlim real at_infinity sequentially" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
59 |
by (simp add: filterlim_at_top_imp_at_infinity filterlim_real_sequentially) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
60 |
|
63546 | 61 |
lemma lim_infinity_imp_sequentially: "(f \<longlongrightarrow> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) \<longlongrightarrow> l) sequentially" |
62 |
by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) |
|
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
63 |
|
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
64 |
|
60758 | 65 |
subsubsection \<open>Boundedness\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
66 |
|
63546 | 67 |
definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" |
68 |
where Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)" |
|
69 |
||
70 |
abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" |
|
71 |
where "Bseq X \<equiv> Bfun X sequentially" |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
72 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
73 |
lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
74 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
75 |
lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
76 |
unfolding Bfun_metric_def by (subst eventually_sequentially_seg) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
77 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
78 |
lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
79 |
unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) |
31355 | 80 |
|
63546 | 81 |
lemma Bfun_def: "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)" |
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
82 |
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
|
83 |
proof safe |
63546 | 84 |
fix y K |
85 |
assume K: "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F" |
|
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
86 |
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
|
87 |
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
|
88 |
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
|
89 |
by eventually_elim auto |
60758 | 90 |
with \<open>0 < K\<close> show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F" |
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
91 |
by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto |
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62369
diff
changeset
|
92 |
qed (force simp del: norm_conv_dist [symmetric]) |
31355 | 93 |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
94 |
lemma BfunI: |
63546 | 95 |
assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" |
96 |
shows "Bfun f F" |
|
97 |
unfolding Bfun_def |
|
31355 | 98 |
proof (intro exI conjI allI) |
99 |
show "0 < max K 1" by simp |
|
44195 | 100 |
show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F" |
63546 | 101 |
using K by (rule eventually_mono) simp |
31355 | 102 |
qed |
103 |
||
104 |
lemma BfunE: |
|
44195 | 105 |
assumes "Bfun f F" |
106 |
obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F" |
|
63546 | 107 |
using assms unfolding Bfun_def by blast |
31355 | 108 |
|
68614 | 109 |
lemma Cauchy_Bseq: |
110 |
assumes "Cauchy X" shows "Bseq X" |
|
111 |
proof - |
|
112 |
have "\<exists>y K. 0 < K \<and> (\<exists>N. \<forall>n\<ge>N. dist (X n) y \<le> K)" |
|
113 |
if "\<And>m n. \<lbrakk>m \<ge> M; n \<ge> M\<rbrakk> \<Longrightarrow> dist (X m) (X n) < 1" for M |
|
114 |
by (meson order.order_iff_strict that zero_less_one) |
|
115 |
with assms show ?thesis |
|
116 |
by (force simp: Cauchy_def Bfun_metric_def eventually_sequentially) |
|
117 |
qed |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
118 |
|
60758 | 119 |
subsubsection \<open>Bounded Sequences\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
120 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
121 |
lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
122 |
by (intro BfunI) (auto simp: eventually_sequentially) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
123 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
124 |
lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
125 |
unfolding Bfun_def eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
126 |
proof safe |
63546 | 127 |
fix N K |
128 |
assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K" |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
129 |
then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K" |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54263
diff
changeset
|
130 |
by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
131 |
(auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
132 |
qed auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
133 |
|
63546 | 134 |
lemma BseqE: "Bseq X \<Longrightarrow> (\<And>K. 0 < K \<Longrightarrow> \<forall>n. norm (X n) \<le> K \<Longrightarrow> Q) \<Longrightarrow> Q" |
135 |
unfolding Bseq_def by auto |
|
136 |
||
137 |
lemma BseqD: "Bseq X \<Longrightarrow> \<exists>K. 0 < K \<and> (\<forall>n. norm (X n) \<le> K)" |
|
138 |
by (simp add: Bseq_def) |
|
139 |
||
140 |
lemma BseqI: "0 < K \<Longrightarrow> \<forall>n. norm (X n) \<le> K \<Longrightarrow> Bseq X" |
|
68615 | 141 |
by (auto simp: Bseq_def) |
63546 | 142 |
|
143 |
lemma Bseq_bdd_above: "Bseq X \<Longrightarrow> bdd_above (range X)" |
|
144 |
for X :: "nat \<Rightarrow> real" |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
145 |
proof (elim BseqE, intro bdd_aboveI2) |
63546 | 146 |
fix K n |
147 |
assume "0 < K" "\<forall>n. norm (X n) \<le> K" |
|
148 |
then show "X n \<le> K" |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
149 |
by (auto elim!: allE[of _ n]) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
150 |
qed |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
151 |
|
63546 | 152 |
lemma Bseq_bdd_above': "Bseq X \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))" |
153 |
for X :: "nat \<Rightarrow> 'a :: real_normed_vector" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
154 |
proof (elim BseqE, intro bdd_aboveI2) |
63546 | 155 |
fix K n |
156 |
assume "0 < K" "\<forall>n. norm (X n) \<le> K" |
|
157 |
then show "norm (X n) \<le> K" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
158 |
by (auto elim!: allE[of _ n]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
159 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
160 |
|
63546 | 161 |
lemma Bseq_bdd_below: "Bseq X \<Longrightarrow> bdd_below (range X)" |
162 |
for X :: "nat \<Rightarrow> real" |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
163 |
proof (elim BseqE, intro bdd_belowI2) |
63546 | 164 |
fix K n |
165 |
assume "0 < K" "\<forall>n. norm (X n) \<le> K" |
|
166 |
then show "- K \<le> X n" |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
167 |
by (auto elim!: allE[of _ n]) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
168 |
qed |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
169 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
170 |
lemma Bseq_eventually_mono: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
171 |
assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) sequentially" "Bseq g" |
63546 | 172 |
shows "Bseq f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
173 |
proof - |
67958
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
174 |
from assms(2) obtain K where "0 < K" and "eventually (\<lambda>n. norm (g n) \<le> K) sequentially" |
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
175 |
unfolding Bfun_def by fast |
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
176 |
with assms(1) have "eventually (\<lambda>n. norm (f n) \<le> K) sequentially" |
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
177 |
by (fast elim: eventually_elim2 order_trans) |
69272 | 178 |
with \<open>0 < K\<close> show "Bseq f" |
67958
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
179 |
unfolding Bfun_def by fast |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
180 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
181 |
|
63546 | 182 |
lemma lemma_NBseq_def: "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
183 |
proof safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
184 |
fix K :: real |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
185 |
from reals_Archimedean2 obtain n :: nat where "K < real n" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
186 |
then have "K \<le> real (Suc n)" by auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
187 |
moreover assume "\<forall>m. norm (X m) \<le> K" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
188 |
ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
189 |
by (blast intro: order_trans) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
190 |
then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" .. |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
191 |
next |
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
192 |
show "\<And>N. \<forall>n. norm (X n) \<le> real (Suc N) \<Longrightarrow> \<exists>K>0. \<forall>n. norm (X n) \<le> K" |
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
193 |
using of_nat_0_less_iff by blast |
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
194 |
qed |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
195 |
|
63546 | 196 |
text \<open>Alternative definition for \<open>Bseq\<close>.\<close> |
197 |
lemma Bseq_iff: "Bseq X \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
|
198 |
by (simp add: Bseq_def) (simp add: lemma_NBseq_def) |
|
199 |
||
200 |
lemma lemma_NBseq_def2: "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
|
68614 | 201 |
proof - |
202 |
have *: "\<And>N. \<forall>n. norm (X n) \<le> 1 + real N \<Longrightarrow> |
|
203 |
\<exists>N. \<forall>n. norm (X n) < 1 + real N" |
|
204 |
by (metis add.commute le_less_trans less_add_one of_nat_Suc) |
|
205 |
then show ?thesis |
|
206 |
unfolding lemma_NBseq_def |
|
207 |
by (metis less_le_not_le not_less_iff_gr_or_eq of_nat_Suc) |
|
208 |
qed |
|
63546 | 209 |
|
210 |
text \<open>Yet another definition for Bseq.\<close> |
|
211 |
lemma Bseq_iff1a: "Bseq X \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) < real (Suc N))" |
|
212 |
by (simp add: Bseq_def lemma_NBseq_def2) |
|
213 |
||
214 |
subsubsection \<open>A Few More Equivalence Theorems for Boundedness\<close> |
|
215 |
||
216 |
text \<open>Alternative formulation for boundedness.\<close> |
|
217 |
lemma Bseq_iff2: "Bseq X \<longleftrightarrow> (\<exists>k > 0. \<exists>x. \<forall>n. norm (X n + - x) \<le> k)" |
|
68614 | 218 |
by (metis BseqE BseqI' add.commute add_cancel_right_left add_uminus_conv_diff norm_add_leD |
219 |
norm_minus_cancel norm_minus_commute) |
|
63546 | 220 |
|
221 |
text \<open>Alternative formulation for boundedness.\<close> |
|
222 |
lemma Bseq_iff3: "Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" |
|
223 |
(is "?P \<longleftrightarrow> ?Q") |
|
53602 | 224 |
proof |
225 |
assume ?P |
|
63546 | 226 |
then obtain K where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" |
68615 | 227 |
by (auto simp: Bseq_def) |
53602 | 228 |
from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
229 |
from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
230 |
by (auto intro: order_trans norm_triangle_ineq4) |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
231 |
then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
232 |
by simp |
60758 | 233 |
with \<open>0 < K + norm (X 0)\<close> show ?Q by blast |
53602 | 234 |
next |
63546 | 235 |
assume ?Q |
68615 | 236 |
then show ?P by (auto simp: Bseq_iff2) |
53602 | 237 |
qed |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
238 |
|
63546 | 239 |
|
240 |
subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close> |
|
241 |
||
242 |
lemma Bseq_minus_iff: "Bseq (\<lambda>n. - (X n) :: 'a::real_normed_vector) \<longleftrightarrow> Bseq X" |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
243 |
by (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
244 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
245 |
lemma Bseq_add: |
63546 | 246 |
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" |
247 |
assumes "Bseq f" |
|
248 |
shows "Bseq (\<lambda>x. f x + c)" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
249 |
proof - |
63546 | 250 |
from assms obtain K where K: "\<And>x. norm (f x) \<le> K" |
251 |
unfolding Bseq_def by blast |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
252 |
{ |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
253 |
fix x :: nat |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
254 |
have "norm (f x + c) \<le> norm (f x) + norm c" by (rule norm_triangle_ineq) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
255 |
also have "norm (f x) \<le> K" by (rule K) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
256 |
finally have "norm (f x + c) \<le> K + norm c" by simp |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
257 |
} |
63546 | 258 |
then show ?thesis by (rule BseqI') |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
259 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
260 |
|
63546 | 261 |
lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq f" |
262 |
for f :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
263 |
using Bseq_add[of f c] Bseq_add[of "\<lambda>x. f x + c" "-c"] by auto |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
264 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
265 |
lemma Bseq_mult: |
63546 | 266 |
fixes f g :: "nat \<Rightarrow> 'a::real_normed_field" |
267 |
assumes "Bseq f" and "Bseq g" |
|
268 |
shows "Bseq (\<lambda>x. f x * g x)" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
269 |
proof - |
63546 | 270 |
from assms obtain K1 K2 where K: "norm (f x) \<le> K1" "K1 > 0" "norm (g x) \<le> K2" "K2 > 0" |
271 |
for x |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
272 |
unfolding Bseq_def by blast |
63546 | 273 |
then have "norm (f x * g x) \<le> K1 * K2" for x |
274 |
by (auto simp: norm_mult intro!: mult_mono) |
|
275 |
then show ?thesis by (rule BseqI') |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
276 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
277 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
278 |
lemma Bfun_const [simp]: "Bfun (\<lambda>_. c) F" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
279 |
unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
280 |
|
63546 | 281 |
lemma Bseq_cmult_iff: |
282 |
fixes c :: "'a::real_normed_field" |
|
283 |
assumes "c \<noteq> 0" |
|
284 |
shows "Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
285 |
proof |
63546 | 286 |
assume "Bseq (\<lambda>x. c * f x)" |
287 |
with Bfun_const have "Bseq (\<lambda>x. inverse c * (c * f x))" |
|
288 |
by (rule Bseq_mult) |
|
289 |
with \<open>c \<noteq> 0\<close> show "Bseq f" |
|
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70804
diff
changeset
|
290 |
by (simp add: field_split_simps) |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
291 |
qed (intro Bseq_mult Bfun_const) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
292 |
|
63546 | 293 |
lemma Bseq_subseq: "Bseq f \<Longrightarrow> Bseq (\<lambda>x. f (g x))" |
294 |
for f :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
295 |
unfolding Bseq_def by auto |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
296 |
|
63546 | 297 |
lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq f" |
298 |
for f :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
299 |
using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
300 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
301 |
lemma increasing_Bseq_subseq_iff: |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
65680
diff
changeset
|
302 |
assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a::real_normed_vector) \<le> norm (f y)" "strict_mono g" |
63546 | 303 |
shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
304 |
proof |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
305 |
assume "Bseq (\<lambda>x. f (g x))" |
63546 | 306 |
then obtain K where K: "\<And>x. norm (f (g x)) \<le> K" |
307 |
unfolding Bseq_def by auto |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
308 |
{ |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
309 |
fix x :: nat |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
310 |
from filterlim_subseq[OF assms(2)] obtain y where "g y \<ge> x" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
311 |
by (auto simp: filterlim_at_top eventually_at_top_linorder) |
63546 | 312 |
then have "norm (f x) \<le> norm (f (g y))" |
313 |
using assms(1) by blast |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
314 |
also have "norm (f (g y)) \<le> K" by (rule K) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
315 |
finally have "norm (f x) \<le> K" . |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
316 |
} |
63546 | 317 |
then show "Bseq f" by (rule BseqI') |
318 |
qed (use Bseq_subseq[of f g] in simp_all) |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
319 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
320 |
lemma nonneg_incseq_Bseq_subseq_iff: |
63546 | 321 |
fixes f :: "nat \<Rightarrow> real" |
322 |
and g :: "nat \<Rightarrow> nat" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
65680
diff
changeset
|
323 |
assumes "\<And>x. f x \<ge> 0" "incseq f" "strict_mono g" |
63546 | 324 |
shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
325 |
using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
326 |
|
63546 | 327 |
lemma Bseq_eq_bounded: "range f \<subseteq> {a..b} \<Longrightarrow> Bseq f" |
328 |
for a b :: real |
|
68614 | 329 |
proof (rule BseqI'[where K="max (norm a) (norm b)"]) |
330 |
fix n assume "range f \<subseteq> {a..b}" |
|
331 |
then have "f n \<in> {a..b}" |
|
332 |
by blast |
|
333 |
then show "norm (f n) \<le> max (norm a) (norm b)" |
|
334 |
by auto |
|
335 |
qed |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
336 |
|
63546 | 337 |
lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> B \<Longrightarrow> Bseq X" |
338 |
for B :: real |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
339 |
by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
340 |
|
63546 | 341 |
lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. B \<le> X i \<Longrightarrow> Bseq X" |
342 |
for B :: real |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
343 |
by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
344 |
|
63546 | 345 |
|
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
346 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Polynomal function extremal theorem, from HOL Light\<close> |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
347 |
|
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
348 |
lemma polyfun_extremal_lemma: |
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
349 |
fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
350 |
assumes "0 < e" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
351 |
shows "\<exists>M. \<forall>z. M \<le> norm(z) \<longrightarrow> norm (\<Sum>i\<le>n. c(i) * z^i) \<le> e * norm(z) ^ (Suc n)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
352 |
proof (induct n) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
353 |
case 0 with assms |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
354 |
show ?case |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
355 |
apply (rule_tac x="norm (c 0) / e" in exI) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
356 |
apply (auto simp: field_simps) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
357 |
done |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
358 |
next |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
359 |
case (Suc n) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
360 |
obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
361 |
using Suc assms by blast |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
362 |
show ?case |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
363 |
proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
364 |
fix z::'a |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
365 |
assume z1: "M \<le> norm z" and "1 + norm (c (Suc n)) / e \<le> norm z" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
366 |
then have z2: "e + norm (c (Suc n)) \<le> e * norm z" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
367 |
using assms by (simp add: field_simps) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
368 |
have "norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
369 |
using M [OF z1] by simp |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
370 |
then have "norm (\<Sum>i\<le>n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
371 |
by simp |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
372 |
then have "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
373 |
by (blast intro: norm_triangle_le elim: ) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
374 |
also have "... \<le> (e + norm (c (Suc n))) * norm z ^ Suc n" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
375 |
by (simp add: norm_power norm_mult algebra_simps) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
376 |
also have "... \<le> (e * norm z) * norm z ^ Suc n" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
377 |
by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
378 |
finally show "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc (Suc n)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
379 |
by simp |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
380 |
qed |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
381 |
qed |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
382 |
|
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
383 |
lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
384 |
fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
385 |
assumes k: "c k \<noteq> 0" "1\<le>k" and kn: "k\<le>n" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
386 |
shows "eventually (\<lambda>z. norm (\<Sum>i\<le>n. c(i) * z^i) \<ge> B) at_infinity" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
387 |
using kn |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
388 |
proof (induction n) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
389 |
case 0 |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
390 |
then show ?case |
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
391 |
using k by simp |
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
392 |
next |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
393 |
case (Suc m) |
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
394 |
show ?case |
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
395 |
proof (cases "c (Suc m) = 0") |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
396 |
case True |
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
397 |
then show ?thesis using Suc k |
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
398 |
by auto (metis antisym_conv less_eq_Suc_le not_le) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
399 |
next |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
400 |
case False |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
401 |
then obtain M where M: |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
402 |
"\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>m. c i * z^i) \<le> norm (c (Suc m)) / 2 * norm z ^ Suc m" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
403 |
using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
404 |
by auto |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
405 |
have "\<exists>b. \<forall>z. b \<le> norm z \<longrightarrow> B \<le> norm (\<Sum>i\<le>Suc m. c i * z^i)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
406 |
proof (rule exI [where x="max M (max 1 (\<bar>B\<bar> / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
407 |
fix z::'a |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
408 |
assume z1: "M \<le> norm z" "1 \<le> norm z" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
409 |
and "\<bar>B\<bar> * 2 / norm (c (Suc m)) \<le> norm z" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
410 |
then have z2: "\<bar>B\<bar> \<le> norm (c (Suc m)) * norm z / 2" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
411 |
using False by (simp add: field_simps) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
412 |
have nz: "norm z \<le> norm z ^ Suc m" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
413 |
by (metis \<open>1 \<le> norm z\<close> One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
414 |
have *: "\<And>y x. norm (c (Suc m)) * norm z / 2 \<le> norm y - norm x \<Longrightarrow> B \<le> norm (x + y)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
415 |
by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
416 |
have "norm z * norm (c (Suc m)) + 2 * norm (\<Sum>i\<le>m. c i * z^i) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
417 |
\<le> norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
418 |
using M [of z] Suc z1 by auto |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
419 |
also have "... \<le> 2 * (norm (c (Suc m)) * norm z ^ Suc m)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
420 |
using nz by (simp add: mult_mono del: power_Suc) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
421 |
finally show "B \<le> norm ((\<Sum>i\<le>m. c i * z^i) + c (Suc m) * z ^ Suc m)" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
422 |
using Suc.IH |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
423 |
apply (auto simp: eventually_at_infinity) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
424 |
apply (rule *) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
425 |
apply (simp add: field_simps norm_mult norm_power) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
426 |
done |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
427 |
qed |
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
428 |
then show ?thesis |
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
429 |
by (simp add: eventually_at_infinity) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
430 |
qed |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
431 |
qed |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
432 |
|
60758 | 433 |
subsection \<open>Convergence to Zero\<close> |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
434 |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
435 |
definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool" |
44195 | 436 |
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
|
437 |
|
63546 | 438 |
lemma ZfunI: "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F" |
439 |
by (simp add: Zfun_def) |
|
440 |
||
441 |
lemma ZfunD: "Zfun f F \<Longrightarrow> 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F" |
|
442 |
by (simp add: Zfun_def) |
|
443 |
||
444 |
lemma Zfun_ssubst: "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
|
445 |
unfolding Zfun_def by (auto elim!: eventually_rev_mp) |
31355 | 446 |
|
44195 | 447 |
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
|
448 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
449 |
|
44195 | 450 |
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
|
451 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
452 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
453 |
lemma Zfun_imp_Zfun: |
44195 | 454 |
assumes f: "Zfun f F" |
63546 | 455 |
and g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F" |
44195 | 456 |
shows "Zfun (\<lambda>x. g x) F" |
63546 | 457 |
proof (cases "0 < K") |
458 |
case K: True |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
459 |
show ?thesis |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
460 |
proof (rule ZfunI) |
63546 | 461 |
fix r :: real |
462 |
assume "0 < r" |
|
463 |
then have "0 < r / K" using K by simp |
|
44195 | 464 |
then have "eventually (\<lambda>x. norm (f x) < r / K) F" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
465 |
using ZfunD [OF f] by blast |
44195 | 466 |
with g show "eventually (\<lambda>x. norm (g x) < r) F" |
46887 | 467 |
proof eventually_elim |
468 |
case (elim x) |
|
63546 | 469 |
then have "norm (f x) * K < r" |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
470 |
by (simp add: pos_less_divide_eq K) |
63546 | 471 |
then show ?case |
46887 | 472 |
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
|
473 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
474 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
475 |
next |
63546 | 476 |
case False |
477 |
then have K: "K \<le> 0" by (simp only: not_less) |
|
31355 | 478 |
show ?thesis |
479 |
proof (rule ZfunI) |
|
480 |
fix r :: real |
|
481 |
assume "0 < r" |
|
44195 | 482 |
from g show "eventually (\<lambda>x. norm (g x) < r) F" |
46887 | 483 |
proof eventually_elim |
484 |
case (elim x) |
|
485 |
also have "norm (f x) * K \<le> norm (f x) * 0" |
|
31355 | 486 |
using K norm_ge_zero by (rule mult_left_mono) |
46887 | 487 |
finally show ?case |
60758 | 488 |
using \<open>0 < r\<close> by simp |
31355 | 489 |
qed |
490 |
qed |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
491 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
492 |
|
63546 | 493 |
lemma Zfun_le: "Zfun g F \<Longrightarrow> \<forall>x. norm (f x) \<le> norm (g x) \<Longrightarrow> Zfun f F" |
494 |
by (erule Zfun_imp_Zfun [where K = 1]) simp |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
495 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
496 |
lemma Zfun_add: |
63546 | 497 |
assumes f: "Zfun f F" |
498 |
and g: "Zfun g F" |
|
44195 | 499 |
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
|
500 |
proof (rule ZfunI) |
63546 | 501 |
fix r :: real |
502 |
assume "0 < r" |
|
503 |
then have r: "0 < r / 2" by simp |
|
44195 | 504 |
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
|
505 |
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
|
506 |
moreover |
44195 | 507 |
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
|
508 |
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
|
509 |
ultimately |
44195 | 510 |
show "eventually (\<lambda>x. norm (f x + g x) < r) F" |
46887 | 511 |
proof eventually_elim |
512 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
513 |
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
|
514 |
by (rule norm_triangle_ineq) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
515 |
also have "\<dots> < r/2 + r/2" |
46887 | 516 |
using elim by (rule add_strict_mono) |
517 |
finally show ?case |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
518 |
by simp |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
519 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
520 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
521 |
|
44195 | 522 |
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
|
523 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
524 |
|
63546 | 525 |
lemma Zfun_diff: "Zfun f F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
526 |
using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
527 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
528 |
lemma (in bounded_linear) Zfun: |
44195 | 529 |
assumes g: "Zfun g F" |
530 |
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
|
531 |
proof - |
63546 | 532 |
obtain K where "norm (f x) \<le> norm x * K" for x |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
533 |
using bounded by blast |
44195 | 534 |
then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F" |
31355 | 535 |
by simp |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
536 |
with g show ?thesis |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
537 |
by (rule Zfun_imp_Zfun) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
538 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
539 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
540 |
lemma (in bounded_bilinear) Zfun: |
44195 | 541 |
assumes f: "Zfun f F" |
63546 | 542 |
and g: "Zfun g F" |
44195 | 543 |
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
|
544 |
proof (rule ZfunI) |
63546 | 545 |
fix r :: real |
546 |
assume r: "0 < r" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
547 |
obtain K where K: "0 < K" |
63546 | 548 |
and norm_le: "norm (x ** y) \<le> norm x * norm y * K" for x y |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
549 |
using pos_bounded by blast |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
550 |
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
|
551 |
by (rule positive_imp_inverse_positive) |
44195 | 552 |
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
|
553 |
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
|
554 |
moreover |
44195 | 555 |
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
|
556 |
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
|
557 |
ultimately |
44195 | 558 |
show "eventually (\<lambda>x. norm (f x ** g x) < r) F" |
46887 | 559 |
proof eventually_elim |
560 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
561 |
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
|
562 |
by (rule norm_le) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
563 |
also have "norm (f x) * norm (g x) * K < r * inverse K * K" |
46887 | 564 |
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
|
565 |
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
|
566 |
by simp |
46887 | 567 |
finally show ?case . |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
568 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
569 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
570 |
|
63546 | 571 |
lemma (in bounded_bilinear) Zfun_left: "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
|
572 |
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
|
573 |
|
63546 | 574 |
lemma (in bounded_bilinear) Zfun_right: "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
|
575 |
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
|
576 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
577 |
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
|
578 |
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
|
579 |
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
|
580 |
|
61973 | 581 |
lemma tendsto_Zfun_iff: "(f \<longlongrightarrow> 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
|
582 |
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
|
583 |
|
63546 | 584 |
lemma tendsto_0_le: |
585 |
"(f \<longlongrightarrow> 0) F \<Longrightarrow> eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F \<Longrightarrow> (g \<longlongrightarrow> 0) F" |
|
56366 | 586 |
by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) |
587 |
||
63546 | 588 |
|
60758 | 589 |
subsubsection \<open>Distance and norms\<close> |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
590 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
591 |
lemma tendsto_dist [tendsto_intros]: |
63546 | 592 |
fixes l m :: "'a::metric_space" |
593 |
assumes f: "(f \<longlongrightarrow> l) F" |
|
594 |
and g: "(g \<longlongrightarrow> m) F" |
|
61973 | 595 |
shows "((\<lambda>x. dist (f x) (g x)) \<longlongrightarrow> dist l m) F" |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
596 |
proof (rule tendstoI) |
63546 | 597 |
fix e :: real |
598 |
assume "0 < e" |
|
599 |
then have e2: "0 < e/2" by simp |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
600 |
from tendstoD [OF f e2] tendstoD [OF g e2] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
601 |
show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
602 |
proof (eventually_elim) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
603 |
case (elim x) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
604 |
then show "dist (dist (f x) (g x)) (dist l m) < e" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
605 |
unfolding dist_real_def |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
606 |
using dist_triangle2 [of "f x" "g x" "l"] |
63546 | 607 |
and dist_triangle2 [of "g x" "l" "m"] |
608 |
and dist_triangle3 [of "l" "m" "f x"] |
|
609 |
and dist_triangle [of "f x" "m" "g x"] |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
610 |
by arith |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
611 |
qed |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
612 |
qed |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
613 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
614 |
lemma continuous_dist[continuous_intros]: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
615 |
fixes f g :: "_ \<Rightarrow> 'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
616 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
617 |
unfolding continuous_def by (rule tendsto_dist) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
618 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
619 |
lemma continuous_on_dist[continuous_intros]: |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
620 |
fixes f g :: "_ \<Rightarrow> 'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
621 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
622 |
unfolding continuous_on_def by (auto intro: tendsto_dist) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
623 |
|
69918
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
624 |
lemma continuous_at_dist: "isCont (dist a) b" |
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
625 |
using continuous_on_dist [OF continuous_on_const continuous_on_id] continuous_on_eq_continuous_within by blast |
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
626 |
|
63546 | 627 |
lemma tendsto_norm [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> norm a) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
628 |
unfolding norm_conv_dist by (intro tendsto_intros) |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
629 |
|
63546 | 630 |
lemma continuous_norm [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))" |
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
|
631 |
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
|
632 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
633 |
lemma continuous_on_norm [continuous_intros]: |
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
|
634 |
"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
|
635 |
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
|
636 |
|
71167
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
637 |
lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
638 |
by (intro continuous_on_id continuous_on_norm) |
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents:
70999
diff
changeset
|
639 |
|
63546 | 640 |
lemma tendsto_norm_zero: "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F" |
641 |
by (drule tendsto_norm) simp |
|
642 |
||
643 |
lemma tendsto_norm_zero_cancel: "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> 0) F" |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
644 |
unfolding tendsto_iff dist_norm by simp |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
645 |
|
63546 | 646 |
lemma tendsto_norm_zero_iff: "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
647 |
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
|
648 |
|
63546 | 649 |
lemma tendsto_rabs [tendsto_intros]: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> \<bar>l\<bar>) F" |
650 |
for l :: real |
|
651 |
by (fold real_norm_def) (rule tendsto_norm) |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
652 |
|
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
|
653 |
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
|
654 |
"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
|
655 |
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
|
656 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
657 |
lemma continuous_on_rabs [continuous_intros]: |
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
|
658 |
"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
|
659 |
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
|
660 |
|
63546 | 661 |
lemma tendsto_rabs_zero: "(f \<longlongrightarrow> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> 0) F" |
662 |
by (fold real_norm_def) (rule tendsto_norm_zero) |
|
663 |
||
664 |
lemma tendsto_rabs_zero_cancel: "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<Longrightarrow> (f \<longlongrightarrow> 0) F" |
|
665 |
by (fold real_norm_def) (rule tendsto_norm_zero_cancel) |
|
666 |
||
667 |
lemma tendsto_rabs_zero_iff: "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F" |
|
668 |
by (fold real_norm_def) (rule tendsto_norm_zero_iff) |
|
669 |
||
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
670 |
|
62368 | 671 |
subsection \<open>Topological Monoid\<close> |
672 |
||
673 |
class topological_monoid_add = topological_space + monoid_add + |
|
674 |
assumes tendsto_add_Pair: "LIM x (nhds a \<times>\<^sub>F nhds b). fst x + snd x :> nhds (a + b)" |
|
675 |
||
676 |
class topological_comm_monoid_add = topological_monoid_add + comm_monoid_add |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
677 |
|
31565 | 678 |
lemma tendsto_add [tendsto_intros]: |
62368 | 679 |
fixes a b :: "'a::topological_monoid_add" |
680 |
shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> a + b) F" |
|
681 |
using filterlim_compose[OF tendsto_add_Pair, of "\<lambda>x. (f x, g x)" a b F] |
|
682 |
by (simp add: nhds_prod[symmetric] tendsto_Pair) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
683 |
|
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
|
684 |
lemma continuous_add [continuous_intros]: |
62368 | 685 |
fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add" |
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
|
686 |
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
|
687 |
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
|
688 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
689 |
lemma continuous_on_add [continuous_intros]: |
62368 | 690 |
fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add" |
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
|
691 |
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
|
692 |
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
|
693 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
694 |
lemma tendsto_add_zero: |
62368 | 695 |
fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add" |
63546 | 696 |
shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> 0) F" |
697 |
by (drule (1) tendsto_add) simp |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
698 |
|
64267 | 699 |
lemma tendsto_sum [tendsto_intros]: |
62368 | 700 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_add" |
63915 | 701 |
shows "(\<And>i. i \<in> I \<Longrightarrow> (f i \<longlongrightarrow> a i) F) \<Longrightarrow> ((\<lambda>x. \<Sum>i\<in>I. f i x) \<longlongrightarrow> (\<Sum>i\<in>I. a i)) F" |
702 |
by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_add) |
|
62368 | 703 |
|
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67399
diff
changeset
|
704 |
lemma tendsto_null_sum: |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67399
diff
changeset
|
705 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_add" |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67399
diff
changeset
|
706 |
assumes "\<And>i. i \<in> I \<Longrightarrow> ((\<lambda>x. f x i) \<longlongrightarrow> 0) F" |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67399
diff
changeset
|
707 |
shows "((\<lambda>i. sum (f i) I) \<longlongrightarrow> 0) F" |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67399
diff
changeset
|
708 |
using tendsto_sum [of I "\<lambda>x y. f y x" "\<lambda>x. 0"] assms by simp |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67399
diff
changeset
|
709 |
|
64267 | 710 |
lemma continuous_sum [continuous_intros]: |
62368 | 711 |
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add" |
63301 | 712 |
shows "(\<And>i. i \<in> I \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>I. f i x)" |
64267 | 713 |
unfolding continuous_def by (rule tendsto_sum) |
714 |
||
715 |
lemma continuous_on_sum [continuous_intros]: |
|
62368 | 716 |
fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::topological_comm_monoid_add" |
63301 | 717 |
shows "(\<And>i. i \<in> I \<Longrightarrow> continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<Sum>i\<in>I. f i x)" |
64267 | 718 |
unfolding continuous_on_def by (auto intro: tendsto_sum) |
62368 | 719 |
|
62369 | 720 |
instance nat :: topological_comm_monoid_add |
63546 | 721 |
by standard |
722 |
(simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) |
|
62369 | 723 |
|
724 |
instance int :: topological_comm_monoid_add |
|
63546 | 725 |
by standard |
726 |
(simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) |
|
727 |
||
62369 | 728 |
|
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
729 |
subsubsection \<open>Topological group\<close> |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
730 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
731 |
class topological_group_add = topological_monoid_add + group_add + |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
732 |
assumes tendsto_uminus_nhds: "(uminus \<longlongrightarrow> - a) (nhds a)" |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
733 |
begin |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
734 |
|
63546 | 735 |
lemma tendsto_minus [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> - a) F" |
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
736 |
by (rule filterlim_compose[OF tendsto_uminus_nhds]) |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
737 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
738 |
end |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
739 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
740 |
class topological_ab_group_add = topological_group_add + ab_group_add |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
741 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
742 |
instance topological_ab_group_add < topological_comm_monoid_add .. |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
743 |
|
63546 | 744 |
lemma continuous_minus [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)" |
745 |
for f :: "'a::t2_space \<Rightarrow> 'b::topological_group_add" |
|
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
746 |
unfolding continuous_def by (rule tendsto_minus) |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
747 |
|
63546 | 748 |
lemma continuous_on_minus [continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)" |
749 |
for f :: "_ \<Rightarrow> 'b::topological_group_add" |
|
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
750 |
unfolding continuous_on_def by (auto intro: tendsto_minus) |
62368 | 751 |
|
63546 | 752 |
lemma tendsto_minus_cancel: "((\<lambda>x. - f x) \<longlongrightarrow> - a) F \<Longrightarrow> (f \<longlongrightarrow> a) F" |
753 |
for a :: "'a::topological_group_add" |
|
754 |
by (drule tendsto_minus) simp |
|
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
755 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
756 |
lemma tendsto_minus_cancel_left: |
63546 | 757 |
"(f \<longlongrightarrow> - (y::_::topological_group_add)) F \<longleftrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> y) F" |
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
758 |
using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
759 |
by auto |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
760 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
761 |
lemma tendsto_diff [tendsto_intros]: |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
762 |
fixes a b :: "'a::topological_group_add" |
63546 | 763 |
shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> a - b) F" |
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
764 |
using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus) |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
765 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
766 |
lemma continuous_diff [continuous_intros]: |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
767 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::topological_group_add" |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
768 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x - g x)" |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
769 |
unfolding continuous_def by (rule tendsto_diff) |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
770 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
771 |
lemma continuous_on_diff [continuous_intros]: |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
772 |
fixes f g :: "_ \<Rightarrow> 'b::topological_group_add" |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
773 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)" |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
774 |
unfolding continuous_on_def by (auto intro: tendsto_diff) |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
775 |
|
67399 | 776 |
lemma continuous_on_op_minus: "continuous_on (s::'a::topological_group_add set) ((-) x)" |
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
777 |
by (rule continuous_intros | simp)+ |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
778 |
|
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
779 |
instance real_normed_vector < topological_ab_group_add |
62368 | 780 |
proof |
63546 | 781 |
fix a b :: 'a |
782 |
show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)" |
|
62368 | 783 |
unfolding tendsto_Zfun_iff add_diff_add |
784 |
using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] |
|
785 |
by (intro Zfun_add) |
|
68615 | 786 |
(auto simp: tendsto_Zfun_iff[symmetric] nhds_prod[symmetric] intro!: tendsto_fst) |
63081
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
787 |
show "(uminus \<longlongrightarrow> - a) (nhds a)" |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
788 |
unfolding tendsto_Zfun_iff minus_diff_minus |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
789 |
using filterlim_ident[of "nhds a"] |
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
immler
parents:
63040
diff
changeset
|
790 |
by (intro Zfun_minus) (simp add: tendsto_Zfun_iff) |
62368 | 791 |
qed |
792 |
||
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65036
diff
changeset
|
793 |
lemmas real_tendsto_sandwich = tendsto_sandwich[where 'a=real] |
50999 | 794 |
|
63546 | 795 |
|
60758 | 796 |
subsubsection \<open>Linear operators and multiplication\<close> |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
797 |
|
70999
5b753486c075
Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents:
70817
diff
changeset
|
798 |
lemma linear_times [simp]: "linear (\<lambda>x. c * x)" |
63546 | 799 |
for c :: "'a::real_algebra" |
61806
d2e62ae01cd8
Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents:
61799
diff
changeset
|
800 |
by (auto simp: linearI distrib_left) |
d2e62ae01cd8
Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents:
61799
diff
changeset
|
801 |
|
63546 | 802 |
lemma (in bounded_linear) tendsto: "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> f a) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
803 |
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
|
804 |
|
63546 | 805 |
lemma (in bounded_linear) continuous: "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))" |
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
|
806 |
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
|
807 |
|
63546 | 808 |
lemma (in bounded_linear) continuous_on: "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))" |
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
|
809 |
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
|
810 |
|
63546 | 811 |
lemma (in bounded_linear) tendsto_zero: "(g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> 0) F" |
812 |
by (drule tendsto) (simp only: zero) |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
813 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
814 |
lemma (in bounded_bilinear) tendsto: |
63546 | 815 |
"(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x ** g x) \<longlongrightarrow> a ** b) F" |
816 |
by (simp only: tendsto_Zfun_iff prod_diff_prod 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
|
817 |
|
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
|
818 |
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
|
819 |
"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
|
820 |
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
|
821 |
|
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
|
822 |
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
|
823 |
"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
|
824 |
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
|
825 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
826 |
lemma (in bounded_bilinear) tendsto_zero: |
61973 | 827 |
assumes f: "(f \<longlongrightarrow> 0) F" |
63546 | 828 |
and g: "(g \<longlongrightarrow> 0) F" |
61973 | 829 |
shows "((\<lambda>x. f x ** g x) \<longlongrightarrow> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
830 |
using tendsto [OF f g] by (simp add: zero_left) |
31355 | 831 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
832 |
lemma (in bounded_bilinear) tendsto_left_zero: |
61973 | 833 |
"(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) \<longlongrightarrow> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
834 |
by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
835 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
836 |
lemma (in bounded_bilinear) tendsto_right_zero: |
61973 | 837 |
"(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) \<longlongrightarrow> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
838 |
by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
839 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
840 |
lemmas tendsto_of_real [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
841 |
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
|
842 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
843 |
lemmas tendsto_scaleR [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
844 |
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
|
845 |
|
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
846 |
|
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
847 |
text\<open>Analogous type class for multiplication\<close> |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
848 |
class topological_semigroup_mult = topological_space + semigroup_mult + |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
849 |
assumes tendsto_mult_Pair: "LIM x (nhds a \<times>\<^sub>F nhds b). fst x * snd x :> nhds (a * b)" |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
850 |
|
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
851 |
instance real_normed_algebra < topological_semigroup_mult |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
852 |
proof |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
853 |
fix a b :: 'a |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
854 |
show "((\<lambda>x. fst x * snd x) \<longlongrightarrow> a * b) (nhds a \<times>\<^sub>F nhds b)" |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
855 |
unfolding nhds_prod[symmetric] |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
856 |
using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
857 |
by (simp add: bounded_bilinear.tendsto [OF bounded_bilinear_mult]) |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
858 |
qed |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
859 |
|
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
860 |
lemma tendsto_mult [tendsto_intros]: |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
861 |
fixes a b :: "'a::topological_semigroup_mult" |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
862 |
shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x * g x) \<longlongrightarrow> a * b) F" |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
863 |
using filterlim_compose[OF tendsto_mult_Pair, of "\<lambda>x. (f x, g x)" a b F] |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
864 |
by (simp add: nhds_prod[symmetric] tendsto_Pair) |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
865 |
|
63546 | 866 |
lemma tendsto_mult_left: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) \<longlongrightarrow> c * l) F" |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
867 |
for c :: "'a::topological_semigroup_mult" |
63546 | 868 |
by (rule tendsto_mult [OF tendsto_const]) |
869 |
||
870 |
lemma tendsto_mult_right: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) \<longlongrightarrow> l * c) F" |
|
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
871 |
for c :: "'a::topological_semigroup_mult" |
63546 | 872 |
by (rule tendsto_mult [OF _ tendsto_const]) |
61806
d2e62ae01cd8
Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents:
61799
diff
changeset
|
873 |
|
70804
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
874 |
lemma tendsto_mult_left_iff [simp]: |
70803
2d658afa1fc0
Generalised two results concerning limits from the real numbers to type classes
paulson <lp15@cam.ac.uk>
parents:
70723
diff
changeset
|
875 |
"c \<noteq> 0 \<Longrightarrow> tendsto(\<lambda>x. c * f x) (c * l) F \<longleftrightarrow> tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" |
70688
3d894e1cfc75
new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents:
70532
diff
changeset
|
876 |
by (auto simp: tendsto_mult_left dest: tendsto_mult_left [where c = "1/c"]) |
3d894e1cfc75
new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents:
70532
diff
changeset
|
877 |
|
70804
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
878 |
lemma tendsto_mult_right_iff [simp]: |
70803
2d658afa1fc0
Generalised two results concerning limits from the real numbers to type classes
paulson <lp15@cam.ac.uk>
parents:
70723
diff
changeset
|
879 |
"c \<noteq> 0 \<Longrightarrow> tendsto(\<lambda>x. f x * c) (l * c) F \<longleftrightarrow> tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" |
2d658afa1fc0
Generalised two results concerning limits from the real numbers to type classes
paulson <lp15@cam.ac.uk>
parents:
70723
diff
changeset
|
880 |
by (auto simp: tendsto_mult_right dest: tendsto_mult_left [where c = "1/c"]) |
70688
3d894e1cfc75
new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents:
70532
diff
changeset
|
881 |
|
70804
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
882 |
lemma tendsto_zero_mult_left_iff [simp]: |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
883 |
fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \<noteq> 0" shows "(\<lambda>n. c * a n)\<longlonglongrightarrow> 0 \<longleftrightarrow> a \<longlonglongrightarrow> 0" |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
884 |
using assms tendsto_mult_left tendsto_mult_left_iff by fastforce |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
885 |
|
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
886 |
lemma tendsto_zero_mult_right_iff [simp]: |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
887 |
fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \<noteq> 0" shows "(\<lambda>n. a n * c)\<longlonglongrightarrow> 0 \<longleftrightarrow> a \<longlonglongrightarrow> 0" |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
888 |
using assms tendsto_mult_right tendsto_mult_right_iff by fastforce |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
889 |
|
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
890 |
lemma tendsto_zero_divide_iff [simp]: |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
891 |
fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \<noteq> 0" shows "(\<lambda>n. a n / c)\<longlonglongrightarrow> 0 \<longleftrightarrow> a \<longlonglongrightarrow> 0" |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
892 |
using tendsto_zero_mult_right_iff [of "1/c" a] assms by (simp add: field_simps) |
4eef7c6ef7bf
More theorems about limits, including cancellation simprules
paulson <lp15@cam.ac.uk>
parents:
70803
diff
changeset
|
893 |
|
70365
4df0628e8545
a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
894 |
lemma lim_const_over_n [tendsto_intros]: |
4df0628e8545
a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
895 |
fixes a :: "'a::real_normed_field" |
4df0628e8545
a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
896 |
shows "(\<lambda>n. a / of_nat n) \<longlonglongrightarrow> 0" |
4df0628e8545
a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
897 |
using tendsto_mult [OF tendsto_const [of a] lim_1_over_n] by simp |
4df0628e8545
a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
898 |
|
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
|
899 |
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
|
900 |
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
|
901 |
|
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
|
902 |
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
|
903 |
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
|
904 |
|
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
|
905 |
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
|
906 |
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
|
907 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
908 |
lemmas continuous_on_of_real [continuous_intros] = |
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
|
909 |
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
|
910 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
911 |
lemmas continuous_on_scaleR [continuous_intros] = |
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
|
912 |
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
|
913 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
914 |
lemmas continuous_on_mult [continuous_intros] = |
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
|
915 |
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
|
916 |
|
44568
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
917 |
lemmas tendsto_mult_zero = |
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
918 |
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
|
919 |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
920 |
lemmas tendsto_mult_left_zero = |
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
921 |
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
|
922 |
|
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
923 |
lemmas tendsto_mult_right_zero = |
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
44342
diff
changeset
|
924 |
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
|
925 |
|
68296
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
926 |
|
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
927 |
lemma continuous_mult_left: |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
928 |
fixes c::"'a::real_normed_algebra" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
929 |
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
930 |
by (rule continuous_mult [OF continuous_const]) |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
931 |
|
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
932 |
lemma continuous_mult_right: |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
933 |
fixes c::"'a::real_normed_algebra" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
934 |
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
935 |
by (rule continuous_mult [OF _ continuous_const]) |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
936 |
|
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
937 |
lemma continuous_on_mult_left: |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
938 |
fixes c::"'a::real_normed_algebra" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
939 |
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
940 |
by (rule continuous_on_mult [OF continuous_on_const]) |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
941 |
|
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
942 |
lemma continuous_on_mult_right: |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
943 |
fixes c::"'a::real_normed_algebra" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
944 |
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)" |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
945 |
by (rule continuous_on_mult [OF _ continuous_on_const]) |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
946 |
|
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
947 |
lemma continuous_on_mult_const [simp]: |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
948 |
fixes c::"'a::real_normed_algebra" |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68860
diff
changeset
|
949 |
shows "continuous_on s ((*) c)" |
68296
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
950 |
by (intro continuous_on_mult_left continuous_on_id) |
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
paulson <lp15@cam.ac.uk>
parents:
68064
diff
changeset
|
951 |
|
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66456
diff
changeset
|
952 |
lemma tendsto_divide_zero: |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66456
diff
changeset
|
953 |
fixes c :: "'a::real_normed_field" |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66456
diff
changeset
|
954 |
shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x / c) \<longlongrightarrow> 0) F" |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66456
diff
changeset
|
955 |
by (cases "c=0") (simp_all add: divide_inverse tendsto_mult_left_zero) |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66456
diff
changeset
|
956 |
|
63546 | 957 |
lemma tendsto_power [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> a ^ n) F" |
958 |
for f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}" |
|
58729
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents:
57512
diff
changeset
|
959 |
by (induct n) (simp_all add: tendsto_mult) |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
960 |
|
65680
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents:
65578
diff
changeset
|
961 |
lemma tendsto_null_power: "\<lbrakk>(f \<longlongrightarrow> 0) F; 0 < n\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> 0) F" |
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents:
65578
diff
changeset
|
962 |
for f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra_1}" |
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents:
65578
diff
changeset
|
963 |
using tendsto_power [of f 0 F n] by (simp add: power_0_left) |
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents:
65578
diff
changeset
|
964 |
|
63546 | 965 |
lemma continuous_power [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)" |
966 |
for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}" |
|
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
|
967 |
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
|
968 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
969 |
lemma continuous_on_power [continuous_intros]: |
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
|
970 |
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
|
971 |
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
|
972 |
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
|
973 |
|
64272 | 974 |
lemma tendsto_prod [tendsto_intros]: |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
975 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}" |
63915 | 976 |
shows "(\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> L i) F) \<Longrightarrow> ((\<lambda>x. \<Prod>i\<in>S. f i x) \<longlongrightarrow> (\<Prod>i\<in>S. L i)) F" |
977 |
by (induct S rule: infinite_finite_induct) (simp_all add: tendsto_mult) |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
978 |
|
64272 | 979 |
lemma continuous_prod [continuous_intros]: |
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
|
980 |
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
|
981 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>S. f i x)" |
64272 | 982 |
unfolding continuous_def by (rule tendsto_prod) |
983 |
||
984 |
lemma continuous_on_prod [continuous_intros]: |
|
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
|
985 |
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
|
986 |
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)" |
64272 | 987 |
unfolding continuous_on_def by (auto intro: tendsto_prod) |
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
|
988 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
989 |
lemma tendsto_of_real_iff: |
63546 | 990 |
"((\<lambda>x. of_real (f x) :: 'a::real_normed_div_algebra) \<longlongrightarrow> of_real c) F \<longleftrightarrow> (f \<longlongrightarrow> c) F" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
991 |
unfolding tendsto_iff by simp |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
992 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
993 |
lemma tendsto_add_const_iff: |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
73885
diff
changeset
|
994 |
"((\<lambda>x. c + f x :: 'a::topological_group_add) \<longlongrightarrow> c + d) F \<longleftrightarrow> (f \<longlongrightarrow> d) F" |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
995 |
using tendsto_add[OF tendsto_const[of c], of f d] |
63546 | 996 |
and tendsto_add[OF tendsto_const[of "-c"], of "\<lambda>x. c + f x" "c + d"] by auto |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
997 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
998 |
|
68860
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
999 |
class topological_monoid_mult = topological_semigroup_mult + monoid_mult |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1000 |
class topological_comm_monoid_mult = topological_monoid_mult + comm_monoid_mult |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1001 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1002 |
lemma tendsto_power_strong [tendsto_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1003 |
fixes f :: "_ \<Rightarrow> 'b :: topological_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1004 |
assumes "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1005 |
shows "((\<lambda>x. f x ^ g x) \<longlongrightarrow> a ^ b) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1006 |
proof - |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1007 |
have "((\<lambda>x. f x ^ b) \<longlongrightarrow> a ^ b) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1008 |
by (induction b) (auto intro: tendsto_intros assms) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1009 |
also from assms(2) have "eventually (\<lambda>x. g x = b) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1010 |
by (simp add: nhds_discrete filterlim_principal) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1011 |
hence "eventually (\<lambda>x. f x ^ b = f x ^ g x) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1012 |
by eventually_elim simp |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1013 |
hence "((\<lambda>x. f x ^ b) \<longlongrightarrow> a ^ b) F \<longleftrightarrow> ((\<lambda>x. f x ^ g x) \<longlongrightarrow> a ^ b) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1014 |
by (intro filterlim_cong refl) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1015 |
finally show ?thesis . |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1016 |
qed |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1017 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1018 |
lemma continuous_mult' [continuous_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1019 |
fixes f g :: "_ \<Rightarrow> 'b::topological_semigroup_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1020 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x * g x)" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1021 |
unfolding continuous_def by (rule tendsto_mult) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1022 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1023 |
lemma continuous_power' [continuous_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1024 |
fixes f :: "_ \<Rightarrow> 'b::topological_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1025 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ^ g x)" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1026 |
unfolding continuous_def by (rule tendsto_power_strong) auto |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1027 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1028 |
lemma continuous_on_mult' [continuous_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1029 |
fixes f g :: "_ \<Rightarrow> 'b::topological_semigroup_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1030 |
shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x * g x)" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1031 |
unfolding continuous_on_def by (auto intro: tendsto_mult) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1032 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1033 |
lemma continuous_on_power' [continuous_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1034 |
fixes f :: "_ \<Rightarrow> 'b::topological_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1035 |
shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x ^ g x)" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1036 |
unfolding continuous_on_def by (auto intro: tendsto_power_strong) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1037 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1038 |
lemma tendsto_mult_one: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1039 |
fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1040 |
shows "(f \<longlongrightarrow> 1) F \<Longrightarrow> (g \<longlongrightarrow> 1) F \<Longrightarrow> ((\<lambda>x. f x * g x) \<longlongrightarrow> 1) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1041 |
by (drule (1) tendsto_mult) simp |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1042 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1043 |
lemma tendsto_prod' [tendsto_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1044 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1045 |
shows "(\<And>i. i \<in> I \<Longrightarrow> (f i \<longlongrightarrow> a i) F) \<Longrightarrow> ((\<lambda>x. \<Prod>i\<in>I. f i x) \<longlongrightarrow> (\<Prod>i\<in>I. a i)) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1046 |
by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_mult) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1047 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1048 |
lemma tendsto_one_prod': |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1049 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1050 |
assumes "\<And>i. i \<in> I \<Longrightarrow> ((\<lambda>x. f x i) \<longlongrightarrow> 1) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1051 |
shows "((\<lambda>i. prod (f i) I) \<longlongrightarrow> 1) F" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1052 |
using tendsto_prod' [of I "\<lambda>x y. f y x" "\<lambda>x. 1"] assms by simp |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1053 |
|
76724
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1054 |
lemma LIMSEQ_prod_0: |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1055 |
fixes f :: "nat \<Rightarrow> 'a::{semidom,topological_space}" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1056 |
assumes "f i = 0" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1057 |
shows "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> 0" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1058 |
proof (subst tendsto_cong) |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1059 |
show "\<forall>\<^sub>F n in sequentially. prod f {..n} = 0" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1060 |
using assms eventually_at_top_linorder by auto |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1061 |
qed auto |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1062 |
|
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1063 |
lemma LIMSEQ_prod_nonneg: |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1064 |
fixes f :: "nat \<Rightarrow> 'a::{linordered_semidom,linorder_topology}" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1065 |
assumes 0: "\<And>n. 0 \<le> f n" and a: "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> a" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1066 |
shows "a \<ge> 0" |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1067 |
by (simp add: "0" prod_nonneg LIMSEQ_le_const [OF a]) |
7ff71bdcf731
Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents:
74513
diff
changeset
|
1068 |
|
68860
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1069 |
lemma continuous_prod' [continuous_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1070 |
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1071 |
shows "(\<And>i. i \<in> I \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>I. f i x)" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1072 |
unfolding continuous_def by (rule tendsto_prod') |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1073 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1074 |
lemma continuous_on_prod' [continuous_intros]: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1075 |
fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::topological_comm_monoid_mult" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1076 |
shows "(\<And>i. i \<in> I \<Longrightarrow> continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<Prod>i\<in>I. f i x)" |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1077 |
unfolding continuous_on_def by (auto intro: tendsto_prod') |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1078 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1079 |
instance nat :: topological_comm_monoid_mult |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1080 |
by standard |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1081 |
(simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1082 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1083 |
instance int :: topological_comm_monoid_mult |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1084 |
by standard |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1085 |
(simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1086 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1087 |
class comm_real_normed_algebra_1 = real_normed_algebra_1 + comm_monoid_mult |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1088 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1089 |
context real_normed_field |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1090 |
begin |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1091 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1092 |
subclass comm_real_normed_algebra_1 |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1093 |
proof |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1094 |
from norm_mult[of "1 :: 'a" 1] show "norm 1 = 1" by simp |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1095 |
qed (simp_all add: norm_mult) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1096 |
|
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1097 |
end |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
68721
diff
changeset
|
1098 |
|
60758 | 1099 |
subsubsection \<open>Inverse and division\<close> |
31355 | 1100 |
|
1101 |
lemma (in bounded_bilinear) Zfun_prod_Bfun: |
|
44195 | 1102 |
assumes f: "Zfun f F" |
63546 | 1103 |
and g: "Bfun g F" |
44195 | 1104 |
shows "Zfun (\<lambda>x. f x ** g x) F" |
31355 | 1105 |
proof - |
1106 |
obtain K where K: "0 \<le> K" |
|
1107 |
and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" |
|
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1108 |
using nonneg_bounded by blast |
31355 | 1109 |
obtain B where B: "0 < B" |
44195 | 1110 |
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
|
1111 |
using g by (rule BfunE) |
44195 | 1112 |
have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) F" |
46887 | 1113 |
using norm_g proof eventually_elim |
1114 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1115 |
have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" |
31355 | 1116 |
by (rule norm_le) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1117 |
also have "\<dots> \<le> norm (f x) * B * K" |
63546 | 1118 |
by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1119 |
also have "\<dots> = norm (f x) * (B * K)" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57447
diff
changeset
|
1120 |
by (rule mult.assoc) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1121 |
finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" . |
31355 | 1122 |
qed |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1123 |
with f show ?thesis |
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1124 |
by (rule Zfun_imp_Zfun) |
31355 | 1125 |
qed |
1126 |
||
1127 |
lemma (in bounded_bilinear) Bfun_prod_Zfun: |
|
44195 | 1128 |
assumes f: "Bfun f F" |
63546 | 1129 |
and g: "Zfun g F" |
44195 | 1130 |
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
|
1131 |
using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) |
31355 | 1132 |
|
1133 |
lemma Bfun_inverse: |
|
1134 |
fixes a :: "'a::real_normed_div_algebra" |
|
61973 | 1135 |
assumes f: "(f \<longlongrightarrow> a) F" |
31355 | 1136 |
assumes a: "a \<noteq> 0" |
44195 | 1137 |
shows "Bfun (\<lambda>x. inverse (f x)) F" |
31355 | 1138 |
proof - |
1139 |
from a have "0 < norm a" by simp |
|
63546 | 1140 |
then have "\<exists>r>0. r < norm a" by (rule dense) |
1141 |
then obtain r where r1: "0 < r" and r2: "r < norm a" |
|
1142 |
by blast |
|
44195 | 1143 |
have "eventually (\<lambda>x. dist (f x) a < r) F" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1144 |
using tendstoD [OF f r1] by blast |
63546 | 1145 |
then have "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F" |
46887 | 1146 |
proof eventually_elim |
1147 |
case (elim x) |
|
63546 | 1148 |
then have 1: "norm (f x - a) < r" |
31355 | 1149 |
by (simp add: dist_norm) |
63546 | 1150 |
then have 2: "f x \<noteq> 0" using r2 by auto |
1151 |
then have "norm (inverse (f x)) = inverse (norm (f x))" |
|
31355 | 1152 |
by (rule nonzero_norm_inverse) |
1153 |
also have "\<dots> \<le> inverse (norm a - r)" |
|
1154 |
proof (rule le_imp_inverse_le) |
|
63546 | 1155 |
show "0 < norm a - r" |
1156 |
using r2 by simp |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1157 |
have "norm a - norm (f x) \<le> norm (a - f x)" |
31355 | 1158 |
by (rule norm_triangle_ineq2) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1159 |
also have "\<dots> = norm (f x - a)" |
31355 | 1160 |
by (rule norm_minus_commute) |
1161 |
also have "\<dots> < r" using 1 . |
|
63546 | 1162 |
finally show "norm a - r \<le> norm (f x)" |
1163 |
by simp |
|
31355 | 1164 |
qed |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
1165 |
finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" . |
31355 | 1166 |
qed |
63546 | 1167 |
then show ?thesis by (rule BfunI) |
31355 | 1168 |
qed |
1169 |
||
31565 | 1170 |
lemma tendsto_inverse [tendsto_intros]: |
31355 | 1171 |
fixes a :: "'a::real_normed_div_algebra" |
61973 | 1172 |
assumes f: "(f \<longlongrightarrow> a) F" |
63546 | 1173 |
and a: "a \<noteq> 0" |
61973 | 1174 |
shows "((\<lambda>x. inverse (f x)) \<longlongrightarrow> inverse a) F" |
31355 | 1175 |
proof - |
1176 |
from a have "0 < norm a" by simp |
|
44195 | 1177 |
with f have "eventually (\<lambda>x. dist (f x) a < norm a) F" |
31355 | 1178 |
by (rule tendstoD) |
44195 | 1179 |
then have "eventually (\<lambda>x. f x \<noteq> 0) F" |
61810 | 1180 |
unfolding dist_norm by (auto elim!: eventually_mono) |
44627 | 1181 |
with a have "eventually (\<lambda>x. inverse (f x) - inverse a = |
1182 |
- (inverse (f x) * (f x - a) * inverse a)) F" |
|
61810 | 1183 |
by (auto elim!: eventually_mono simp: inverse_diff_inverse) |
44627 | 1184 |
moreover have "Zfun (\<lambda>x. - (inverse (f x) * (f x - a) * inverse a)) F" |
1185 |
by (intro Zfun_minus Zfun_mult_left |
|
1186 |
bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] |
|
1187 |
Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) |
|
1188 |
ultimately show ?thesis |
|
1189 |
unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) |
|
31355 | 1190 |
qed |
1191 |
||
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
|
1192 |
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
|
1193 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
63546 | 1194 |
assumes "continuous F f" |
1195 |
and "f (Lim F (\<lambda>x. x)) \<noteq> 0" |
|
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
|
1196 |
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
|
1197 |
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
|
1198 |
|
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
|
1199 |
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
|
1200 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
63546 | 1201 |
assumes "continuous (at a within s) f" |
1202 |
and "f a \<noteq> 0" |
|
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
|
1203 |
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
|
1204 |
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
|
1205 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
1206 |
lemma continuous_on_inverse[continuous_intros]: |
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
|
1207 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra" |
63546 | 1208 |
assumes "continuous_on s f" |
1209 |
and "\<forall>x\<in>s. f x \<noteq> 0" |
|
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
|
1210 |
shows "continuous_on s (\<lambda>x. inverse (f x))" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1211 |
using assms unfolding continuous_on_def by (blast intro: tendsto_inverse) |
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
|
1212 |
|
31565 | 1213 |
lemma tendsto_divide [tendsto_intros]: |
31355 | 1214 |
fixes a b :: "'a::real_normed_field" |
63546 | 1215 |
shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> ((\<lambda>x. f x / g x) \<longlongrightarrow> a / b) F" |
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
1216 |
by (simp add: tendsto_mult tendsto_inverse divide_inverse) |
31355 | 1217 |
|
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
|
1218 |
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
|
1219 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" |
63546 | 1220 |
assumes "continuous F f" |
1221 |
and "continuous F g" |
|
1222 |
and "g (Lim F (\<lambda>x. x)) \<noteq> 0" |
|
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
|
1223 |
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
|
1224 |
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
|
1225 |
|
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
|
1226 |
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
|
1227 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" |
63546 | 1228 |
assumes "continuous (at a within s) f" "continuous (at a within s) g" |
1229 |
and "g a \<noteq> 0" |
|
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
|
1230 |
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
|
1231 |
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
|
1232 |
|
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
|
1233 |
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
|
1234 |
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
|
1235 |
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
|
1236 |
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
|
1237 |
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
|
1238 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
1239 |
lemma continuous_on_divide[continuous_intros]: |
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
|
1240 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_field" |
63546 | 1241 |
assumes "continuous_on s f" "continuous_on s g" |
1242 |
and "\<forall>x\<in>s. g x \<noteq> 0" |
|
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
|
1243 |
shows "continuous_on s (\<lambda>x. (f x) / (g x))" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1244 |
using assms unfolding continuous_on_def by (blast intro: tendsto_divide) |
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
|
1245 |
|
71837
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1246 |
lemma tendsto_power_int [tendsto_intros]: |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1247 |
fixes a :: "'a::real_normed_div_algebra" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1248 |
assumes f: "(f \<longlongrightarrow> a) F" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1249 |
and a: "a \<noteq> 0" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1250 |
shows "((\<lambda>x. power_int (f x) n) \<longlongrightarrow> power_int a n) F" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1251 |
using assms by (cases n rule: int_cases4) (auto intro!: tendsto_intros simp: power_int_minus) |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1252 |
|
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1253 |
lemma continuous_power_int: |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1254 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1255 |
assumes "continuous F f" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1256 |
and "f (Lim F (\<lambda>x. x)) \<noteq> 0" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1257 |
shows "continuous F (\<lambda>x. power_int (f x) n)" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1258 |
using assms unfolding continuous_def by (rule tendsto_power_int) |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1259 |
|
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1260 |
lemma continuous_at_within_power_int[continuous_intros]: |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1261 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1262 |
assumes "continuous (at a within s) f" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1263 |
and "f a \<noteq> 0" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1264 |
shows "continuous (at a within s) (\<lambda>x. power_int (f x) n)" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1265 |
using assms unfolding continuous_within by (rule tendsto_power_int) |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1266 |
|
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1267 |
lemma continuous_on_power_int [continuous_intros]: |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1268 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1269 |
assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1270 |
shows "continuous_on s (\<lambda>x. power_int (f x) n)" |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1271 |
using assms unfolding continuous_on_def by (blast intro: tendsto_power_int) |
dca11678c495
new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents:
71827
diff
changeset
|
1272 |
|
63546 | 1273 |
lemma tendsto_sgn [tendsto_intros]: "(f \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. sgn (f x)) \<longlongrightarrow> sgn l) F" |
1274 |
for l :: "'a::real_normed_vector" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
1275 |
unfolding sgn_div_norm by (simp add: tendsto_intros) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
1276 |
|
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
|
1277 |
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
|
1278 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
63546 | 1279 |
assumes "continuous F f" |
1280 |
and "f (Lim F (\<lambda>x. x)) \<noteq> 0" |
|
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
|
1281 |
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
|
1282 |
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
|
1283 |
|
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
|
1284 |
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
|
1285 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
63546 | 1286 |
assumes "continuous (at a within s) f" |
1287 |
and "f a \<noteq> 0" |
|
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
|
1288 |
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
|
1289 |
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
|
1290 |
|
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
|
1291 |
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
|
1292 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
63546 | 1293 |
assumes "isCont f a" |
1294 |
and "f a \<noteq> 0" |
|
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
|
1295 |
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
|
1296 |
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
|
1297 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
1298 |
lemma continuous_on_sgn[continuous_intros]: |
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
|
1299 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
63546 | 1300 |
assumes "continuous_on s f" |
1301 |
and "\<forall>x\<in>s. f x \<noteq> 0" |
|
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
|
1302 |
shows "continuous_on s (\<lambda>x. sgn (f x))" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1303 |
using assms unfolding continuous_on_def by (blast intro: tendsto_sgn) |
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
|
1304 |
|
50325 | 1305 |
lemma filterlim_at_infinity: |
61076 | 1306 |
fixes f :: "_ \<Rightarrow> 'a::real_normed_vector" |
50325 | 1307 |
assumes "0 \<le> c" |
1308 |
shows "(LIM x F. f x :> at_infinity) \<longleftrightarrow> (\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F)" |
|
1309 |
unfolding filterlim_iff eventually_at_infinity |
|
1310 |
proof safe |
|
63546 | 1311 |
fix P :: "'a \<Rightarrow> bool" |
1312 |
fix b |
|
50325 | 1313 |
assume *: "\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F" |
63546 | 1314 |
assume P: "\<forall>x. b \<le> norm x \<longrightarrow> P x" |
50325 | 1315 |
have "max b (c + 1) > c" by auto |
1316 |
with * have "eventually (\<lambda>x. max b (c + 1) \<le> norm (f x)) F" |
|
1317 |
by auto |
|
1318 |
then show "eventually (\<lambda>x. P (f x)) F" |
|
1319 |
proof eventually_elim |
|
63546 | 1320 |
case (elim x) |
50325 | 1321 |
with P show "P (f x)" by auto |
1322 |
qed |
|
1323 |
qed force |
|
1324 |
||
67371
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1325 |
lemma filterlim_at_infinity_imp_norm_at_top: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1326 |
fixes F |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1327 |
assumes "filterlim f at_infinity F" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1328 |
shows "filterlim (\<lambda>x. norm (f x)) at_top F" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1329 |
proof - |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1330 |
{ |
68611 | 1331 |
fix r :: real |
1332 |
have "\<forall>\<^sub>F x in F. r \<le> norm (f x)" using filterlim_at_infinity[of 0 f F] assms |
|
1333 |
by (cases "r > 0") |
|
67371
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1334 |
(auto simp: not_less intro: always_eventually order.trans[OF _ norm_ge_zero]) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1335 |
} |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1336 |
thus ?thesis by (auto simp: filterlim_at_top) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1337 |
qed |
68611 | 1338 |
|
67371
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1339 |
lemma filterlim_norm_at_top_imp_at_infinity: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1340 |
fixes F |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1341 |
assumes "filterlim (\<lambda>x. norm (f x)) at_top F" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1342 |
shows "filterlim f at_infinity F" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1343 |
using filterlim_at_infinity[of 0 f F] assms by (auto simp: filterlim_at_top) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1344 |
|
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1345 |
lemma filterlim_norm_at_top: "filterlim norm at_top at_infinity" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1346 |
by (rule filterlim_at_infinity_imp_norm_at_top) (rule filterlim_ident) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1347 |
|
67950
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1348 |
lemma filterlim_at_infinity_conv_norm_at_top: |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1349 |
"filterlim f at_infinity G \<longleftrightarrow> filterlim (\<lambda>x. norm (f x)) at_top G" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1350 |
by (auto simp: filterlim_at_infinity[OF order.refl] filterlim_at_top_gt[of _ _ 0]) |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1351 |
|
67371
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1352 |
lemma eventually_not_equal_at_infinity: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1353 |
"eventually (\<lambda>x. x \<noteq> (a :: 'a :: {real_normed_vector})) at_infinity" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1354 |
proof - |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1355 |
from filterlim_norm_at_top[where 'a = 'a] |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1356 |
have "\<forall>\<^sub>F x in at_infinity. norm a < norm (x::'a)" by (auto simp: filterlim_at_top_dense) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1357 |
thus ?thesis by eventually_elim auto |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1358 |
qed |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1359 |
|
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1360 |
lemma filterlim_int_of_nat_at_topD: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1361 |
fixes F |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1362 |
assumes "filterlim (\<lambda>x. f (int x)) F at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1363 |
shows "filterlim f F at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1364 |
proof - |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1365 |
have "filterlim (\<lambda>x. f (int (nat x))) F at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1366 |
by (rule filterlim_compose[OF assms filterlim_nat_sequentially]) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1367 |
also have "?this \<longleftrightarrow> filterlim f F at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1368 |
by (intro filterlim_cong refl eventually_mono [OF eventually_ge_at_top[of "0::int"]]) auto |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1369 |
finally show ?thesis . |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1370 |
qed |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1371 |
|
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1372 |
lemma filterlim_int_sequentially [tendsto_intros]: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1373 |
"filterlim int at_top sequentially" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1374 |
unfolding filterlim_at_top |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1375 |
proof |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1376 |
fix C :: int |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1377 |
show "eventually (\<lambda>n. int n \<ge> C) at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1378 |
using eventually_ge_at_top[of "nat \<lceil>C\<rceil>"] by eventually_elim linarith |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1379 |
qed |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1380 |
|
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1381 |
lemma filterlim_real_of_int_at_top [tendsto_intros]: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1382 |
"filterlim real_of_int at_top at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1383 |
unfolding filterlim_at_top |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1384 |
proof |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1385 |
fix C :: real |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1386 |
show "eventually (\<lambda>n. real_of_int n \<ge> C) at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1387 |
using eventually_ge_at_top[of "\<lceil>C\<rceil>"] by eventually_elim linarith |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1388 |
qed |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1389 |
|
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1390 |
lemma filterlim_abs_real: "filterlim (abs::real \<Rightarrow> real) at_top at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1391 |
proof (subst filterlim_cong[OF refl refl]) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1392 |
from eventually_ge_at_top[of "0::real"] show "eventually (\<lambda>x::real. \<bar>x\<bar> = x) at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1393 |
by eventually_elim simp |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1394 |
qed (simp_all add: filterlim_ident) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1395 |
|
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1396 |
lemma filterlim_of_real_at_infinity [tendsto_intros]: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1397 |
"filterlim (of_real :: real \<Rightarrow> 'a :: real_normed_algebra_1) at_infinity at_top" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1398 |
by (intro filterlim_norm_at_top_imp_at_infinity) (auto simp: filterlim_abs_real) |
68611 | 1399 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1400 |
lemma not_tendsto_and_filterlim_at_infinity: |
63546 | 1401 |
fixes c :: "'a::real_normed_vector" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1402 |
assumes "F \<noteq> bot" |
63546 | 1403 |
and "(f \<longlongrightarrow> c) F" |
1404 |
and "filterlim f at_infinity F" |
|
1405 |
shows False |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1406 |
proof - |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
1407 |
from tendstoD[OF assms(2), of "1/2"] |
63546 | 1408 |
have "eventually (\<lambda>x. dist (f x) c < 1/2) F" |
1409 |
by simp |
|
1410 |
moreover |
|
1411 |
from filterlim_at_infinity[of "norm c" f F] assms(3) |
|
1412 |
have "eventually (\<lambda>x. norm (f x) \<ge> norm c + 1) F" by simp |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1413 |
ultimately have "eventually (\<lambda>x. False) F" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1414 |
proof eventually_elim |
63546 | 1415 |
fix x |
1416 |
assume A: "dist (f x) c < 1/2" |
|
1417 |
assume "norm (f x) \<ge> norm c + 1" |
|
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62369
diff
changeset
|
1418 |
also have "norm (f x) = dist (f x) 0" by simp |
63546 | 1419 |
also have "\<dots> \<le> dist (f x) c + dist c 0" by (rule dist_triangle) |
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62369
diff
changeset
|
1420 |
finally show False using A by simp |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1421 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1422 |
with assms show False by simp |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1423 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1424 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1425 |
lemma filterlim_at_infinity_imp_not_convergent: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1426 |
assumes "filterlim f at_infinity sequentially" |
63546 | 1427 |
shows "\<not> convergent f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1428 |
by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1429 |
(simp_all add: convergent_LIMSEQ_iff) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1430 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1431 |
lemma filterlim_at_infinity_imp_eventually_ne: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1432 |
assumes "filterlim f at_infinity F" |
63546 | 1433 |
shows "eventually (\<lambda>z. f z \<noteq> c) F" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1434 |
proof - |
63546 | 1435 |
have "norm c + 1 > 0" |
1436 |
by (intro add_nonneg_pos) simp_all |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1437 |
with filterlim_at_infinity[OF order.refl, of f F] assms |
63546 | 1438 |
have "eventually (\<lambda>z. norm (f z) \<ge> norm c + 1) F" |
1439 |
by blast |
|
1440 |
then show ?thesis |
|
1441 |
by eventually_elim auto |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1442 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1443 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
1444 |
lemma tendsto_of_nat [tendsto_intros]: |
63546 | 1445 |
"filterlim (of_nat :: nat \<Rightarrow> 'a::real_normed_algebra_1) at_infinity sequentially" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1446 |
proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) |
63040 | 1447 |
fix r :: real |
1448 |
assume r: "r > 0" |
|
1449 |
define n where "n = nat \<lceil>r\<rceil>" |
|
63546 | 1450 |
from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r" |
1451 |
unfolding n_def by linarith |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1452 |
from eventually_ge_at_top[of n] show "eventually (\<lambda>m. norm (of_nat m :: 'a) \<ge> r) sequentially" |
63546 | 1453 |
by eventually_elim (use n in simp_all) |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1454 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1455 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1456 |
|
69593 | 1457 |
subsection \<open>Relate \<^const>\<open>at\<close>, \<^const>\<open>at_left\<close> and \<^const>\<open>at_right\<close>\<close> |
50347 | 1458 |
|
60758 | 1459 |
text \<open> |
69593 | 1460 |
This lemmas are useful for conversion between \<^term>\<open>at x\<close> to \<^term>\<open>at_left x\<close> and |
1461 |
\<^term>\<open>at_right x\<close> and also \<^term>\<open>at_right 0\<close>. |
|
60758 | 1462 |
\<close> |
50347 | 1463 |
|
51471 | 1464 |
lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] |
50323 | 1465 |
|
63546 | 1466 |
lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d)" |
1467 |
for a d :: "'a::real_normed_vector" |
|
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1468 |
by (rule filtermap_fun_inverse[where g="\<lambda>x. x + d"]) |
63546 | 1469 |
(auto intro!: tendsto_eq_intros filterlim_ident) |
1470 |
||
1471 |
lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a)" |
|
1472 |
for a :: "'a::real_normed_vector" |
|
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1473 |
by (rule filtermap_fun_inverse[where g=uminus]) |
63546 | 1474 |
(auto intro!: tendsto_eq_intros filterlim_ident) |
1475 |
||
1476 |
lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d)" |
|
1477 |
for a d :: "'a::real_normed_vector" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1478 |
by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) |
50347 | 1479 |
|
63546 | 1480 |
lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d)" |
1481 |
for a d :: "real" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1482 |
by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) |
50323 | 1483 |
|
73795 | 1484 |
lemma filterlim_shift: |
1485 |
fixes d :: "'a::real_normed_vector" |
|
1486 |
assumes "filterlim f F (at a)" |
|
1487 |
shows "filterlim (f \<circ> (+) d) F (at (a - d))" |
|
1488 |
unfolding filterlim_iff |
|
1489 |
proof (intro strip) |
|
1490 |
fix P |
|
1491 |
assume "eventually P F" |
|
1492 |
then have "\<forall>\<^sub>F x in filtermap (\<lambda>y. y - d) (at a). P (f (d + x))" |
|
1493 |
using assms by (force simp add: filterlim_iff eventually_filtermap) |
|
1494 |
then show "(\<forall>\<^sub>F x in at (a - d). P ((f \<circ> (+) d) x))" |
|
1495 |
by (force simp add: filtermap_at_shift) |
|
1496 |
qed |
|
1497 |
||
1498 |
lemma filterlim_shift_iff: |
|
1499 |
fixes d :: "'a::real_normed_vector" |
|
1500 |
shows "filterlim (f \<circ> (+) d) F (at (a - d)) = filterlim f F (at a)" (is "?lhs = ?rhs") |
|
1501 |
proof |
|
1502 |
assume L: ?lhs show ?rhs |
|
1503 |
using filterlim_shift [OF L, of "-d"] by (simp add: filterlim_iff) |
|
1504 |
qed (metis filterlim_shift) |
|
1505 |
||
63546 | 1506 |
lemma at_right_to_0: "at_right a = filtermap (\<lambda>x. x + a) (at_right 0)" |
1507 |
for a :: real |
|
50347 | 1508 |
using filtermap_at_right_shift[of "-a" 0] by simp |
1509 |
||
1510 |
lemma filterlim_at_right_to_0: |
|
63546 | 1511 |
"filterlim f F (at_right a) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)" |
1512 |
for a :: real |
|
50347 | 1513 |
unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. |
1514 |
||
1515 |
lemma eventually_at_right_to_0: |
|
63546 | 1516 |
"eventually P (at_right a) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)" |
1517 |
for a :: real |
|
50347 | 1518 |
unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) |
1519 |
||
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1520 |
lemma at_to_0: "at a = filtermap (\<lambda>x. x + a) (at 0)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1521 |
for a :: "'a::real_normed_vector" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1522 |
using filtermap_at_shift[of "-a" 0] by simp |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1523 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1524 |
lemma filterlim_at_to_0: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1525 |
"filterlim f F (at a) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at 0)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1526 |
for a :: "'a::real_normed_vector" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1527 |
unfolding filterlim_def filtermap_filtermap at_to_0[of a] .. |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1528 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1529 |
lemma eventually_at_to_0: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1530 |
"eventually P (at a) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at 0)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1531 |
for a :: "'a::real_normed_vector" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1532 |
unfolding at_to_0[of a] by (simp add: eventually_filtermap) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1533 |
|
63546 | 1534 |
lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a)" |
1535 |
for a :: "'a::real_normed_vector" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1536 |
by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) |
50347 | 1537 |
|
63546 | 1538 |
lemma at_left_minus: "at_left a = filtermap (\<lambda>x. - x) (at_right (- a))" |
1539 |
for a :: real |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1540 |
by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) |
50323 | 1541 |
|
63546 | 1542 |
lemma at_right_minus: "at_right a = filtermap (\<lambda>x. - x) (at_left (- a))" |
1543 |
for a :: real |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1544 |
by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) |
50347 | 1545 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1546 |
|
50347 | 1547 |
lemma filterlim_at_left_to_right: |
63546 | 1548 |
"filterlim f F (at_left a) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))" |
1549 |
for a :: real |
|
50347 | 1550 |
unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. |
1551 |
||
1552 |
lemma eventually_at_left_to_right: |
|
63546 | 1553 |
"eventually P (at_left a) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))" |
1554 |
for a :: real |
|
50347 | 1555 |
unfolding at_left_minus[of a] by (simp add: eventually_filtermap) |
1556 |
||
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1557 |
lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1558 |
unfolding filterlim_at_top eventually_at_bot_dense |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1559 |
by (metis leI minus_less_iff order_less_asym) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1560 |
|
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1561 |
lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1562 |
unfolding filterlim_at_bot eventually_at_top_dense |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1563 |
by (metis leI less_minus_iff order_less_asym) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1564 |
|
68611 | 1565 |
lemma at_bot_mirror : |
1566 |
shows "(at_bot::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_top" |
|
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
1567 |
proof (rule filtermap_fun_inverse[symmetric]) |
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
1568 |
show "filterlim uminus at_top (at_bot::'a filter)" |
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
1569 |
using eventually_at_bot_linorder filterlim_at_top le_minus_iff by force |
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
1570 |
show "filterlim uminus (at_bot::'a filter) at_top" |
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
1571 |
by (simp add: filterlim_at_bot minus_le_iff) |
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
1572 |
qed auto |
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68296
diff
changeset
|
1573 |
|
68611 | 1574 |
lemma at_top_mirror : |
1575 |
shows "(at_top::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_bot" |
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68296
diff
changeset
|
1576 |
apply (subst at_bot_mirror) |
68615 | 1577 |
by (auto simp: filtermap_filtermap) |
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
|
1578 |
|
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
|
1579 |
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
|
1580 |
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
|
1581 |
|
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
|
1582 |
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
|
1583 |
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
|
1584 |
|
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
|
1585 |
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
|
1586 |
using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] |
63546 | 1587 |
and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F] |
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
|
1588 |
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
|
1589 |
|
68721 | 1590 |
lemma tendsto_at_botI_sequentially: |
1591 |
fixes f :: "real \<Rightarrow> 'b::first_countable_topology" |
|
1592 |
assumes *: "\<And>X. filterlim X at_bot sequentially \<Longrightarrow> (\<lambda>n. f (X n)) \<longlonglongrightarrow> y" |
|
1593 |
shows "(f \<longlongrightarrow> y) at_bot" |
|
1594 |
unfolding filterlim_at_bot_mirror |
|
1595 |
proof (rule tendsto_at_topI_sequentially) |
|
1596 |
fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially" |
|
1597 |
thus "(\<lambda>n. f (-X n)) \<longlonglongrightarrow> y" by (intro *) (auto simp: filterlim_uminus_at_top) |
|
1598 |
qed |
|
1599 |
||
67950
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1600 |
lemma filterlim_at_infinity_imp_filterlim_at_top: |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1601 |
assumes "filterlim (f :: 'a \<Rightarrow> real) at_infinity F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1602 |
assumes "eventually (\<lambda>x. f x > 0) F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1603 |
shows "filterlim f at_top F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1604 |
proof - |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1605 |
from assms(2) have *: "eventually (\<lambda>x. norm (f x) = f x) F" by eventually_elim simp |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1606 |
from assms(1) show ?thesis unfolding filterlim_at_infinity_conv_norm_at_top |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1607 |
by (subst (asm) filterlim_cong[OF refl refl *]) |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1608 |
qed |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1609 |
|
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1610 |
lemma filterlim_at_infinity_imp_filterlim_at_bot: |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1611 |
assumes "filterlim (f :: 'a \<Rightarrow> real) at_infinity F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1612 |
assumes "eventually (\<lambda>x. f x < 0) F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1613 |
shows "filterlim f at_bot F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1614 |
proof - |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1615 |
from assms(2) have *: "eventually (\<lambda>x. norm (f x) = -f x) F" by eventually_elim simp |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1616 |
from assms(1) have "filterlim (\<lambda>x. - f x) at_top F" |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1617 |
unfolding filterlim_at_infinity_conv_norm_at_top |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1618 |
by (subst (asm) filterlim_cong[OF refl refl *]) |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1619 |
thus ?thesis by (simp add: filterlim_uminus_at_top) |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1620 |
qed |
99eaa5cedbb7
Added some simple facts about limits
Manuel Eberl <eberlm@in.tum.de>
parents:
67707
diff
changeset
|
1621 |
|
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
|
1622 |
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
|
1623 |
unfolding filterlim_uminus_at_top by simp |
50323 | 1624 |
|
50347 | 1625 |
lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" |
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1626 |
unfolding filterlim_at_top_gt[where c=0] eventually_at_filter |
50347 | 1627 |
proof safe |
63546 | 1628 |
fix Z :: real |
1629 |
assume [arith]: "0 < Z" |
|
50347 | 1630 |
then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)" |
68615 | 1631 |
by (auto simp: eventually_nhds_metric dist_real_def intro!: exI[of _ "\<bar>inverse Z\<bar>"]) |
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1632 |
then show "eventually (\<lambda>x. x \<noteq> 0 \<longrightarrow> x \<in> {0<..} \<longrightarrow> Z \<le> inverse x) (nhds 0)" |
61810 | 1633 |
by (auto elim!: eventually_mono simp: inverse_eq_divide field_simps) |
50347 | 1634 |
qed |
1635 |
||
50325 | 1636 |
lemma tendsto_inverse_0: |
61076 | 1637 |
fixes x :: "_ \<Rightarrow> 'a::real_normed_div_algebra" |
61973 | 1638 |
shows "(inverse \<longlongrightarrow> (0::'a)) at_infinity" |
50325 | 1639 |
unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity |
1640 |
proof safe |
|
63546 | 1641 |
fix r :: real |
1642 |
assume "0 < r" |
|
50325 | 1643 |
show "\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> norm (inverse x :: 'a) < r" |
1644 |
proof (intro exI[of _ "inverse (r / 2)"] allI impI) |
|
1645 |
fix x :: 'a |
|
60758 | 1646 |
from \<open>0 < r\<close> have "0 < inverse (r / 2)" by simp |
50325 | 1647 |
also assume *: "inverse (r / 2) \<le> norm x" |
1648 |
finally show "norm (inverse x) < r" |
|
63546 | 1649 |
using * \<open>0 < r\<close> |
1650 |
by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) |
|
50325 | 1651 |
qed |
1652 |
qed |
|
1653 |
||
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1654 |
lemma tendsto_add_filterlim_at_infinity: |
63546 | 1655 |
fixes c :: "'b::real_normed_vector" |
1656 |
and F :: "'a filter" |
|
1657 |
assumes "(f \<longlongrightarrow> c) F" |
|
1658 |
and "filterlim g at_infinity F" |
|
1659 |
shows "filterlim (\<lambda>x. f x + g x) at_infinity F" |
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1660 |
proof (subst filterlim_at_infinity[OF order_refl], safe) |
63546 | 1661 |
fix r :: real |
1662 |
assume r: "r > 0" |
|
1663 |
from assms(1) have "((\<lambda>x. norm (f x)) \<longlongrightarrow> norm c) F" |
|
1664 |
by (rule tendsto_norm) |
|
1665 |
then have "eventually (\<lambda>x. norm (f x) < norm c + 1) F" |
|
1666 |
by (rule order_tendstoD) simp_all |
|
1667 |
moreover from r have "r + norm c + 1 > 0" |
|
1668 |
by (intro add_pos_nonneg) simp_all |
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1669 |
with assms(2) have "eventually (\<lambda>x. norm (g x) \<ge> r + norm c + 1) F" |
63546 | 1670 |
unfolding filterlim_at_infinity[OF order_refl] |
1671 |
by (elim allE[of _ "r + norm c + 1"]) simp_all |
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1672 |
ultimately show "eventually (\<lambda>x. norm (f x + g x) \<ge> r) F" |
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1673 |
proof eventually_elim |
63546 | 1674 |
fix x :: 'a |
1675 |
assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \<le> norm (g x)" |
|
1676 |
from A B have "r \<le> norm (g x) - norm (f x)" |
|
1677 |
by simp |
|
1678 |
also have "norm (g x) - norm (f x) \<le> norm (g x + f x)" |
|
1679 |
by (rule norm_diff_ineq) |
|
1680 |
finally show "r \<le> norm (f x + g x)" |
|
1681 |
by (simp add: add_ac) |
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1682 |
qed |
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1683 |
qed |
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1684 |
|
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1685 |
lemma tendsto_add_filterlim_at_infinity': |
63546 | 1686 |
fixes c :: "'b::real_normed_vector" |
1687 |
and F :: "'a filter" |
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1688 |
assumes "filterlim f at_infinity F" |
63546 | 1689 |
and "(g \<longlongrightarrow> c) F" |
1690 |
shows "filterlim (\<lambda>x. f x + g x) at_infinity F" |
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1691 |
by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+ |
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
1692 |
|
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1693 |
lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)" |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1694 |
unfolding filterlim_at |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1695 |
by (auto simp: eventually_at_top_dense) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1696 |
(metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1697 |
|
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1698 |
lemma filterlim_inverse_at_top: |
61973 | 1699 |
"(f \<longlongrightarrow> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> LIM x F. inverse (f x) :> at_top" |
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1700 |
by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) |
61810 | 1701 |
(simp add: filterlim_def eventually_filtermap eventually_mono at_within_def le_principal) |
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1702 |
|
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1703 |
lemma filterlim_inverse_at_bot_neg: |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1704 |
"LIM x (at_left (0::real)). inverse x :> at_bot" |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1705 |
by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1706 |
|
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1707 |
lemma filterlim_inverse_at_bot: |
61973 | 1708 |
"(f \<longlongrightarrow> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x < 0) F \<Longrightarrow> LIM x F. inverse (f x) :> at_bot" |
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1709 |
unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1710 |
by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1711 |
|
50347 | 1712 |
lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" |
60721
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1713 |
by (intro filtermap_fun_inverse[symmetric, where g=inverse]) |
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents:
60182
diff
changeset
|
1714 |
(auto intro: filterlim_inverse_at_top_right filterlim_inverse_at_right_top) |
50347 | 1715 |
|
1716 |
lemma eventually_at_right_to_top: |
|
1717 |
"eventually P (at_right (0::real)) \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) at_top" |
|
1718 |
unfolding at_right_to_top eventually_filtermap .. |
|
1719 |
||
1720 |
lemma filterlim_at_right_to_top: |
|
1721 |
"filterlim f F (at_right (0::real)) \<longleftrightarrow> (LIM x at_top. f (inverse x) :> F)" |
|
1722 |
unfolding filterlim_def at_right_to_top filtermap_filtermap .. |
|
1723 |
||
1724 |
lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" |
|
1725 |
unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. |
|
1726 |
||
1727 |
lemma eventually_at_top_to_right: |
|
1728 |
"eventually P at_top \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) (at_right (0::real))" |
|
1729 |
unfolding at_top_to_right eventually_filtermap .. |
|
1730 |
||
1731 |
lemma filterlim_at_top_to_right: |
|
1732 |
"filterlim f F at_top \<longleftrightarrow> (LIM x (at_right (0::real)). f (inverse x) :> F)" |
|
1733 |
unfolding filterlim_def at_top_to_right filtermap_filtermap .. |
|
1734 |
||
50325 | 1735 |
lemma filterlim_inverse_at_infinity: |
61076 | 1736 |
fixes x :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}" |
50325 | 1737 |
shows "filterlim inverse at_infinity (at (0::'a))" |
1738 |
unfolding filterlim_at_infinity[OF order_refl] |
|
1739 |
proof safe |
|
63546 | 1740 |
fix r :: real |
1741 |
assume "0 < r" |
|
50325 | 1742 |
then show "eventually (\<lambda>x::'a. r \<le> norm (inverse x)) (at 0)" |
1743 |
unfolding eventually_at norm_inverse |
|
1744 |
by (intro exI[of _ "inverse r"]) |
|
1745 |
(auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) |
|
1746 |
qed |
|
1747 |
||
1748 |
lemma filterlim_inverse_at_iff: |
|
61076 | 1749 |
fixes g :: "'a \<Rightarrow> 'b::{real_normed_div_algebra, division_ring}" |
50325 | 1750 |
shows "(LIM x F. inverse (g x) :> at 0) \<longleftrightarrow> (LIM x F. g x :> at_infinity)" |
1751 |
unfolding filterlim_def filtermap_filtermap[symmetric] |
|
1752 |
proof |
|
1753 |
assume "filtermap g F \<le> at_infinity" |
|
1754 |
then have "filtermap inverse (filtermap g F) \<le> filtermap inverse at_infinity" |
|
1755 |
by (rule filtermap_mono) |
|
1756 |
also have "\<dots> \<le> at 0" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1757 |
using tendsto_inverse_0[where 'a='b] |
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
1758 |
by (auto intro!: exI[of _ 1] |
63546 | 1759 |
simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity) |
50325 | 1760 |
finally show "filtermap inverse (filtermap g F) \<le> at 0" . |
1761 |
next |
|
1762 |
assume "filtermap inverse (filtermap g F) \<le> at 0" |
|
1763 |
then have "filtermap inverse (filtermap inverse (filtermap g F)) \<le> filtermap inverse (at 0)" |
|
1764 |
by (rule filtermap_mono) |
|
1765 |
with filterlim_inverse_at_infinity show "filtermap g F \<le> at_infinity" |
|
1766 |
by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) |
|
1767 |
qed |
|
1768 |
||
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1769 |
lemma tendsto_mult_filterlim_at_infinity: |
63546 | 1770 |
fixes c :: "'a::real_normed_field" |
64394 | 1771 |
assumes "(f \<longlongrightarrow> c) F" "c \<noteq> 0" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1772 |
assumes "filterlim g at_infinity F" |
63546 | 1773 |
shows "filterlim (\<lambda>x. f x * g x) at_infinity F" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1774 |
proof - |
61973 | 1775 |
have "((\<lambda>x. inverse (f x) * inverse (g x)) \<longlongrightarrow> inverse c * 0) F" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1776 |
by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) |
63546 | 1777 |
then have "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" |
1778 |
unfolding filterlim_at |
|
1779 |
using assms |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1780 |
by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) |
63546 | 1781 |
then show ?thesis |
1782 |
by (subst filterlim_inverse_at_iff[symmetric]) simp_all |
|
68611 | 1783 |
qed |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1784 |
|
61973 | 1785 |
lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) \<longlongrightarrow> 0) F" |
73885
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1786 |
by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff) |
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1787 |
|
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1788 |
lemma filterlim_inverse_at_top_iff: |
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1789 |
"eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> (LIM x F. inverse (f x) :> at_top) \<longleftrightarrow> (f \<longlongrightarrow> (0 :: real)) F" |
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1790 |
by (auto dest: tendsto_inverse_0_at_top filterlim_inverse_at_top) |
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1791 |
|
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1792 |
lemma filterlim_at_top_iff_inverse_0: |
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1793 |
"eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> (LIM x F. f x :> at_top) \<longleftrightarrow> ((inverse \<circ> f) \<longlongrightarrow> (0 :: real)) F" |
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1794 |
using filterlim_inverse_at_top_iff [of "inverse \<circ> f"] by auto |
50419 | 1795 |
|
63556 | 1796 |
lemma real_tendsto_divide_at_top: |
1797 |
fixes c::"real" |
|
1798 |
assumes "(f \<longlongrightarrow> c) F" |
|
1799 |
assumes "filterlim g at_top F" |
|
1800 |
shows "((\<lambda>x. f x / g x) \<longlongrightarrow> 0) F" |
|
1801 |
by (auto simp: divide_inverse_commute |
|
1802 |
intro!: tendsto_mult[THEN tendsto_eq_rhs] tendsto_inverse_0_at_top assms) |
|
1803 |
||
63546 | 1804 |
lemma mult_nat_left_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. c * x) at_top sequentially" |
1805 |
for c :: nat |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
65680
diff
changeset
|
1806 |
by (rule filterlim_subseq) (auto simp: strict_mono_def) |
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1807 |
|
63546 | 1808 |
lemma mult_nat_right_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. x * c) at_top sequentially" |
1809 |
for c :: nat |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
65680
diff
changeset
|
1810 |
by (rule filterlim_subseq) (auto simp: strict_mono_def) |
63546 | 1811 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1812 |
lemma filterlim_times_pos: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1813 |
"LIM x F1. c * f x :> at_right l" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1814 |
if "filterlim f (at_right p) F1" "0 < c" "l = c * p" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1815 |
for c::"'a::{linordered_field, linorder_topology}" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1816 |
unfolding filterlim_iff |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1817 |
proof safe |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1818 |
fix P |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1819 |
assume "\<forall>\<^sub>F x in at_right l. P x" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1820 |
then obtain d where "c * p < d" "\<And>y. y > c * p \<Longrightarrow> y < d \<Longrightarrow> P y" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1821 |
unfolding \<open>l = _ \<close> eventually_at_right_field |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1822 |
by auto |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1823 |
then have "\<forall>\<^sub>F a in at_right p. P (c * a)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1824 |
by (auto simp: eventually_at_right_field \<open>0 < c\<close> field_simps intro!: exI[where x="d/c"]) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1825 |
from that(1)[unfolded filterlim_iff, rule_format, OF this] |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1826 |
show "\<forall>\<^sub>F x in F1. P (c * f x)" . |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1827 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1828 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1829 |
lemma filtermap_nhds_times: "c \<noteq> 0 \<Longrightarrow> filtermap (times c) (nhds a) = nhds (c * a)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1830 |
for a c :: "'a::real_normed_field" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1831 |
by (rule filtermap_fun_inverse[where g="\<lambda>x. inverse c * x"]) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1832 |
(auto intro!: tendsto_eq_intros filterlim_ident) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1833 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1834 |
lemma filtermap_times_pos_at_right: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1835 |
fixes c::"'a::{linordered_field, linorder_topology}" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1836 |
assumes "c > 0" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1837 |
shows "filtermap (times c) (at_right p) = at_right (c * p)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1838 |
using assms |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1839 |
by (intro filtermap_fun_inverse[where g="\<lambda>x. inverse c * x"]) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1840 |
(auto intro!: filterlim_ident filterlim_times_pos) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
1841 |
|
63546 | 1842 |
lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity" |
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1843 |
proof (rule antisym) |
61973 | 1844 |
have "(inverse \<longlongrightarrow> (0::'a)) at_infinity" |
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1845 |
by (fact tendsto_inverse_0) |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1846 |
then show "filtermap inverse at_infinity \<le> at (0::'a)" |
68615 | 1847 |
using filterlim_def filterlim_ident filterlim_inverse_at_iff by fastforce |
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1848 |
next |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1849 |
have "filtermap inverse (filtermap inverse (at (0::'a))) \<le> filtermap inverse at_infinity" |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1850 |
using filterlim_inverse_at_infinity unfolding filterlim_def |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1851 |
by (rule filtermap_mono) |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1852 |
then show "at (0::'a) \<le> filtermap inverse at_infinity" |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1853 |
by (simp add: filtermap_ident filtermap_filtermap) |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1854 |
qed |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1855 |
|
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1856 |
lemma lim_at_infinity_0: |
63546 | 1857 |
fixes l :: "'a::{real_normed_field,field}" |
1858 |
shows "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> ((f \<circ> inverse) \<longlongrightarrow> l) (at (0::'a))" |
|
1859 |
by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) |
|
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1860 |
|
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1861 |
lemma lim_zero_infinity: |
63546 | 1862 |
fixes l :: "'a::{real_normed_field,field}" |
61973 | 1863 |
shows "((\<lambda>x. f(1 / x)) \<longlongrightarrow> l) (at (0::'a)) \<Longrightarrow> (f \<longlongrightarrow> l) at_infinity" |
63546 | 1864 |
by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) |
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1865 |
|
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
1866 |
|
60758 | 1867 |
text \<open> |
63546 | 1868 |
We only show rules for multiplication and addition when the functions are either against a real |
1869 |
value or against infinity. Further rules are easy to derive by using @{thm |
|
1870 |
filterlim_uminus_at_top}. |
|
60758 | 1871 |
\<close> |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1872 |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
1873 |
lemma filterlim_tendsto_pos_mult_at_top: |
63546 | 1874 |
assumes f: "(f \<longlongrightarrow> c) F" |
1875 |
and c: "0 < c" |
|
1876 |
and g: "LIM x F. g x :> at_top" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1877 |
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
|
1878 |
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
|
1879 |
proof safe |
63546 | 1880 |
fix Z :: real |
1881 |
assume "0 < Z" |
|
60758 | 1882 |
from f \<open>0 < c\<close> have "eventually (\<lambda>x. c / 2 < f x) F" |
61810 | 1883 |
by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono |
63546 | 1884 |
simp: dist_real_def abs_real_def split: if_split_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
|
1885 |
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
|
1886 |
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
|
1887 |
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
|
1888 |
proof eventually_elim |
63546 | 1889 |
case (elim x) |
60758 | 1890 |
with \<open>0 < Z\<close> \<open>0 < c\<close> have "c / 2 * (Z / c * 2) \<le> f x * g x" |
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
|
1891 |
by (intro mult_mono) (auto simp: zero_le_divide_iff) |
60758 | 1892 |
with \<open>0 < c\<close> 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
|
1893 |
by simp |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1894 |
qed |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1895 |
qed |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1896 |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
1897 |
lemma filterlim_at_top_mult_at_top: |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1898 |
assumes f: "LIM x F. f x :> at_top" |
63546 | 1899 |
and g: "LIM x F. g x :> at_top" |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1900 |
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
|
1901 |
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
|
1902 |
proof safe |
63546 | 1903 |
fix Z :: real |
1904 |
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
|
1905 |
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
|
1906 |
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
|
1907 |
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
|
1908 |
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
|
1909 |
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
|
1910 |
proof eventually_elim |
63546 | 1911 |
case (elim x) |
60758 | 1912 |
with \<open>0 < Z\<close> have "1 * Z \<le> f x * g x" |
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
|
1913 |
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
|
1914 |
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
|
1915 |
by simp |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1916 |
qed |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1917 |
qed |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1918 |
|
63556 | 1919 |
lemma filterlim_at_top_mult_tendsto_pos: |
1920 |
assumes f: "(f \<longlongrightarrow> c) F" |
|
1921 |
and c: "0 < c" |
|
1922 |
and g: "LIM x F. g x :> at_top" |
|
1923 |
shows "LIM x F. (g x * f x:: real) :> at_top" |
|
1924 |
by (auto simp: mult.commute intro!: filterlim_tendsto_pos_mult_at_top f c g) |
|
1925 |
||
50419 | 1926 |
lemma filterlim_tendsto_pos_mult_at_bot: |
63546 | 1927 |
fixes c :: real |
1928 |
assumes "(f \<longlongrightarrow> c) F" "0 < c" "filterlim g at_bot F" |
|
50419 | 1929 |
shows "LIM x F. f x * g x :> at_bot" |
1930 |
using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\<lambda>x. - g x"] assms(3) |
|
1931 |
unfolding filterlim_uminus_at_bot by simp |
|
1932 |
||
60182
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
hoelzl
parents:
60141
diff
changeset
|
1933 |
lemma filterlim_tendsto_neg_mult_at_bot: |
63546 | 1934 |
fixes c :: real |
1935 |
assumes c: "(f \<longlongrightarrow> c) F" "c < 0" and g: "filterlim g at_top F" |
|
60182
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
hoelzl
parents:
60141
diff
changeset
|
1936 |
shows "LIM x F. f x * g x :> at_bot" |
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
hoelzl
parents:
60141
diff
changeset
|
1937 |
using c filterlim_tendsto_pos_mult_at_top[of "\<lambda>x. - f x" "- c" F, OF _ _ g] |
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
hoelzl
parents:
60141
diff
changeset
|
1938 |
unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp |
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
hoelzl
parents:
60141
diff
changeset
|
1939 |
|
56330 | 1940 |
lemma filterlim_pow_at_top: |
63721 | 1941 |
fixes f :: "'a \<Rightarrow> real" |
63546 | 1942 |
assumes "0 < n" |
1943 |
and f: "LIM x F. f x :> at_top" |
|
56330 | 1944 |
shows "LIM x F. (f x)^n :: real :> at_top" |
63546 | 1945 |
using \<open>0 < n\<close> |
1946 |
proof (induct n) |
|
1947 |
case 0 |
|
1948 |
then show ?case by simp |
|
1949 |
next |
|
56330 | 1950 |
case (Suc n) with f show ?case |
1951 |
by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top) |
|
63546 | 1952 |
qed |
56330 | 1953 |
|
1954 |
lemma filterlim_pow_at_bot_even: |
|
1955 |
fixes f :: "real \<Rightarrow> real" |
|
1956 |
shows "0 < n \<Longrightarrow> LIM x F. f x :> at_bot \<Longrightarrow> even n \<Longrightarrow> LIM x F. (f x)^n :> at_top" |
|
1957 |
using filterlim_pow_at_top[of n "\<lambda>x. - f x" F] by (simp add: filterlim_uminus_at_top) |
|
1958 |
||
1959 |
lemma filterlim_pow_at_bot_odd: |
|
1960 |
fixes f :: "real \<Rightarrow> real" |
|
1961 |
shows "0 < n \<Longrightarrow> LIM x F. f x :> at_bot \<Longrightarrow> odd n \<Longrightarrow> LIM x F. (f x)^n :> at_bot" |
|
1962 |
using filterlim_pow_at_top[of n "\<lambda>x. - f x" F] by (simp add: filterlim_uminus_at_bot) |
|
1963 |
||
67371
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1964 |
lemma filterlim_power_at_infinity [tendsto_intros]: |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1965 |
fixes F and f :: "'a \<Rightarrow> 'b :: real_normed_div_algebra" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1966 |
assumes "filterlim f at_infinity F" "n > 0" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1967 |
shows "filterlim (\<lambda>x. f x ^ n) at_infinity F" |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1968 |
by (rule filterlim_norm_at_top_imp_at_infinity) |
68611 | 1969 |
(auto simp: norm_power intro!: filterlim_pow_at_top assms |
67371
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1970 |
intro: filterlim_at_infinity_imp_norm_at_top) |
2d9cf74943e1
moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
1971 |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
1972 |
lemma filterlim_tendsto_add_at_top: |
61973 | 1973 |
assumes f: "(f \<longlongrightarrow> c) F" |
63546 | 1974 |
and g: "LIM x F. g x :> at_top" |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1975 |
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
|
1976 |
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
|
1977 |
proof safe |
63546 | 1978 |
fix Z :: real |
1979 |
assume "0 < Z" |
|
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1980 |
from f have "eventually (\<lambda>x. c - 1 < f x) F" |
61810 | 1981 |
by (auto dest!: tendstoD[where e=1] elim!: eventually_mono 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
|
1982 |
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
|
1983 |
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
|
1984 |
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
|
1985 |
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
|
1986 |
qed |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1987 |
|
50347 | 1988 |
lemma LIM_at_top_divide: |
1989 |
fixes f g :: "'a \<Rightarrow> real" |
|
61973 | 1990 |
assumes f: "(f \<longlongrightarrow> a) F" "0 < a" |
63546 | 1991 |
and g: "(g \<longlongrightarrow> 0) F" "eventually (\<lambda>x. 0 < g x) F" |
50347 | 1992 |
shows "LIM x F. f x / g x :> at_top" |
1993 |
unfolding divide_inverse |
|
1994 |
by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) |
|
1995 |
||
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
1996 |
lemma filterlim_at_top_add_at_top: |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1997 |
assumes f: "LIM x F. f x :> at_top" |
63546 | 1998 |
and g: "LIM x F. g x :> at_top" |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
1999 |
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
|
2000 |
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
|
2001 |
proof safe |
63546 | 2002 |
fix Z :: real |
2003 |
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
|
2004 |
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
|
2005 |
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
|
2006 |
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
|
2007 |
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
|
2008 |
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
|
2009 |
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
|
2010 |
qed |
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
2011 |
|
50331 | 2012 |
lemma tendsto_divide_0: |
61076 | 2013 |
fixes f :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}" |
61973 | 2014 |
assumes f: "(f \<longlongrightarrow> c) F" |
63546 | 2015 |
and g: "LIM x F. g x :> at_infinity" |
61973 | 2016 |
shows "((\<lambda>x. f x / g x) \<longlongrightarrow> 0) F" |
63546 | 2017 |
using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] |
2018 |
by (simp add: divide_inverse) |
|
50331 | 2019 |
|
2020 |
lemma linear_plus_1_le_power: |
|
2021 |
fixes x :: real |
|
2022 |
assumes x: "0 \<le> x" |
|
2023 |
shows "real n * x + 1 \<le> (x + 1) ^ n" |
|
2024 |
proof (induct n) |
|
63546 | 2025 |
case 0 |
2026 |
then show ?case by simp |
|
2027 |
next |
|
50331 | 2028 |
case (Suc n) |
63546 | 2029 |
from x have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)" |
2030 |
by (simp add: field_simps) |
|
50331 | 2031 |
also have "\<dots> \<le> (x + 1)^Suc n" |
2032 |
using Suc x by (simp add: mult_left_mono) |
|
2033 |
finally show ?case . |
|
63546 | 2034 |
qed |
50331 | 2035 |
|
2036 |
lemma filterlim_realpow_sequentially_gt1: |
|
2037 |
fixes x :: "'a :: real_normed_div_algebra" |
|
2038 |
assumes x[arith]: "1 < norm x" |
|
2039 |
shows "LIM n sequentially. x ^ n :> at_infinity" |
|
2040 |
proof (intro filterlim_at_infinity[THEN iffD2] allI impI) |
|
63546 | 2041 |
fix y :: real |
2042 |
assume "0 < y" |
|
72219
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
2043 |
obtain N :: nat where "y < real N * (norm x - 1)" |
0f38c96a0a74
tidying up some theorem statements
paulson <lp15@cam.ac.uk>
parents:
71837
diff
changeset
|
2044 |
by (meson diff_gt_0_iff_gt reals_Archimedean3 x) |
63546 | 2045 |
also have "\<dots> \<le> real N * (norm x - 1) + 1" |
2046 |
by simp |
|
2047 |
also have "\<dots> \<le> (norm x - 1 + 1) ^ N" |
|
2048 |
by (rule linear_plus_1_le_power) simp |
|
2049 |
also have "\<dots> = norm x ^ N" |
|
2050 |
by simp |
|
50331 | 2051 |
finally have "\<forall>n\<ge>N. y \<le> norm x ^ n" |
2052 |
by (metis order_less_le_trans power_increasing order_less_imp_le x) |
|
2053 |
then show "eventually (\<lambda>n. y \<le> norm (x ^ n)) sequentially" |
|
2054 |
unfolding eventually_sequentially |
|
2055 |
by (auto simp: norm_power) |
|
2056 |
qed simp |
|
2057 |
||
51471 | 2058 |
|
66456
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2059 |
lemma filterlim_divide_at_infinity: |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2060 |
fixes f g :: "'a \<Rightarrow> 'a :: real_normed_field" |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2061 |
assumes "filterlim f (nhds c) F" "filterlim g (at 0) F" "c \<noteq> 0" |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2062 |
shows "filterlim (\<lambda>x. f x / g x) at_infinity F" |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2063 |
proof - |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2064 |
have "filterlim (\<lambda>x. f x * inverse (g x)) at_infinity F" |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2065 |
by (intro tendsto_mult_filterlim_at_infinity[OF assms(1,3)] |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2066 |
filterlim_compose [OF filterlim_inverse_at_infinity assms(2)]) |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2067 |
thus ?thesis by (simp add: field_simps) |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2068 |
qed |
621897f47fab
Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
2069 |
|
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2070 |
subsection \<open>Floor and Ceiling\<close> |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2071 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2072 |
lemma eventually_floor_less: |
63546 | 2073 |
fixes f :: "'a \<Rightarrow> 'b::{order_topology,floor_ceiling}" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2074 |
assumes f: "(f \<longlongrightarrow> l) F" |
63546 | 2075 |
and l: "l \<notin> \<int>" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2076 |
shows "\<forall>\<^sub>F x in F. of_int (floor l) < f x" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2077 |
by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2078 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2079 |
lemma eventually_less_ceiling: |
63546 | 2080 |
fixes f :: "'a \<Rightarrow> 'b::{order_topology,floor_ceiling}" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2081 |
assumes f: "(f \<longlongrightarrow> l) F" |
63546 | 2082 |
and l: "l \<notin> \<int>" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2083 |
shows "\<forall>\<^sub>F x in F. f x < of_int (ceiling l)" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2084 |
by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2085 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2086 |
lemma eventually_floor_eq: |
63546 | 2087 |
fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2088 |
assumes f: "(f \<longlongrightarrow> l) F" |
63546 | 2089 |
and l: "l \<notin> \<int>" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2090 |
shows "\<forall>\<^sub>F x in F. floor (f x) = floor l" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2091 |
using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2092 |
by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2093 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2094 |
lemma eventually_ceiling_eq: |
63546 | 2095 |
fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2096 |
assumes f: "(f \<longlongrightarrow> l) F" |
63546 | 2097 |
and l: "l \<notin> \<int>" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2098 |
shows "\<forall>\<^sub>F x in F. ceiling (f x) = ceiling l" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2099 |
using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2100 |
by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2101 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2102 |
lemma tendsto_of_int_floor: |
63546 | 2103 |
fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2104 |
assumes "(f \<longlongrightarrow> l) F" |
63546 | 2105 |
and "l \<notin> \<int>" |
2106 |
shows "((\<lambda>x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \<longlongrightarrow> of_int (floor l)) F" |
|
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2107 |
using eventually_floor_eq[OF assms] |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2108 |
by (simp add: eventually_mono topological_tendstoI) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2109 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2110 |
lemma tendsto_of_int_ceiling: |
63546 | 2111 |
fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2112 |
assumes "(f \<longlongrightarrow> l) F" |
63546 | 2113 |
and "l \<notin> \<int>" |
2114 |
shows "((\<lambda>x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \<longlongrightarrow> of_int (ceiling l)) F" |
|
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2115 |
using eventually_ceiling_eq[OF assms] |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2116 |
by (simp add: eventually_mono topological_tendstoI) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2117 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2118 |
lemma continuous_on_of_int_floor: |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2119 |
"continuous_on (UNIV - \<int>::'a::{order_topology, floor_ceiling} set) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2120 |
(\<lambda>x. of_int (floor x)::'b::{ring_1, topological_space})" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2121 |
unfolding continuous_on_def |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2122 |
by (auto intro!: tendsto_of_int_floor) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2123 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2124 |
lemma continuous_on_of_int_ceiling: |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2125 |
"continuous_on (UNIV - \<int>::'a::{order_topology, floor_ceiling} set) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2126 |
(\<lambda>x. of_int (ceiling x)::'b::{ring_1, topological_space})" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2127 |
unfolding continuous_on_def |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2128 |
by (auto intro!: tendsto_of_int_ceiling) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2129 |
|
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63104
diff
changeset
|
2130 |
|
60758 | 2131 |
subsection \<open>Limits of Sequences\<close> |
51526 | 2132 |
|
62368 | 2133 |
lemma [trans]: "X = Y \<Longrightarrow> Y \<longlonglongrightarrow> z \<Longrightarrow> X \<longlonglongrightarrow> z" |
51526 | 2134 |
by simp |
2135 |
||
2136 |
lemma LIMSEQ_iff: |
|
2137 |
fixes L :: "'a::real_normed_vector" |
|
61969 | 2138 |
shows "(X \<longlonglongrightarrow> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)" |
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
59867
diff
changeset
|
2139 |
unfolding lim_sequentially dist_norm .. |
51526 | 2140 |
|
63546 | 2141 |
lemma LIMSEQ_I: "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X \<longlonglongrightarrow> L" |
2142 |
for L :: "'a::real_normed_vector" |
|
2143 |
by (simp add: LIMSEQ_iff) |
|
2144 |
||
2145 |
lemma LIMSEQ_D: "X \<longlonglongrightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r" |
|
2146 |
for L :: "'a::real_normed_vector" |
|
2147 |
by (simp add: LIMSEQ_iff) |
|
2148 |
||
2149 |
lemma LIMSEQ_linear: "X \<longlonglongrightarrow> x \<Longrightarrow> l > 0 \<Longrightarrow> (\<lambda> n. X (n * l)) \<longlonglongrightarrow> x" |
|
51526 | 2150 |
unfolding tendsto_def eventually_sequentially |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57447
diff
changeset
|
2151 |
by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute) |
51526 | 2152 |
|
63546 | 2153 |
|
2154 |
text \<open>Transformation of limit.\<close> |
|
2155 |
||
2156 |
lemma Lim_transform: "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F" |
|
2157 |
for a b :: "'a::real_normed_vector" |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2158 |
using tendsto_add [of g a F "\<lambda>x. f x - g x" 0] by simp |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2159 |
|
63546 | 2160 |
lemma Lim_transform2: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> a) F" |
2161 |
for a b :: "'a::real_normed_vector" |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2162 |
by (erule Lim_transform) (simp add: tendsto_minus_cancel) |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2163 |
|
63546 | 2164 |
proposition Lim_transform_eq: "((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F \<longleftrightarrow> (g \<longlongrightarrow> a) F" |
2165 |
for a :: "'a::real_normed_vector" |
|
2166 |
using Lim_transform Lim_transform2 by blast |
|
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62369
diff
changeset
|
2167 |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2168 |
lemma Lim_transform_eventually: |
70532
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents:
70365
diff
changeset
|
2169 |
"\<lbrakk>(f \<longlongrightarrow> l) F; eventually (\<lambda>x. f x = g x) F\<rbrakk> \<Longrightarrow> (g \<longlongrightarrow> l) F" |
68615 | 2170 |
using eventually_elim2 by (fastforce simp add: tendsto_def) |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2171 |
|
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2172 |
lemma Lim_transform_within: |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2173 |
assumes "(f \<longlongrightarrow> l) (at x within S)" |
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2174 |
and "0 < d" |
63546 | 2175 |
and "\<And>x'. x'\<in>S \<Longrightarrow> 0 < dist x' x \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x'" |
61973 | 2176 |
shows "(g \<longlongrightarrow> l) (at x within S)" |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2177 |
proof (rule Lim_transform_eventually) |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2178 |
show "eventually (\<lambda>x. f x = g x) (at x within S)" |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2179 |
using assms by (auto simp: eventually_at) |
63546 | 2180 |
show "(f \<longlongrightarrow> l) (at x within S)" |
2181 |
by fact |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2182 |
qed |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2183 |
|
67706
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2184 |
lemma filterlim_transform_within: |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2185 |
assumes "filterlim g G (at x within S)" |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2186 |
assumes "G \<le> F" "0<d" "(\<And>x'. x' \<in> S \<Longrightarrow> 0 < dist x' x \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') " |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2187 |
shows "filterlim f F (at x within S)" |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2188 |
using assms |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2189 |
apply (elim filterlim_mono_eventually) |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2190 |
unfolding eventually_at by auto |
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents:
67673
diff
changeset
|
2191 |
|
63546 | 2192 |
text \<open>Common case assuming being away from some crucial point like 0.\<close> |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2193 |
lemma Lim_transform_away_within: |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2194 |
fixes a b :: "'a::t1_space" |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2195 |
assumes "a \<noteq> b" |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2196 |
and "\<forall>x\<in>S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" |
61973 | 2197 |
and "(f \<longlongrightarrow> l) (at a within S)" |
2198 |
shows "(g \<longlongrightarrow> l) (at a within S)" |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2199 |
proof (rule Lim_transform_eventually) |
63546 | 2200 |
show "(f \<longlongrightarrow> l) (at a within S)" |
2201 |
by fact |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2202 |
show "eventually (\<lambda>x. f x = g x) (at a within S)" |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2203 |
unfolding eventually_at_topological |
63546 | 2204 |
by (rule exI [where x="- {b}"]) (simp add: open_Compl assms) |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2205 |
qed |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2206 |
|
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2207 |
lemma Lim_transform_away_at: |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2208 |
fixes a b :: "'a::t1_space" |
63546 | 2209 |
assumes ab: "a \<noteq> b" |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2210 |
and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" |
61973 | 2211 |
and fl: "(f \<longlongrightarrow> l) (at a)" |
2212 |
shows "(g \<longlongrightarrow> l) (at a)" |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2213 |
using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2214 |
|
63546 | 2215 |
text \<open>Alternatively, within an open set.\<close> |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2216 |
lemma Lim_transform_within_open: |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2217 |
assumes "(f \<longlongrightarrow> l) (at a within T)" |
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2218 |
and "open s" and "a \<in> s" |
63546 | 2219 |
and "\<And>x. x\<in>s \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x" |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2220 |
shows "(g \<longlongrightarrow> l) (at a within T)" |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2221 |
proof (rule Lim_transform_eventually) |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2222 |
show "eventually (\<lambda>x. f x = g x) (at a within T)" |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2223 |
unfolding eventually_at_topological |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2224 |
using assms by auto |
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2225 |
show "(f \<longlongrightarrow> l) (at a within T)" by fact |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2226 |
qed |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2227 |
|
63546 | 2228 |
|
2229 |
text \<open>A congruence rule allowing us to transform limits assuming not at point.\<close> |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2230 |
|
72220 | 2231 |
lemma Lim_cong_within: |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2232 |
assumes "a = b" |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2233 |
and "x = y" |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2234 |
and "S = T" |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2235 |
and "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x" |
61973 | 2236 |
shows "(f \<longlongrightarrow> x) (at a within S) \<longleftrightarrow> (g \<longlongrightarrow> y) (at b within T)" |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2237 |
unfolding tendsto_def eventually_at_topological |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2238 |
using assms by simp |
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2239 |
|
63546 | 2240 |
text \<open>An unbounded sequence's inverse tends to 0.\<close> |
65578
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
2241 |
lemma LIMSEQ_inverse_zero: |
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
2242 |
assumes "\<And>r::real. \<exists>N. \<forall>n\<ge>N. r < X n" |
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
2243 |
shows "(\<lambda>n. inverse (X n)) \<longlonglongrightarrow> 0" |
51526 | 2244 |
apply (rule filterlim_compose[OF tendsto_inverse_0]) |
68615 | 2245 |
by (metis assms eventually_at_top_linorderI filterlim_at_top_dense filterlim_at_top_imp_at_infinity) |
51526 | 2246 |
|
69593 | 2247 |
text \<open>The sequence \<^term>\<open>1/n\<close> tends to 0 as \<^term>\<open>n\<close> tends to infinity.\<close> |
63546 | 2248 |
lemma LIMSEQ_inverse_real_of_nat: "(\<lambda>n. inverse (real (Suc n))) \<longlonglongrightarrow> 0" |
51526 | 2249 |
by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc |
63546 | 2250 |
filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity) |
2251 |
||
2252 |
text \<open> |
|
69593 | 2253 |
The sequence \<^term>\<open>r + 1/n\<close> tends to \<^term>\<open>r\<close> as \<^term>\<open>n\<close> tends to |
63546 | 2254 |
infinity is now easily proved. |
2255 |
\<close> |
|
2256 |
||
2257 |
lemma LIMSEQ_inverse_real_of_nat_add: "(\<lambda>n. r + inverse (real (Suc n))) \<longlonglongrightarrow> r" |
|
51526 | 2258 |
using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto |
2259 |
||
63546 | 2260 |
lemma LIMSEQ_inverse_real_of_nat_add_minus: "(\<lambda>n. r + -inverse (real (Suc n))) \<longlonglongrightarrow> r" |
51526 | 2261 |
using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]] |
2262 |
by auto |
|
2263 |
||
63546 | 2264 |
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow> r" |
51526 | 2265 |
using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] |
2266 |
by auto |
|
2267 |
||
61973 | 2268 |
lemma lim_inverse_n: "((\<lambda>n. inverse(of_nat n)) \<longlongrightarrow> (0::'a::real_normed_field)) sequentially" |
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2269 |
using lim_1_over_n by (simp add: inverse_eq_divide) |
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2270 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
2271 |
lemma lim_inverse_n': "((\<lambda>n. 1 / n) \<longlongrightarrow> 0) sequentially" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
2272 |
using lim_inverse_n |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
2273 |
by (simp add: inverse_eq_divide) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67673
diff
changeset
|
2274 |
|
61969 | 2275 |
lemma LIMSEQ_Suc_n_over_n: "(\<lambda>n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \<longlonglongrightarrow> 1" |
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2276 |
proof (rule Lim_transform_eventually) |
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2277 |
show "eventually (\<lambda>n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially" |
63546 | 2278 |
using eventually_gt_at_top[of "0::nat"] |
2279 |
by eventually_elim (simp add: field_simps) |
|
61969 | 2280 |
have "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1 + 0" |
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2281 |
by (intro tendsto_add tendsto_const lim_inverse_n) |
63546 | 2282 |
then show "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1" |
2283 |
by simp |
|
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2284 |
qed |
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2285 |
|
61969 | 2286 |
lemma LIMSEQ_n_over_Suc_n: "(\<lambda>n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \<longlonglongrightarrow> 1" |
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2287 |
proof (rule Lim_transform_eventually) |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2288 |
show "eventually (\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a) = |
63546 | 2289 |
of_nat n / of_nat (Suc n)) sequentially" |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2290 |
using eventually_gt_at_top[of "0::nat"] |
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2291 |
by eventually_elim (simp add: field_simps del: of_nat_Suc) |
61969 | 2292 |
have "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> inverse 1" |
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2293 |
by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all |
63546 | 2294 |
then show "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> 1" |
2295 |
by simp |
|
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2296 |
qed |
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61169
diff
changeset
|
2297 |
|
63546 | 2298 |
|
60758 | 2299 |
subsection \<open>Convergence on sequences\<close> |
51526 | 2300 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2301 |
lemma convergent_cong: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2302 |
assumes "eventually (\<lambda>x. f x = g x) sequentially" |
63546 | 2303 |
shows "convergent f \<longleftrightarrow> convergent g" |
2304 |
unfolding convergent_def |
|
2305 |
by (subst filterlim_cong[OF refl refl assms]) (rule refl) |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2306 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2307 |
lemma convergent_Suc_iff: "convergent (\<lambda>n. f (Suc n)) \<longleftrightarrow> convergent f" |
71827 | 2308 |
by (auto simp: convergent_def filterlim_sequentially_Suc) |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2309 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2310 |
lemma convergent_ignore_initial_segment: "convergent (\<lambda>n. f (n + m)) = convergent f" |
63546 | 2311 |
proof (induct m arbitrary: f) |
2312 |
case 0 |
|
2313 |
then show ?case by simp |
|
2314 |
next |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2315 |
case (Suc m) |
63546 | 2316 |
have "convergent (\<lambda>n. f (n + Suc m)) \<longleftrightarrow> convergent (\<lambda>n. f (Suc n + m))" |
2317 |
by simp |
|
2318 |
also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. f (n + m))" |
|
2319 |
by (rule convergent_Suc_iff) |
|
2320 |
also have "\<dots> \<longleftrightarrow> convergent f" |
|
2321 |
by (rule Suc) |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2322 |
finally show ?case . |
63546 | 2323 |
qed |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2324 |
|
51526 | 2325 |
lemma convergent_add: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2326 |
fixes X Y :: "nat \<Rightarrow> 'a::topological_monoid_add" |
51526 | 2327 |
assumes "convergent (\<lambda>n. X n)" |
63546 | 2328 |
and "convergent (\<lambda>n. Y n)" |
51526 | 2329 |
shows "convergent (\<lambda>n. X n + Y n)" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
2330 |
using assms unfolding convergent_def by (blast intro: tendsto_add) |
51526 | 2331 |
|
64267 | 2332 |
lemma convergent_sum: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2333 |
fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::topological_comm_monoid_add" |
63915 | 2334 |
shows "(\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)) \<Longrightarrow> convergent (\<lambda>n. \<Sum>i\<in>A. X i n)" |
2335 |
by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add) |
|
51526 | 2336 |
|
2337 |
lemma (in bounded_linear) convergent: |
|
2338 |
assumes "convergent (\<lambda>n. X n)" |
|
2339 |
shows "convergent (\<lambda>n. f (X n))" |
|
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
2340 |
using assms unfolding convergent_def by (blast intro: tendsto) |
51526 | 2341 |
|
2342 |
lemma (in bounded_bilinear) convergent: |
|
63546 | 2343 |
assumes "convergent (\<lambda>n. X n)" |
2344 |
and "convergent (\<lambda>n. Y n)" |
|
51526 | 2345 |
shows "convergent (\<lambda>n. X n ** Y n)" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
2346 |
using assms unfolding convergent_def by (blast intro: tendsto) |
51526 | 2347 |
|
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2348 |
lemma convergent_minus_iff: |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2349 |
fixes X :: "nat \<Rightarrow> 'a::topological_group_add" |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2350 |
shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)" |
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2351 |
unfolding convergent_def by (force dest: tendsto_minus) |
51526 | 2352 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2353 |
lemma convergent_diff: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2354 |
fixes X Y :: "nat \<Rightarrow> 'a::topological_group_add" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2355 |
assumes "convergent (\<lambda>n. X n)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2356 |
assumes "convergent (\<lambda>n. Y n)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2357 |
shows "convergent (\<lambda>n. X n - Y n)" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
2358 |
using assms unfolding convergent_def by (blast intro: tendsto_diff) |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2359 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2360 |
lemma convergent_norm: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2361 |
assumes "convergent f" |
63546 | 2362 |
shows "convergent (\<lambda>n. norm (f n))" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2363 |
proof - |
63546 | 2364 |
from assms have "f \<longlonglongrightarrow> lim f" |
2365 |
by (simp add: convergent_LIMSEQ_iff) |
|
2366 |
then have "(\<lambda>n. norm (f n)) \<longlonglongrightarrow> norm (lim f)" |
|
2367 |
by (rule tendsto_norm) |
|
2368 |
then show ?thesis |
|
2369 |
by (auto simp: convergent_def) |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2370 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2371 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2372 |
lemma convergent_of_real: |
63546 | 2373 |
"convergent f \<Longrightarrow> convergent (\<lambda>n. of_real (f n) :: 'a::real_normed_algebra_1)" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2374 |
unfolding convergent_def by (blast intro!: tendsto_of_real) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2375 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2376 |
lemma convergent_add_const_iff: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2377 |
"convergent (\<lambda>n. c + f n :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2378 |
proof |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2379 |
assume "convergent (\<lambda>n. c + f n)" |
63546 | 2380 |
from convergent_diff[OF this convergent_const[of c]] show "convergent f" |
2381 |
by simp |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2382 |
next |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2383 |
assume "convergent f" |
63546 | 2384 |
from convergent_add[OF convergent_const[of c] this] show "convergent (\<lambda>n. c + f n)" |
2385 |
by simp |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2386 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2387 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2388 |
lemma convergent_add_const_right_iff: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2389 |
"convergent (\<lambda>n. f n + c :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2390 |
using convergent_add_const_iff[of c f] by (simp add: add_ac) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2391 |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2392 |
lemma convergent_diff_const_right_iff: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2393 |
"convergent (\<lambda>n. f n - c :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2394 |
using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2395 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2396 |
lemma convergent_mult: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2397 |
fixes X Y :: "nat \<Rightarrow> 'a::topological_semigroup_mult" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2398 |
assumes "convergent (\<lambda>n. X n)" |
63546 | 2399 |
and "convergent (\<lambda>n. Y n)" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2400 |
shows "convergent (\<lambda>n. X n * Y n)" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
2401 |
using assms unfolding convergent_def by (blast intro: tendsto_mult) |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2402 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2403 |
lemma convergent_mult_const_iff: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2404 |
assumes "c \<noteq> 0" |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2405 |
shows "convergent (\<lambda>n. c * f n :: 'a::{field,topological_semigroup_mult}) \<longleftrightarrow> convergent f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2406 |
proof |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2407 |
assume "convergent (\<lambda>n. c * f n)" |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61976
diff
changeset
|
2408 |
from assms convergent_mult[OF this convergent_const[of "inverse c"]] |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2409 |
show "convergent f" by (simp add: field_simps) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2410 |
next |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2411 |
assume "convergent f" |
63546 | 2412 |
from convergent_mult[OF convergent_const[of c] this] show "convergent (\<lambda>n. c * f n)" |
2413 |
by simp |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2414 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2415 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2416 |
lemma convergent_mult_const_right_iff: |
68064
b249fab48c76
type class generalisations; some work on infinite products
paulson <lp15@cam.ac.uk>
parents:
67958
diff
changeset
|
2417 |
fixes c :: "'a::{field,topological_semigroup_mult}" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2418 |
assumes "c \<noteq> 0" |
63546 | 2419 |
shows "convergent (\<lambda>n. f n * c) \<longleftrightarrow> convergent f" |
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2420 |
using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2421 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2422 |
lemma convergent_imp_Bseq: "convergent f \<Longrightarrow> Bseq f" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2423 |
by (simp add: Cauchy_Bseq convergent_Cauchy) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2424 |
|
51526 | 2425 |
|
60758 | 2426 |
text \<open>A monotone sequence converges to its least upper bound.\<close> |
51526 | 2427 |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2428 |
lemma LIMSEQ_incseq_SUP: |
63546 | 2429 |
fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder,linorder_topology}" |
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2430 |
assumes u: "bdd_above (range X)" |
63546 | 2431 |
and X: "incseq X" |
61969 | 2432 |
shows "X \<longlonglongrightarrow> (SUP i. X i)" |
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2433 |
by (rule order_tendstoI) |
63546 | 2434 |
(auto simp: eventually_sequentially u less_cSUP_iff |
2435 |
intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) |
|
51526 | 2436 |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2437 |
lemma LIMSEQ_decseq_INF: |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2438 |
fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2439 |
assumes u: "bdd_below (range X)" |
63546 | 2440 |
and X: "decseq X" |
61969 | 2441 |
shows "X \<longlonglongrightarrow> (INF i. X i)" |
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2442 |
by (rule order_tendstoI) |
63546 | 2443 |
(auto simp: eventually_sequentially u cINF_less_iff |
2444 |
intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) |
|
2445 |
||
2446 |
text \<open>Main monotonicity theorem.\<close> |
|
2447 |
||
2448 |
lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent X" |
|
2449 |
for X :: "nat \<Rightarrow> real" |
|
2450 |
by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP |
|
2451 |
dest: Bseq_bdd_above Bseq_bdd_below) |
|
2452 |
||
2453 |
lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent X" |
|
2454 |
for X :: "nat \<Rightarrow> real" |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
2455 |
by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def) |
51526 | 2456 |
|
63546 | 2457 |
lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \<Longrightarrow> convergent f \<longleftrightarrow> Bseq f" |
2458 |
for f :: "nat \<Rightarrow> real" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2459 |
using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2460 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2461 |
lemma Bseq_monoseq_convergent'_inc: |
63546 | 2462 |
fixes f :: "nat \<Rightarrow> real" |
2463 |
shows "Bseq (\<lambda>n. f (n + M)) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<le> f n) \<Longrightarrow> convergent f" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2464 |
by (subst convergent_ignore_initial_segment [symmetric, of _ M]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2465 |
(auto intro!: Bseq_monoseq_convergent simp: monoseq_def) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2466 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2467 |
lemma Bseq_monoseq_convergent'_dec: |
63546 | 2468 |
fixes f :: "nat \<Rightarrow> real" |
2469 |
shows "Bseq (\<lambda>n. f (n + M)) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<ge> f n) \<Longrightarrow> convergent f" |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2470 |
by (subst convergent_ignore_initial_segment [symmetric, of _ M]) |
63546 | 2471 |
(auto intro!: Bseq_monoseq_convergent simp: monoseq_def) |
2472 |
||
2473 |
lemma Cauchy_iff: "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)" |
|
2474 |
for X :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
51526 | 2475 |
unfolding Cauchy_def dist_norm .. |
2476 |
||
63546 | 2477 |
lemma CauchyI: "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X" |
2478 |
for X :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
2479 |
by (simp add: Cauchy_iff) |
|
2480 |
||
2481 |
lemma CauchyD: "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e" |
|
2482 |
for X :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
2483 |
by (simp add: Cauchy_iff) |
|
51526 | 2484 |
|
2485 |
lemma incseq_convergent: |
|
2486 |
fixes X :: "nat \<Rightarrow> real" |
|
63546 | 2487 |
assumes "incseq X" |
2488 |
and "\<forall>i. X i \<le> B" |
|
61969 | 2489 |
obtains L where "X \<longlonglongrightarrow> L" "\<forall>i. X i \<le> L" |
51526 | 2490 |
proof atomize_elim |
60758 | 2491 |
from incseq_bounded[OF assms] \<open>incseq X\<close> Bseq_monoseq_convergent[of X] |
61969 | 2492 |
obtain L where "X \<longlonglongrightarrow> L" |
51526 | 2493 |
by (auto simp: convergent_def monoseq_def incseq_def) |
61969 | 2494 |
with \<open>incseq X\<close> show "\<exists>L. X \<longlonglongrightarrow> L \<and> (\<forall>i. X i \<le> L)" |
51526 | 2495 |
by (auto intro!: exI[of _ L] incseq_le) |
2496 |
qed |
|
2497 |
||
2498 |
lemma decseq_convergent: |
|
2499 |
fixes X :: "nat \<Rightarrow> real" |
|
63546 | 2500 |
assumes "decseq X" |
2501 |
and "\<forall>i. B \<le> X i" |
|
61969 | 2502 |
obtains L where "X \<longlonglongrightarrow> L" "\<forall>i. L \<le> X i" |
51526 | 2503 |
proof atomize_elim |
60758 | 2504 |
from decseq_bounded[OF assms] \<open>decseq X\<close> Bseq_monoseq_convergent[of X] |
61969 | 2505 |
obtain L where "X \<longlonglongrightarrow> L" |
51526 | 2506 |
by (auto simp: convergent_def monoseq_def decseq_def) |
61969 | 2507 |
with \<open>decseq X\<close> show "\<exists>L. X \<longlonglongrightarrow> L \<and> (\<forall>i. L \<le> X i)" |
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68296
diff
changeset
|
2508 |
by (auto intro!: exI[of _ L] decseq_ge) |
51526 | 2509 |
qed |
2510 |
||
70694
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2511 |
lemma monoseq_convergent: |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2512 |
fixes X :: "nat \<Rightarrow> real" |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2513 |
assumes X: "monoseq X" and B: "\<And>i. \<bar>X i\<bar> \<le> B" |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2514 |
obtains L where "X \<longlonglongrightarrow> L" |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2515 |
using X unfolding monoseq_iff |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2516 |
proof |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2517 |
assume "incseq X" |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2518 |
show thesis |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2519 |
using abs_le_D1 [OF B] incseq_convergent [OF \<open>incseq X\<close>] that by meson |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2520 |
next |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2521 |
assume "decseq X" |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2522 |
show thesis |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2523 |
using decseq_convergent [OF \<open>decseq X\<close>] that |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2524 |
by (metis B abs_le_iff add.inverse_inverse neg_le_iff_le) |
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
paulson <lp15@cam.ac.uk>
parents:
70688
diff
changeset
|
2525 |
qed |
63546 | 2526 |
|
77102
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2527 |
subsection \<open>More about @{term filterlim} (thanks to Wenda Li)\<close> |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2528 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2529 |
lemma filterlim_at_infinity_times: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2530 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_field" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2531 |
assumes "filterlim f at_infinity F" "filterlim g at_infinity F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2532 |
shows "filterlim (\<lambda>x. f x * g x) at_infinity F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2533 |
proof - |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2534 |
have "((\<lambda>x. inverse (f x) * inverse (g x)) \<longlongrightarrow> 0 * 0) F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2535 |
by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2536 |
then have "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at 0) F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2537 |
unfolding filterlim_at using assms |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2538 |
by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2539 |
then show ?thesis |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2540 |
by (subst filterlim_inverse_at_iff[symmetric]) simp_all |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2541 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2542 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2543 |
lemma filterlim_at_top_at_bot[elim]: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2544 |
fixes f::"'a \<Rightarrow> 'b::unbounded_dense_linorder" and F::"'a filter" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2545 |
assumes top:"filterlim f at_top F" and bot: "filterlim f at_bot F" and "F\<noteq>bot" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2546 |
shows False |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2547 |
proof - |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2548 |
obtain c::'b where True by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2549 |
have "\<forall>\<^sub>F x in F. c < f x" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2550 |
using top unfolding filterlim_at_top_dense by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2551 |
moreover have "\<forall>\<^sub>F x in F. f x < c" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2552 |
using bot unfolding filterlim_at_bot_dense by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2553 |
ultimately have "\<forall>\<^sub>F x in F. c < f x \<and> f x < c" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2554 |
using eventually_conj by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2555 |
then have "\<forall>\<^sub>F x in F. False" by (auto elim:eventually_mono) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2556 |
then show False using \<open>F\<noteq>bot\<close> by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2557 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2558 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2559 |
lemma filterlim_at_top_nhds[elim]: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2560 |
fixes f::"'a \<Rightarrow> 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2561 |
assumes top:"filterlim f at_top F" and tendsto: "(f \<longlongrightarrow> c) F" and "F\<noteq>bot" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2562 |
shows False |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2563 |
proof - |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2564 |
obtain c'::'b where "c'>c" using gt_ex by blast |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2565 |
have "\<forall>\<^sub>F x in F. c' < f x" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2566 |
using top unfolding filterlim_at_top_dense by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2567 |
moreover have "\<forall>\<^sub>F x in F. f x < c'" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2568 |
using order_tendstoD[OF tendsto,of c'] \<open>c'>c\<close> by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2569 |
ultimately have "\<forall>\<^sub>F x in F. c' < f x \<and> f x < c'" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2570 |
using eventually_conj by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2571 |
then have "\<forall>\<^sub>F x in F. False" by (auto elim:eventually_mono) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2572 |
then show False using \<open>F\<noteq>bot\<close> by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2573 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2574 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2575 |
lemma filterlim_at_bot_nhds[elim]: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2576 |
fixes f::"'a \<Rightarrow> 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2577 |
assumes top:"filterlim f at_bot F" and tendsto: "(f \<longlongrightarrow> c) F" and "F\<noteq>bot" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2578 |
shows False |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2579 |
proof - |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2580 |
obtain c'::'b where "c'<c" using lt_ex by blast |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2581 |
have "\<forall>\<^sub>F x in F. c' > f x" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2582 |
using top unfolding filterlim_at_bot_dense by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2583 |
moreover have "\<forall>\<^sub>F x in F. f x > c'" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2584 |
using order_tendstoD[OF tendsto,of c'] \<open>c'<c\<close> by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2585 |
ultimately have "\<forall>\<^sub>F x in F. c' < f x \<and> f x < c'" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2586 |
using eventually_conj by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2587 |
then have "\<forall>\<^sub>F x in F. False" by (auto elim:eventually_mono) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2588 |
then show False using \<open>F\<noteq>bot\<close> by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2589 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2590 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2591 |
lemma eventually_times_inverse_1: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2592 |
fixes f::"'a \<Rightarrow> 'b::{field,t2_space}" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2593 |
assumes "(f \<longlongrightarrow> c) F" "c\<noteq>0" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2594 |
shows "\<forall>\<^sub>F x in F. inverse (f x) * f x = 1" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2595 |
by (smt (verit) assms eventually_mono mult.commute right_inverse tendsto_imp_eventually_ne) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2596 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2597 |
lemma filterlim_at_infinity_divide_iff: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2598 |
fixes f::"'a \<Rightarrow> 'b::real_normed_field" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2599 |
assumes "(f \<longlongrightarrow> c) F" "c\<noteq>0" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2600 |
shows "(LIM x F. f x / g x :> at_infinity) \<longleftrightarrow> (LIM x F. g x :> at 0)" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2601 |
proof |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2602 |
assume "LIM x F. f x / g x :> at_infinity" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2603 |
then have "LIM x F. inverse (f x) * (f x / g x) :> at_infinity" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2604 |
using assms tendsto_inverse tendsto_mult_filterlim_at_infinity by fastforce |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2605 |
then have "LIM x F. inverse (g x) :> at_infinity" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2606 |
apply (elim filterlim_mono_eventually) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2607 |
using eventually_times_inverse_1[OF assms] |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2608 |
by (auto elim:eventually_mono simp add:field_simps) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2609 |
then show "filterlim g (at 0) F" using filterlim_inverse_at_iff[symmetric] by force |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2610 |
next |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2611 |
assume "filterlim g (at 0) F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2612 |
then have "filterlim (\<lambda>x. inverse (g x)) at_infinity F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2613 |
using filterlim_compose filterlim_inverse_at_infinity by blast |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2614 |
then have "LIM x F. f x * inverse (g x) :> at_infinity" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2615 |
using tendsto_mult_filterlim_at_infinity[OF assms, of "\<lambda>x. inverse(g x)"] |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2616 |
by simp |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2617 |
then show "LIM x F. f x / g x :> at_infinity" by (simp add: divide_inverse) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2618 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2619 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2620 |
lemma filterlim_tendsto_pos_mult_at_top_iff: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2621 |
fixes f::"'a \<Rightarrow> real" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2622 |
assumes "(f \<longlongrightarrow> c) F" and "0 < c" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2623 |
shows "(LIM x F. (f x * g x) :> at_top) \<longleftrightarrow> (LIM x F. g x :> at_top)" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2624 |
proof |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2625 |
assume "filterlim g at_top F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2626 |
then show "LIM x F. f x * g x :> at_top" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2627 |
using filterlim_tendsto_pos_mult_at_top[OF assms] by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2628 |
next |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2629 |
assume asm:"LIM x F. f x * g x :> at_top" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2630 |
have "((\<lambda>x. inverse (f x)) \<longlongrightarrow> inverse c) F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2631 |
using tendsto_inverse[OF assms(1)] \<open>0<c\<close> by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2632 |
moreover have "inverse c >0" using assms(2) by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2633 |
ultimately have "LIM x F. inverse (f x) * (f x * g x) :> at_top" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2634 |
using filterlim_tendsto_pos_mult_at_top[OF _ _ asm,of "\<lambda>x. inverse (f x)" "inverse c"] by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2635 |
then show "LIM x F. g x :> at_top" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2636 |
apply (elim filterlim_mono_eventually) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2637 |
apply simp_all[2] |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2638 |
using eventually_times_inverse_1[OF assms(1)] \<open>c>0\<close> eventually_mono by fastforce |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2639 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2640 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2641 |
lemma filterlim_tendsto_pos_mult_at_bot_iff: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2642 |
fixes c :: real |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2643 |
assumes "(f \<longlongrightarrow> c) F" "0 < c" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2644 |
shows "(LIM x F. f x * g x :> at_bot) \<longleftrightarrow> filterlim g at_bot F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2645 |
using filterlim_tendsto_pos_mult_at_top_iff[OF assms(1,2), of "\<lambda>x. - g x"] |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2646 |
unfolding filterlim_uminus_at_bot by simp |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2647 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2648 |
lemma filterlim_tendsto_neg_mult_at_top_iff: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2649 |
fixes f::"'a \<Rightarrow> real" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2650 |
assumes "(f \<longlongrightarrow> c) F" and "c < 0" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2651 |
shows "(LIM x F. (f x * g x) :> at_top) \<longleftrightarrow> (LIM x F. g x :> at_bot)" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2652 |
proof - |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2653 |
have "(LIM x F. f x * g x :> at_top) = (LIM x F. - g x :> at_top)" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2654 |
apply (rule filterlim_tendsto_pos_mult_at_top_iff[of "\<lambda>x. - f x" "-c" F "\<lambda>x. - g x", simplified]) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2655 |
using assms by (auto intro: tendsto_intros ) |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2656 |
also have "... = (LIM x F. g x :> at_bot)" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2657 |
using filterlim_uminus_at_bot[symmetric] by auto |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2658 |
finally show ?thesis . |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2659 |
qed |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2660 |
|
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2661 |
lemma filterlim_tendsto_neg_mult_at_bot_iff: |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2662 |
fixes c :: real |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2663 |
assumes "(f \<longlongrightarrow> c) F" "0 > c" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2664 |
shows "(LIM x F. f x * g x :> at_bot) \<longleftrightarrow> filterlim g at_top F" |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2665 |
using filterlim_tendsto_neg_mult_at_top_iff[OF assms(1,2), of "\<lambda>x. - g x"] |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2666 |
unfolding filterlim_uminus_at_top by simp |
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
paulson <lp15@cam.ac.uk>
parents:
76724
diff
changeset
|
2667 |
|
60758 | 2668 |
subsection \<open>Power Sequences\<close> |
51526 | 2669 |
|
63546 | 2670 |
lemma Bseq_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> Bseq (\<lambda>n. x ^ n)" |
2671 |
for x :: real |
|
68615 | 2672 |
by (metis decseq_bounded decseq_def power_decreasing zero_le_power) |
63546 | 2673 |
|
2674 |
lemma monoseq_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> monoseq (\<lambda>n. x ^ n)" |
|
2675 |
for x :: real |
|
68615 | 2676 |
using monoseq_def power_decreasing by blast |
63546 | 2677 |
|
2678 |
lemma convergent_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> convergent (\<lambda>n. x ^ n)" |
|
2679 |
for x :: real |
|
2680 |
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) |
|
2681 |
||
2682 |
lemma LIMSEQ_inverse_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) \<longlonglongrightarrow> 0" |
|
2683 |
for x :: real |
|
51526 | 2684 |
by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp |
2685 |
||
2686 |
lemma LIMSEQ_realpow_zero: |
|
63546 | 2687 |
fixes x :: real |
2688 |
assumes "0 \<le> x" "x < 1" |
|
2689 |
shows "(\<lambda>n. x ^ n) \<longlonglongrightarrow> 0" |
|
2690 |
proof (cases "x = 0") |
|
2691 |
case False |
|
2692 |
with \<open>0 \<le> x\<close> have x0: "0 < x" by simp |
|
2693 |
then have "1 < inverse x" |
|
2694 |
using \<open>x < 1\<close> by (rule one_less_inverse) |
|
2695 |
then have "(\<lambda>n. inverse (inverse x ^ n)) \<longlonglongrightarrow> 0" |
|
51526 | 2696 |
by (rule LIMSEQ_inverse_realpow_zero) |
63546 | 2697 |
then show ?thesis by (simp add: power_inverse) |
2698 |
next |
|
2699 |
case True |
|
2700 |
show ?thesis |
|
2701 |
by (rule LIMSEQ_imp_Suc) (simp add: True) |
|
2702 |
qed |
|
2703 |
||
70723
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
paulson <lp15@cam.ac.uk>
parents:
70694
diff
changeset
|
2704 |
lemma LIMSEQ_power_zero [tendsto_intros]: "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0" |
63546 | 2705 |
for x :: "'a::real_normed_algebra_1" |
2706 |
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) |
|
68615 | 2707 |
by (simp add: Zfun_le norm_power_ineq tendsto_Zfun_iff) |
51526 | 2708 |
|
61969 | 2709 |
lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) \<longlonglongrightarrow> 0" |
51526 | 2710 |
by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp |
2711 |
||
63556 | 2712 |
lemma |
2713 |
tendsto_power_zero: |
|
2714 |
fixes x::"'a::real_normed_algebra_1" |
|
2715 |
assumes "filterlim f at_top F" |
|
2716 |
assumes "norm x < 1" |
|
2717 |
shows "((\<lambda>y. x ^ (f y)) \<longlongrightarrow> 0) F" |
|
2718 |
proof (rule tendstoI) |
|
2719 |
fix e::real assume "0 < e" |
|
2720 |
from tendstoD[OF LIMSEQ_power_zero[OF \<open>norm x < 1\<close>] \<open>0 < e\<close>] |
|
2721 |
have "\<forall>\<^sub>F xa in sequentially. norm (x ^ xa) < e" |
|
2722 |
by simp |
|
2723 |
then obtain N where N: "norm (x ^ n) < e" if "n \<ge> N" for n |
|
2724 |
by (auto simp: eventually_sequentially) |
|
2725 |
have "\<forall>\<^sub>F i in F. f i \<ge> N" |
|
2726 |
using \<open>filterlim f sequentially F\<close> |
|
2727 |
by (simp add: filterlim_at_top) |
|
2728 |
then show "\<forall>\<^sub>F i in F. dist (x ^ f i) 0 < e" |
|
68615 | 2729 |
by eventually_elim (auto simp: N) |
63556 | 2730 |
qed |
2731 |
||
69593 | 2732 |
text \<open>Limit of \<^term>\<open>c^n\<close> for \<^term>\<open>\<bar>c\<bar> < 1\<close>.\<close> |
51526 | 2733 |
|
68614 | 2734 |
lemma LIMSEQ_abs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) \<longlonglongrightarrow> 0" |
51526 | 2735 |
by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) |
2736 |
||
68614 | 2737 |
lemma LIMSEQ_abs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) \<longlonglongrightarrow> 0" |
51526 | 2738 |
by (rule LIMSEQ_power_zero) simp |
2739 |
||
2740 |
||
60758 | 2741 |
subsection \<open>Limits of Functions\<close> |
51526 | 2742 |
|
63546 | 2743 |
lemma LIM_eq: "f \<midarrow>a\<rightarrow> L = (\<forall>r>0. \<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r)" |
2744 |
for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" |
|
2745 |
by (simp add: LIM_def dist_norm) |
|
51526 | 2746 |
|
2747 |
lemma LIM_I: |
|
63546 | 2748 |
"(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r) \<Longrightarrow> f \<midarrow>a\<rightarrow> L" |
2749 |
for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" |
|
2750 |
by (simp add: LIM_eq) |
|
2751 |
||
2752 |
lemma LIM_D: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>s>0.\<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r" |
|
2753 |
for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" |
|
2754 |
by (simp add: LIM_eq) |
|
2755 |
||
2756 |
lemma LIM_offset: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. f (x + k)) \<midarrow>(a - k)\<rightarrow> L" |
|
2757 |
for a :: "'a::real_normed_vector" |
|
2758 |
by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap) |
|
2759 |
||
2760 |
lemma LIM_offset_zero: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L" |
|
2761 |
for a :: "'a::real_normed_vector" |
|
2762 |
by (drule LIM_offset [where k = a]) (simp add: add.commute) |
|
2763 |
||
2764 |
lemma LIM_offset_zero_cancel: "(\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L \<Longrightarrow> f \<midarrow>a\<rightarrow> L" |
|
2765 |
for a :: "'a::real_normed_vector" |
|
2766 |
by (drule LIM_offset [where k = "- a"]) simp |
|
2767 |
||
72245 | 2768 |
lemma LIM_offset_zero_iff: "NO_MATCH 0 a \<Longrightarrow> f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L" |
63546 | 2769 |
for f :: "'a :: real_normed_vector \<Rightarrow> _" |
51642
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
hoelzl
parents:
51641
diff
changeset
|
2770 |
using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto |
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
hoelzl
parents:
51641
diff
changeset
|
2771 |
|
70999
5b753486c075
Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents:
70817
diff
changeset
|
2772 |
lemma tendsto_offset_zero_iff: |
5b753486c075
Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents:
70817
diff
changeset
|
2773 |
fixes f :: "'a :: real_normed_vector \<Rightarrow> _" |
72245 | 2774 |
assumes " NO_MATCH 0 a" "a \<in> S" "open S" |
70999
5b753486c075
Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents:
70817
diff
changeset
|
2775 |
shows "(f \<longlongrightarrow> L) (at a within S) \<longleftrightarrow> ((\<lambda>h. f (a + h)) \<longlongrightarrow> L) (at 0)" |
72245 | 2776 |
using assms by (simp add: tendsto_within_open_NO_MATCH LIM_offset_zero_iff) |
70999
5b753486c075
Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents:
70817
diff
changeset
|
2777 |
|
63546 | 2778 |
lemma LIM_zero: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. f x - l) \<longlongrightarrow> 0) F" |
65578
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
2779 |
for f :: "'a \<Rightarrow> 'b::real_normed_vector" |
63546 | 2780 |
unfolding tendsto_iff dist_norm by simp |
51526 | 2781 |
|
2782 |
lemma LIM_zero_cancel: |
|
65578
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
2783 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
61973 | 2784 |
shows "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> l) F" |
51526 | 2785 |
unfolding tendsto_iff dist_norm by simp |
2786 |
||
63546 | 2787 |
lemma LIM_zero_iff: "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F = (f \<longlongrightarrow> l) F" |
65578
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
2788 |
for f :: "'a \<Rightarrow> 'b::real_normed_vector" |
63546 | 2789 |
unfolding tendsto_iff dist_norm by simp |
51526 | 2790 |
|
2791 |
lemma LIM_imp_LIM: |
|
2792 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
|
2793 |
fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector" |
|
61976 | 2794 |
assumes f: "f \<midarrow>a\<rightarrow> l" |
63546 | 2795 |
and le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)" |
61976 | 2796 |
shows "g \<midarrow>a\<rightarrow> m" |
63546 | 2797 |
by (rule metric_LIM_imp_LIM [OF f]) (simp add: dist_norm le) |
51526 | 2798 |
|
2799 |
lemma LIM_equal2: |
|
2800 |
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" |
|
63546 | 2801 |
assumes "0 < R" |
2802 |
and "\<And>x. x \<noteq> a \<Longrightarrow> norm (x - a) < R \<Longrightarrow> f x = g x" |
|
61976 | 2803 |
shows "g \<midarrow>a\<rightarrow> l \<Longrightarrow> f \<midarrow>a\<rightarrow> l" |
68594 | 2804 |
by (rule metric_LIM_equal2 [OF _ assms]) (simp_all add: dist_norm) |
51526 | 2805 |
|
2806 |
lemma LIM_compose2: |
|
2807 |
fixes a :: "'a::real_normed_vector" |
|
61976 | 2808 |
assumes f: "f \<midarrow>a\<rightarrow> b" |
63546 | 2809 |
and g: "g \<midarrow>b\<rightarrow> c" |
2810 |
and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b" |
|
61976 | 2811 |
shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> c" |
63546 | 2812 |
by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) |
51526 | 2813 |
|
2814 |
lemma real_LIM_sandwich_zero: |
|
2815 |
fixes f g :: "'a::topological_space \<Rightarrow> real" |
|
61976 | 2816 |
assumes f: "f \<midarrow>a\<rightarrow> 0" |
63546 | 2817 |
and 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x" |
2818 |
and 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x" |
|
61976 | 2819 |
shows "g \<midarrow>a\<rightarrow> 0" |
51526 | 2820 |
proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) |
63546 | 2821 |
fix x |
2822 |
assume x: "x \<noteq> a" |
|
2823 |
with 1 have "norm (g x - 0) = g x" by simp |
|
51526 | 2824 |
also have "g x \<le> f x" by (rule 2 [OF x]) |
2825 |
also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self) |
|
2826 |
also have "\<bar>f x\<bar> = norm (f x - 0)" by simp |
|
2827 |
finally show "norm (g x - 0) \<le> norm (f x - 0)" . |
|
2828 |
qed |
|
2829 |
||
2830 |
||
60758 | 2831 |
subsection \<open>Continuity\<close> |
51526 | 2832 |
|
63546 | 2833 |
lemma LIM_isCont_iff: "(f \<midarrow>a\<rightarrow> f a) = ((\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> f a)" |
2834 |
for f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" |
|
2835 |
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) |
|
2836 |
||
2837 |
lemma isCont_iff: "isCont f x = (\<lambda>h. f (x + h)) \<midarrow>0\<rightarrow> f x" |
|
2838 |
for f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" |
|
2839 |
by (simp add: isCont_def LIM_isCont_iff) |
|
51526 | 2840 |
|
2841 |
lemma isCont_LIM_compose2: |
|
2842 |
fixes a :: "'a::real_normed_vector" |
|
2843 |
assumes f [unfolded isCont_def]: "isCont f a" |
|
63546 | 2844 |
and g: "g \<midarrow>f a\<rightarrow> l" |
2845 |
and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a" |
|
61976 | 2846 |
shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> l" |
63546 | 2847 |
by (rule LIM_compose2 [OF f g inj]) |
2848 |
||
2849 |
lemma isCont_norm [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a" |
|
2850 |
for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
51526 | 2851 |
by (fact continuous_norm) |
2852 |
||
63546 | 2853 |
lemma isCont_rabs [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a" |
2854 |
for f :: "'a::t2_space \<Rightarrow> real" |
|
51526 | 2855 |
by (fact continuous_rabs) |
2856 |
||
63546 | 2857 |
lemma isCont_add [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x + g x) a" |
2858 |
for f :: "'a::t2_space \<Rightarrow> 'b::topological_monoid_add" |
|
51526 | 2859 |
by (fact continuous_add) |
2860 |
||
63546 | 2861 |
lemma isCont_minus [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a" |
2862 |
for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
51526 | 2863 |
by (fact continuous_minus) |
2864 |
||
63546 | 2865 |
lemma isCont_diff [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x - g x) a" |
2866 |
for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
|
51526 | 2867 |
by (fact continuous_diff) |
2868 |
||
63546 | 2869 |
lemma isCont_mult [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x * g x) a" |
2870 |
for f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra" |
|
51526 | 2871 |
by (fact continuous_mult) |
2872 |
||
63546 | 2873 |
lemma (in bounded_linear) isCont: "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a" |
51526 | 2874 |
by (fact continuous) |
2875 |
||
63546 | 2876 |
lemma (in bounded_bilinear) isCont: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a" |
51526 | 2877 |
by (fact continuous) |
2878 |
||
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2879 |
lemmas isCont_scaleR [simp] = |
51526 | 2880 |
bounded_bilinear.isCont [OF bounded_bilinear_scaleR] |
2881 |
||
2882 |
lemmas isCont_of_real [simp] = |
|
2883 |
bounded_linear.isCont [OF bounded_linear_of_real] |
|
2884 |
||
63546 | 2885 |
lemma isCont_power [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a" |
2886 |
for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}" |
|
51526 | 2887 |
by (fact continuous_power) |
2888 |
||
64267 | 2889 |
lemma isCont_sum [simp]: "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a" |
63546 | 2890 |
for f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add" |
64267 | 2891 |
by (auto intro: continuous_sum) |
51526 | 2892 |
|
63546 | 2893 |
|
60758 | 2894 |
subsection \<open>Uniform Continuity\<close> |
51526 | 2895 |
|
63104 | 2896 |
lemma uniformly_continuous_on_def: |
2897 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
2898 |
shows "uniformly_continuous_on s f \<longleftrightarrow> |
|
2899 |
(\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)" |
|
2900 |
unfolding uniformly_continuous_on_uniformity |
|
2901 |
uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal |
|
2902 |
by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric) |
|
2903 |
||
63546 | 2904 |
abbreviation isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" |
2905 |
where "isUCont f \<equiv> uniformly_continuous_on UNIV f" |
|
2906 |
||
2907 |
lemma isUCont_def: "isUCont f \<longleftrightarrow> (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)" |
|
63104 | 2908 |
by (auto simp: uniformly_continuous_on_def dist_commute) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
2909 |
|
63546 | 2910 |
lemma isUCont_isCont: "isUCont f \<Longrightarrow> isCont f x" |
63104 | 2911 |
by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at) |
2912 |
||
2913 |
lemma uniformly_continuous_on_Cauchy: |
|
63546 | 2914 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
63104 | 2915 |
assumes "uniformly_continuous_on S f" "Cauchy X" "\<And>n. X n \<in> S" |
2916 |
shows "Cauchy (\<lambda>n. f (X n))" |
|
2917 |
using assms |
|
68594 | 2918 |
unfolding uniformly_continuous_on_def by (meson Cauchy_def) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
2919 |
|
63546 | 2920 |
lemma isUCont_Cauchy: "isUCont f \<Longrightarrow> Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" |
63104 | 2921 |
by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all |
68611 | 2922 |
|
64287 | 2923 |
lemma uniformly_continuous_imp_Cauchy_continuous: |
2924 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
67091 | 2925 |
shows "\<lbrakk>uniformly_continuous_on S f; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f \<circ> \<sigma>)" |
64287 | 2926 |
by (simp add: uniformly_continuous_on_def Cauchy_def) meson |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
2927 |
|
51526 | 2928 |
lemma (in bounded_linear) isUCont: "isUCont f" |
63546 | 2929 |
unfolding isUCont_def dist_norm |
51526 | 2930 |
proof (intro allI impI) |
63546 | 2931 |
fix r :: real |
2932 |
assume r: "0 < r" |
|
2933 |
obtain K where K: "0 < K" and norm_le: "norm (f x) \<le> norm x * K" for x |
|
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
2934 |
using pos_bounded by blast |
51526 | 2935 |
show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r" |
2936 |
proof (rule exI, safe) |
|
56541 | 2937 |
from r K show "0 < r / K" by simp |
51526 | 2938 |
next |
2939 |
fix x y :: 'a |
|
2940 |
assume xy: "norm (x - y) < r / K" |
|
2941 |
have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) |
|
2942 |
also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le) |
|
2943 |
also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq) |
|
2944 |
finally show "norm (f x - f y) < r" . |
|
2945 |
qed |
|
2946 |
qed |
|
2947 |
||
2948 |
lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" |
|
63546 | 2949 |
by (rule isUCont [THEN isUCont_Cauchy]) |
51526 | 2950 |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
2951 |
lemma LIM_less_bound: |
51526 | 2952 |
fixes f :: "real \<Rightarrow> real" |
2953 |
assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x" |
|
2954 |
shows "0 \<le> f x" |
|
63952
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents:
63915
diff
changeset
|
2955 |
proof (rule tendsto_lowerbound) |
61973 | 2956 |
show "(f \<longlongrightarrow> f x) (at_left x)" |
60758 | 2957 |
using \<open>isCont f x\<close> by (simp add: filterlim_at_split isCont_def) |
51526 | 2958 |
show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)" |
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51531
diff
changeset
|
2959 |
using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"]) |
51526 | 2960 |
qed simp |
51471 | 2961 |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2962 |
|
60758 | 2963 |
subsection \<open>Nested Intervals and Bisection -- Needed for Compactness\<close> |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2964 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2965 |
lemma nested_sequence_unique: |
61969 | 2966 |
assumes "\<forall>n. f n \<le> f (Suc n)" "\<forall>n. g (Suc n) \<le> g n" "\<forall>n. f n \<le> g n" "(\<lambda>n. f n - g n) \<longlonglongrightarrow> 0" |
2967 |
shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f \<longlonglongrightarrow> l) \<and> ((\<forall>n. l \<le> g n) \<and> g \<longlonglongrightarrow> l)" |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2968 |
proof - |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2969 |
have "incseq f" unfolding incseq_Suc_iff by fact |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2970 |
have "decseq g" unfolding decseq_Suc_iff by fact |
63546 | 2971 |
have "f n \<le> g 0" for n |
2972 |
proof - |
|
2973 |
from \<open>decseq g\<close> have "g n \<le> g 0" |
|
2974 |
by (rule decseqD) simp |
|
2975 |
with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] show ?thesis |
|
2976 |
by auto |
|
2977 |
qed |
|
61969 | 2978 |
then obtain u where "f \<longlonglongrightarrow> u" "\<forall>i. f i \<le> u" |
60758 | 2979 |
using incseq_convergent[OF \<open>incseq f\<close>] by auto |
63546 | 2980 |
moreover have "f 0 \<le> g n" for n |
2981 |
proof - |
|
60758 | 2982 |
from \<open>incseq f\<close> have "f 0 \<le> f n" by (rule incseqD) simp |
63546 | 2983 |
with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] show ?thesis |
2984 |
by simp |
|
2985 |
qed |
|
61969 | 2986 |
then obtain l where "g \<longlonglongrightarrow> l" "\<forall>i. l \<le> g i" |
60758 | 2987 |
using decseq_convergent[OF \<open>decseq g\<close>] by auto |
61969 | 2988 |
moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \<open>f \<longlonglongrightarrow> u\<close> \<open>g \<longlonglongrightarrow> l\<close>]] |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2989 |
ultimately show ?thesis by auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2990 |
qed |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2991 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2992 |
lemma Bolzano[consumes 1, case_names trans local]: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2993 |
fixes P :: "real \<Rightarrow> real \<Rightarrow> bool" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2994 |
assumes [arith]: "a \<le> b" |
63546 | 2995 |
and trans: "\<And>a b c. P a b \<Longrightarrow> P b c \<Longrightarrow> a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> P a c" |
2996 |
and local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b" |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2997 |
shows "P a b" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
2998 |
proof - |
74513
67d87d224e00
A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents:
74475
diff
changeset
|
2999 |
define bisect where "bisect \<equiv> \<lambda>(x,y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2)" |
67d87d224e00
A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents:
74475
diff
changeset
|
3000 |
define l u where "l n \<equiv> fst ((bisect^^n)(a,b))" and "u n \<equiv> snd ((bisect^^n)(a,b))" for n |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3001 |
have l[simp]: "l 0 = a" "\<And>n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3002 |
and u[simp]: "u 0 = b" "\<And>n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3003 |
by (simp_all add: l_def u_def bisect_def split: prod.split) |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3004 |
|
63546 | 3005 |
have [simp]: "l n \<le> u n" for n by (induct n) auto |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3006 |
|
61969 | 3007 |
have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l \<longlonglongrightarrow> x) \<and> ((\<forall>n. x \<le> u n) \<and> u \<longlonglongrightarrow> x)" |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3008 |
proof (safe intro!: nested_sequence_unique) |
63546 | 3009 |
show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" for n |
3010 |
by (induct n) auto |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3011 |
next |
63546 | 3012 |
have "l n - u n = (a - b) / 2^n" for n |
3013 |
by (induct n) (auto simp: field_simps) |
|
3014 |
then show "(\<lambda>n. l n - u n) \<longlonglongrightarrow> 0" |
|
3015 |
by (simp add: LIMSEQ_divide_realpow_zero) |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3016 |
qed fact |
63546 | 3017 |
then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l \<longlonglongrightarrow> x" "u \<longlonglongrightarrow> x" |
3018 |
by auto |
|
3019 |
obtain d where "0 < d" and d: "a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b" for a b |
|
60758 | 3020 |
using \<open>l 0 \<le> x\<close> \<open>x \<le> u 0\<close> local[of x] by auto |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3021 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3022 |
show "P a b" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3023 |
proof (rule ccontr) |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
3024 |
assume "\<not> P a b" |
63546 | 3025 |
have "\<not> P (l n) (u n)" for n |
3026 |
proof (induct n) |
|
3027 |
case 0 |
|
3028 |
then show ?case |
|
3029 |
by (simp add: \<open>\<not> P a b\<close>) |
|
3030 |
next |
|
3031 |
case (Suc n) |
|
3032 |
with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case |
|
3033 |
by auto |
|
3034 |
qed |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3035 |
moreover |
63546 | 3036 |
{ |
3037 |
have "eventually (\<lambda>n. x - d / 2 < l n) sequentially" |
|
61969 | 3038 |
using \<open>0 < d\<close> \<open>l \<longlonglongrightarrow> x\<close> by (intro order_tendstoD[of _ x]) auto |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3039 |
moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially" |
61969 | 3040 |
using \<open>0 < d\<close> \<open>u \<longlonglongrightarrow> x\<close> by (intro order_tendstoD[of _ x]) auto |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3041 |
ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3042 |
proof eventually_elim |
63546 | 3043 |
case (elim n) |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3044 |
from add_strict_mono[OF this] have "u n - l n < d" by simp |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3045 |
with x show "P (l n) (u n)" by (rule d) |
63546 | 3046 |
qed |
3047 |
} |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3048 |
ultimately show False by simp |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3049 |
qed |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3050 |
qed |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3051 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3052 |
lemma compact_Icc[simp, intro]: "compact {a .. b::real}" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3053 |
proof (cases "a \<le> b", rule compactI) |
63546 | 3054 |
fix C |
3055 |
assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C" |
|
63040 | 3056 |
define T where "T = {a .. b}" |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3057 |
from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3058 |
proof (induct rule: Bolzano) |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3059 |
case (trans a b c) |
63546 | 3060 |
then have *: "{a..c} = {a..b} \<union> {b..c}" |
3061 |
by auto |
|
3062 |
with trans obtain C1 C2 |
|
3063 |
where "C1\<subseteq>C" "finite C1" "{a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C" "finite C2" "{b..c} \<subseteq> \<Union>C2" |
|
3064 |
by auto |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3065 |
with trans show ?case |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3066 |
unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3067 |
next |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3068 |
case (local x) |
63546 | 3069 |
with C have "x \<in> \<Union>C" by auto |
3070 |
with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" |
|
3071 |
by auto |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3072 |
then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c" |
62101 | 3073 |
by (auto simp: open_dist dist_real_def subset_eq Ball_def abs_less_iff) |
60758 | 3074 |
with \<open>c \<in> C\<close> show ?case |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3075 |
by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3076 |
qed |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3077 |
qed simp |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3078 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3079 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3080 |
lemma continuous_image_closed_interval: |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3081 |
fixes a b and f :: "real \<Rightarrow> real" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3082 |
defines "S \<equiv> {a..b}" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3083 |
assumes "a \<le> b" and f: "continuous_on S f" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3084 |
shows "\<exists>c d. f`S = {c..d} \<and> c \<le> d" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3085 |
proof - |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3086 |
have S: "compact S" "S \<noteq> {}" |
60758 | 3087 |
using \<open>a \<le> b\<close> by (auto simp: S_def) |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3088 |
obtain c where "c \<in> S" "\<forall>d\<in>S. f d \<le> f c" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3089 |
using continuous_attains_sup[OF S f] by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3090 |
moreover obtain d where "d \<in> S" "\<forall>c\<in>S. f d \<le> f c" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3091 |
using continuous_attains_inf[OF S f] by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3092 |
moreover have "connected (f`S)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3093 |
using connected_continuous_image[OF f] connected_Icc by (auto simp: S_def) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3094 |
ultimately have "f ` S = {f d .. f c} \<and> f d \<le> f c" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3095 |
by (auto simp: connected_iff_interval) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3096 |
then show ?thesis |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3097 |
by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3098 |
qed |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57276
diff
changeset
|
3099 |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60758
diff
changeset
|
3100 |
lemma open_Collect_positive: |
67958
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
3101 |
fixes f :: "'a::topological_space \<Rightarrow> real" |
63546 | 3102 |
assumes f: "continuous_on s f" |
3103 |
shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < f x}" |
|
3104 |
using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"] |
|
3105 |
by (auto simp: Int_def field_simps) |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60758
diff
changeset
|
3106 |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60758
diff
changeset
|
3107 |
lemma open_Collect_less_Int: |
67958
732c0b059463
tuned proofs and generalized some lemmas about limits
huffman
parents:
67950
diff
changeset
|
3108 |
fixes f g :: "'a::topological_space \<Rightarrow> real" |
63546 | 3109 |
assumes f: "continuous_on s f" |
3110 |
and g: "continuous_on s g" |
|
3111 |
shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. f x < g x}" |
|
3112 |
using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60758
diff
changeset
|
3113 |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60758
diff
changeset
|
3114 |
|
60758 | 3115 |
subsection \<open>Boundedness of continuous functions\<close> |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3116 |
|
60758 | 3117 |
text\<open>By bisection, function continuous on closed interval is bounded above\<close> |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3118 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3119 |
lemma isCont_eq_Ub: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3120 |
fixes f :: "real \<Rightarrow> 'a::linorder_topology" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3121 |
shows "a \<le> b \<Longrightarrow> \<forall>x::real. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3122 |
\<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)" |
63546 | 3123 |
using continuous_attains_sup[of "{a..b}" f] |
68615 | 3124 |
by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3125 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3126 |
lemma isCont_eq_Lb: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3127 |
fixes f :: "real \<Rightarrow> 'a::linorder_topology" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3128 |
shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3129 |
\<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)" |
63546 | 3130 |
using continuous_attains_inf[of "{a..b}" f] |
68615 | 3131 |
by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3132 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3133 |
lemma isCont_bounded: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3134 |
fixes f :: "real \<Rightarrow> 'a::linorder_topology" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3135 |
shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> \<exists>M. \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3136 |
using isCont_eq_Ub[of a b f] by auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3137 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3138 |
lemma isCont_has_Ub: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3139 |
fixes f :: "real \<Rightarrow> 'a::linorder_topology" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3140 |
shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3141 |
\<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<forall>N. N < M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x))" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3142 |
using isCont_eq_Ub[of a b f] by auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3143 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3144 |
lemma isCont_Lb_Ub: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3145 |
fixes f :: "real \<Rightarrow> real" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3146 |
assumes "a \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x" |
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
3147 |
shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and> |
63546 | 3148 |
(\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))" |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3149 |
proof - |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3150 |
obtain M where M: "a \<le> M" "M \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> f M" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3151 |
using isCont_eq_Ub[OF assms] by auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3152 |
obtain L where L: "a \<le> L" "L \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f L \<le> f x" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3153 |
using isCont_eq_Lb[OF assms] by auto |
68615 | 3154 |
have "(\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f L \<le> f x \<and> f x \<le> f M)" |
3155 |
using M L by simp |
|
3156 |
moreover |
|
3157 |
have "(\<forall>y. f L \<le> y \<and> y \<le> f M \<longrightarrow> (\<exists>x\<ge>a. x \<le> b \<and> f x = y))" |
|
3158 |
proof (cases "L \<le> M") |
|
3159 |
case True then show ?thesis |
|
3160 |
using IVT[of f L _ M] M L assms by (metis order.trans) |
|
3161 |
next |
|
3162 |
case False then show ?thesis |
|
3163 |
using IVT2[of f L _ M] |
|
3164 |
by (metis L(2) M(1) assms(2) le_cases order.trans) |
|
3165 |
qed |
|
3166 |
ultimately show ?thesis |
|
3167 |
by blast |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3168 |
qed |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3169 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3170 |
|
63546 | 3171 |
text \<open>Continuity of inverse function.\<close> |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3172 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3173 |
lemma isCont_inverse_function: |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3174 |
fixes f g :: "real \<Rightarrow> real" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3175 |
assumes d: "0 < d" |
68611 | 3176 |
and inj: "\<And>z. \<bar>z-x\<bar> \<le> d \<Longrightarrow> g (f z) = z" |
3177 |
and cont: "\<And>z. \<bar>z-x\<bar> \<le> d \<Longrightarrow> isCont f z" |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3178 |
shows "isCont g (f x)" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3179 |
proof - |
63546 | 3180 |
let ?A = "f (x - d)" |
3181 |
let ?B = "f (x + d)" |
|
3182 |
let ?D = "{x - d..x + d}" |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3183 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3184 |
have f: "continuous_on ?D f" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3185 |
using cont by (intro continuous_at_imp_continuous_on ballI) auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3186 |
then have g: "continuous_on (f`?D) g" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3187 |
using inj by (intro continuous_on_inv) auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3188 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3189 |
from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3190 |
by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3191 |
with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3192 |
by (rule continuous_on_subset) |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3193 |
moreover |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3194 |
have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3195 |
using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3196 |
then have "f x \<in> {min ?A ?B <..< max ?A ?B}" |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3197 |
by auto |
68615 | 3198 |
ultimately show ?thesis |
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3199 |
by (simp add: continuous_on_eq_continuous_at) |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3200 |
qed |
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3201 |
|
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3202 |
lemma isCont_inverse_function2: |
63546 | 3203 |
fixes f g :: "real \<Rightarrow> real" |
3204 |
shows |
|
68611 | 3205 |
"\<lbrakk>a < x; x < b; |
3206 |
\<And>z. \<lbrakk>a \<le> z; z \<le> b\<rbrakk> \<Longrightarrow> g (f z) = z; |
|
3207 |
\<And>z. \<lbrakk>a \<le> z; z \<le> b\<rbrakk> \<Longrightarrow> isCont f z\<rbrakk> \<Longrightarrow> isCont g (f x)" |
|
63546 | 3208 |
apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"]) |
3209 |
apply (simp_all add: abs_le_iff) |
|
3210 |
done |
|
51529
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents:
51526
diff
changeset
|
3211 |
|
63546 | 3212 |
text \<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110.\<close> |
3213 |
lemma LIM_fun_gt_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> 0 < f x)" |
|
3214 |
for f :: "real \<Rightarrow> real" |
|
68615 | 3215 |
by (force simp: dest: LIM_D) |
63546 | 3216 |
|
3217 |
lemma LIM_fun_less_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x < 0)" |
|
3218 |
for f :: "real \<Rightarrow> real" |
|
68615 | 3219 |
by (drule LIM_D [where r="-l"]) force+ |
63546 | 3220 |
|
3221 |
lemma LIM_fun_not_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x \<noteq> 0)" |
|
3222 |
for f :: "real \<Rightarrow> real" |
|
68615 | 3223 |
using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
3224 |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
3225 |
end |