Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
1 (* Title: HOL/Limits.thy
3 Author: Jacques D. Fleuriot, University of Cambridge
4 Author: Lawrence C Paulson
8 section \<open>Limits on Real Vector Spaces\<close>
11 imports Real_Vector_Spaces
14 subsection \<open>Filter going to infinity norm\<close>
16 definition at_infinity :: "'a::real_normed_vector filter" where
17 "at_infinity = (INF r. principal {x. r \<le> norm x})"
19 lemma eventually_at_infinity: "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)"
20 unfolding at_infinity_def
21 by (subst eventually_INF_base)
22 (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b])
24 lemma at_infinity_eq_at_top_bot:
25 "(at_infinity :: real filter) = sup at_top at_bot"
26 apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity
27 eventually_at_top_linorder eventually_at_bot_linorder)
29 apply (rule_tac x="b" in exI, simp)
30 apply (rule_tac x="- b" in exI, simp)
31 apply (rule_tac x="max (- Na) N" in exI, auto simp: abs_real_def)
34 lemma at_top_le_at_infinity: "at_top \<le> (at_infinity :: real filter)"
35 unfolding at_infinity_eq_at_top_bot by simp
37 lemma at_bot_le_at_infinity: "at_bot \<le> (at_infinity :: real filter)"
38 unfolding at_infinity_eq_at_top_bot by simp
40 lemma filterlim_at_top_imp_at_infinity:
41 fixes f :: "_ \<Rightarrow> real"
42 shows "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F"
43 by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl])
45 lemma lim_infinity_imp_sequentially:
46 "(f ---> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) ---> l) sequentially"
47 by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially)
50 subsubsection \<open>Boundedness\<close>
52 definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
53 Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
55 abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
56 "Bseq X \<equiv> Bfun X sequentially"
58 lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" ..
60 lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
61 unfolding Bfun_metric_def by (subst eventually_sequentially_seg)
63 lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
64 unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
67 "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
68 unfolding Bfun_metric_def norm_conv_dist
70 fix y K assume "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F"
71 moreover have "eventually (\<lambda>x. dist (f x) 0 \<le> dist (f x) y + dist 0 y) F"
72 by (intro always_eventually) (metis dist_commute dist_triangle)
73 with * have "eventually (\<lambda>x. dist (f x) 0 \<le> K + dist 0 y) F"
74 by eventually_elim auto
75 with \<open>0 < K\<close> show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F"
76 by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto
80 assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F"
82 proof (intro exI conjI allI)
83 show "0 < max K 1" by simp
85 show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F"
86 using K by (rule eventually_elim1, simp)
91 obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F"
92 using assms unfolding Bfun_def by blast
94 lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
95 unfolding Cauchy_def Bfun_metric_def eventually_sequentially
96 apply (erule_tac x=1 in allE)
99 apply (rule_tac x="X M" in exI)
100 apply (rule_tac x=1 in exI)
101 apply (erule_tac x=M in allE)
103 apply (rule_tac x=M in exI)
104 apply (auto simp: dist_commute)
108 subsubsection \<open>Bounded Sequences\<close>
110 lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X"
111 by (intro BfunI) (auto simp: eventually_sequentially)
113 lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X"
114 by (intro BfunI) (auto simp: eventually_sequentially)
116 lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)"
117 unfolding Bfun_def eventually_sequentially
119 fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K"
120 then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K"
121 by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2)
122 (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj)
125 lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
126 unfolding Bseq_def by auto
128 lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
129 by (simp add: Bseq_def)
131 lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
132 by (auto simp add: Bseq_def)
134 lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)"
135 proof (elim BseqE, intro bdd_aboveI2)
136 fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K"
137 by (auto elim!: allE[of _ n])
140 lemma Bseq_bdd_above':
141 "Bseq (X::nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))"
142 proof (elim BseqE, intro bdd_aboveI2)
143 fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "norm (X n) \<le> K"
144 by (auto elim!: allE[of _ n])
147 lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)"
148 proof (elim BseqE, intro bdd_belowI2)
149 fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n"
150 by (auto elim!: allE[of _ n])
153 lemma Bseq_eventually_mono:
154 assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) sequentially" "Bseq g"
157 from assms(1) obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm (f n) \<le> norm (g n)"
158 by (auto simp: eventually_at_top_linorder)
159 moreover from assms(2) obtain K where K: "\<And>n. norm (g n) \<le> K" by (blast elim!: BseqE)
160 ultimately have "norm (f n) \<le> max K (Max {norm (f n) |n. n < N})" for n
161 apply (cases "n < N")
162 apply (rule max.coboundedI2, rule Max.coboundedI, auto) []
163 apply (rule max.coboundedI1, force intro: order.trans[OF N K])
165 thus ?thesis by (blast intro: BseqI')
168 lemma lemma_NBseq_def:
169 "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
172 from reals_Archimedean2 obtain n :: nat where "K < real n" ..
173 then have "K \<le> real (Suc n)" by auto
174 moreover assume "\<forall>m. norm (X m) \<le> K"
175 ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)"
176 by (blast intro: order_trans)
177 then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
179 show "\<And>N. \<forall>n. norm (X n) \<le> real (Suc N) \<Longrightarrow> \<exists>K>0. \<forall>n. norm (X n) \<le> K"
180 using of_nat_0_less_iff by blast
183 text\<open>alternative definition for Bseq\<close>
184 lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
185 apply (simp add: Bseq_def)
186 apply (simp (no_asm) add: lemma_NBseq_def)
189 lemma lemma_NBseq_def2:
190 "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
191 apply (subst lemma_NBseq_def, auto)
192 apply (rule_tac x = "Suc N" in exI)
193 apply (rule_tac [2] x = N in exI)
194 apply (auto simp add: of_nat_Suc)
195 prefer 2 apply (blast intro: order_less_imp_le)
196 apply (drule_tac x = n in spec, simp)
199 (* yet another definition for Bseq *)
200 lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
201 by (simp add: Bseq_def lemma_NBseq_def2)
203 subsubsection\<open>A Few More Equivalence Theorems for Boundedness\<close>
205 text\<open>alternative formulation for boundedness\<close>
206 lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
207 apply (unfold Bseq_def, safe)
208 apply (rule_tac [2] x = "k + norm x" in exI)
209 apply (rule_tac x = K in exI, simp)
210 apply (rule exI [where x = 0], auto)
211 apply (erule order_less_le_trans, simp)
212 apply (drule_tac x=n in spec)
213 apply (drule order_trans [OF norm_triangle_ineq2])
217 text\<open>alternative formulation for boundedness\<close>
219 "Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" (is "?P \<longleftrightarrow> ?Q")
223 where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def)
224 from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp
225 from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)"
226 by (auto intro: order_trans norm_triangle_ineq4)
227 then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
229 with \<open>0 < K + norm (X 0)\<close> show ?Q by blast
231 assume ?Q then show ?P by (auto simp add: Bseq_iff2)
234 lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
235 apply (simp add: Bseq_def)
236 apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
237 apply (drule_tac x = n in spec, arith)
241 subsubsection\<open>Upper Bounds and Lubs of Bounded Sequences\<close>
243 lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
244 by (simp add: Bseq_def)
247 assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
248 shows "Bseq (\<lambda>x. f x + c)"
250 from assms obtain K where K: "\<And>x. norm (f x) \<le> K" unfolding Bseq_def by blast
253 have "norm (f x + c) \<le> norm (f x) + norm c" by (rule norm_triangle_ineq)
254 also have "norm (f x) \<le> K" by (rule K)
255 finally have "norm (f x + c) \<le> K + norm c" by simp
257 thus ?thesis by (rule BseqI')
260 lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
261 using Bseq_add[of f c] Bseq_add[of "\<lambda>x. f x + c" "-c"] by auto
264 assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_field)"
265 assumes "Bseq (g :: nat \<Rightarrow> 'a :: real_normed_field)"
266 shows "Bseq (\<lambda>x. f x * g x)"
268 from assms obtain K1 K2 where K: "\<And>x. norm (f x) \<le> K1" "K1 > 0" "\<And>x. norm (g x) \<le> K2" "K2 > 0"
269 unfolding Bseq_def by blast
270 hence "\<And>x. norm (f x * g x) \<le> K1 * K2" by (auto simp: norm_mult intro!: mult_mono)
271 thus ?thesis by (rule BseqI')
274 lemma Bfun_const [simp]: "Bfun (\<lambda>_. c) F"
275 unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"])
277 lemma Bseq_cmult_iff: "(c :: 'a :: real_normed_field) \<noteq> 0 \<Longrightarrow> Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f"
279 assume "c \<noteq> 0" "Bseq (\<lambda>x. c * f x)"
280 find_theorems "Bfun (\<lambda>_. ?c) _"
281 from Bfun_const this(2) have "Bseq (\<lambda>x. inverse c * (c * f x))" by (rule Bseq_mult)
282 with `c \<noteq> 0` show "Bseq f" by (simp add: divide_simps)
283 qed (intro Bseq_mult Bfun_const)
285 lemma Bseq_subseq: "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> Bseq (\<lambda>x. f (g x))"
286 unfolding Bseq_def by auto
288 lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
289 using Bseq_offset[of f 1] by (auto intro: Bseq_subseq)
291 lemma increasing_Bseq_subseq_iff:
292 assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a :: real_normed_vector) \<le> norm (f y)" "subseq g"
293 shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
295 assume "Bseq (\<lambda>x. f (g x))"
296 then obtain K where K: "\<And>x. norm (f (g x)) \<le> K" unfolding Bseq_def by auto
299 from filterlim_subseq[OF assms(2)] obtain y where "g y \<ge> x"
300 by (auto simp: filterlim_at_top eventually_at_top_linorder)
301 hence "norm (f x) \<le> norm (f (g y))" using assms(1) by blast
302 also have "norm (f (g y)) \<le> K" by (rule K)
303 finally have "norm (f x) \<le> K" .
305 thus "Bseq f" by (rule BseqI')
306 qed (insert Bseq_subseq[of f g], simp_all)
308 lemma nonneg_incseq_Bseq_subseq_iff:
309 assumes "\<And>x. f x \<ge> 0" "incseq (f :: nat \<Rightarrow> real)" "subseq g"
310 shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
311 using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def)
313 lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f"
314 apply (simp add: subset_eq)
315 apply (rule BseqI'[where K="max (norm a) (norm b)"])
316 apply (erule_tac x=n in allE)
320 lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X"
321 by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def)
323 lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
324 by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def)
326 subsection \<open>Bounded Monotonic Sequences\<close>
328 subsubsection\<open>A Bounded and Monotonic Sequence Converges\<close>
331 (* FIXME: one use in NSA/HSEQ.thy *)
332 lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
333 apply (rule_tac x="X m" in exI)
334 apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const])
335 unfolding eventually_sequentially
339 subsection \<open>Convergence to Zero\<close>
341 definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
342 where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)"
345 "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F"
346 unfolding Zfun_def by simp
349 "\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F"
350 unfolding Zfun_def by simp
353 "eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F"
354 unfolding Zfun_def by (auto elim!: eventually_rev_mp)
356 lemma Zfun_zero: "Zfun (\<lambda>x. 0) F"
357 unfolding Zfun_def by simp
359 lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) F = Zfun (\<lambda>x. f x) F"
360 unfolding Zfun_def by simp
363 assumes f: "Zfun f F"
364 assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F"
365 shows "Zfun (\<lambda>x. g x) F"
370 fix r::real assume "0 < r"
371 hence "0 < r / K" using K by simp
372 then have "eventually (\<lambda>x. norm (f x) < r / K) F"
373 using ZfunD [OF f] by blast
374 with g show "eventually (\<lambda>x. norm (g x) < r) F"
375 proof eventually_elim
377 hence "norm (f x) * K < r"
378 by (simp add: pos_less_divide_eq K)
380 by (simp add: order_le_less_trans [OF elim(1)])
384 assume "\<not> 0 < K"
385 hence K: "K \<le> 0" by (simp only: not_less)
390 from g show "eventually (\<lambda>x. norm (g x) < r) F"
391 proof eventually_elim
393 also have "norm (f x) * K \<le> norm (f x) * 0"
394 using K norm_ge_zero by (rule mult_left_mono)
396 using \<open>0 < r\<close> by simp
401 lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F"
402 by (erule_tac K="1" in Zfun_imp_Zfun, simp)
405 assumes f: "Zfun f F" and g: "Zfun g F"
406 shows "Zfun (\<lambda>x. f x + g x) F"
408 fix r::real assume "0 < r"
409 hence r: "0 < r / 2" by simp
410 have "eventually (\<lambda>x. norm (f x) < r/2) F"
411 using f r by (rule ZfunD)
413 have "eventually (\<lambda>x. norm (g x) < r/2) F"
414 using g r by (rule ZfunD)
416 show "eventually (\<lambda>x. norm (f x + g x) < r) F"
417 proof eventually_elim
419 have "norm (f x + g x) \<le> norm (f x) + norm (g x)"
420 by (rule norm_triangle_ineq)
421 also have "\<dots> < r/2 + r/2"
422 using elim by (rule add_strict_mono)
428 lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F"
429 unfolding Zfun_def by simp
431 lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
432 using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus)
434 lemma (in bounded_linear) Zfun:
435 assumes g: "Zfun g F"
436 shows "Zfun (\<lambda>x. f (g x)) F"
438 obtain K where "\<And>x. norm (f x) \<le> norm x * K"
439 using bounded by blast
440 then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F"
443 by (rule Zfun_imp_Zfun)
446 lemma (in bounded_bilinear) Zfun:
447 assumes f: "Zfun f F"
448 assumes g: "Zfun g F"
449 shows "Zfun (\<lambda>x. f x ** g x) F"
451 fix r::real assume r: "0 < r"
452 obtain K where K: "0 < K"
453 and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
454 using pos_bounded by blast
455 from K have K': "0 < inverse K"
456 by (rule positive_imp_inverse_positive)
457 have "eventually (\<lambda>x. norm (f x) < r) F"
458 using f r by (rule ZfunD)
460 have "eventually (\<lambda>x. norm (g x) < inverse K) F"
461 using g K' by (rule ZfunD)
463 show "eventually (\<lambda>x. norm (f x ** g x) < r) F"
464 proof eventually_elim
466 have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K"
468 also have "norm (f x) * norm (g x) * K < r * inverse K * K"
469 by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K)
470 also from K have "r * inverse K * K = r"
476 lemma (in bounded_bilinear) Zfun_left:
477 "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F"
478 by (rule bounded_linear_left [THEN bounded_linear.Zfun])
480 lemma (in bounded_bilinear) Zfun_right:
481 "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F"
482 by (rule bounded_linear_right [THEN bounded_linear.Zfun])
484 lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult]
485 lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult]
486 lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult]
488 lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\<lambda>x. f x - a) F"
489 by (simp only: tendsto_iff Zfun_def dist_norm)
491 lemma tendsto_0_le: "\<lbrakk>(f ---> 0) F; eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F\<rbrakk>
492 \<Longrightarrow> (g ---> 0) F"
493 by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff)
495 subsubsection \<open>Distance and norms\<close>
497 lemma tendsto_dist [tendsto_intros]:
498 fixes l m :: "'a :: metric_space"
499 assumes f: "(f ---> l) F" and g: "(g ---> m) F"
500 shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F"
501 proof (rule tendstoI)
502 fix e :: real assume "0 < e"
503 hence e2: "0 < e/2" by simp
504 from tendstoD [OF f e2] tendstoD [OF g e2]
505 show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
506 proof (eventually_elim)
508 then show "dist (dist (f x) (g x)) (dist l m) < e"
509 unfolding dist_real_def
510 using dist_triangle2 [of "f x" "g x" "l"]
511 using dist_triangle2 [of "g x" "l" "m"]
512 using dist_triangle3 [of "l" "m" "f x"]
513 using dist_triangle [of "f x" "m" "g x"]
518 lemma continuous_dist[continuous_intros]:
519 fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
520 shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))"
521 unfolding continuous_def by (rule tendsto_dist)
523 lemma continuous_on_dist[continuous_intros]:
524 fixes f g :: "_ \<Rightarrow> 'a :: metric_space"
525 shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))"
526 unfolding continuous_on_def by (auto intro: tendsto_dist)
528 lemma tendsto_norm [tendsto_intros]:
529 "(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F"
530 unfolding norm_conv_dist by (intro tendsto_intros)
532 lemma continuous_norm [continuous_intros]:
533 "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))"
534 unfolding continuous_def by (rule tendsto_norm)
536 lemma continuous_on_norm [continuous_intros]:
537 "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. norm (f x))"
538 unfolding continuous_on_def by (auto intro: tendsto_norm)
540 lemma tendsto_norm_zero:
541 "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) F"
542 by (drule tendsto_norm, simp)
544 lemma tendsto_norm_zero_cancel:
545 "((\<lambda>x. norm (f x)) ---> 0) F \<Longrightarrow> (f ---> 0) F"
546 unfolding tendsto_iff dist_norm by simp
548 lemma tendsto_norm_zero_iff:
549 "((\<lambda>x. norm (f x)) ---> 0) F \<longleftrightarrow> (f ---> 0) F"
550 unfolding tendsto_iff dist_norm by simp
552 lemma tendsto_rabs [tendsto_intros]:
553 "(f ---> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) F"
554 by (fold real_norm_def, rule tendsto_norm)
556 lemma continuous_rabs [continuous_intros]:
557 "continuous F f \<Longrightarrow> continuous F (\<lambda>x. \<bar>f x :: real\<bar>)"
558 unfolding real_norm_def[symmetric] by (rule continuous_norm)
560 lemma continuous_on_rabs [continuous_intros]:
561 "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. \<bar>f x :: real\<bar>)"
562 unfolding real_norm_def[symmetric] by (rule continuous_on_norm)
564 lemma tendsto_rabs_zero:
565 "(f ---> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) F"
566 by (fold real_norm_def, rule tendsto_norm_zero)
568 lemma tendsto_rabs_zero_cancel:
569 "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<Longrightarrow> (f ---> 0) F"
570 by (fold real_norm_def, rule tendsto_norm_zero_cancel)
572 lemma tendsto_rabs_zero_iff:
573 "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<longleftrightarrow> (f ---> 0) F"
574 by (fold real_norm_def, rule tendsto_norm_zero_iff)
576 subsubsection \<open>Addition and subtraction\<close>
578 lemma tendsto_add [tendsto_intros]:
579 fixes a b :: "'a::real_normed_vector"
580 shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) F"
581 by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add)
583 lemma continuous_add [continuous_intros]:
584 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
585 shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x + g x)"
586 unfolding continuous_def by (rule tendsto_add)
588 lemma continuous_on_add [continuous_intros]:
589 fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
590 shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)"
591 unfolding continuous_on_def by (auto intro: tendsto_add)
593 lemma tendsto_add_zero:
594 fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
595 shows "\<lbrakk>(f ---> 0) F; (g ---> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) F"
596 by (drule (1) tendsto_add, simp)
598 lemma tendsto_minus [tendsto_intros]:
599 fixes a :: "'a::real_normed_vector"
600 shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) F"
601 by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus)
603 lemma continuous_minus [continuous_intros]:
604 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
605 shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
606 unfolding continuous_def by (rule tendsto_minus)
608 lemma continuous_on_minus [continuous_intros]:
609 fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
610 shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
611 unfolding continuous_on_def by (auto intro: tendsto_minus)
613 lemma tendsto_minus_cancel:
614 fixes a :: "'a::real_normed_vector"
615 shows "((\<lambda>x. - f x) ---> - a) F \<Longrightarrow> (f ---> a) F"
616 by (drule tendsto_minus, simp)
618 lemma tendsto_minus_cancel_left:
619 "(f ---> - (y::_::real_normed_vector)) F \<longleftrightarrow> ((\<lambda>x. - f x) ---> y) F"
620 using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F]
623 lemma tendsto_diff [tendsto_intros]:
624 fixes a b :: "'a::real_normed_vector"
625 shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F"
626 using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus)
628 lemma continuous_diff [continuous_intros]:
629 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
630 shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x - g x)"
631 unfolding continuous_def by (rule tendsto_diff)
633 lemma continuous_on_diff [continuous_intros]:
634 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
635 shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)"
636 unfolding continuous_on_def by (auto intro: tendsto_diff)
638 lemma continuous_on_op_minus: "continuous_on (s::'a::real_normed_vector set) (op - x)"
639 by (rule continuous_intros | simp)+
641 lemma tendsto_setsum [tendsto_intros]:
642 fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector"
643 assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) F"
644 shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) F"
645 proof (cases "finite S")
646 assume "finite S" thus ?thesis using assms
647 by (induct, simp, simp add: tendsto_add)
650 lemma continuous_setsum [continuous_intros]:
651 fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
652 shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>S. f i x)"
653 unfolding continuous_def by (rule tendsto_setsum)
655 lemma continuous_on_setsum [continuous_intros]:
656 fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::real_normed_vector"
657 shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Sum>i\<in>S. f i x)"
658 unfolding continuous_on_def by (auto intro: tendsto_setsum)
660 lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real]
662 subsubsection \<open>Linear operators and multiplication\<close>
664 lemma (in bounded_linear) tendsto:
665 "(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F"
666 by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun)
668 lemma (in bounded_linear) continuous:
669 "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))"
670 using tendsto[of g _ F] by (auto simp: continuous_def)
672 lemma (in bounded_linear) continuous_on:
673 "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))"
674 using tendsto[of g] by (auto simp: continuous_on_def)
676 lemma (in bounded_linear) tendsto_zero:
677 "(g ---> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) F"
678 by (drule tendsto, simp only: zero)
680 lemma (in bounded_bilinear) tendsto:
681 "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) F"
682 by (simp only: tendsto_Zfun_iff prod_diff_prod
683 Zfun_add Zfun Zfun_left Zfun_right)
685 lemma (in bounded_bilinear) continuous:
686 "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ** g x)"
687 using tendsto[of f _ F g] by (auto simp: continuous_def)
689 lemma (in bounded_bilinear) continuous_on:
690 "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x ** g x)"
691 using tendsto[of f _ _ g] by (auto simp: continuous_on_def)
693 lemma (in bounded_bilinear) tendsto_zero:
694 assumes f: "(f ---> 0) F"
695 assumes g: "(g ---> 0) F"
696 shows "((\<lambda>x. f x ** g x) ---> 0) F"
697 using tendsto [OF f g] by (simp add: zero_left)
699 lemma (in bounded_bilinear) tendsto_left_zero:
700 "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) F"
701 by (rule bounded_linear.tendsto_zero [OF bounded_linear_left])
703 lemma (in bounded_bilinear) tendsto_right_zero:
704 "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) F"
705 by (rule bounded_linear.tendsto_zero [OF bounded_linear_right])
707 lemmas tendsto_of_real [tendsto_intros] =
708 bounded_linear.tendsto [OF bounded_linear_of_real]
710 lemmas tendsto_scaleR [tendsto_intros] =
711 bounded_bilinear.tendsto [OF bounded_bilinear_scaleR]
713 lemmas tendsto_mult [tendsto_intros] =
714 bounded_bilinear.tendsto [OF bounded_bilinear_mult]
716 lemmas continuous_of_real [continuous_intros] =
717 bounded_linear.continuous [OF bounded_linear_of_real]
719 lemmas continuous_scaleR [continuous_intros] =
720 bounded_bilinear.continuous [OF bounded_bilinear_scaleR]
722 lemmas continuous_mult [continuous_intros] =
723 bounded_bilinear.continuous [OF bounded_bilinear_mult]
725 lemmas continuous_on_of_real [continuous_intros] =
726 bounded_linear.continuous_on [OF bounded_linear_of_real]
728 lemmas continuous_on_scaleR [continuous_intros] =
729 bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR]
731 lemmas continuous_on_mult [continuous_intros] =
732 bounded_bilinear.continuous_on [OF bounded_bilinear_mult]
734 lemmas tendsto_mult_zero =
735 bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult]
737 lemmas tendsto_mult_left_zero =
738 bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult]
740 lemmas tendsto_mult_right_zero =
741 bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult]
743 lemma tendsto_power [tendsto_intros]:
744 fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
745 shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) F"
746 by (induct n) (simp_all add: tendsto_mult)
748 lemma continuous_power [continuous_intros]:
749 fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
750 shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)"
751 unfolding continuous_def by (rule tendsto_power)
753 lemma continuous_on_power [continuous_intros]:
754 fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
755 shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. (f x)^n)"
756 unfolding continuous_on_def by (auto intro: tendsto_power)
758 lemma tendsto_setprod [tendsto_intros]:
759 fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
760 assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) F"
761 shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) F"
762 proof (cases "finite S")
763 assume "finite S" thus ?thesis using assms
764 by (induct, simp, simp add: tendsto_mult)
767 lemma continuous_setprod [continuous_intros]:
768 fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
769 shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>S. f i x)"
770 unfolding continuous_def by (rule tendsto_setprod)
772 lemma continuous_on_setprod [continuous_intros]:
773 fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
774 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)"
775 unfolding continuous_on_def by (auto intro: tendsto_setprod)
777 lemma tendsto_of_real_iff:
778 "((\<lambda>x. of_real (f x) :: 'a :: real_normed_div_algebra) ---> of_real c) F \<longleftrightarrow> (f ---> c) F"
779 unfolding tendsto_iff by simp
781 lemma tendsto_add_const_iff:
782 "((\<lambda>x. c + f x :: 'a :: real_normed_vector) ---> c + d) F \<longleftrightarrow> (f ---> d) F"
783 using tendsto_add[OF tendsto_const[of c], of f d]
784 tendsto_add[OF tendsto_const[of "-c"], of "\<lambda>x. c + f x" "c + d"] by auto
787 subsubsection \<open>Inverse and division\<close>
789 lemma (in bounded_bilinear) Zfun_prod_Bfun:
790 assumes f: "Zfun f F"
791 assumes g: "Bfun g F"
792 shows "Zfun (\<lambda>x. f x ** g x) F"
794 obtain K where K: "0 \<le> K"
795 and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
796 using nonneg_bounded by blast
797 obtain B where B: "0 < B"
798 and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) F"
799 using g by (rule BfunE)
800 have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) F"
801 using norm_g proof eventually_elim
803 have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K"
805 also have "\<dots> \<le> norm (f x) * B * K"
806 by (intro mult_mono' order_refl norm_g norm_ge_zero
807 mult_nonneg_nonneg K elim)
808 also have "\<dots> = norm (f x) * (B * K)"
810 finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" .
813 by (rule Zfun_imp_Zfun)
816 lemma (in bounded_bilinear) flip:
817 "bounded_bilinear (\<lambda>x y. y ** x)"
819 apply (rule add_right)
820 apply (rule add_left)
821 apply (rule scaleR_right)
822 apply (rule scaleR_left)
823 apply (subst mult.commute)
828 lemma (in bounded_bilinear) Bfun_prod_Zfun:
829 assumes f: "Bfun f F"
830 assumes g: "Zfun g F"
831 shows "Zfun (\<lambda>x. f x ** g x) F"
832 using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun)
834 lemma Bfun_inverse_lemma:
835 fixes x :: "'a::real_normed_div_algebra"
836 shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
837 apply (subst nonzero_norm_inverse, clarsimp)
838 apply (erule (1) le_imp_inverse_le)
842 fixes a :: "'a::real_normed_div_algebra"
843 assumes f: "(f ---> a) F"
844 assumes a: "a \<noteq> 0"
845 shows "Bfun (\<lambda>x. inverse (f x)) F"
847 from a have "0 < norm a" by simp
848 hence "\<exists>r>0. r < norm a" by (rule dense)
849 then obtain r where r1: "0 < r" and r2: "r < norm a" by blast
850 have "eventually (\<lambda>x. dist (f x) a < r) F"
851 using tendstoD [OF f r1] by blast
852 hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F"
853 proof eventually_elim
855 hence 1: "norm (f x - a) < r"
856 by (simp add: dist_norm)
857 hence 2: "f x \<noteq> 0" using r2 by auto
858 hence "norm (inverse (f x)) = inverse (norm (f x))"
859 by (rule nonzero_norm_inverse)
860 also have "\<dots> \<le> inverse (norm a - r)"
861 proof (rule le_imp_inverse_le)
862 show "0 < norm a - r" using r2 by simp
864 have "norm a - norm (f x) \<le> norm (a - f x)"
865 by (rule norm_triangle_ineq2)
866 also have "\<dots> = norm (f x - a)"
867 by (rule norm_minus_commute)
868 also have "\<dots> < r" using 1 .
869 finally show "norm a - r \<le> norm (f x)" by simp
871 finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" .
873 thus ?thesis by (rule BfunI)
876 lemma tendsto_inverse [tendsto_intros]:
877 fixes a :: "'a::real_normed_div_algebra"
878 assumes f: "(f ---> a) F"
879 assumes a: "a \<noteq> 0"
880 shows "((\<lambda>x. inverse (f x)) ---> inverse a) F"
882 from a have "0 < norm a" by simp
883 with f have "eventually (\<lambda>x. dist (f x) a < norm a) F"
885 then have "eventually (\<lambda>x. f x \<noteq> 0) F"
886 unfolding dist_norm by (auto elim!: eventually_elim1)
887 with a have "eventually (\<lambda>x. inverse (f x) - inverse a =
888 - (inverse (f x) * (f x - a) * inverse a)) F"
889 by (auto elim!: eventually_elim1 simp: inverse_diff_inverse)
890 moreover have "Zfun (\<lambda>x. - (inverse (f x) * (f x - a) * inverse a)) F"
891 by (intro Zfun_minus Zfun_mult_left
892 bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult]
893 Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff])
894 ultimately show ?thesis
895 unfolding tendsto_Zfun_iff by (rule Zfun_ssubst)
898 lemma continuous_inverse:
899 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
900 assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
901 shows "continuous F (\<lambda>x. inverse (f x))"
902 using assms unfolding continuous_def by (rule tendsto_inverse)
904 lemma continuous_at_within_inverse[continuous_intros]:
905 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
906 assumes "continuous (at a within s) f" and "f a \<noteq> 0"
907 shows "continuous (at a within s) (\<lambda>x. inverse (f x))"
908 using assms unfolding continuous_within by (rule tendsto_inverse)
910 lemma isCont_inverse[continuous_intros, simp]:
911 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
912 assumes "isCont f a" and "f a \<noteq> 0"
913 shows "isCont (\<lambda>x. inverse (f x)) a"
914 using assms unfolding continuous_at by (rule tendsto_inverse)
916 lemma continuous_on_inverse[continuous_intros]:
917 fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
918 assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
919 shows "continuous_on s (\<lambda>x. inverse (f x))"
920 using assms unfolding continuous_on_def by (blast intro: tendsto_inverse)
922 lemma tendsto_divide [tendsto_intros]:
923 fixes a b :: "'a::real_normed_field"
924 shows "\<lbrakk>(f ---> a) F; (g ---> b) F; b \<noteq> 0\<rbrakk>
925 \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) F"
926 by (simp add: tendsto_mult tendsto_inverse divide_inverse)
928 lemma continuous_divide:
929 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field"
930 assumes "continuous F f" and "continuous F g" and "g (Lim F (\<lambda>x. x)) \<noteq> 0"
931 shows "continuous F (\<lambda>x. (f x) / (g x))"
932 using assms unfolding continuous_def by (rule tendsto_divide)
934 lemma continuous_at_within_divide[continuous_intros]:
935 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field"
936 assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \<noteq> 0"
937 shows "continuous (at a within s) (\<lambda>x. (f x) / (g x))"
938 using assms unfolding continuous_within by (rule tendsto_divide)
940 lemma isCont_divide[continuous_intros, simp]:
941 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field"
942 assumes "isCont f a" "isCont g a" "g a \<noteq> 0"
943 shows "isCont (\<lambda>x. (f x) / g x) a"
944 using assms unfolding continuous_at by (rule tendsto_divide)
946 lemma continuous_on_divide[continuous_intros]:
947 fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_field"
948 assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. g x \<noteq> 0"
949 shows "continuous_on s (\<lambda>x. (f x) / (g x))"
950 using assms unfolding continuous_on_def by (blast intro: tendsto_divide)
952 lemma tendsto_sgn [tendsto_intros]:
953 fixes l :: "'a::real_normed_vector"
954 shows "\<lbrakk>(f ---> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) F"
955 unfolding sgn_div_norm by (simp add: tendsto_intros)
957 lemma continuous_sgn:
958 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
959 assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
960 shows "continuous F (\<lambda>x. sgn (f x))"
961 using assms unfolding continuous_def by (rule tendsto_sgn)
963 lemma continuous_at_within_sgn[continuous_intros]:
964 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
965 assumes "continuous (at a within s) f" and "f a \<noteq> 0"
966 shows "continuous (at a within s) (\<lambda>x. sgn (f x))"
967 using assms unfolding continuous_within by (rule tendsto_sgn)
969 lemma isCont_sgn[continuous_intros]:
970 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
971 assumes "isCont f a" and "f a \<noteq> 0"
972 shows "isCont (\<lambda>x. sgn (f x)) a"
973 using assms unfolding continuous_at by (rule tendsto_sgn)
975 lemma continuous_on_sgn[continuous_intros]:
976 fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
977 assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
978 shows "continuous_on s (\<lambda>x. sgn (f x))"
979 using assms unfolding continuous_on_def by (blast intro: tendsto_sgn)
981 lemma filterlim_at_infinity:
982 fixes f :: "_ \<Rightarrow> 'a::real_normed_vector"
984 shows "(LIM x F. f x :> at_infinity) \<longleftrightarrow> (\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F)"
985 unfolding filterlim_iff eventually_at_infinity
987 fix P :: "'a \<Rightarrow> bool" and b
988 assume *: "\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F"
989 and P: "\<forall>x. b \<le> norm x \<longrightarrow> P x"
990 have "max b (c + 1) > c" by auto
991 with * have "eventually (\<lambda>x. max b (c + 1) \<le> norm (f x)) F"
993 then show "eventually (\<lambda>x. P (f x)) F"
994 proof eventually_elim
995 fix x assume "max b (c + 1) \<le> norm (f x)"
996 with P show "P (f x)" by auto
1000 lemma not_tendsto_and_filterlim_at_infinity:
1001 assumes "F \<noteq> bot"
1002 assumes "(f ---> (c :: 'a :: real_normed_vector)) F"
1003 assumes "filterlim f at_infinity F"
1006 from tendstoD[OF assms(2), of "1/2"]
1007 have "eventually (\<lambda>x. dist (f x) c < 1/2) F" by simp
1008 moreover from filterlim_at_infinity[of "norm c" f F] assms(3)
1009 have "eventually (\<lambda>x. norm (f x) \<ge> norm c + 1) F" by simp
1010 ultimately have "eventually (\<lambda>x. False) F"
1011 proof eventually_elim
1012 fix x assume A: "dist (f x) c < 1/2" and B: "norm (f x) \<ge> norm c + 1"
1014 also have "norm (f x) = dist (f x) 0" by (simp add: norm_conv_dist)
1015 also have "... \<le> dist (f x) c + dist c 0" by (rule dist_triangle)
1017 finally show False by (simp add: norm_conv_dist)
1019 with assms show False by simp
1022 lemma filterlim_at_infinity_imp_not_convergent:
1023 assumes "filterlim f at_infinity sequentially"
1024 shows "\<not>convergent f"
1025 by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms])
1026 (simp_all add: convergent_LIMSEQ_iff)
1028 lemma filterlim_at_infinity_imp_eventually_ne:
1029 assumes "filterlim f at_infinity F"
1030 shows "eventually (\<lambda>z. f z \<noteq> c) F"
1032 have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all
1033 with filterlim_at_infinity[OF order.refl, of f F] assms
1034 have "eventually (\<lambda>z. norm (f z) \<ge> norm c + 1) F" by blast
1035 thus ?thesis by eventually_elim auto
1038 lemma tendsto_of_nat [tendsto_intros]:
1039 "filterlim (of_nat :: nat \<Rightarrow> 'a :: real_normed_algebra_1) at_infinity sequentially"
1040 proof (subst filterlim_at_infinity[OF order.refl], intro allI impI)
1041 fix r :: real assume r: "r > 0"
1042 def n \<equiv> "nat \<lceil>r\<rceil>"
1043 from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r" unfolding n_def by linarith
1044 from eventually_ge_at_top[of n] show "eventually (\<lambda>m. norm (of_nat m :: 'a) \<ge> r) sequentially"
1045 by eventually_elim (insert n, simp_all)
1049 subsection \<open>Relate @{const at}, @{const at_left} and @{const at_right}\<close>
1053 This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
1054 @{term "at_right x"} and also @{term "at_right 0"}.
1058 lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real]
1060 lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d::'a::real_normed_vector)"
1061 by (rule filtermap_fun_inverse[where g="\<lambda>x. x + d"])
1062 (auto intro!: tendsto_eq_intros filterlim_ident)
1064 lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a::'a::real_normed_vector)"
1065 by (rule filtermap_fun_inverse[where g=uminus])
1066 (auto intro!: tendsto_eq_intros filterlim_ident)
1068 lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d::'a::real_normed_vector)"
1069 by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric])
1071 lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d::real)"
1072 by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric])
1074 lemma at_right_to_0: "at_right (a::real) = filtermap (\<lambda>x. x + a) (at_right 0)"
1075 using filtermap_at_right_shift[of "-a" 0] by simp
1077 lemma filterlim_at_right_to_0:
1078 "filterlim f F (at_right (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)"
1079 unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] ..
1081 lemma eventually_at_right_to_0:
1082 "eventually P (at_right (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)"
1083 unfolding at_right_to_0[of a] by (simp add: eventually_filtermap)
1085 lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a::'a::real_normed_vector)"
1086 by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric])
1088 lemma at_left_minus: "at_left (a::real) = filtermap (\<lambda>x. - x) (at_right (- a))"
1089 by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric])
1091 lemma at_right_minus: "at_right (a::real) = filtermap (\<lambda>x. - x) (at_left (- a))"
1092 by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric])
1094 lemma filterlim_at_left_to_right:
1095 "filterlim f F (at_left (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))"
1096 unfolding filterlim_def filtermap_filtermap at_left_minus[of a] ..
1098 lemma eventually_at_left_to_right:
1099 "eventually P (at_left (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))"
1100 unfolding at_left_minus[of a] by (simp add: eventually_filtermap)
1102 lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top"
1103 unfolding filterlim_at_top eventually_at_bot_dense
1104 by (metis leI minus_less_iff order_less_asym)
1106 lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot"
1107 unfolding filterlim_at_bot eventually_at_top_dense
1108 by (metis leI less_minus_iff order_less_asym)
1110 lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: real filter)"
1111 by (rule filtermap_fun_inverse[symmetric, of uminus])
1112 (auto intro: filterlim_uminus_at_bot_at_top filterlim_uminus_at_top_at_bot)
1114 lemma at_bot_mirror: "at_bot = filtermap uminus (at_top :: real filter)"
1115 unfolding at_top_mirror filtermap_filtermap by (simp add: filtermap_ident)
1117 lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \<longleftrightarrow> (LIM x at_bot. f (-x::real) :> F)"
1118 unfolding filterlim_def at_top_mirror filtermap_filtermap ..
1120 lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \<longleftrightarrow> (LIM x at_top. f (-x::real) :> F)"
1121 unfolding filterlim_def at_bot_mirror filtermap_filtermap ..
1123 lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_bot)"
1124 using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F]
1125 using filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F]
1128 lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_top)"
1129 unfolding filterlim_uminus_at_top by simp
1131 lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top"
1132 unfolding filterlim_at_top_gt[where c=0] eventually_at_filter
1134 fix Z :: real assume [arith]: "0 < Z"
1135 then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)"
1136 by (auto simp add: eventually_nhds_metric dist_real_def intro!: exI[of _ "\<bar>inverse Z\<bar>"])
1137 then show "eventually (\<lambda>x. x \<noteq> 0 \<longrightarrow> x \<in> {0<..} \<longrightarrow> Z \<le> inverse x) (nhds 0)"
1138 by (auto elim!: eventually_elim1 simp: inverse_eq_divide field_simps)
1141 lemma tendsto_inverse_0:
1142 fixes x :: "_ \<Rightarrow> 'a::real_normed_div_algebra"
1143 shows "(inverse ---> (0::'a)) at_infinity"
1144 unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity
1146 fix r :: real assume "0 < r"
1147 show "\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> norm (inverse x :: 'a) < r"
1148 proof (intro exI[of _ "inverse (r / 2)"] allI impI)
1150 from \<open>0 < r\<close> have "0 < inverse (r / 2)" by simp
1151 also assume *: "inverse (r / 2) \<le> norm x"
1152 finally show "norm (inverse x) < r"
1153 using * \<open>0 < r\<close> by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps)
1157 lemma tendsto_add_filterlim_at_infinity:
1158 assumes "(f ---> (c :: 'b :: real_normed_vector)) (F :: 'a filter)"
1159 assumes "filterlim g at_infinity F"
1160 shows "filterlim (\<lambda>x. f x + g x) at_infinity F"
1161 proof (subst filterlim_at_infinity[OF order_refl], safe)
1162 fix r :: real assume r: "r > 0"
1163 from assms(1) have "((\<lambda>x. norm (f x)) ---> norm c) F" by (rule tendsto_norm)
1164 hence "eventually (\<lambda>x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all
1165 moreover from r have "r + norm c + 1 > 0" by (intro add_pos_nonneg) simp_all
1166 with assms(2) have "eventually (\<lambda>x. norm (g x) \<ge> r + norm c + 1) F"
1167 unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all
1168 ultimately show "eventually (\<lambda>x. norm (f x + g x) \<ge> r) F"
1169 proof eventually_elim
1170 fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \<le> norm (g x)"
1171 from A B have "r \<le> norm (g x) - norm (f x)" by simp
1172 also have "norm (g x) - norm (f x) \<le> norm (g x + f x)" by (rule norm_diff_ineq)
1173 finally show "r \<le> norm (f x + g x)" by (simp add: add_ac)
1177 lemma tendsto_add_filterlim_at_infinity':
1178 assumes "filterlim f at_infinity F"
1179 assumes "(g ---> (c :: 'b :: real_normed_vector)) (F :: 'a filter)"
1180 shows "filterlim (\<lambda>x. f x + g x) at_infinity F"
1181 by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+
1183 lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)"
1184 unfolding filterlim_at
1185 by (auto simp: eventually_at_top_dense)
1186 (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl)
1188 lemma filterlim_inverse_at_top:
1189 "(f ---> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> LIM x F. inverse (f x) :> at_top"
1190 by (intro filterlim_compose[OF filterlim_inverse_at_top_right])
1191 (simp add: filterlim_def eventually_filtermap eventually_elim1 at_within_def le_principal)
1193 lemma filterlim_inverse_at_bot_neg:
1194 "LIM x (at_left (0::real)). inverse x :> at_bot"
1195 by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right)
1197 lemma filterlim_inverse_at_bot:
1198 "(f ---> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x < 0) F \<Longrightarrow> LIM x F. inverse (f x) :> at_bot"
1199 unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric]
1200 by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric])
1202 lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top"
1203 by (intro filtermap_fun_inverse[symmetric, where g=inverse])
1204 (auto intro: filterlim_inverse_at_top_right filterlim_inverse_at_right_top)
1206 lemma eventually_at_right_to_top:
1207 "eventually P (at_right (0::real)) \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) at_top"
1208 unfolding at_right_to_top eventually_filtermap ..
1210 lemma filterlim_at_right_to_top:
1211 "filterlim f F (at_right (0::real)) \<longleftrightarrow> (LIM x at_top. f (inverse x) :> F)"
1212 unfolding filterlim_def at_right_to_top filtermap_filtermap ..
1214 lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))"
1215 unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident ..
1217 lemma eventually_at_top_to_right:
1218 "eventually P at_top \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) (at_right (0::real))"
1219 unfolding at_top_to_right eventually_filtermap ..
1221 lemma filterlim_at_top_to_right:
1222 "filterlim f F at_top \<longleftrightarrow> (LIM x (at_right (0::real)). f (inverse x) :> F)"
1223 unfolding filterlim_def at_top_to_right filtermap_filtermap ..
1225 lemma filterlim_inverse_at_infinity:
1226 fixes x :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
1227 shows "filterlim inverse at_infinity (at (0::'a))"
1228 unfolding filterlim_at_infinity[OF order_refl]
1230 fix r :: real assume "0 < r"
1231 then show "eventually (\<lambda>x::'a. r \<le> norm (inverse x)) (at 0)"
1232 unfolding eventually_at norm_inverse
1233 by (intro exI[of _ "inverse r"])
1234 (auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide)
1237 lemma filterlim_inverse_at_iff:
1238 fixes g :: "'a \<Rightarrow> 'b::{real_normed_div_algebra, division_ring}"
1239 shows "(LIM x F. inverse (g x) :> at 0) \<longleftrightarrow> (LIM x F. g x :> at_infinity)"
1240 unfolding filterlim_def filtermap_filtermap[symmetric]
1242 assume "filtermap g F \<le> at_infinity"
1243 then have "filtermap inverse (filtermap g F) \<le> filtermap inverse at_infinity"
1244 by (rule filtermap_mono)
1245 also have "\<dots> \<le> at 0"
1246 using tendsto_inverse_0[where 'a='b]
1247 by (auto intro!: exI[of _ 1]
1248 simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity)
1249 finally show "filtermap inverse (filtermap g F) \<le> at 0" .
1251 assume "filtermap inverse (filtermap g F) \<le> at 0"
1252 then have "filtermap inverse (filtermap inverse (filtermap g F)) \<le> filtermap inverse (at 0)"
1253 by (rule filtermap_mono)
1254 with filterlim_inverse_at_infinity show "filtermap g F \<le> at_infinity"
1255 by (auto intro: order_trans simp: filterlim_def filtermap_filtermap)
1258 lemma tendsto_mult_filterlim_at_infinity:
1259 assumes "F \<noteq> bot" "(f ---> (c :: 'a :: real_normed_field)) F" "c \<noteq> 0"
1260 assumes "filterlim g at_infinity F"
1261 shows "filterlim (\<lambda>x. f x * g x) at_infinity F"
1263 have "((\<lambda>x. inverse (f x) * inverse (g x)) ---> inverse c * 0) F"
1264 by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0])
1265 hence "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F"
1266 unfolding filterlim_at using assms
1267 by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj)
1268 thus ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all
1271 lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) ---> 0) F"
1272 by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff)
1274 lemma mult_nat_left_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. c * x :: nat) at_top sequentially"
1275 by (rule filterlim_subseq) (auto simp: subseq_def)
1277 lemma mult_nat_right_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. x * c :: nat) at_top sequentially"
1278 by (rule filterlim_subseq) (auto simp: subseq_def)
1280 lemma at_to_infinity:
1281 fixes x :: "'a :: {real_normed_field,field}"
1282 shows "(at (0::'a)) = filtermap inverse at_infinity"
1283 proof (rule antisym)
1284 have "(inverse ---> (0::'a)) at_infinity"
1285 by (fact tendsto_inverse_0)
1286 then show "filtermap inverse at_infinity \<le> at (0::'a)"
1287 apply (simp add: le_principal eventually_filtermap eventually_at_infinity filterlim_def at_within_def)
1288 apply (rule_tac x="1" in exI, auto)
1291 have "filtermap inverse (filtermap inverse (at (0::'a))) \<le> filtermap inverse at_infinity"
1292 using filterlim_inverse_at_infinity unfolding filterlim_def
1293 by (rule filtermap_mono)
1294 then show "at (0::'a) \<le> filtermap inverse at_infinity"
1295 by (simp add: filtermap_ident filtermap_filtermap)
1298 lemma lim_at_infinity_0:
1299 fixes l :: "'a :: {real_normed_field,field}"
1300 shows "(f ---> l) at_infinity \<longleftrightarrow> ((f o inverse) ---> l) (at (0::'a))"
1301 by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap)
1303 lemma lim_zero_infinity:
1304 fixes l :: "'a :: {real_normed_field,field}"
1305 shows "((\<lambda>x. f(1 / x)) ---> l) (at (0::'a)) \<Longrightarrow> (f ---> l) at_infinity"
1306 by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def)
1311 We only show rules for multiplication and addition when the functions are either against a real
1312 value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}.
1316 lemma filterlim_tendsto_pos_mult_at_top:
1317 assumes f: "(f ---> c) F" and c: "0 < c"
1318 assumes g: "LIM x F. g x :> at_top"
1319 shows "LIM x F. (f x * g x :: real) :> at_top"
1320 unfolding filterlim_at_top_gt[where c=0]
1322 fix Z :: real assume "0 < Z"
1323 from f \<open>0 < c\<close> have "eventually (\<lambda>x. c / 2 < f x) F"
1324 by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_elim1
1325 simp: dist_real_def abs_real_def split: split_if_asm)
1326 moreover from g have "eventually (\<lambda>x. (Z / c * 2) \<le> g x) F"
1327 unfolding filterlim_at_top by auto
1328 ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F"
1329 proof eventually_elim
1330 fix x assume "c / 2 < f x" "Z / c * 2 \<le> g x"
1331 with \<open>0 < Z\<close> \<open>0 < c\<close> have "c / 2 * (Z / c * 2) \<le> f x * g x"
1332 by (intro mult_mono) (auto simp: zero_le_divide_iff)
1333 with \<open>0 < c\<close> show "Z \<le> f x * g x"
1338 lemma filterlim_at_top_mult_at_top:
1339 assumes f: "LIM x F. f x :> at_top"
1340 assumes g: "LIM x F. g x :> at_top"
1341 shows "LIM x F. (f x * g x :: real) :> at_top"
1342 unfolding filterlim_at_top_gt[where c=0]
1344 fix Z :: real assume "0 < Z"
1345 from f have "eventually (\<lambda>x. 1 \<le> f x) F"
1346 unfolding filterlim_at_top by auto
1347 moreover from g have "eventually (\<lambda>x. Z \<le> g x) F"
1348 unfolding filterlim_at_top by auto
1349 ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F"
1350 proof eventually_elim
1351 fix x assume "1 \<le> f x" "Z \<le> g x"
1352 with \<open>0 < Z\<close> have "1 * Z \<le> f x * g x"
1353 by (intro mult_mono) (auto simp: zero_le_divide_iff)
1354 then show "Z \<le> f x * g x"
1359 lemma filterlim_tendsto_pos_mult_at_bot:
1360 assumes "(f ---> c) F" "0 < (c::real)" "filterlim g at_bot F"
1361 shows "LIM x F. f x * g x :> at_bot"
1362 using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\<lambda>x. - g x"] assms(3)
1363 unfolding filterlim_uminus_at_bot by simp
1365 lemma filterlim_tendsto_neg_mult_at_bot:
1366 assumes c: "(f ---> c) F" "(c::real) < 0" and g: "filterlim g at_top F"
1367 shows "LIM x F. f x * g x :> at_bot"
1368 using c filterlim_tendsto_pos_mult_at_top[of "\<lambda>x. - f x" "- c" F, OF _ _ g]
1369 unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp
1371 lemma filterlim_pow_at_top:
1372 fixes f :: "real \<Rightarrow> real"
1373 assumes "0 < n" and f: "LIM x F. f x :> at_top"
1374 shows "LIM x F. (f x)^n :: real :> at_top"
1375 using \<open>0 < n\<close> proof (induct n)
1376 case (Suc n) with f show ?case
1377 by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top)
1380 lemma filterlim_pow_at_bot_even:
1381 fixes f :: "real \<Rightarrow> real"
1382 shows "0 < n \<Longrightarrow> LIM x F. f x :> at_bot \<Longrightarrow> even n \<Longrightarrow> LIM x F. (f x)^n :> at_top"
1383 using filterlim_pow_at_top[of n "\<lambda>x. - f x" F] by (simp add: filterlim_uminus_at_top)
1385 lemma filterlim_pow_at_bot_odd:
1386 fixes f :: "real \<Rightarrow> real"
1387 shows "0 < n \<Longrightarrow> LIM x F. f x :> at_bot \<Longrightarrow> odd n \<Longrightarrow> LIM x F. (f x)^n :> at_bot"
1388 using filterlim_pow_at_top[of n "\<lambda>x. - f x" F] by (simp add: filterlim_uminus_at_bot)
1390 lemma filterlim_tendsto_add_at_top:
1391 assumes f: "(f ---> c) F"
1392 assumes g: "LIM x F. g x :> at_top"
1393 shows "LIM x F. (f x + g x :: real) :> at_top"
1394 unfolding filterlim_at_top_gt[where c=0]
1396 fix Z :: real assume "0 < Z"
1397 from f have "eventually (\<lambda>x. c - 1 < f x) F"
1398 by (auto dest!: tendstoD[where e=1] elim!: eventually_elim1 simp: dist_real_def)
1399 moreover from g have "eventually (\<lambda>x. Z - (c - 1) \<le> g x) F"
1400 unfolding filterlim_at_top by auto
1401 ultimately show "eventually (\<lambda>x. Z \<le> f x + g x) F"
1402 by eventually_elim simp
1405 lemma LIM_at_top_divide:
1406 fixes f g :: "'a \<Rightarrow> real"
1407 assumes f: "(f ---> a) F" "0 < a"
1408 assumes g: "(g ---> 0) F" "eventually (\<lambda>x. 0 < g x) F"
1409 shows "LIM x F. f x / g x :> at_top"
1410 unfolding divide_inverse
1411 by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g])
1413 lemma filterlim_at_top_add_at_top:
1414 assumes f: "LIM x F. f x :> at_top"
1415 assumes g: "LIM x F. g x :> at_top"
1416 shows "LIM x F. (f x + g x :: real) :> at_top"
1417 unfolding filterlim_at_top_gt[where c=0]
1419 fix Z :: real assume "0 < Z"
1420 from f have "eventually (\<lambda>x. 0 \<le> f x) F"
1421 unfolding filterlim_at_top by auto
1422 moreover from g have "eventually (\<lambda>x. Z \<le> g x) F"
1423 unfolding filterlim_at_top by auto
1424 ultimately show "eventually (\<lambda>x. Z \<le> f x + g x) F"
1425 by eventually_elim simp
1428 lemma tendsto_divide_0:
1429 fixes f :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
1430 assumes f: "(f ---> c) F"
1431 assumes g: "LIM x F. g x :> at_infinity"
1432 shows "((\<lambda>x. f x / g x) ---> 0) F"
1433 using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse)
1435 lemma linear_plus_1_le_power:
1437 assumes x: "0 \<le> x"
1438 shows "real n * x + 1 \<le> (x + 1) ^ n"
1441 have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)"
1442 by (simp add: field_simps of_nat_Suc x)
1443 also have "\<dots> \<le> (x + 1)^Suc n"
1444 using Suc x by (simp add: mult_left_mono)
1445 finally show ?case .
1448 lemma filterlim_realpow_sequentially_gt1:
1449 fixes x :: "'a :: real_normed_div_algebra"
1450 assumes x[arith]: "1 < norm x"
1451 shows "LIM n sequentially. x ^ n :> at_infinity"
1452 proof (intro filterlim_at_infinity[THEN iffD2] allI impI)
1453 fix y :: real assume "0 < y"
1454 have "0 < norm x - 1" by simp
1455 then obtain N::nat where "y < real N * (norm x - 1)" by (blast dest: reals_Archimedean3)
1456 also have "\<dots> \<le> real N * (norm x - 1) + 1" by simp
1457 also have "\<dots> \<le> (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp
1458 also have "\<dots> = norm x ^ N" by simp
1459 finally have "\<forall>n\<ge>N. y \<le> norm x ^ n"
1460 by (metis order_less_le_trans power_increasing order_less_imp_le x)
1461 then show "eventually (\<lambda>n. y \<le> norm (x ^ n)) sequentially"
1462 unfolding eventually_sequentially
1463 by (auto simp: norm_power)
1467 subsection \<open>Limits of Sequences\<close>
1469 lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
1473 fixes L :: "'a::real_normed_vector"
1474 shows "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
1475 unfolding lim_sequentially dist_norm ..
1478 fixes L :: "'a::real_normed_vector"
1479 shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
1480 by (simp add: LIMSEQ_iff)
1483 fixes L :: "'a::real_normed_vector"
1484 shows "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
1485 by (simp add: LIMSEQ_iff)
1487 lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
1488 unfolding tendsto_def eventually_sequentially
1489 by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute)
1491 lemma Bseq_inverse_lemma:
1492 fixes x :: "'a::real_normed_div_algebra"
1493 shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
1494 apply (subst nonzero_norm_inverse, clarsimp)
1495 apply (erule (1) le_imp_inverse_le)
1499 fixes a :: "'a::real_normed_div_algebra"
1500 shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
1501 by (rule Bfun_inverse)
1503 text\<open>Transformation of limit.\<close>
1505 lemma eventually_at2:
1506 "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
1507 unfolding eventually_at dist_nz by auto
1509 lemma Lim_transform:
1510 fixes a b :: "'a::real_normed_vector"
1511 shows "\<lbrakk>(g ---> a) F; ((\<lambda>x. f x - g x) ---> 0) F\<rbrakk> \<Longrightarrow> (f ---> a) F"
1512 using tendsto_add [of g a F "\<lambda>x. f x - g x" 0] by simp
1514 lemma Lim_transform2:
1515 fixes a b :: "'a::real_normed_vector"
1516 shows "\<lbrakk>(f ---> a) F; ((\<lambda>x. f x - g x) ---> 0) F\<rbrakk> \<Longrightarrow> (g ---> a) F"
1517 by (erule Lim_transform) (simp add: tendsto_minus_cancel)
1519 lemma Lim_transform_eventually:
1520 "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net \<Longrightarrow> (g ---> l) net"
1521 apply (rule topological_tendstoI)
1522 apply (drule (2) topological_tendstoD)
1523 apply (erule (1) eventually_elim2, simp)
1526 lemma Lim_transform_within:
1528 and "\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"
1529 and "(f ---> l) (at x within S)"
1530 shows "(g ---> l) (at x within S)"
1531 proof (rule Lim_transform_eventually)
1532 show "eventually (\<lambda>x. f x = g x) (at x within S)"
1533 using assms(1,2) by (auto simp: dist_nz eventually_at)
1534 show "(f ---> l) (at x within S)" by fact
1537 lemma Lim_transform_at:
1539 and "\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"
1540 and "(f ---> l) (at x)"
1541 shows "(g ---> l) (at x)"
1543 proof (rule Lim_transform_eventually)
1544 show "eventually (\<lambda>x. f x = g x) (at x)"
1545 unfolding eventually_at2
1546 using assms(1,2) by auto
1549 text\<open>Common case assuming being away from some crucial point like 0.\<close>
1551 lemma Lim_transform_away_within:
1552 fixes a b :: "'a::t1_space"
1553 assumes "a \<noteq> b"
1554 and "\<forall>x\<in>S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
1555 and "(f ---> l) (at a within S)"
1556 shows "(g ---> l) (at a within S)"
1557 proof (rule Lim_transform_eventually)
1558 show "(f ---> l) (at a within S)" by fact
1559 show "eventually (\<lambda>x. f x = g x) (at a within S)"
1560 unfolding eventually_at_topological
1561 by (rule exI [where x="- {b}"], simp add: open_Compl assms)
1564 lemma Lim_transform_away_at:
1565 fixes a b :: "'a::t1_space"
1566 assumes ab: "a\<noteq>b"
1567 and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
1568 and fl: "(f ---> l) (at a)"
1569 shows "(g ---> l) (at a)"
1570 using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp
1572 text\<open>Alternatively, within an open set.\<close>
1574 lemma Lim_transform_within_open:
1575 assumes "open S" and "a \<in> S"
1576 and "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"
1577 and "(f ---> l) (at a)"
1578 shows "(g ---> l) (at a)"
1579 proof (rule Lim_transform_eventually)
1580 show "eventually (\<lambda>x. f x = g x) (at a)"
1581 unfolding eventually_at_topological
1582 using assms(1,2,3) by auto
1583 show "(f ---> l) (at a)" by fact
1586 text\<open>A congruence rule allowing us to transform limits assuming not at point.\<close>
1588 (* FIXME: Only one congruence rule for tendsto can be used at a time! *)
1590 lemma Lim_cong_within(*[cong add]*):
1594 and "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x"
1595 shows "(f ---> x) (at a within S) \<longleftrightarrow> (g ---> y) (at b within T)"
1596 unfolding tendsto_def eventually_at_topological
1599 lemma Lim_cong_at(*[cong add]*):
1600 assumes "a = b" "x = y"
1601 and "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
1602 shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
1603 unfolding tendsto_def eventually_at_topological
1605 text\<open>An unbounded sequence's inverse tends to 0\<close>
1607 lemma LIMSEQ_inverse_zero:
1608 "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
1609 apply (rule filterlim_compose[OF tendsto_inverse_0])
1610 apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially)
1611 apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
1614 text\<open>The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity\<close>
1616 lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
1617 by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
1618 filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
1620 text\<open>The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
1621 infinity is now easily proved\<close>
1623 lemma LIMSEQ_inverse_real_of_nat_add:
1624 "(%n. r + inverse(real(Suc n))) ----> r"
1625 using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto
1627 lemma LIMSEQ_inverse_real_of_nat_add_minus:
1628 "(%n. r + -inverse(real(Suc n))) ----> r"
1629 using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]]
1632 lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
1633 "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
1634 using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]]
1637 lemma lim_1_over_n: "((\<lambda>n. 1 / of_nat n) ---> (0::'a::real_normed_field)) sequentially"
1638 proof (subst lim_sequentially, intro allI impI exI)
1639 fix e :: real assume e: "e > 0"
1640 fix n :: nat assume n: "n \<ge> nat \<lceil>inverse e + 1\<rceil>"
1641 have "inverse e < of_nat (nat \<lceil>inverse e + 1\<rceil>)" by linarith
1643 finally show "dist (1 / of_nat n :: 'a) 0 < e" using e
1644 by (simp add: divide_simps mult.commute norm_conv_dist[symmetric] norm_divide)
1647 lemma lim_inverse_n: "((\<lambda>n. inverse(of_nat n)) ---> (0::'a::real_normed_field)) sequentially"
1648 using lim_1_over_n by (simp add: inverse_eq_divide)
1650 lemma LIMSEQ_Suc_n_over_n: "(\<lambda>n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) ----> 1"
1651 proof (rule Lim_transform_eventually)
1652 show "eventually (\<lambda>n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially"
1653 using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps)
1654 have "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) ----> 1 + 0"
1655 by (intro tendsto_add tendsto_const lim_inverse_n)
1656 thus "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) ----> 1" by simp
1659 lemma LIMSEQ_n_over_Suc_n: "(\<lambda>n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) ----> 1"
1660 proof (rule Lim_transform_eventually)
1661 show "eventually (\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a) =
1662 of_nat n / of_nat (Suc n)) sequentially"
1663 using eventually_gt_at_top[of "0::nat"]
1664 by eventually_elim (simp add: field_simps del: of_nat_Suc)
1665 have "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) ----> inverse 1"
1666 by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all
1667 thus "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) ----> 1" by simp
1670 subsection \<open>Convergence on sequences\<close>
1672 lemma convergent_cong:
1673 assumes "eventually (\<lambda>x. f x = g x) sequentially"
1674 shows "convergent f \<longleftrightarrow> convergent g"
1675 unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl)
1677 lemma convergent_Suc_iff: "convergent (\<lambda>n. f (Suc n)) \<longleftrightarrow> convergent f"
1678 by (auto simp: convergent_def LIMSEQ_Suc_iff)
1680 lemma convergent_ignore_initial_segment: "convergent (\<lambda>n. f (n + m)) = convergent f"
1681 proof (induction m arbitrary: f)
1683 have "convergent (\<lambda>n. f (n + Suc m)) \<longleftrightarrow> convergent (\<lambda>n. f (Suc n + m))" by simp
1684 also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. f (n + m))" by (rule convergent_Suc_iff)
1685 also have "\<dots> \<longleftrightarrow> convergent f" by (rule Suc)
1686 finally show ?case .
1689 lemma convergent_add:
1690 fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
1691 assumes "convergent (\<lambda>n. X n)"
1692 assumes "convergent (\<lambda>n. Y n)"
1693 shows "convergent (\<lambda>n. X n + Y n)"
1694 using assms unfolding convergent_def by (blast intro: tendsto_add)
1696 lemma convergent_setsum:
1697 fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
1698 assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
1699 shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
1700 proof (cases "finite A")
1701 case True from this and assms show ?thesis
1702 by (induct A set: finite) (simp_all add: convergent_const convergent_add)
1703 qed (simp add: convergent_const)
1705 lemma (in bounded_linear) convergent:
1706 assumes "convergent (\<lambda>n. X n)"
1707 shows "convergent (\<lambda>n. f (X n))"
1708 using assms unfolding convergent_def by (blast intro: tendsto)
1710 lemma (in bounded_bilinear) convergent:
1711 assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
1712 shows "convergent (\<lambda>n. X n ** Y n)"
1713 using assms unfolding convergent_def by (blast intro: tendsto)
1715 lemma convergent_minus_iff:
1716 fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
1717 shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
1718 apply (simp add: convergent_def)
1719 apply (auto dest: tendsto_minus)
1720 apply (drule tendsto_minus, auto)
1723 lemma convergent_diff:
1724 fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
1725 assumes "convergent (\<lambda>n. X n)"
1726 assumes "convergent (\<lambda>n. Y n)"
1727 shows "convergent (\<lambda>n. X n - Y n)"
1728 using assms unfolding convergent_def by (blast intro: tendsto_diff)
1730 lemma convergent_norm:
1731 assumes "convergent f"
1732 shows "convergent (\<lambda>n. norm (f n))"
1734 from assms have "f ----> lim f" by (simp add: convergent_LIMSEQ_iff)
1735 hence "(\<lambda>n. norm (f n)) ----> norm (lim f)" by (rule tendsto_norm)
1736 thus ?thesis by (auto simp: convergent_def)
1739 lemma convergent_of_real:
1740 "convergent f \<Longrightarrow> convergent (\<lambda>n. of_real (f n) :: 'a :: real_normed_algebra_1)"
1741 unfolding convergent_def by (blast intro!: tendsto_of_real)
1743 lemma convergent_add_const_iff:
1744 "convergent (\<lambda>n. c + f n :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
1746 assume "convergent (\<lambda>n. c + f n)"
1747 from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp
1749 assume "convergent f"
1750 from convergent_add[OF convergent_const[of c] this] show "convergent (\<lambda>n. c + f n)" by simp
1753 lemma convergent_add_const_right_iff:
1754 "convergent (\<lambda>n. f n + c :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
1755 using convergent_add_const_iff[of c f] by (simp add: add_ac)
1757 lemma convergent_diff_const_right_iff:
1758 "convergent (\<lambda>n. f n - c :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
1759 using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac)
1761 lemma convergent_mult:
1762 fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
1763 assumes "convergent (\<lambda>n. X n)"
1764 assumes "convergent (\<lambda>n. Y n)"
1765 shows "convergent (\<lambda>n. X n * Y n)"
1766 using assms unfolding convergent_def by (blast intro: tendsto_mult)
1768 lemma convergent_mult_const_iff:
1769 assumes "c \<noteq> 0"
1770 shows "convergent (\<lambda>n. c * f n :: 'a :: real_normed_field) \<longleftrightarrow> convergent f"
1772 assume "convergent (\<lambda>n. c * f n)"
1773 from assms convergent_mult[OF this convergent_const[of "inverse c"]]
1774 show "convergent f" by (simp add: field_simps)
1776 assume "convergent f"
1777 from convergent_mult[OF convergent_const[of c] this] show "convergent (\<lambda>n. c * f n)" by simp
1780 lemma convergent_mult_const_right_iff:
1781 assumes "c \<noteq> 0"
1782 shows "convergent (\<lambda>n. (f n :: 'a :: real_normed_field) * c) \<longleftrightarrow> convergent f"
1783 using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac)
1785 lemma convergent_imp_Bseq: "convergent f \<Longrightarrow> Bseq f"
1786 by (simp add: Cauchy_Bseq convergent_Cauchy)
1789 text \<open>A monotone sequence converges to its least upper bound.\<close>
1791 lemma LIMSEQ_incseq_SUP:
1792 fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
1793 assumes u: "bdd_above (range X)"
1794 assumes X: "incseq X"
1795 shows "X ----> (SUP i. X i)"
1796 by (rule order_tendstoI)
1797 (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u])
1799 lemma LIMSEQ_decseq_INF:
1800 fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
1801 assumes u: "bdd_below (range X)"
1802 assumes X: "decseq X"
1803 shows "X ----> (INF i. X i)"
1804 by (rule order_tendstoI)
1805 (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
1807 text\<open>Main monotonicity theorem\<close>
1809 lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
1810 by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below)
1812 lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
1813 by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def)
1815 lemma monoseq_imp_convergent_iff_Bseq: "monoseq (f :: nat \<Rightarrow> real) \<Longrightarrow> convergent f \<longleftrightarrow> Bseq f"
1816 using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast
1818 lemma Bseq_monoseq_convergent'_inc:
1819 "Bseq (\<lambda>n. f (n + M) :: real) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<le> f n) \<Longrightarrow> convergent f"
1820 by (subst convergent_ignore_initial_segment [symmetric, of _ M])
1821 (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
1823 lemma Bseq_monoseq_convergent'_dec:
1824 "Bseq (\<lambda>n. f (n + M) :: real) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<ge> f n) \<Longrightarrow> convergent f"
1825 by (subst convergent_ignore_initial_segment [symmetric, of _ M])
1826 (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
1829 fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
1830 shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
1831 unfolding Cauchy_def dist_norm ..
1834 fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
1835 shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
1836 by (simp add: Cauchy_iff)
1839 fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
1840 shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
1841 by (simp add: Cauchy_iff)
1843 lemma incseq_convergent:
1844 fixes X :: "nat \<Rightarrow> real"
1845 assumes "incseq X" and "\<forall>i. X i \<le> B"
1846 obtains L where "X ----> L" "\<forall>i. X i \<le> L"
1848 from incseq_bounded[OF assms] \<open>incseq X\<close> Bseq_monoseq_convergent[of X]
1849 obtain L where "X ----> L"
1850 by (auto simp: convergent_def monoseq_def incseq_def)
1851 with \<open>incseq X\<close> show "\<exists>L. X ----> L \<and> (\<forall>i. X i \<le> L)"
1852 by (auto intro!: exI[of _ L] incseq_le)
1855 lemma decseq_convergent:
1856 fixes X :: "nat \<Rightarrow> real"
1857 assumes "decseq X" and "\<forall>i. B \<le> X i"
1858 obtains L where "X ----> L" "\<forall>i. L \<le> X i"
1860 from decseq_bounded[OF assms] \<open>decseq X\<close> Bseq_monoseq_convergent[of X]
1861 obtain L where "X ----> L"
1862 by (auto simp: convergent_def monoseq_def decseq_def)
1863 with \<open>decseq X\<close> show "\<exists>L. X ----> L \<and> (\<forall>i. L \<le> X i)"
1864 by (auto intro!: exI[of _ L] decseq_le)
1867 subsubsection \<open>Cauchy Sequences are Bounded\<close>
1869 text\<open>A Cauchy sequence is bounded -- this is the standard
1870 proof mechanization rather than the nonstandard proof\<close>
1872 lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
1873 ==> \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
1874 apply (clarify, drule spec, drule (1) mp)
1875 apply (simp only: norm_minus_commute)
1876 apply (drule order_le_less_trans [OF norm_triangle_ineq2])
1880 subsection \<open>Power Sequences\<close>
1882 text\<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
1883 "x<1"}. Proof will use (NS) Cauchy equivalence for convergence and
1884 also fact that bounded and monotonic sequence converges.\<close>
1886 lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
1887 apply (simp add: Bseq_def)
1888 apply (rule_tac x = 1 in exI)
1889 apply (simp add: power_abs)
1890 apply (auto dest: power_mono)
1893 lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
1894 apply (clarify intro!: mono_SucI2)
1895 apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
1898 lemma convergent_realpow:
1899 "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
1900 by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
1902 lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
1903 by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp
1905 lemma LIMSEQ_realpow_zero:
1906 "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
1908 assume "0 \<le> x" and "x \<noteq> 0"
1909 hence x0: "0 < x" by simp
1911 from x0 x1 have "1 < inverse x"
1912 by (rule one_less_inverse)
1913 hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
1914 by (rule LIMSEQ_inverse_realpow_zero)
1915 thus ?thesis by (simp add: power_inverse)
1916 qed (rule LIMSEQ_imp_Suc, simp)
1918 lemma LIMSEQ_power_zero:
1919 fixes x :: "'a::{real_normed_algebra_1}"
1920 shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
1921 apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
1922 apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
1923 apply (simp add: power_abs norm_power_ineq)
1926 lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) ----> 0"
1927 by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp
1929 text\<open>Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}\<close>
1931 lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) ----> 0"
1932 by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
1934 lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) ----> 0"
1935 by (rule LIMSEQ_power_zero) simp
1938 subsection \<open>Limits of Functions\<close>
1941 fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
1942 shows "f -- a --> L =
1943 (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
1944 by (simp add: LIM_def dist_norm)
1947 fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
1948 shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
1950 by (simp add: LIM_eq)
1953 fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
1954 shows "[| f -- a --> L; 0<r |]
1955 ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
1956 by (simp add: LIM_eq)
1959 fixes a :: "'a::real_normed_vector"
1960 shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"
1961 unfolding filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap by simp
1963 lemma LIM_offset_zero:
1964 fixes a :: "'a::real_normed_vector"
1965 shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
1966 by (drule_tac k="a" in LIM_offset, simp add: add.commute)
1968 lemma LIM_offset_zero_cancel:
1969 fixes a :: "'a::real_normed_vector"
1970 shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"
1971 by (drule_tac k="- a" in LIM_offset, simp)
1973 lemma LIM_offset_zero_iff:
1974 fixes f :: "'a :: real_normed_vector \<Rightarrow> _"
1975 shows "f -- a --> L \<longleftrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
1976 using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto
1979 fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
1980 shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"
1981 unfolding tendsto_iff dist_norm by simp
1983 lemma LIM_zero_cancel:
1984 fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
1985 shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"
1986 unfolding tendsto_iff dist_norm by simp
1989 fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
1990 shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"
1991 unfolding tendsto_iff dist_norm by simp
1994 fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
1995 fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
1996 assumes f: "f -- a --> l"
1997 assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
1998 shows "g -- a --> m"
1999 by (rule metric_LIM_imp_LIM [OF f],
2000 simp add: dist_norm le)
2003 fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
2005 assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
2006 shows "g -- a --> l \<Longrightarrow> f -- a --> l"
2007 by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
2010 fixes a :: "'a::real_normed_vector"
2011 assumes f: "f -- a --> b"
2012 assumes g: "g -- b --> c"
2013 assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
2014 shows "(\<lambda>x. g (f x)) -- a --> c"
2015 by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
2017 lemma real_LIM_sandwich_zero:
2018 fixes f g :: "'a::topological_space \<Rightarrow> real"
2019 assumes f: "f -- a --> 0"
2020 assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
2021 assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
2022 shows "g -- a --> 0"
2023 proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
2024 fix x assume x: "x \<noteq> a"
2025 have "norm (g x - 0) = g x" by (simp add: 1 x)
2026 also have "g x \<le> f x" by (rule 2 [OF x])
2027 also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)
2028 also have "\<bar>f x\<bar> = norm (f x - 0)" by simp
2029 finally show "norm (g x - 0) \<le> norm (f x - 0)" .
2033 subsection \<open>Continuity\<close>
2035 lemma LIM_isCont_iff:
2036 fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
2037 shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"
2038 by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
2041 fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
2042 shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"
2043 by (simp add: isCont_def LIM_isCont_iff)
2045 lemma isCont_LIM_compose2:
2046 fixes a :: "'a::real_normed_vector"
2047 assumes f [unfolded isCont_def]: "isCont f a"
2048 assumes g: "g -- f a --> l"
2049 assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
2050 shows "(\<lambda>x. g (f x)) -- a --> l"
2051 by (rule LIM_compose2 [OF f g inj])
2054 lemma isCont_norm [simp]:
2055 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
2056 shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
2057 by (fact continuous_norm)
2059 lemma isCont_rabs [simp]:
2060 fixes f :: "'a::t2_space \<Rightarrow> real"
2061 shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
2062 by (fact continuous_rabs)
2064 lemma isCont_add [simp]:
2065 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
2066 shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
2067 by (fact continuous_add)
2069 lemma isCont_minus [simp]:
2070 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
2071 shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
2072 by (fact continuous_minus)
2074 lemma isCont_diff [simp]:
2075 fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
2076 shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
2077 by (fact continuous_diff)
2079 lemma isCont_mult [simp]:
2080 fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
2081 shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
2082 by (fact continuous_mult)
2084 lemma (in bounded_linear) isCont:
2085 "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
2086 by (fact continuous)
2088 lemma (in bounded_bilinear) isCont:
2089 "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
2090 by (fact continuous)
2092 lemmas isCont_scaleR [simp] =
2093 bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
2095 lemmas isCont_of_real [simp] =
2096 bounded_linear.isCont [OF bounded_linear_of_real]
2098 lemma isCont_power [simp]:
2099 fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
2100 shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
2101 by (fact continuous_power)
2103 lemma isCont_setsum [simp]:
2104 fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
2105 shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
2106 by (auto intro: continuous_setsum)
2108 subsection \<open>Uniform Continuity\<close>
2111 isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
2112 "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
2114 lemma isUCont_isCont: "isUCont f ==> isCont f x"
2115 by (simp add: isUCont_def isCont_def LIM_def, force)
2117 lemma isUCont_Cauchy:
2118 "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
2119 unfolding isUCont_def
2120 apply (rule metric_CauchyI)
2121 apply (drule_tac x=e in spec, safe)
2122 apply (drule_tac e=s in metric_CauchyD, safe)
2123 apply (rule_tac x=M in exI, simp)
2126 lemma (in bounded_linear) isUCont: "isUCont f"
2127 unfolding isUCont_def dist_norm
2128 proof (intro allI impI)
2129 fix r::real assume r: "0 < r"
2130 obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
2131 using pos_bounded by blast
2132 show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
2133 proof (rule exI, safe)
2134 from r K show "0 < r / K" by simp
2137 assume xy: "norm (x - y) < r / K"
2138 have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff)
2139 also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le)
2140 also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq)
2141 finally show "norm (f x - f y) < r" .
2145 lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
2146 by (rule isUCont [THEN isUCont_Cauchy])
2148 lemma LIM_less_bound:
2149 fixes f :: "real \<Rightarrow> real"
2150 assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
2152 proof (rule tendsto_le_const)
2153 show "(f ---> f x) (at_left x)"
2154 using \<open>isCont f x\<close> by (simp add: filterlim_at_split isCont_def)
2155 show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
2156 using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"])
2160 subsection \<open>Nested Intervals and Bisection -- Needed for Compactness\<close>
2162 lemma nested_sequence_unique:
2163 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) ----> 0"
2164 shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f ----> l) \<and> ((\<forall>n. l \<le> g n) \<and> g ----> l)"
2166 have "incseq f" unfolding incseq_Suc_iff by fact
2167 have "decseq g" unfolding decseq_Suc_iff by fact
2170 from \<open>decseq g\<close> have "g n \<le> g 0" by (rule decseqD) simp
2171 with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f n \<le> g 0" by auto }
2172 then obtain u where "f ----> u" "\<forall>i. f i \<le> u"
2173 using incseq_convergent[OF \<open>incseq f\<close>] by auto
2176 from \<open>incseq f\<close> have "f 0 \<le> f n" by (rule incseqD) simp
2177 with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f 0 \<le> g n" by simp }
2178 then obtain l where "g ----> l" "\<forall>i. l \<le> g i"
2179 using decseq_convergent[OF \<open>decseq g\<close>] by auto
2180 moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \<open>f ----> u\<close> \<open>g ----> l\<close>]]
2181 ultimately show ?thesis by auto
2184 lemma Bolzano[consumes 1, case_names trans local]:
2185 fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
2186 assumes [arith]: "a \<le> b"
2187 assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
2188 assumes 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"
2191 def bisect \<equiv> "rec_nat (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
2192 def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
2193 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)"
2194 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)"
2195 by (simp_all add: l_def u_def bisect_def split: prod.split)
2197 { fix n have "l n \<le> u n" by (induct n) auto } note this[simp]
2199 have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l ----> x) \<and> ((\<forall>n. x \<le> u n) \<and> u ----> x)"
2200 proof (safe intro!: nested_sequence_unique)
2201 fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) auto
2203 { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
2204 then show "(\<lambda>n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero)
2206 then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l ----> x" "u ----> x" by auto
2207 obtain d where "0 < d" and d: "\<And>a b. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b"
2208 using \<open>l 0 \<le> x\<close> \<open>x \<le> u 0\<close> local[of x] by auto
2212 assume "\<not> P a b"
2213 { fix n have "\<not> P (l n) (u n)"
2215 case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
2216 qed (simp add: \<open>\<not> P a b\<close>) }
2218 { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
2219 using \<open>0 < d\<close> \<open>l ----> x\<close> by (intro order_tendstoD[of _ x]) auto
2220 moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
2221 using \<open>0 < d\<close> \<open>u ----> x\<close> by (intro order_tendstoD[of _ x]) auto
2222 ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
2223 proof eventually_elim
2224 fix n assume "x - d / 2 < l n" "u n < x + d / 2"
2225 from add_strict_mono[OF this] have "u n - l n < d" by simp
2226 with x show "P (l n) (u n)" by (rule d)
2228 ultimately show False by simp
2232 lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
2233 proof (cases "a \<le> b", rule compactI)
2234 fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
2236 from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
2237 proof (induct rule: Bolzano)
2239 then have *: "{a .. c} = {a .. b} \<union> {b .. c}" by auto
2240 from trans obtain C1 C2 where "C1\<subseteq>C \<and> finite C1 \<and> {a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C \<and> finite C2 \<and> {b..c} \<subseteq> \<Union>C2"
2242 with trans show ?case
2243 unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto
2246 then have "x \<in> \<Union>C" using C by auto
2247 with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
2248 then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
2249 by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff)
2250 with \<open>c \<in> C\<close> show ?case
2251 by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
2256 lemma continuous_image_closed_interval:
2257 fixes a b and f :: "real \<Rightarrow> real"
2258 defines "S \<equiv> {a..b}"
2259 assumes "a \<le> b" and f: "continuous_on S f"
2260 shows "\<exists>c d. f`S = {c..d} \<and> c \<le> d"
2262 have S: "compact S" "S \<noteq> {}"
2263 using \<open>a \<le> b\<close> by (auto simp: S_def)
2264 obtain c where "c \<in> S" "\<forall>d\<in>S. f d \<le> f c"
2265 using continuous_attains_sup[OF S f] by auto
2266 moreover obtain d where "d \<in> S" "\<forall>c\<in>S. f d \<le> f c"
2267 using continuous_attains_inf[OF S f] by auto
2268 moreover have "connected (f`S)"
2269 using connected_continuous_image[OF f] connected_Icc by (auto simp: S_def)
2270 ultimately have "f ` S = {f d .. f c} \<and> f d \<le> f c"
2271 by (auto simp: connected_iff_interval)
2276 lemma open_Collect_positive:
2277 fixes f :: "'a::t2_space \<Rightarrow> real"
2278 assumes f: "continuous_on s f"
2279 shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < f x}"
2280 using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"]
2281 by (auto simp: Int_def field_simps)
2283 lemma open_Collect_less_Int:
2284 fixes f g :: "'a::t2_space \<Rightarrow> real"
2285 assumes f: "continuous_on s f" and g: "continuous_on s g"
2286 shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. f x < g x}"
2287 using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps)
2290 subsection \<open>Boundedness of continuous functions\<close>
2292 text\<open>By bisection, function continuous on closed interval is bounded above\<close>
2295 fixes f :: "real \<Rightarrow> 'a::linorder_topology"
2296 shows "a \<le> b \<Longrightarrow> \<forall>x::real. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
2297 \<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)"
2298 using continuous_attains_sup[of "{a .. b}" f]
2299 by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
2302 fixes f :: "real \<Rightarrow> 'a::linorder_topology"
2303 shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
2304 \<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)"
2305 using continuous_attains_inf[of "{a .. b}" f]
2306 by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
2308 lemma isCont_bounded:
2309 fixes f :: "real \<Rightarrow> 'a::linorder_topology"
2310 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"
2311 using isCont_eq_Ub[of a b f] by auto
2313 lemma isCont_has_Ub:
2314 fixes f :: "real \<Rightarrow> 'a::linorder_topology"
2315 shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
2316 \<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))"
2317 using isCont_eq_Ub[of a b f] by auto
2319 (*HOL style here: object-level formulations*)
2320 lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
2321 (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
2322 --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
2323 by (blast intro: IVT)
2325 lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
2326 (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
2327 --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
2328 by (blast intro: IVT2)
2331 fixes f :: "real \<Rightarrow> real"
2332 assumes "a \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
2333 shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and>
2334 (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))"
2336 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"
2337 using isCont_eq_Ub[OF assms] by auto
2338 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"
2339 using isCont_eq_Lb[OF assms] by auto
2341 using IVT[of f L _ M] IVT2[of f L _ M] M L assms
2342 apply (rule_tac x="f L" in exI)
2343 apply (rule_tac x="f M" in exI)
2344 apply (cases "L \<le> M")
2345 apply (simp, metis order_trans)
2346 apply (simp, metis order_trans)
2351 text\<open>Continuity of inverse function\<close>
2353 lemma isCont_inverse_function:
2354 fixes f g :: "real \<Rightarrow> real"
2356 and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
2357 and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
2358 shows "isCont g (f x)"
2360 let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}"
2362 have f: "continuous_on ?D f"
2363 using cont by (intro continuous_at_imp_continuous_on ballI) auto
2364 then have g: "continuous_on (f`?D) g"
2365 using inj by (intro continuous_on_inv) auto
2367 from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D"
2368 by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max)
2369 with g have "continuous_on {min ?A ?B <..< max ?A ?B} g"
2370 by (rule continuous_on_subset)
2372 have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)"
2373 using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto
2374 then have "f x \<in> {min ?A ?B <..< max ?A ?B}"
2378 by (simp add: continuous_on_eq_continuous_at)
2381 lemma isCont_inverse_function2:
2382 fixes f g :: "real \<Rightarrow> real" shows
2383 "\<lbrakk>a < x; x < b;
2384 \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
2385 \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
2386 \<Longrightarrow> isCont g (f x)"
2387 apply (rule isCont_inverse_function
2388 [where f=f and d="min (x - a) (b - x)"])
2389 apply (simp_all add: abs_le_iff)
2392 (* need to rename second isCont_inverse *)
2394 lemma isCont_inv_fun:
2395 fixes f g :: "real \<Rightarrow> real"
2396 shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
2397 \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]
2399 by (rule isCont_inverse_function)
2401 text\<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110\<close>
2402 lemma LIM_fun_gt_zero:
2403 fixes f :: "real \<Rightarrow> real"
2404 shows "f -- c --> 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)"
2405 apply (drule (1) LIM_D, clarify)
2406 apply (rule_tac x = s in exI)
2407 apply (simp add: abs_less_iff)
2410 lemma LIM_fun_less_zero:
2411 fixes f :: "real \<Rightarrow> real"
2412 shows "f -- c --> 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)"
2413 apply (drule LIM_D [where r="-l"], simp, clarify)
2414 apply (rule_tac x = s in exI)
2415 apply (simp add: abs_less_iff)
2418 lemma LIM_fun_not_zero:
2419 fixes f :: "real \<Rightarrow> real"
2420 shows "f -- c --> 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)"
2421 using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff)