(* Title: HOL/Limits.thy Author: Brian Huffman Author: Jacques D. Fleuriot, University of Cambridge Author: Lawrence C Paulson Author: Jeremy Avigad *) header {* Limits on Real Vector Spaces *} theory Limits imports Real_Vector_Spaces begin subsection {* Filter going to infinity norm *} definition at_infinity :: "'a::real_normed_vector filter" where "at_infinity = Abs_filter (\P. \r. \x. r \ norm x \ P x)" lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def proof (rule eventually_Abs_filter, rule is_filter.intro) fix P Q :: "'a \ bool" assume "\r. \x. r \ norm x \ P x" and "\s. \x. s \ norm x \ Q x" then obtain r s where "\x. r \ norm x \ P x" and "\x. s \ norm x \ Q x" by auto then have "\x. max r s \ norm x \ P x \ Q x" by simp then show "\r. \x. r \ norm x \ P x \ Q x" .. qed auto lemma at_infinity_eq_at_top_bot: "(at_infinity \ real filter) = sup at_top at_bot" unfolding sup_filter_def at_infinity_def eventually_at_top_linorder eventually_at_bot_linorder proof (intro arg_cong[where f=Abs_filter] ext iffI) fix P :: "real \ bool" assume "\r. \x. r \ norm x \ P x" then guess r .. then have "(\x\r. P x) \ (\x\-r. P x)" by auto then show "(\r. \x\r. P x) \ (\r. \x\r. P x)" by auto next fix P :: "real \ bool" assume "(\r. \x\r. P x) \ (\r. \x\r. P x)" then obtain p q where "\x\p. P x" "\x\q. P x" by auto then show "\r. \x. r \ norm x \ P x" by (intro exI[of _ "max p (-q)"]) (auto simp: abs_real_def) qed lemma at_top_le_at_infinity: "at_top \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma at_bot_le_at_infinity: "at_bot \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp subsubsection {* Boundedness *} definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where "Bseq X \ Bfun X sequentially" lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" unfolding Bfun_metric_def by (subst eventually_sequentially_seg) lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) lemma Bfun_def: "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" unfolding Bfun_metric_def norm_conv_dist proof safe fix y K assume "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" moreover have "eventually (\x. dist (f x) 0 \ dist (f x) y + dist 0 y) F" by (intro always_eventually) (metis dist_commute dist_triangle) with * have "eventually (\x. dist (f x) 0 \ K + dist 0 y) F" by eventually_elim auto with `0 < K` show "\K>0. eventually (\x. dist (f x) 0 \ K) F" by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto qed auto lemma BfunI: assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp next show "eventually (\x. norm (f x) \ max K 1) F" using K by (rule eventually_elim1, simp) qed lemma BfunE: assumes "Bfun f F" obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by fast lemma Cauchy_Bseq: "Cauchy X \ Bseq X" unfolding Cauchy_def Bfun_metric_def eventually_sequentially apply (erule_tac x=1 in allE) apply simp apply safe apply (rule_tac x="X M" in exI) apply (rule_tac x=1 in exI) apply (erule_tac x=M in allE) apply simp apply (rule_tac x=M in exI) apply (auto simp: dist_commute) done subsubsection {* Bounded Sequences *} lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" by (intro BfunI) (auto simp: eventually_sequentially) lemma BseqI2': "\n\N. norm (X n) \ K \ Bseq X" by (intro BfunI) (auto simp: eventually_sequentially) lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" unfolding Bfun_def eventually_sequentially proof safe fix N K assume "0 < K" "\n\N. norm (X n) \ K" then show "\K>0. \n. norm (X n) \ K" by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2) (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) qed auto lemma BseqE: "\Bseq X; \K. \0 < K; \n. norm (X n) \ K\ \ Q\ \ Q" unfolding Bseq_def by auto lemma BseqD: "Bseq X ==> \K. 0 < K & (\n. norm (X n) \ K)" by (simp add: Bseq_def) lemma BseqI: "[| 0 < K; \n. norm (X n) \ K |] ==> Bseq X" by (auto simp add: Bseq_def) lemma lemma_NBseq_def: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) \ real(Suc N))" proof safe fix K :: real from reals_Archimedean2 obtain n :: nat where "K < real n" .. then have "K \ real (Suc n)" by auto moreover assume "\m. norm (X m) \ K" ultimately have "\m. norm (X m) \ real (Suc n)" by (blast intro: order_trans) then show "\N. \n. norm (X n) \ real (Suc N)" .. qed (force simp add: real_of_nat_Suc) text{* alternative definition for Bseq *} lemma Bseq_iff: "Bseq X = (\N. \n. norm (X n) \ real(Suc N))" apply (simp add: Bseq_def) apply (simp (no_asm) add: lemma_NBseq_def) done lemma lemma_NBseq_def2: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" apply (subst lemma_NBseq_def, auto) apply (rule_tac x = "Suc N" in exI) apply (rule_tac [2] x = N in exI) apply (auto simp add: real_of_nat_Suc) prefer 2 apply (blast intro: order_less_imp_le) apply (drule_tac x = n in spec, simp) done (* yet another definition for Bseq *) lemma Bseq_iff1a: "Bseq X = (\N. \n. norm (X n) < real(Suc N))" by (simp add: Bseq_def lemma_NBseq_def2) subsubsection{*A Few More Equivalence Theorems for Boundedness*} text{*alternative formulation for boundedness*} lemma Bseq_iff2: "Bseq X = (\k > 0. \x. \n. norm (X(n) + -x) \ k)" apply (unfold Bseq_def, safe) apply (rule_tac [2] x = "k + norm x" in exI) apply (rule_tac x = K in exI, simp) apply (rule exI [where x = 0], auto) apply (erule order_less_le_trans, simp) apply (drule_tac x=n in spec, fold diff_minus) apply (drule order_trans [OF norm_triangle_ineq2]) apply simp done text{*alternative formulation for boundedness*} lemma Bseq_iff3: "Bseq X = (\k > 0. \N. \n. norm(X(n) + -X(N)) \ k)" apply safe apply (simp add: Bseq_def, safe) apply (rule_tac x = "K + norm (X N)" in exI) apply auto apply (erule order_less_le_trans, simp) apply (rule_tac x = N in exI, safe) apply (drule_tac x = n in spec) apply (rule order_trans [OF norm_triangle_ineq], simp) apply (auto simp add: Bseq_iff2) done lemma BseqI2: "(\n. k \ f n & f n \ (K::real)) ==> Bseq f" apply (simp add: Bseq_def) apply (rule_tac x = " (\k\ + \K\) + 1" in exI, auto) apply (drule_tac x = n in spec, arith) done subsubsection{*Upper Bounds and Lubs of Bounded Sequences*} lemma Bseq_isUb: "!!(X::nat=>real). Bseq X ==> \U. isUb (UNIV::real set) {x. \n. X n = x} U" by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff) text{* Use completeness of reals (supremum property) to show that any bounded sequence has a least upper bound*} lemma Bseq_isLub: "!!(X::nat=>real). Bseq X ==> \U. isLub (UNIV::real set) {x. \n. X n = x} U" by (blast intro: reals_complete Bseq_isUb) lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" by (simp add: Bseq_def) lemma Bseq_eq_bounded: "range f \ {a .. b::real} \ Bseq f" apply (simp add: subset_eq) apply (rule BseqI'[where K="max (norm a) (norm b)"]) apply (erule_tac x=n in allE) apply auto done lemma incseq_bounded: "incseq X \ \i. X i \ (B::real) \ Bseq X" by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) lemma decseq_bounded: "decseq X \ \i. (B::real) \ X i \ Bseq X" by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) subsection {* Bounded Monotonic Sequences *} subsubsection{*A Bounded and Monotonic Sequence Converges*} (* TODO: delete *) (* FIXME: one use in NSA/HSEQ.thy *) lemma Bmonoseq_LIMSEQ: "\n. m \ n --> X n = X m ==> \L. (X ----> L)" apply (rule_tac x="X m" in exI) apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const]) unfolding eventually_sequentially apply blast done subsection {* Convergence to Zero *} definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" where "Zfun f F = (\r>0. eventually (\x. norm (f x) < r) F)" lemma ZfunI: "(\r. 0 < r \ eventually (\x. norm (f x) < r) F) \ Zfun f F" unfolding Zfun_def by simp lemma ZfunD: "\Zfun f F; 0 < r\ \ eventually (\x. norm (f x) < r) F" unfolding Zfun_def by simp lemma Zfun_ssubst: "eventually (\x. f x = g x) F \ Zfun g F \ Zfun f F" unfolding Zfun_def by (auto elim!: eventually_rev_mp) lemma Zfun_zero: "Zfun (\x. 0) F" unfolding Zfun_def by simp lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) F = Zfun (\x. f x) F" unfolding Zfun_def by simp lemma Zfun_imp_Zfun: assumes f: "Zfun f F" assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) F" shows "Zfun (\x. g x) F" proof (cases) assume K: "0 < K" show ?thesis proof (rule ZfunI) fix r::real assume "0 < r" hence "0 < r / K" using K by (rule divide_pos_pos) then have "eventually (\x. norm (f x) < r / K) F" using ZfunD [OF f] by fast with g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) hence "norm (f x) * K < r" by (simp add: pos_less_divide_eq K) thus ?case by (simp add: order_le_less_trans [OF elim(1)]) qed qed next assume "\ 0 < K" hence K: "K \ 0" by (simp only: not_less) show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" from g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) also have "norm (f x) * K \ norm (f x) * 0" using K norm_ge_zero by (rule mult_left_mono) finally show ?case using `0 < r` by simp qed qed qed lemma Zfun_le: "\Zfun g F; \x. norm (f x) \ norm (g x)\ \ Zfun f F" by (erule_tac K="1" in Zfun_imp_Zfun, simp) lemma Zfun_add: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x + g x) F" proof (rule ZfunI) fix r::real assume "0 < r" hence r: "0 < r / 2" by simp have "eventually (\x. norm (f x) < r/2) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < r/2) F" using g r by (rule ZfunD) ultimately show "eventually (\x. norm (f x + g x) < r) F" proof eventually_elim case (elim x) have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) also have "\ < r/2 + r/2" using elim by (rule add_strict_mono) finally show ?case by simp qed qed lemma Zfun_minus: "Zfun f F \ Zfun (\x. - f x) F" unfolding Zfun_def by simp lemma Zfun_diff: "\Zfun f F; Zfun g F\ \ Zfun (\x. f x - g x) F" by (simp only: diff_minus Zfun_add Zfun_minus) lemma (in bounded_linear) Zfun: assumes g: "Zfun g F" shows "Zfun (\x. f (g x)) F" proof - obtain K where "\x. norm (f x) \ norm x * K" using bounded by fast then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) F" by simp with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: assumes f: "Zfun f F" assumes g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" proof (rule ZfunI) fix r::real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using pos_bounded by fast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) have "eventually (\x. norm (f x) < r) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < inverse K) F" using g K' by (rule ZfunD) ultimately show "eventually (\x. norm (f x ** g x) < r) F" proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "norm (f x) * norm (g x) * K < r * inverse K * K" by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) also from K have "r * inverse K * K = r" by simp finally show ?case . qed qed lemma (in bounded_bilinear) Zfun_left: "Zfun f F \ Zfun (\x. f x ** a) F" by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: "Zfun f F \ Zfun (\x. a ** f x) F" by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) subsubsection {* Distance and norms *} lemma tendsto_dist [tendsto_intros]: fixes l m :: "'a :: metric_space" assumes f: "(f ---> l) F" and g: "(g ---> m) F" shows "((\x. dist (f x) (g x)) ---> dist l m) F" proof (rule tendstoI) fix e :: real assume "0 < e" hence e2: "0 < e/2" by simp from tendstoD [OF f e2] tendstoD [OF g e2] show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" proof (eventually_elim) case (elim x) then show "dist (dist (f x) (g x)) (dist l m) < e" unfolding dist_real_def using dist_triangle2 [of "f x" "g x" "l"] using dist_triangle2 [of "g x" "l" "m"] using dist_triangle3 [of "l" "m" "f x"] using dist_triangle [of "f x" "m" "g x"] by arith qed qed lemma continuous_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" unfolding continuous_def by (rule tendsto_dist) lemma continuous_on_dist[continuous_on_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" unfolding continuous_on_def by (auto intro: tendsto_dist) lemma tendsto_norm [tendsto_intros]: "(f ---> a) F \ ((\x. norm (f x)) ---> norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) lemma continuous_norm [continuous_intros]: "continuous F f \ continuous F (\x. norm (f x))" unfolding continuous_def by (rule tendsto_norm) lemma continuous_on_norm [continuous_on_intros]: "continuous_on s f \ continuous_on s (\x. norm (f x))" unfolding continuous_on_def by (auto intro: tendsto_norm) lemma tendsto_norm_zero: "(f ---> 0) F \ ((\x. norm (f x)) ---> 0) F" by (drule tendsto_norm, simp) lemma tendsto_norm_zero_cancel: "((\x. norm (f x)) ---> 0) F \ (f ---> 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_norm_zero_iff: "((\x. norm (f x)) ---> 0) F \ (f ---> 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_rabs [tendsto_intros]: "(f ---> (l::real)) F \ ((\x. \f x\) ---> \l\) F" by (fold real_norm_def, rule tendsto_norm) lemma continuous_rabs [continuous_intros]: "continuous F f \ continuous F (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_norm) lemma continuous_on_rabs [continuous_on_intros]: "continuous_on s f \ continuous_on s (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_on_norm) lemma tendsto_rabs_zero: "(f ---> (0::real)) F \ ((\x. \f x\) ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero) lemma tendsto_rabs_zero_cancel: "((\x. \f x\) ---> (0::real)) F \ (f ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero_cancel) lemma tendsto_rabs_zero_iff: "((\x. \f x\) ---> (0::real)) F \ (f ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero_iff) subsubsection {* Addition and subtraction *} lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::real_normed_vector" shows "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x + g x) ---> a + b) F" by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) lemma continuous_add [continuous_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_vector" shows "continuous F f \ continuous F g \ continuous F (\x. f x + g x)" unfolding continuous_def by (rule tendsto_add) lemma continuous_on_add [continuous_on_intros]: fixes f g :: "_ \ 'b::real_normed_vector" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x + g x)" unfolding continuous_on_def by (auto intro: tendsto_add) lemma tendsto_add_zero: fixes f g :: "_ \ 'b::real_normed_vector" shows "\(f ---> 0) F; (g ---> 0) F\ \ ((\x. f x + g x) ---> 0) F" by (drule (1) tendsto_add, simp) lemma tendsto_minus [tendsto_intros]: fixes a :: "'a::real_normed_vector" shows "(f ---> a) F \ ((\x. - f x) ---> - a) F" by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) lemma continuous_minus [continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" shows "continuous F f \ continuous F (\x. - f x)" unfolding continuous_def by (rule tendsto_minus) lemma continuous_on_minus [continuous_on_intros]: fixes f :: "_ \ 'b::real_normed_vector" shows "continuous_on s f \ continuous_on s (\x. - f x)" unfolding continuous_on_def by (auto intro: tendsto_minus) lemma tendsto_minus_cancel: fixes a :: "'a::real_normed_vector" shows "((\x. - f x) ---> - a) F \ (f ---> a) F" by (drule tendsto_minus, simp) lemma tendsto_minus_cancel_left: "(f ---> - (y::_::real_normed_vector)) F \ ((\x. - f x) ---> y) F" using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] by auto lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::real_normed_vector" shows "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x - g x) ---> a - b) F" by (simp add: diff_minus tendsto_add tendsto_minus) lemma continuous_diff [continuous_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_vector" shows "continuous F f \ continuous F g \ continuous F (\x. f x - g x)" unfolding continuous_def by (rule tendsto_diff) lemma continuous_on_diff [continuous_on_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_vector" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x - g x)" unfolding continuous_on_def by (auto intro: tendsto_diff) lemma tendsto_setsum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::real_normed_vector" assumes "\i. i \ S \ (f i ---> a i) F" shows "((\x. \i\S. f i x) ---> (\i\S. a i)) F" proof (cases "finite S") assume "finite S" thus ?thesis using assms by (induct, simp add: tendsto_const, simp add: tendsto_add) next assume "\ finite S" thus ?thesis by (simp add: tendsto_const) qed lemma continuous_setsum [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::real_normed_vector" shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" unfolding continuous_def by (rule tendsto_setsum) lemma continuous_on_setsum [continuous_intros]: fixes f :: "'a \ _ \ 'c::real_normed_vector" shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" unfolding continuous_on_def by (auto intro: tendsto_setsum) lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real] subsubsection {* Linear operators and multiplication *} lemma (in bounded_linear) tendsto: "(g ---> a) F \ ((\x. f (g x)) ---> f a) F" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_linear) continuous: "continuous F g \ continuous F (\x. f (g x))" using tendsto[of g _ F] by (auto simp: continuous_def) lemma (in bounded_linear) continuous_on: "continuous_on s g \ continuous_on s (\x. f (g x))" using tendsto[of g] by (auto simp: continuous_on_def) lemma (in bounded_linear) tendsto_zero: "(g ---> 0) F \ ((\x. f (g x)) ---> 0) F" by (drule tendsto, simp only: zero) lemma (in bounded_bilinear) tendsto: "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x ** g x) ---> a ** b) F" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) lemma (in bounded_bilinear) continuous: "continuous F f \ continuous F g \ continuous F (\x. f x ** g x)" using tendsto[of f _ F g] by (auto simp: continuous_def) lemma (in bounded_bilinear) continuous_on: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x ** g x)" using tendsto[of f _ _ g] by (auto simp: continuous_on_def) lemma (in bounded_bilinear) tendsto_zero: assumes f: "(f ---> 0) F" assumes g: "(g ---> 0) F" shows "((\x. f x ** g x) ---> 0) F" using tendsto [OF f g] by (simp add: zero_left) lemma (in bounded_bilinear) tendsto_left_zero: "(f ---> 0) F \ ((\x. f x ** c) ---> 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) lemma (in bounded_bilinear) tendsto_right_zero: "(f ---> 0) F \ ((\x. c ** f x) ---> 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) lemmas tendsto_of_real [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_of_real] lemmas tendsto_scaleR [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] lemmas tendsto_mult [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_mult] lemmas continuous_of_real [continuous_intros] = bounded_linear.continuous [OF bounded_linear_of_real] lemmas continuous_scaleR [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_scaleR] lemmas continuous_mult [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_mult] lemmas continuous_on_of_real [continuous_on_intros] = bounded_linear.continuous_on [OF bounded_linear_of_real] lemmas continuous_on_scaleR [continuous_on_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] lemmas continuous_on_mult [continuous_on_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_mult] lemmas tendsto_mult_zero = bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_left_zero = bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_right_zero = bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] lemma tendsto_power [tendsto_intros]: fixes f :: "'a \ 'b::{power,real_normed_algebra}" shows "(f ---> a) F \ ((\x. f x ^ n) ---> a ^ n) F" by (induct n) (simp_all add: tendsto_const tendsto_mult) lemma continuous_power [continuous_intros]: fixes f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" shows "continuous F f \ continuous F (\x. (f x)^n)" unfolding continuous_def by (rule tendsto_power) lemma continuous_on_power [continuous_on_intros]: fixes f :: "_ \ 'b::{power,real_normed_algebra}" shows "continuous_on s f \ continuous_on s (\x. (f x)^n)" unfolding continuous_on_def by (auto intro: tendsto_power) lemma tendsto_setprod [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::{real_normed_algebra,comm_ring_1}" assumes "\i. i \ S \ (f i ---> L i) F" shows "((\x. \i\S. f i x) ---> (\i\S. L i)) F" proof (cases "finite S") assume "finite S" thus ?thesis using assms by (induct, simp add: tendsto_const, simp add: tendsto_mult) next assume "\ finite S" thus ?thesis by (simp add: tendsto_const) qed lemma continuous_setprod [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" unfolding continuous_def by (rule tendsto_setprod) lemma continuous_on_setprod [continuous_intros]: fixes f :: "'a \ _ \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" unfolding continuous_on_def by (auto intro: tendsto_setprod) subsubsection {* Inverse and division *} lemma (in bounded_bilinear) Zfun_prod_Bfun: assumes f: "Zfun f F" assumes g: "Bfun g F" shows "Zfun (\x. f x ** g x) F" proof - obtain K where K: "0 \ K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using nonneg_bounded by fast obtain B where B: "0 < B" and norm_g: "eventually (\x. norm (g x) \ B) F" using g by (rule BfunE) have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) F" using norm_g proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "\ \ norm (f x) * B * K" by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) also have "\ = norm (f x) * (B * K)" by (rule mult_assoc) finally show "norm (f x ** g x) \ norm (f x) * (B * K)" . qed with f show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) flip: "bounded_bilinear (\x y. y ** x)" apply default apply (rule add_right) apply (rule add_left) apply (rule scaleR_right) apply (rule scaleR_left) apply (subst mult_commute) using bounded by fast lemma (in bounded_bilinear) Bfun_prod_Zfun: assumes f: "Bfun f F" assumes g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma Bfun_inverse_lemma: fixes x :: "'a::real_normed_div_algebra" shows "\r \ norm x; 0 < r\ \ norm (inverse x) \ inverse r" apply (subst nonzero_norm_inverse, clarsimp) apply (erule (1) le_imp_inverse_le) done lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f ---> a) F" assumes a: "a \ 0" shows "Bfun (\x. inverse (f x)) F" proof - from a have "0 < norm a" by simp hence "\r>0. r < norm a" by (rule dense) then obtain r where r1: "0 < r" and r2: "r < norm a" by fast have "eventually (\x. dist (f x) a < r) F" using tendstoD [OF f r1] by fast hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" proof eventually_elim case (elim x) hence 1: "norm (f x - a) < r" by (simp add: dist_norm) hence 2: "f x \ 0" using r2 by auto hence "norm (inverse (f x)) = inverse (norm (f x))" by (rule nonzero_norm_inverse) also have "\ \ inverse (norm a - r)" proof (rule le_imp_inverse_le) show "0 < norm a - r" using r2 by simp next have "norm a - norm (f x) \ norm (a - f x)" by (rule norm_triangle_ineq2) also have "\ = norm (f x - a)" by (rule norm_minus_commute) also have "\ < r" using 1 . finally show "norm a - r \ norm (f x)" by simp qed finally show "norm (inverse (f x)) \ inverse (norm a - r)" . qed thus ?thesis by (rule BfunI) qed lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f ---> a) F" assumes a: "a \ 0" shows "((\x. inverse (f x)) ---> inverse a) F" proof - from a have "0 < norm a" by simp with f have "eventually (\x. dist (f x) a < norm a) F" by (rule tendstoD) then have "eventually (\x. f x \ 0) F" unfolding dist_norm by (auto elim!: eventually_elim1) with a have "eventually (\x. inverse (f x) - inverse a = - (inverse (f x) * (f x - a) * inverse a)) F" by (auto elim!: eventually_elim1 simp: inverse_diff_inverse) moreover have "Zfun (\x. - (inverse (f x) * (f x - a) * inverse a)) F" by (intro Zfun_minus Zfun_mult_left bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) ultimately show ?thesis unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) qed lemma continuous_inverse: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. inverse (f x))" using assms unfolding continuous_def by (rule tendsto_inverse) lemma continuous_at_within_inverse[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. inverse (f x))" using assms unfolding continuous_within by (rule tendsto_inverse) lemma isCont_inverse[continuous_intros, simp]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "isCont f a" and "f a \ 0" shows "isCont (\x. inverse (f x)) a" using assms unfolding continuous_at by (rule tendsto_inverse) lemma continuous_on_inverse[continuous_on_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. inverse (f x))" using assms unfolding continuous_on_def by (fast intro: tendsto_inverse) lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" shows "\(f ---> a) F; (g ---> b) F; b \ 0\ \ ((\x. f x / g x) ---> a / b) F" by (simp add: tendsto_mult tendsto_inverse divide_inverse) lemma continuous_divide: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous F f" and "continuous F g" and "g (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) / (g x))" using assms unfolding continuous_def by (rule tendsto_divide) lemma continuous_at_within_divide[continuous_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \ 0" shows "continuous (at a within s) (\x. (f x) / (g x))" using assms unfolding continuous_within by (rule tendsto_divide) lemma isCont_divide[continuous_intros, simp]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "isCont f a" "isCont g a" "g a \ 0" shows "isCont (\x. (f x) / g x) a" using assms unfolding continuous_at by (rule tendsto_divide) lemma continuous_on_divide[continuous_on_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_field" assumes "continuous_on s f" "continuous_on s g" and "\x\s. g x \ 0" shows "continuous_on s (\x. (f x) / (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_divide) lemma tendsto_sgn [tendsto_intros]: fixes l :: "'a::real_normed_vector" shows "\(f ---> l) F; l \ 0\ \ ((\x. sgn (f x)) ---> sgn l) F" unfolding sgn_div_norm by (simp add: tendsto_intros) lemma continuous_sgn: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. sgn (f x))" using assms unfolding continuous_def by (rule tendsto_sgn) lemma continuous_at_within_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. sgn (f x))" using assms unfolding continuous_within by (rule tendsto_sgn) lemma isCont_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "isCont f a" and "f a \ 0" shows "isCont (\x. sgn (f x)) a" using assms unfolding continuous_at by (rule tendsto_sgn) lemma continuous_on_sgn[continuous_on_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. sgn (f x))" using assms unfolding continuous_on_def by (fast intro: tendsto_sgn) lemma filterlim_at_infinity: fixes f :: "_ \ 'a\real_normed_vector" assumes "0 \ c" shows "(LIM x F. f x :> at_infinity) \ (\r>c. eventually (\x. r \ norm (f x)) F)" unfolding filterlim_iff eventually_at_infinity proof safe fix P :: "'a \ bool" and b assume *: "\r>c. eventually (\x. r \ norm (f x)) F" and P: "\x. b \ norm x \ P x" have "max b (c + 1) > c" by auto with * have "eventually (\x. max b (c + 1) \ norm (f x)) F" by auto then show "eventually (\x. P (f x)) F" proof eventually_elim fix x assume "max b (c + 1) \ norm (f x)" with P show "P (f x)" by auto qed qed force subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *} text {* This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and @{term "at_right x"} and also @{term "at_right 0"}. *} lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] lemma filtermap_homeomorph: assumes f: "continuous (at a) f" assumes g: "continuous (at (f a)) g" assumes bij1: "\x. f (g x) = x" and bij2: "\x. g (f x) = x" shows "filtermap f (nhds a) = nhds (f a)" unfolding filter_eq_iff eventually_filtermap eventually_nhds proof safe fix P S assume S: "open S" "f a \ S" and P: "\x\S. P x" from continuous_within_topological[THEN iffD1, rule_format, OF f S] P show "\S. open S \ a \ S \ (\x\S. P (f x))" by auto next fix P S assume S: "open S" "a \ S" and P: "\x\S. P (f x)" with continuous_within_topological[THEN iffD1, rule_format, OF g, of S] bij2 obtain A where "open A" "f a \ A" "(\y\A. g y \ S)" by (metis UNIV_I) with P bij1 show "\S. open S \ f a \ S \ (\x\S. P x)" by (force intro!: exI[of _ A]) qed lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d::'a::real_normed_vector)" by (rule filtermap_homeomorph[where g="\x. x + d"]) (auto intro: continuous_intros) lemma filtermap_nhds_minus: "filtermap (\x. - x) (nhds a) = nhds (- a::'a::real_normed_vector)" by (rule filtermap_homeomorph[where g=uminus]) (auto intro: continuous_minus) lemma filtermap_at_shift: "filtermap (\x. x - d) (at a) = at (a - d::'a::real_normed_vector)" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma filtermap_at_right_shift: "filtermap (\x. x - d) (at_right a) = at_right (a - d::real)" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma at_right_to_0: "at_right (a::real) = filtermap (\x. x + a) (at_right 0)" using filtermap_at_right_shift[of "-a" 0] by simp lemma filterlim_at_right_to_0: "filterlim f F (at_right (a::real)) \ filterlim (\x. f (x + a)) F (at_right 0)" unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. lemma eventually_at_right_to_0: "eventually P (at_right (a::real)) \ eventually (\x. P (x + a)) (at_right 0)" unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) lemma filtermap_at_minus: "filtermap (\x. - x) (at a) = at (- a::'a::real_normed_vector)" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_left_minus: "at_left (a::real) = filtermap (\x. - x) (at_right (- a))" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_right_minus: "at_right (a::real) = filtermap (\x. - x) (at_left (- a))" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma filterlim_at_left_to_right: "filterlim f F (at_left (a::real)) \ filterlim (\x. f (- x)) F (at_right (-a))" unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. lemma eventually_at_left_to_right: "eventually P (at_left (a::real)) \ eventually (\x. P (- x)) (at_right (-a))" unfolding at_left_minus[of a] by (simp add: eventually_filtermap) lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: real filter)" unfolding filter_eq_iff eventually_filtermap eventually_at_top_linorder eventually_at_bot_linorder by (metis le_minus_iff minus_minus) lemma at_bot_mirror: "at_bot = filtermap uminus (at_top :: real filter)" unfolding at_top_mirror filtermap_filtermap by (simp add: filtermap_ident) lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \ (LIM x at_bot. f (-x::real) :> F)" unfolding filterlim_def at_top_mirror filtermap_filtermap .. lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \ (LIM x at_top. f (-x::real) :> F)" unfolding filterlim_def at_bot_mirror filtermap_filtermap .. lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" unfolding filterlim_at_top eventually_at_bot_dense by (metis leI minus_less_iff order_less_asym) lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" unfolding filterlim_at_bot eventually_at_top_dense by (metis leI less_minus_iff order_less_asym) lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \ (LIM x F. - (f x) :: real :> at_bot)" using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] using filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\x. - f x" F] by auto lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \ (LIM x F. - (f x) :: real :> at_top)" unfolding filterlim_uminus_at_top by simp lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" unfolding filterlim_at_top_gt[where c=0] eventually_at_filter proof safe fix Z :: real assume [arith]: "0 < Z" then have "eventually (\x. x < inverse Z) (nhds 0)" by (auto simp add: eventually_nhds_metric dist_real_def intro!: exI[of _ "\inverse Z\"]) then show "eventually (\x. x \ 0 \ x \ {0<..} \ Z \ inverse x) (nhds 0)" by (auto elim!: eventually_elim1 simp: inverse_eq_divide field_simps) qed lemma filterlim_inverse_at_top: "(f ---> (0 :: real)) F \ eventually (\x. 0 < f x) F \ LIM x F. inverse (f x) :> at_top" by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) (simp add: filterlim_def eventually_filtermap eventually_elim1 at_within_def le_principal) lemma filterlim_inverse_at_bot_neg: "LIM x (at_left (0::real)). inverse x :> at_bot" by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) lemma filterlim_inverse_at_bot: "(f ---> (0 :: real)) F \ eventually (\x. f x < 0) F \ LIM x F. inverse (f x) :> at_bot" unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) lemma tendsto_inverse_0: fixes x :: "_ \ 'a\real_normed_div_algebra" shows "(inverse ---> (0::'a)) at_infinity" unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity proof safe fix r :: real assume "0 < r" show "\b. \x. b \ norm x \ norm (inverse x :: 'a) < r" proof (intro exI[of _ "inverse (r / 2)"] allI impI) fix x :: 'a from `0 < r` have "0 < inverse (r / 2)" by simp also assume *: "inverse (r / 2) \ norm x" finally show "norm (inverse x) < r" using * `0 < r` by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) qed qed lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" proof (rule antisym) have "(inverse ---> (0::real)) at_top" by (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) then show "filtermap inverse at_top \ at_right (0::real)" by (simp add: le_principal eventually_filtermap eventually_gt_at_top filterlim_def at_within_def) next have "filtermap inverse (filtermap inverse (at_right (0::real))) \ filtermap inverse at_top" using filterlim_inverse_at_top_right unfolding filterlim_def by (rule filtermap_mono) then show "at_right (0::real) \ filtermap inverse at_top" by (simp add: filtermap_ident filtermap_filtermap) qed lemma eventually_at_right_to_top: "eventually P (at_right (0::real)) \ eventually (\x. P (inverse x)) at_top" unfolding at_right_to_top eventually_filtermap .. lemma filterlim_at_right_to_top: "filterlim f F (at_right (0::real)) \ (LIM x at_top. f (inverse x) :> F)" unfolding filterlim_def at_right_to_top filtermap_filtermap .. lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. lemma eventually_at_top_to_right: "eventually P at_top \ eventually (\x. P (inverse x)) (at_right (0::real))" unfolding at_top_to_right eventually_filtermap .. lemma filterlim_at_top_to_right: "filterlim f F at_top \ (LIM x (at_right (0::real)). f (inverse x) :> F)" unfolding filterlim_def at_top_to_right filtermap_filtermap .. lemma filterlim_inverse_at_infinity: fixes x :: "_ \ 'a\{real_normed_div_algebra, division_ring_inverse_zero}" shows "filterlim inverse at_infinity (at (0::'a))" unfolding filterlim_at_infinity[OF order_refl] proof safe fix r :: real assume "0 < r" then show "eventually (\x::'a. r \ norm (inverse x)) (at 0)" unfolding eventually_at norm_inverse by (intro exI[of _ "inverse r"]) (auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) qed lemma filterlim_inverse_at_iff: fixes g :: "'a \ 'b\{real_normed_div_algebra, division_ring_inverse_zero}" shows "(LIM x F. inverse (g x) :> at 0) \ (LIM x F. g x :> at_infinity)" unfolding filterlim_def filtermap_filtermap[symmetric] proof assume "filtermap g F \ at_infinity" then have "filtermap inverse (filtermap g F) \ filtermap inverse at_infinity" by (rule filtermap_mono) also have "\ \ at 0" using tendsto_inverse_0[where 'a='b] by (auto intro!: exI[of _ 1] simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity) finally show "filtermap inverse (filtermap g F) \ at 0" . next assume "filtermap inverse (filtermap g F) \ at 0" then have "filtermap inverse (filtermap inverse (filtermap g F)) \ filtermap inverse (at 0)" by (rule filtermap_mono) with filterlim_inverse_at_infinity show "filtermap g F \ at_infinity" by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) qed lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \ ((\x. inverse (f x) :: real) ---> 0) F" by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff) text {* We only show rules for multiplication and addition when the functions are either against a real value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}. *} lemma filterlim_tendsto_pos_mult_at_top: assumes f: "(f ---> c) F" and c: "0 < c" assumes g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f `0 < c` have "eventually (\x. c / 2 < f x) F" by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_elim1 simp: dist_real_def abs_real_def split: split_if_asm) moreover from g have "eventually (\x. (Z / c * 2) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim fix x assume "c / 2 < f x" "Z / c * 2 \ g x" with `0 < Z` `0 < c` have "c / 2 * (Z / c * 2) \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) with `0 < c` show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_at_top: assumes f: "LIM x F. f x :> at_top" assumes g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 1 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim fix x assume "1 \ f x" "Z \ g x" with `0 < Z` have "1 * Z \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) then show "Z \ f x * g x" by simp qed qed lemma filterlim_tendsto_pos_mult_at_bot: assumes "(f ---> c) F" "0 < (c::real)" "filterlim g at_bot F" shows "LIM x F. f x * g x :> at_bot" using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\x. - g x"] assms(3) unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_add_at_top: assumes f: "(f ---> c) F" assumes g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. c - 1 < f x) F" by (auto dest!: tendstoD[where e=1] elim!: eventually_elim1 simp: dist_real_def) moreover from g have "eventually (\x. Z - (c - 1) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma LIM_at_top_divide: fixes f g :: "'a \ real" assumes f: "(f ---> a) F" "0 < a" assumes g: "(g ---> 0) F" "eventually (\x. 0 < g x) F" shows "LIM x F. f x / g x :> at_top" unfolding divide_inverse by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) lemma filterlim_at_top_add_at_top: assumes f: "LIM x F. f x :> at_top" assumes g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 0 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma tendsto_divide_0: fixes f :: "_ \ 'a\{real_normed_div_algebra, division_ring_inverse_zero}" assumes f: "(f ---> c) F" assumes g: "LIM x F. g x :> at_infinity" shows "((\x. f x / g x) ---> 0) F" using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse) lemma linear_plus_1_le_power: fixes x :: real assumes x: "0 \ x" shows "real n * x + 1 \ (x + 1) ^ n" proof (induct n) case (Suc n) have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" by (simp add: field_simps real_of_nat_Suc mult_nonneg_nonneg x) also have "\ \ (x + 1)^Suc n" using Suc x by (simp add: mult_left_mono) finally show ?case . qed simp lemma filterlim_realpow_sequentially_gt1: fixes x :: "'a :: real_normed_div_algebra" assumes x[arith]: "1 < norm x" shows "LIM n sequentially. x ^ n :> at_infinity" proof (intro filterlim_at_infinity[THEN iffD2] allI impI) fix y :: real assume "0 < y" have "0 < norm x - 1" by simp then obtain N::nat where "y < real N * (norm x - 1)" by (blast dest: reals_Archimedean3) also have "\ \ real N * (norm x - 1) + 1" by simp also have "\ \ (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp also have "\ = norm x ^ N" by simp finally have "\n\N. y \ norm x ^ n" by (metis order_less_le_trans power_increasing order_less_imp_le x) then show "eventually (\n. y \ norm (x ^ n)) sequentially" unfolding eventually_sequentially by (auto simp: norm_power) qed simp subsection {* Limits of Sequences *} lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z" by simp lemma LIMSEQ_iff: fixes L :: "'a::real_normed_vector" shows "(X ----> L) = (\r>0. \no. \n \ no. norm (X n - L) < r)" unfolding LIMSEQ_def dist_norm .. lemma LIMSEQ_I: fixes L :: "'a::real_normed_vector" shows "(\r. 0 < r \ \no. \n\no. norm (X n - L) < r) \ X ----> L" by (simp add: LIMSEQ_iff) lemma LIMSEQ_D: fixes L :: "'a::real_normed_vector" shows "\X ----> L; 0 < r\ \ \no. \n\no. norm (X n - L) < r" by (simp add: LIMSEQ_iff) lemma LIMSEQ_linear: "\ X ----> x ; l > 0 \ \ (\ n. X (n * l)) ----> x" unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute) lemma Bseq_inverse_lemma: fixes x :: "'a::real_normed_div_algebra" shows "\r \ norm x; 0 < r\ \ norm (inverse x) \ inverse r" apply (subst nonzero_norm_inverse, clarsimp) apply (erule (1) le_imp_inverse_le) done lemma Bseq_inverse: fixes a :: "'a::real_normed_div_algebra" shows "\X ----> a; a \ 0\ \ Bseq (\n. inverse (X n))" by (rule Bfun_inverse) lemma LIMSEQ_diff_approach_zero: fixes L :: "'a::real_normed_vector" shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L" by (drule (1) tendsto_add, simp) lemma LIMSEQ_diff_approach_zero2: fixes L :: "'a::real_normed_vector" shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L" by (drule (1) tendsto_diff, simp) text{*An unbounded sequence's inverse tends to 0*} lemma LIMSEQ_inverse_zero: "\r::real. \N. \n\N. r < X n \ (\n. inverse (X n)) ----> 0" apply (rule filterlim_compose[OF tendsto_inverse_0]) apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially) apply (metis abs_le_D1 linorder_le_cases linorder_not_le) done text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*} lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0" by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity) text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to infinity is now easily proved*} lemma LIMSEQ_inverse_real_of_nat_add: "(%n. r + inverse(real(Suc n))) ----> r" using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus: "(%n. r + -inverse(real(Suc n))) ----> r" using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r" using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] by auto subsection {* Convergence on sequences *} lemma convergent_add: fixes X Y :: "nat \ 'a::real_normed_vector" assumes "convergent (\n. X n)" assumes "convergent (\n. Y n)" shows "convergent (\n. X n + Y n)" using assms unfolding convergent_def by (fast intro: tendsto_add) lemma convergent_setsum: fixes X :: "'a \ nat \ 'b::real_normed_vector" assumes "\i. i \ A \ convergent (\n. X i n)" shows "convergent (\n. \i\A. X i n)" proof (cases "finite A") case True from this and assms show ?thesis by (induct A set: finite) (simp_all add: convergent_const convergent_add) qed (simp add: convergent_const) lemma (in bounded_linear) convergent: assumes "convergent (\n. X n)" shows "convergent (\n. f (X n))" using assms unfolding convergent_def by (fast intro: tendsto) lemma (in bounded_bilinear) convergent: assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n ** Y n)" using assms unfolding convergent_def by (fast intro: tendsto) lemma convergent_minus_iff: fixes X :: "nat \ 'a::real_normed_vector" shows "convergent X \ convergent (\n. - X n)" apply (simp add: convergent_def) apply (auto dest: tendsto_minus) apply (drule tendsto_minus, auto) done text {* A monotone sequence converges to its least upper bound. *} lemma isLub_mono_imp_LIMSEQ: fixes X :: "nat \ real" assumes u: "isLub UNIV {x. \n. X n = x} u" (* FIXME: use 'range X' *) assumes X: "\m n. m \ n \ X m \ X n" shows "X ----> u" proof (rule LIMSEQ_I) have 1: "\n. X n \ u" using isLubD2 [OF u] by auto have "\y. (\n. X n \ y) \ u \ y" using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def) hence 2: "\yn. y < X n" by (metis not_le) fix r :: real assume "0 < r" hence "u - r < u" by simp hence "\m. u - r < X m" using 2 by simp then obtain m where "u - r < X m" .. with X have "\n\m. u - r < X n" by (fast intro: less_le_trans) hence "\m. \n\m. u - r < X n" .. thus "\m. \n\m. norm (X n - u) < r" using 1 by (simp add: diff_less_eq add_commute) qed text{*A standard proof of the theorem for monotone increasing sequence*} lemma Bseq_mono_convergent: "Bseq X \ \m. \n \ m. X m \ X n \ convergent (X::nat=>real)" by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI) text{*Main monotonicity theorem*} lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent (X::nat\real)" by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff Bseq_mono_convergent) lemma Cauchy_iff: fixes X :: "nat \ 'a::real_normed_vector" shows "Cauchy X \ (\e>0. \M. \m\M. \n\M. norm (X m - X n) < e)" unfolding Cauchy_def dist_norm .. lemma CauchyI: fixes X :: "nat \ 'a::real_normed_vector" shows "(\e. 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e) \ Cauchy X" by (simp add: Cauchy_iff) lemma CauchyD: fixes X :: "nat \ 'a::real_normed_vector" shows "\Cauchy X; 0 < e\ \ \M. \m\M. \n\M. norm (X m - X n) < e" by (simp add: Cauchy_iff) lemma incseq_convergent: fixes X :: "nat \ real" assumes "incseq X" and "\i. X i \ B" obtains L where "X ----> L" "\i. X i \ L" proof atomize_elim from incseq_bounded[OF assms] `incseq X` Bseq_monoseq_convergent[of X] obtain L where "X ----> L" by (auto simp: convergent_def monoseq_def incseq_def) with `incseq X` show "\L. X ----> L \ (\i. X i \ L)" by (auto intro!: exI[of _ L] incseq_le) qed lemma decseq_convergent: fixes X :: "nat \ real" assumes "decseq X" and "\i. B \ X i" obtains L where "X ----> L" "\i. L \ X i" proof atomize_elim from decseq_bounded[OF assms] `decseq X` Bseq_monoseq_convergent[of X] obtain L where "X ----> L" by (auto simp: convergent_def monoseq_def decseq_def) with `decseq X` show "\L. X ----> L \ (\i. L \ X i)" by (auto intro!: exI[of _ L] decseq_le) qed subsubsection {* Cauchy Sequences are Bounded *} text{*A Cauchy sequence is bounded -- this is the standard proof mechanization rather than the nonstandard proof*} lemma lemmaCauchy: "\n \ M. norm (X M - X n) < (1::real) ==> \n \ M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)" apply (clarify, drule spec, drule (1) mp) apply (simp only: norm_minus_commute) apply (drule order_le_less_trans [OF norm_triangle_ineq2]) apply simp done subsection {* Power Sequences *} text{*The sequence @{term "x^n"} tends to 0 if @{term "0\x"} and @{term "x<1"}. Proof will use (NS) Cauchy equivalence for convergence and also fact that bounded and monotonic sequence converges.*} lemma Bseq_realpow: "[| 0 \ (x::real); x \ 1 |] ==> Bseq (%n. x ^ n)" apply (simp add: Bseq_def) apply (rule_tac x = 1 in exI) apply (simp add: power_abs) apply (auto dest: power_mono) done lemma monoseq_realpow: fixes x :: real shows "[| 0 \ x; x \ 1 |] ==> monoseq (%n. x ^ n)" apply (clarify intro!: mono_SucI2) apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto) done lemma convergent_realpow: "[| 0 \ (x::real); x \ 1 |] ==> convergent (%n. x ^ n)" by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \ (\n. inverse (x ^ n)) ----> 0" by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp lemma LIMSEQ_realpow_zero: "\0 \ (x::real); x < 1\ \ (\n. x ^ n) ----> 0" proof cases assume "0 \ x" and "x \ 0" hence x0: "0 < x" by simp assume x1: "x < 1" from x0 x1 have "1 < inverse x" by (rule one_less_inverse) hence "(\n. inverse (inverse x ^ n)) ----> 0" by (rule LIMSEQ_inverse_realpow_zero) thus ?thesis by (simp add: power_inverse) qed (rule LIMSEQ_imp_Suc, simp add: tendsto_const) lemma LIMSEQ_power_zero: fixes x :: "'a::{real_normed_algebra_1}" shows "norm x < 1 \ (\n. x ^ n) ----> 0" apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) apply (simp only: tendsto_Zfun_iff, erule Zfun_le) apply (simp add: power_abs norm_power_ineq) done lemma LIMSEQ_divide_realpow_zero: "1 < x \ (\n. a / (x ^ n) :: real) ----> 0" by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp text{*Limit of @{term "c^n"} for @{term"\c\ < 1"}*} lemma LIMSEQ_rabs_realpow_zero: "\c\ < 1 \ (\n. \c\ ^ n :: real) ----> 0" by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) lemma LIMSEQ_rabs_realpow_zero2: "\c\ < 1 \ (\n. c ^ n :: real) ----> 0" by (rule LIMSEQ_power_zero) simp subsection {* Limits of Functions *} lemma LIM_eq: fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" shows "f -- a --> L = (\r>0.\s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r)" by (simp add: LIM_def dist_norm) lemma LIM_I: fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" shows "(!!r. 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r) ==> f -- a --> L" by (simp add: LIM_eq) lemma LIM_D: fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" shows "[| f -- a --> L; 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r" by (simp add: LIM_eq) lemma LIM_offset: fixes a :: "'a::real_normed_vector" shows "f -- a --> L \ (\x. f (x + k)) -- a - k --> L" unfolding filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap by simp lemma LIM_offset_zero: fixes a :: "'a::real_normed_vector" shows "f -- a --> L \ (\h. f (a + h)) -- 0 --> L" by (drule_tac k="a" in LIM_offset, simp add: add_commute) lemma LIM_offset_zero_cancel: fixes a :: "'a::real_normed_vector" shows "(\h. f (a + h)) -- 0 --> L \ f -- a --> L" by (drule_tac k="- a" in LIM_offset, simp) lemma LIM_offset_zero_iff: fixes f :: "'a :: real_normed_vector \ _" shows "f -- a --> L \ (\h. f (a + h)) -- 0 --> L" using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto lemma LIM_zero: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "(f ---> l) F \ ((\x. f x - l) ---> 0) F" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_cancel: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "((\x. f x - l) ---> 0) F \ (f ---> l) F" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_iff: fixes f :: "'a::metric_space \ 'b::real_normed_vector" shows "((\x. f x - l) ---> 0) F = (f ---> l) F" unfolding tendsto_iff dist_norm by simp lemma LIM_imp_LIM: fixes f :: "'a::topological_space \ 'b::real_normed_vector" fixes g :: "'a::topological_space \ 'c::real_normed_vector" assumes f: "f -- a --> l" assumes le: "\x. x \ a \ norm (g x - m) \ norm (f x - l)" shows "g -- a --> m" by (rule metric_LIM_imp_LIM [OF f], simp add: dist_norm le) lemma LIM_equal2: fixes f g :: "'a::real_normed_vector \ 'b::topological_space" assumes 1: "0 < R" assumes 2: "\x. \x \ a; norm (x - a) < R\ \ f x = g x" shows "g -- a --> l \ f -- a --> l" by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm) lemma LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f: "f -- a --> b" assumes g: "g -- b --> c" assumes inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ b" shows "(\x. g (f x)) -- a --> c" by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" assumes f: "f -- a --> 0" assumes 1: "\x. x \ a \ 0 \ g x" assumes 2: "\x. x \ a \ g x \ f x" shows "g -- a --> 0" proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) fix x assume x: "x \ a" have "norm (g x - 0) = g x" by (simp add: 1 x) also have "g x \ f x" by (rule 2 [OF x]) also have "f x \ \f x\" by (rule abs_ge_self) also have "\f x\ = norm (f x - 0)" by simp finally show "norm (g x - 0) \ norm (f x - 0)" . qed subsection {* Continuity *} lemma LIM_isCont_iff: fixes f :: "'a::real_normed_vector \ 'b::topological_space" shows "(f -- a --> f a) = ((\h. f (a + h)) -- 0 --> f a)" by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) lemma isCont_iff: fixes f :: "'a::real_normed_vector \ 'b::topological_space" shows "isCont f x = (\h. f (x + h)) -- 0 --> f x" by (simp add: isCont_def LIM_isCont_iff) lemma isCont_LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f [unfolded isCont_def]: "isCont f a" assumes g: "g -- f a --> l" assumes inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ f a" shows "(\x. g (f x)) -- a --> l" by (rule LIM_compose2 [OF f g inj]) lemma isCont_norm [simp]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" shows "isCont f a \ isCont (\x. norm (f x)) a" by (fact continuous_norm) lemma isCont_rabs [simp]: fixes f :: "'a::t2_space \ real" shows "isCont f a \ isCont (\x. \f x\) a" by (fact continuous_rabs) lemma isCont_add [simp]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" shows "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" by (fact continuous_add) lemma isCont_minus [simp]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" shows "isCont f a \ isCont (\x. - f x) a" by (fact continuous_minus) lemma isCont_diff [simp]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" shows "\isCont f a; isCont g a\ \ isCont (\x. f x - g x) a" by (fact continuous_diff) lemma isCont_mult [simp]: fixes f g :: "'a::t2_space \ 'b::real_normed_algebra" shows "\isCont f a; isCont g a\ \ isCont (\x. f x * g x) a" by (fact continuous_mult) lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" by (fact continuous) lemma (in bounded_bilinear) isCont: "\isCont f a; isCont g a\ \ isCont (\x. f x ** g x) a" by (fact continuous) lemmas isCont_scaleR [simp] = bounded_bilinear.isCont [OF bounded_bilinear_scaleR] lemmas isCont_of_real [simp] = bounded_linear.isCont [OF bounded_linear_of_real] lemma isCont_power [simp]: fixes f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" shows "isCont f a \ isCont (\x. f x ^ n) a" by (fact continuous_power) lemma isCont_setsum [simp]: fixes f :: "'a \ 'b::t2_space \ 'c::real_normed_vector" shows "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" by (auto intro: continuous_setsum) lemmas isCont_intros = isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR isCont_of_real isCont_power isCont_sgn isCont_setsum subsection {* Uniform Continuity *} definition isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" lemma isUCont_isCont: "isUCont f ==> isCont f x" by (simp add: isUCont_def isCont_def LIM_def, force) lemma isUCont_Cauchy: "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" unfolding isUCont_def apply (rule metric_CauchyI) apply (drule_tac x=e in spec, safe) apply (drule_tac e=s in metric_CauchyD, safe) apply (rule_tac x=M in exI, simp) done lemma (in bounded_linear) isUCont: "isUCont f" unfolding isUCont_def dist_norm proof (intro allI impI) fix r::real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "\x. norm (f x) \ norm x * K" using pos_bounded by fast show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" proof (rule exI, safe) from r K show "0 < r / K" by (rule divide_pos_pos) next fix x y :: 'a assume xy: "norm (x - y) < r / K" have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) also have "\ \ norm (x - y) * K" by (rule norm_le) also from K xy have "\ < r" by (simp only: pos_less_divide_eq) finally show "norm (f x - f y) < r" . qed qed lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" by (rule isUCont [THEN isUCont_Cauchy]) lemma LIM_less_bound: fixes f :: "real \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" shows "0 \ f x" proof (rule tendsto_le_const) show "(f ---> f x) (at_left x)" using `isCont f x` by (simp add: filterlim_at_split isCont_def) show "eventually (\x. 0 \ f x) (at_left x)" using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"]) qed simp subsection {* Nested Intervals and Bisection -- Needed for Compactness *} lemma nested_sequence_unique: assumes "\n. f n \ f (Suc n)" "\n. g (Suc n) \ g n" "\n. f n \ g n" "(\n. f n - g n) ----> 0" shows "\l::real. ((\n. f n \ l) \ f ----> l) \ ((\n. l \ g n) \ g ----> l)" proof - have "incseq f" unfolding incseq_Suc_iff by fact have "decseq g" unfolding decseq_Suc_iff by fact { fix n from `decseq g` have "g n \ g 0" by (rule decseqD) simp with `\n. f n \ g n`[THEN spec, of n] have "f n \ g 0" by auto } then obtain u where "f ----> u" "\i. f i \ u" using incseq_convergent[OF `incseq f`] by auto moreover { fix n from `incseq f` have "f 0 \ f n" by (rule incseqD) simp with `\n. f n \ g n`[THEN spec, of n] have "f 0 \ g n" by simp } then obtain l where "g ----> l" "\i. l \ g i" using decseq_convergent[OF `decseq g`] by auto moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]] ultimately show ?thesis by auto qed lemma Bolzano[consumes 1, case_names trans local]: fixes P :: "real \ real \ bool" assumes [arith]: "a \ b" assumes trans: "\a b c. \P a b; P b c; a \ b; b \ c\ \ P a c" assumes local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" shows "P a b" proof - def bisect \ "nat_rec (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" def l \ "\n. fst (bisect n)" and u \ "\n. snd (bisect n)" have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" by (simp_all add: l_def u_def bisect_def split: prod.split) { fix n have "l n \ u n" by (induct n) auto } note this[simp] have "\x. ((\n. l n \ x) \ l ----> x) \ ((\n. x \ u n) \ u ----> x)" proof (safe intro!: nested_sequence_unique) fix n show "l n \ l (Suc n)" "u (Suc n) \ u n" by (induct n) auto next { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) } then show "(\n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero) qed fact then obtain x where x: "\n. l n \ x" "\n. x \ u n" and "l ----> x" "u ----> x" by auto obtain d where "0 < d" and d: "\a b. a \ x \ x \ b \ b - a < d \ P a b" using `l 0 \ x` `x \ u 0` local[of x] by auto show "P a b" proof (rule ccontr) assume "\ P a b" { fix n have "\ P (l n) (u n)" proof (induct n) case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto qed (simp add: `\ P a b`) } moreover { have "eventually (\n. x - d / 2 < l n) sequentially" using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto moreover have "eventually (\n. u n < x + d / 2) sequentially" using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto ultimately have "eventually (\n. P (l n) (u n)) sequentially" proof eventually_elim fix n assume "x - d / 2 < l n" "u n < x + d / 2" from add_strict_mono[OF this] have "u n - l n < d" by simp with x show "P (l n) (u n)" by (rule d) qed } ultimately show False by simp qed qed lemma compact_Icc[simp, intro]: "compact {a .. b::real}" proof (cases "a \ b", rule compactI) fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" def T == "{a .. b}" from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" proof (induct rule: Bolzano) case (trans a b c) then have *: "{a .. c} = {a .. b} \ {b .. c}" by auto from trans obtain C1 C2 where "C1\C \ finite C1 \ {a..b} \ \C1" "C2\C \ finite C2 \ {b..c} \ \C2" by (auto simp: *) with trans show ?case unfolding * by (intro exI[of _ "C1 \ C2"]) auto next case (local x) then have "x \ \C" using C by auto with C(2) obtain c where "x \ c" "open c" "c \ C" by auto then obtain e where "0 < e" "{x - e <..< x + e} \ c" by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff) with `c \ C` show ?case by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto qed qed simp subsection {* Boundedness of continuous functions *} text{*By bisection, function continuous on closed interval is bounded above*} lemma isCont_eq_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_sup[of "{a .. b}" f] by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_eq_Lb: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_inf[of "{a .. b}" f] by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_bounded: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" using isCont_eq_Ub[of a b f] by auto lemma isCont_has_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" using isCont_eq_Ub[of a b f] by auto (*HOL style here: object-level formulations*) lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & (\x. a \ x & x \ b --> isCont f x)) --> (\x. a \ x & x \ b & f(x) = y)" by (blast intro: IVT) lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & (\x. a \ x & x \ b --> isCont f x)) --> (\x. a \ x & x \ b & f(x) = y)" by (blast intro: IVT2) lemma isCont_Lb_Ub: fixes f :: "real \ real" assumes "a \ b" "\x. a \ x \ x \ b \ isCont f x" shows "\L M. (\x. a \ x \ x \ b \ L \ f x \ f x \ M) \ (\y. L \ y \ y \ M \ (\x. a \ x \ x \ b \ (f x = y)))" proof - obtain M where M: "a \ M" "M \ b" "\x. a \ x \ x \ b \ f x \ f M" using isCont_eq_Ub[OF assms] by auto obtain L where L: "a \ L" "L \ b" "\x. a \ x \ x \ b \ f L \ f x" using isCont_eq_Lb[OF assms] by auto show ?thesis using IVT[of f L _ M] IVT2[of f L _ M] M L assms apply (rule_tac x="f L" in exI) apply (rule_tac x="f M" in exI) apply (cases "L \ M") apply (simp, metis order_trans) apply (simp, metis order_trans) done qed text{*Continuity of inverse function*} lemma isCont_inverse_function: fixes f g :: "real \ real" assumes d: "0 < d" and inj: "\z. \z-x\ \ d \ g (f z) = z" and cont: "\z. \z-x\ \ d \ isCont f z" shows "isCont g (f x)" proof - let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}" have f: "continuous_on ?D f" using cont by (intro continuous_at_imp_continuous_on ballI) auto then have g: "continuous_on (f`?D) g" using inj by (intro continuous_on_inv) auto from d f have "{min ?A ?B <..< max ?A ?B} \ f ` ?D" by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" by (rule continuous_on_subset) moreover have "(?A < f x \ f x < ?B) \ (?B < f x \ f x < ?A)" using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto then have "f x \ {min ?A ?B <..< max ?A ?B}" by auto ultimately show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma isCont_inverse_function2: fixes f g :: "real \ real" shows "\a < x; x < b; \z. a \ z \ z \ b \ g (f z) = z; \z. a \ z \ z \ b \ isCont f z\ \ isCont g (f x)" apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"]) apply (simp_all add: abs_le_iff) done (* need to rename second isCont_inverse *) lemma isCont_inv_fun: fixes f g :: "real \ real" shows "[| 0 < d; \z. \z - x\ \ d --> g(f(z)) = z; \z. \z - x\ \ d --> isCont f z |] ==> isCont g (f x)" by (rule isCont_inverse_function) text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} lemma LIM_fun_gt_zero: fixes f :: "real \ real" shows "f -- c --> l \ 0 < l \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ 0 < f x)" apply (drule (1) LIM_D, clarify) apply (rule_tac x = s in exI) apply (simp add: abs_less_iff) done lemma LIM_fun_less_zero: fixes f :: "real \ real" shows "f -- c --> l \ l < 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x < 0)" apply (drule LIM_D [where r="-l"], simp, clarify) apply (rule_tac x = s in exI) apply (simp add: abs_less_iff) done lemma LIM_fun_not_zero: fixes f :: "real \ real" shows "f -- c --> l \ l \ 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x \ 0)" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff) end