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