# HG changeset patch # User wenzelm # Date 1469224547 -7200 # Node ID 5f097087fa1e898362e0d9c245425298675a3540 # Parent c2f69dac03535fe0be00c111a69ec992d97a7549 misc tuning and modernization; diff -r c2f69dac0353 -r 5f097087fa1e src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Jul 22 21:43:56 2016 +0200 +++ b/src/HOL/Limits.thy Fri Jul 22 23:55:47 2016 +0200 @@ -8,13 +8,13 @@ section \Limits on Real Vector Spaces\ theory Limits -imports Real_Vector_Spaces + imports Real_Vector_Spaces begin subsection \Filter going to infinity norm\ -definition at_infinity :: "'a::real_normed_vector filter" where - "at_infinity = (INF r. principal {x. r \ norm x})" +definition at_infinity :: "'a::real_normed_vector filter" + where "at_infinity = (INF r. principal {x. r \ norm x})" lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def @@ -22,21 +22,24 @@ (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) corollary eventually_at_infinity_pos: - "eventually p at_infinity \ (\b. 0 < b \ (\x. norm x \ b \ p x))" -apply (simp add: eventually_at_infinity, auto) -apply (case_tac "b \ 0") -using norm_ge_zero order_trans zero_less_one apply blast -apply (force simp:) -done - -lemma at_infinity_eq_at_top_bot: - "(at_infinity :: real filter) = sup at_top at_bot" + "eventually p at_infinity \ (\b. 0 < b \ (\x. norm x \ b \ p x))" + apply (simp add: eventually_at_infinity) + apply auto + apply (case_tac "b \ 0") + using norm_ge_zero order_trans zero_less_one apply blast + apply force + done + +lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity - eventually_at_top_linorder eventually_at_bot_linorder) + eventually_at_top_linorder eventually_at_bot_linorder) apply safe - apply (rule_tac x="b" in exI, simp) - apply (rule_tac x="- b" in exI, simp) - apply (rule_tac x="max (- Na) N" in exI, auto simp: abs_real_def) + apply (rule_tac x="b" in exI) + apply simp + apply (rule_tac x="- b" in exI) + apply simp + apply (rule_tac x="max (- Na) N" in exI) + apply (auto simp: abs_real_def) done lemma at_top_le_at_infinity: "at_top \ (at_infinity :: real filter)" @@ -45,23 +48,21 @@ lemma at_bot_le_at_infinity: "at_bot \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp -lemma filterlim_at_top_imp_at_infinity: - fixes f :: "_ \ real" - shows "filterlim f at_top F \ filterlim f at_infinity F" +lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \ filterlim f at_infinity F" + for f :: "_ \ real" by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) -lemma lim_infinity_imp_sequentially: - "(f \ l) at_infinity \ ((\n. f(n)) \ l) sequentially" -by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) +lemma lim_infinity_imp_sequentially: "(f \ l) at_infinity \ ((\n. f(n)) \ l) sequentially" + by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) 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" +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" .. @@ -71,11 +72,11 @@ 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)" +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 K: "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" + fix y K + assume K: "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" @@ -85,19 +86,19 @@ qed (force simp del: norm_conv_dist [symmetric]) lemma BfunI: - assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" -unfolding Bfun_def + 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_mono, simp) + using K by (rule eventually_mono) 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 blast + using assms unfolding Bfun_def by blast lemma Cauchy_Bseq: "Cauchy X \ Bseq X" unfolding Cauchy_def Bfun_metric_def eventually_sequentially @@ -124,57 +125,66 @@ 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" + 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"] max.strict_coboundedI2) (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 Bseq_bdd_above: "Bseq (X::nat \ real) \ bdd_above (range X)" +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 Bseq_bdd_above: "Bseq X \ bdd_above (range X)" + for X :: "nat \ real" proof (elim BseqE, intro bdd_aboveI2) - fix K n assume "0 < K" "\n. norm (X n) \ K" then show "X n \ K" + fix K n + assume "0 < K" "\n. norm (X n) \ K" + then show "X n \ K" by (auto elim!: allE[of _ n]) qed -lemma Bseq_bdd_above': - "Bseq (X::nat \ 'a :: real_normed_vector) \ bdd_above (range (\n. norm (X n)))" +lemma Bseq_bdd_above': "Bseq X \ bdd_above (range (\n. norm (X n)))" + for X :: "nat \ 'a :: real_normed_vector" proof (elim BseqE, intro bdd_aboveI2) - fix K n assume "0 < K" "\n. norm (X n) \ K" then show "norm (X n) \ K" + fix K n + assume "0 < K" "\n. norm (X n) \ K" + then show "norm (X n) \ K" by (auto elim!: allE[of _ n]) qed -lemma Bseq_bdd_below: "Bseq (X::nat \ real) \ bdd_below (range X)" +lemma Bseq_bdd_below: "Bseq X \ bdd_below (range X)" + for X :: "nat \ real" proof (elim BseqE, intro bdd_belowI2) - fix K n assume "0 < K" "\n. norm (X n) \ K" then show "- K \ X n" + fix K n + assume "0 < K" "\n. norm (X n) \ K" + then show "- K \ X n" by (auto elim!: allE[of _ n]) qed lemma Bseq_eventually_mono: assumes "eventually (\n. norm (f n) \ norm (g n)) sequentially" "Bseq g" - shows "Bseq f" + shows "Bseq f" proof - from assms(1) obtain N where N: "\n. n \ N \ norm (f n) \ norm (g n)" by (auto simp: eventually_at_top_linorder) - moreover from assms(2) obtain K where K: "\n. norm (g n) \ K" by (blast elim!: BseqE) + moreover from assms(2) obtain K where K: "\n. norm (g n) \ K" + by (blast elim!: BseqE) ultimately have "norm (f n) \ max K (Max {norm (f n) |n. n < N})" for n apply (cases "n < N") - apply (rule max.coboundedI2, rule Max.coboundedI, auto) [] - apply (rule max.coboundedI1, force intro: order.trans[OF N K]) + subgoal by (rule max.coboundedI2, rule Max.coboundedI) auto + subgoal by (rule max.coboundedI1) (force intro: order.trans[OF N K]) done - thus ?thesis by (blast intro: BseqI') + then show ?thesis by (blast intro: BseqI') qed -lemma lemma_NBseq_def: - "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) \ real(Suc N))" +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" .. @@ -188,47 +198,50 @@ using of_nat_0_less_iff by blast qed -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: 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) -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)" (is "?P \ ?Q") +text \Alternative definition for \Bseq\.\ +lemma Bseq_iff: "Bseq X \ (\N. \n. norm (X n) \ real(Suc N))" + by (simp add: Bseq_def) (simp add: lemma_NBseq_def) + +lemma lemma_NBseq_def2: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" + apply (subst lemma_NBseq_def) + apply auto + apply (rule_tac x = "Suc N" in exI) + apply (rule_tac [2] x = N in exI) + apply auto + prefer 2 apply (blast intro: order_less_imp_le) + apply (drule_tac x = n in spec) + apply simp + done + +text \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) + apply safe + apply (rule_tac [2] x = "k + norm x" in exI) + apply (rule_tac x = K in exI) + apply simp + apply (rule exI [where x = 0]) + apply auto + apply (erule order_less_le_trans) + apply simp + apply (drule_tac x=n in spec) + 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)" + (is "?P \ ?Q") proof assume ?P - then obtain K - where *: "0 < K" and **: "\n. norm (X n) \ K" by (auto simp add: Bseq_def) + then obtain K where *: "0 < K" and **: "\n. norm (X n) \ K" + by (auto simp add: Bseq_def) from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp from ** have "\n. norm (X n - X 0) \ K + norm (X 0)" by (auto intro: order_trans norm_triangle_ineq4) @@ -236,129 +249,150 @@ by simp with \0 < K + norm (X 0)\ show ?Q by blast next - assume ?Q then show ?P by (auto simp add: Bseq_iff2) + assume ?Q + then show ?P by (auto simp add: Bseq_iff2) qed -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_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" +lemma BseqI2: "\n. k \ f n \ f n \ K \ Bseq f" + for k K :: real + apply (simp add: Bseq_def) + apply (rule_tac x = "(\k\ + \K\) + 1" in exI) + apply auto + apply (drule_tac x = n in spec) + apply arith + done + + +subsubsection \Upper Bounds and Lubs of Bounded Sequences\ + +lemma Bseq_minus_iff: "Bseq (\n. - (X n) :: 'a::real_normed_vector) \ Bseq X" by (simp add: Bseq_def) lemma Bseq_add: - assumes "Bseq (f :: nat \ 'a :: real_normed_vector)" - shows "Bseq (\x. f x + c)" + fixes f :: "nat \ 'a::real_normed_vector" + assumes "Bseq f" + shows "Bseq (\x. f x + c)" proof - - from assms obtain K where K: "\x. norm (f x) \ K" unfolding Bseq_def by blast + from assms obtain K where K: "\x. norm (f x) \ K" + unfolding Bseq_def by blast { fix x :: nat have "norm (f x + c) \ norm (f x) + norm c" by (rule norm_triangle_ineq) also have "norm (f x) \ K" by (rule K) finally have "norm (f x + c) \ K + norm c" by simp } - thus ?thesis by (rule BseqI') + then show ?thesis by (rule BseqI') qed -lemma Bseq_add_iff: "Bseq (\x. f x + c) \ Bseq (f :: nat \ 'a :: real_normed_vector)" +lemma Bseq_add_iff: "Bseq (\x. f x + c) \ Bseq f" + for f :: "nat \ 'a::real_normed_vector" using Bseq_add[of f c] Bseq_add[of "\x. f x + c" "-c"] by auto lemma Bseq_mult: - assumes "Bseq (f :: nat \ 'a :: real_normed_field)" - assumes "Bseq (g :: nat \ 'a :: real_normed_field)" - shows "Bseq (\x. f x * g x)" + fixes f g :: "nat \ 'a::real_normed_field" + assumes "Bseq f" and "Bseq g" + shows "Bseq (\x. f x * g x)" proof - - from assms obtain K1 K2 where K: "\x. norm (f x) \ K1" "K1 > 0" "\x. norm (g x) \ K2" "K2 > 0" + from assms obtain K1 K2 where K: "norm (f x) \ K1" "K1 > 0" "norm (g x) \ K2" "K2 > 0" + for x unfolding Bseq_def by blast - hence "\x. norm (f x * g x) \ K1 * K2" by (auto simp: norm_mult intro!: mult_mono) - thus ?thesis by (rule BseqI') + then have "norm (f x * g x) \ K1 * K2" for x + by (auto simp: norm_mult intro!: mult_mono) + then show ?thesis by (rule BseqI') qed lemma Bfun_const [simp]: "Bfun (\_. c) F" unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) -lemma Bseq_cmult_iff: "(c :: 'a :: real_normed_field) \ 0 \ Bseq (\x. c * f x) \ Bseq f" +lemma Bseq_cmult_iff: + fixes c :: "'a::real_normed_field" + assumes "c \ 0" + shows "Bseq (\x. c * f x) \ Bseq f" proof - assume "c \ 0" "Bseq (\x. c * f x)" - find_theorems "Bfun (\_. ?c) _" - from Bfun_const this(2) have "Bseq (\x. inverse c * (c * f x))" by (rule Bseq_mult) - with \c \ 0\ show "Bseq f" by (simp add: divide_simps) + assume "Bseq (\x. c * f x)" + with Bfun_const have "Bseq (\x. inverse c * (c * f x))" + by (rule Bseq_mult) + with \c \ 0\ show "Bseq f" + by (simp add: divide_simps) qed (intro Bseq_mult Bfun_const) -lemma Bseq_subseq: "Bseq (f :: nat \ 'a :: real_normed_vector) \ Bseq (\x. f (g x))" +lemma Bseq_subseq: "Bseq f \ Bseq (\x. f (g x))" + for f :: "nat \ 'a::real_normed_vector" unfolding Bseq_def by auto -lemma Bseq_Suc_iff: "Bseq (\n. f (Suc n)) \ Bseq (f :: nat \ 'a :: real_normed_vector)" +lemma Bseq_Suc_iff: "Bseq (\n. f (Suc n)) \ Bseq f" + for f :: "nat \ 'a::real_normed_vector" using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) lemma increasing_Bseq_subseq_iff: - assumes "\x y. x \ y \ norm (f x :: 'a :: real_normed_vector) \ norm (f y)" "subseq g" - shows "Bseq (\x. f (g x)) \ Bseq f" + assumes "\x y. x \ y \ norm (f x :: 'a::real_normed_vector) \ norm (f y)" "subseq g" + shows "Bseq (\x. f (g x)) \ Bseq f" proof assume "Bseq (\x. f (g x))" - then obtain K where K: "\x. norm (f (g x)) \ K" unfolding Bseq_def by auto + then obtain K where K: "\x. norm (f (g x)) \ K" + unfolding Bseq_def by auto { fix x :: nat from filterlim_subseq[OF assms(2)] obtain y where "g y \ x" by (auto simp: filterlim_at_top eventually_at_top_linorder) - hence "norm (f x) \ norm (f (g y))" using assms(1) by blast + then have "norm (f x) \ norm (f (g y))" + using assms(1) by blast also have "norm (f (g y)) \ K" by (rule K) finally have "norm (f x) \ K" . } - thus "Bseq f" by (rule BseqI') -qed (insert Bseq_subseq[of f g], simp_all) + then show "Bseq f" by (rule BseqI') +qed (use Bseq_subseq[of f g] in simp_all) lemma nonneg_incseq_Bseq_subseq_iff: - assumes "\x. f x \ 0" "incseq (f :: nat \ real)" "subseq g" - shows "Bseq (\x. f (g x)) \ Bseq f" + fixes f :: "nat \ real" + and g :: "nat \ nat" + assumes "\x. f x \ 0" "incseq f" "subseq g" + shows "Bseq (\x. f (g x)) \ Bseq f" using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) -lemma Bseq_eq_bounded: "range f \ {a .. b::real} \ Bseq f" +lemma Bseq_eq_bounded: "range f \ {a..b} \ Bseq f" + for a b :: real 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" +lemma incseq_bounded: "incseq X \ \i. X i \ B \ Bseq X" + for B :: real 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" +lemma decseq_bounded: "decseq X \ \i. B \ X i \ Bseq X" + for B :: real 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\ +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)" +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" +lemma ZfunI: "(\r. 0 < r \ eventually (\x. norm (f x) < r) F) \ Zfun f F" + by (simp add: Zfun_def) + +lemma ZfunD: "Zfun f F \ 0 < r \ eventually (\x. norm (f x) < r) F" + by (simp add: Zfun_def) + +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" @@ -369,28 +403,29 @@ lemma Zfun_imp_Zfun: assumes f: "Zfun f F" - assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) F" + and g: "eventually (\x. norm (g x) \ norm (f x) * K) F" shows "Zfun (\x. g x) F" -proof (cases) - assume K: "0 < K" +proof (cases "0 < K") + case K: True show ?thesis proof (rule ZfunI) - fix r::real assume "0 < r" - hence "0 < r / K" using K by simp + fix r :: real + assume "0 < r" + then have "0 < r / K" using K by simp then have "eventually (\x. norm (f x) < r / K) F" using ZfunD [OF f] by blast with g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) - hence "norm (f x) * K < r" + then have "norm (f x) * K < r" by (simp add: pos_less_divide_eq K) - thus ?case + then show ?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) + case False + then have K: "K \ 0" by (simp only: not_less) show ?thesis proof (rule ZfunI) fix r :: real @@ -406,15 +441,17 @@ 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_le: "Zfun g F \ \x. norm (f x) \ norm (g x) \ Zfun f F" + by (erule Zfun_imp_Zfun [where K = 1]) simp lemma Zfun_add: - assumes f: "Zfun f F" and g: "Zfun g F" + 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 + fix r :: real + assume "0 < r" + then have r: "0 < r / 2" by simp have "eventually (\x. norm (f x) < r/2) F" using f r by (rule ZfunD) moreover @@ -436,14 +473,14 @@ 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" +lemma Zfun_diff: "Zfun f F \ Zfun g F \ Zfun (\x. f x - g x) F" using Zfun_add [of f F "\x. - g x"] by (simp 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" + obtain K where "norm (f x) \ norm x * K" for x using bounded by blast then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) F" by simp @@ -453,12 +490,13 @@ lemma (in bounded_bilinear) Zfun: assumes f: "Zfun f F" - assumes g: "Zfun g F" + and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" proof (rule ZfunI) - fix r::real assume r: "0 < r" + 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" + and norm_le: "norm (x ** y) \ norm x * norm y * K" for x y using pos_bounded by blast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) @@ -481,12 +519,10 @@ qed qed -lemma (in bounded_bilinear) Zfun_left: - "Zfun f F \ Zfun (\x. f x ** a) F" +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" +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] @@ -496,19 +532,22 @@ lemma tendsto_Zfun_iff: "(f \ a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) -lemma tendsto_0_le: "\(f \ 0) F; eventually (\x. norm (g x) \ norm (f x) * K) F\ - \ (g \ 0) F" +lemma tendsto_0_le: + "(f \ 0) F \ eventually (\x. norm (g x) \ norm (f x) * K) F \ (g \ 0) F" by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) + 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" + 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 + fix e :: real + assume "0 < e" + then have 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) @@ -516,9 +555,9 @@ 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"] + and dist_triangle2 [of "g x" "l" "m"] + and dist_triangle3 [of "l" "m" "f x"] + and dist_triangle [of "f x" "m" "g x"] by arith qed qed @@ -533,33 +572,28 @@ 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" +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))" +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_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" +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" +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 tendsto_rabs [tendsto_intros]: "(f \ l) F \ ((\x. \f x\) \ \l\) F" + for l :: real + by (fold real_norm_def) (rule tendsto_norm) lemma continuous_rabs [continuous_intros]: "continuous F f \ continuous F (\x. \f x :: real\)" @@ -569,17 +603,15 @@ "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) +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) + subsection \Topological Monoid\ @@ -606,17 +638,22 @@ lemma tendsto_add_zero: fixes f g :: "_ \ 'b::topological_monoid_add" - shows "\(f \ 0) F; (g \ 0) F\ \ ((\x. f x + g x) \ 0) F" - by (drule (1) tendsto_add, simp) + shows "(f \ 0) F \ (g \ 0) F \ ((\x. f x + g x) \ 0) F" + by (drule (1) tendsto_add) simp lemma tendsto_setsum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" assumes "\i. i \ I \ (f i \ a i) F" shows "((\x. \i\I. f i x) \ (\i\I. a i)) F" proof (cases "finite I") - assume "finite I" thus ?thesis using assms - by (induct, simp, simp add: tendsto_add) -qed simp + case True + then show ?thesis + using assms by induct (simp_all add: tendsto_add) +next + case False + then show ?thesis + by simp +qed lemma continuous_setsum [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" @@ -629,10 +666,13 @@ unfolding continuous_on_def by (auto intro: tendsto_setsum) instance nat :: topological_comm_monoid_add - proof qed (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) + by standard + (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_add - proof qed (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) + by standard + (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) + subsubsection \Topological group\ @@ -640,7 +680,7 @@ assumes tendsto_uminus_nhds: "(uminus \ - a) (nhds a)" begin -lemma tendsto_minus [tendsto_intros]: "(f \ a) F \ ((\x. - f x) \ -a) F" +lemma tendsto_minus [tendsto_intros]: "(f \ a) F \ ((\x. - f x) \ - a) F" by (rule filterlim_compose[OF tendsto_uminus_nhds]) end @@ -649,29 +689,26 @@ instance topological_ab_group_add < topological_comm_monoid_add .. -lemma continuous_minus [continuous_intros]: - fixes f :: "'a::t2_space \ 'b::topological_group_add" - shows "continuous F f \ continuous F (\x. - f x)" +lemma continuous_minus [continuous_intros]: "continuous F f \ continuous F (\x. - f x)" + for f :: "'a::t2_space \ 'b::topological_group_add" unfolding continuous_def by (rule tendsto_minus) -lemma continuous_on_minus [continuous_intros]: - fixes f :: "_ \ 'b::topological_group_add" - shows "continuous_on s f \ continuous_on s (\x. - f x)" +lemma continuous_on_minus [continuous_intros]: "continuous_on s f \ continuous_on s (\x. - f x)" + for f :: "_ \ 'b::topological_group_add" unfolding continuous_on_def by (auto intro: tendsto_minus) -lemma tendsto_minus_cancel: - fixes a :: "'a::topological_group_add" - shows "((\x. - f x) \ - a) F \ (f \ a) F" - by (drule tendsto_minus, simp) +lemma tendsto_minus_cancel: "((\x. - f x) \ - a) F \ (f \ a) F" + for a :: "'a::topological_group_add" + by (drule tendsto_minus) simp lemma tendsto_minus_cancel_left: - "(f \ - (y::_::topological_group_add)) F \ ((\x. - f x) \ y) F" + "(f \ - (y::_::topological_group_add)) 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::topological_group_add" - shows "\(f \ a) F; (g \ b) F\ \ ((\x. f x - g x) \ a - b) F" + shows "(f \ a) F \ (g \ b) F \ ((\x. f x - g x) \ a - b) F" using tendsto_add [of f a F "\x. - g x" "- b"] by (simp add: tendsto_minus) lemma continuous_diff [continuous_intros]: @@ -689,7 +726,8 @@ instance real_normed_vector < topological_ab_group_add proof - fix a b :: 'a show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" + fix a b :: 'a + show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" unfolding tendsto_Zfun_iff add_diff_add using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (intro Zfun_add) @@ -702,32 +740,28 @@ lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real] + subsubsection \Linear operators and multiplication\ -lemma linear_times: - fixes c::"'a::real_algebra" shows "linear (\x. c * x)" +lemma linear_times: "linear (\x. c * x)" + for c :: "'a::real_algebra" by (auto simp: linearI distrib_left) -lemma (in bounded_linear) tendsto: - "(g \ a) F \ ((\x. f (g x)) \ f a) F" +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))" +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))" +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_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) + "(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)" @@ -739,7 +773,7 @@ lemma (in bounded_bilinear) tendsto_zero: assumes f: "(f \ 0) F" - assumes g: "(g \ 0) F" + and g: "(g \ 0) F" shows "((\x. f x ** g x) \ 0) F" using tendsto [OF f g] by (simp add: zero_left) @@ -760,15 +794,13 @@ lemmas tendsto_mult [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_mult] -lemma tendsto_mult_left: - fixes c::"'a::real_normed_algebra" - shows "(f \ l) F \ ((\x. c * (f x)) \ c * l) F" -by (rule tendsto_mult [OF tendsto_const]) - -lemma tendsto_mult_right: - fixes c::"'a::real_normed_algebra" - shows "(f \ l) F \ ((\x. (f x) * c) \ l * c) F" -by (rule tendsto_mult [OF _ tendsto_const]) +lemma tendsto_mult_left: "(f \ l) F \ ((\x. c * (f x)) \ c * l) F" + for c :: "'a::real_normed_algebra" + by (rule tendsto_mult [OF tendsto_const]) + +lemma tendsto_mult_right: "(f \ l) F \ ((\x. (f x) * c) \ l * c) F" + for c :: "'a::real_normed_algebra" + by (rule tendsto_mult [OF _ tendsto_const]) lemmas continuous_of_real [continuous_intros] = bounded_linear.continuous [OF bounded_linear_of_real] @@ -797,14 +829,12 @@ 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" +lemma tendsto_power [tendsto_intros]: "(f \ a) F \ ((\x. f x ^ n) \ a ^ n) F" + for f :: "'a \ 'b::{power,real_normed_algebra}" by (induct n) (simp_all add: 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)" +lemma continuous_power [continuous_intros]: "continuous F f \ continuous F (\x. (f x)^n)" + for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" unfolding continuous_def by (rule tendsto_power) lemma continuous_on_power [continuous_intros]: @@ -817,9 +847,13 @@ 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, simp add: tendsto_mult) -qed simp + case True + then show ?thesis using assms + by induct (simp_all add: tendsto_mult) +next + case False + then show ?thesis by simp +qed lemma continuous_setprod [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::{real_normed_algebra,comm_ring_1}" @@ -832,20 +866,20 @@ unfolding continuous_on_def by (auto intro: tendsto_setprod) lemma tendsto_of_real_iff: - "((\x. of_real (f x) :: 'a :: real_normed_div_algebra) \ of_real c) F \ (f \ c) F" + "((\x. of_real (f x) :: 'a::real_normed_div_algebra) \ of_real c) F \ (f \ c) F" unfolding tendsto_iff by simp lemma tendsto_add_const_iff: - "((\x. c + f x :: 'a :: real_normed_vector) \ c + d) F \ (f \ d) F" + "((\x. c + f x :: 'a::real_normed_vector) \ c + d) F \ (f \ d) F" using tendsto_add[OF tendsto_const[of c], of f d] - tendsto_add[OF tendsto_const[of "-c"], of "\x. c + f x" "c + d"] by auto + and tendsto_add[OF tendsto_const[of "-c"], of "\x. c + f x" "c + d"] by auto subsubsection \Inverse and division\ lemma (in bounded_bilinear) Zfun_prod_Bfun: assumes f: "Zfun f F" - assumes g: "Bfun g F" + and g: "Bfun g F" shows "Zfun (\x. f x ** g x) F" proof - obtain K where K: "0 \ K" @@ -860,8 +894,7 @@ 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) + 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)" . @@ -872,14 +905,15 @@ lemma (in bounded_bilinear) Bfun_prod_Zfun: assumes f: "Bfun f F" - assumes g: "Zfun g F" + and 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) + shows "r \ norm x \ 0 < r \ norm (inverse x) \ inverse r" + apply (subst nonzero_norm_inverse) + apply clarsimp apply (erule (1) le_imp_inverse_le) done @@ -890,38 +924,40 @@ 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 blast + then have "\r>0. r < norm a" by (rule dense) + then obtain r where r1: "0 < r" and r2: "r < norm a" + by blast have "eventually (\x. dist (f x) a < r) F" using tendstoD [OF f r1] by blast - hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" + then have "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" proof eventually_elim case (elim x) - hence 1: "norm (f x - a) < r" + then have 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))" + then have 2: "f x \ 0" using r2 by auto + then have "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 + show "0 < norm a - r" + using r2 by simp 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 + 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) + then show ?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" + and a: "a \ 0" shows "((\x. inverse (f x)) \ inverse a) F" proof - from a have "0 < norm a" by simp @@ -942,43 +978,49 @@ lemma continuous_inverse: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" - assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" + 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" + 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" + 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_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" - assumes "continuous_on s f" and "\x\s. f x \ 0" + 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 (blast 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" + 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" + 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" + 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) @@ -990,36 +1032,40 @@ lemma continuous_on_divide[continuous_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" + 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 (blast 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" +lemma tendsto_sgn [tendsto_intros]: "(f \ l) F \ l \ 0 \ ((\x. sgn (f x)) \ sgn l) F" + for l :: "'a::real_normed_vector" 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" + 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" + 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" + 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_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" - assumes "continuous_on s f" and "\x\s. f x \ 0" + 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 (blast intro: tendsto_sgn) @@ -1029,35 +1075,40 @@ 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 + fix P :: "'a \ bool" + fix b assume *: "\r>c. eventually (\x. r \ norm (f x)) F" - and P: "\x. b \ norm x \ P x" + assume 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)" + case (elim x) with P show "P (f x)" by auto qed qed force lemma not_tendsto_and_filterlim_at_infinity: + fixes c :: "'a::real_normed_vector" assumes "F \ bot" - assumes "(f \ (c :: 'a :: real_normed_vector)) F" - assumes "filterlim f at_infinity F" - shows False + and "(f \ c) F" + and "filterlim f at_infinity F" + shows False proof - from tendstoD[OF assms(2), of "1/2"] - have "eventually (\x. dist (f x) c < 1/2) F" by simp - moreover from filterlim_at_infinity[of "norm c" f F] assms(3) - have "eventually (\x. norm (f x) \ norm c + 1) F" by simp + have "eventually (\x. dist (f x) c < 1/2) F" + by simp + moreover + from filterlim_at_infinity[of "norm c" f F] assms(3) + have "eventually (\x. norm (f x) \ norm c + 1) F" by simp ultimately have "eventually (\x. False) F" proof eventually_elim - fix x assume A: "dist (f x) c < 1/2" and B: "norm (f x) \ norm c + 1" - note B + fix x + assume A: "dist (f x) c < 1/2" + assume "norm (f x) \ norm c + 1" also have "norm (f x) = dist (f x) 0" by simp - also have "... \ dist (f x) c + dist c 0" by (rule dist_triangle) + also have "\ \ dist (f x) c + dist c 0" by (rule dist_triangle) finally show False using A by simp qed with assms show False by simp @@ -1065,83 +1116,97 @@ lemma filterlim_at_infinity_imp_not_convergent: assumes "filterlim f at_infinity sequentially" - shows "\convergent f" + shows "\ convergent f" by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms]) (simp_all add: convergent_LIMSEQ_iff) lemma filterlim_at_infinity_imp_eventually_ne: assumes "filterlim f at_infinity F" - shows "eventually (\z. f z \ c) F" + shows "eventually (\z. f z \ c) F" proof - - have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all + have "norm c + 1 > 0" + by (intro add_nonneg_pos) simp_all with filterlim_at_infinity[OF order.refl, of f F] assms - have "eventually (\z. norm (f z) \ norm c + 1) F" by blast - thus ?thesis by eventually_elim auto + have "eventually (\z. norm (f z) \ norm c + 1) F" + by blast + then show ?thesis + by eventually_elim auto qed lemma tendsto_of_nat [tendsto_intros]: - "filterlim (of_nat :: nat \ 'a :: real_normed_algebra_1) at_infinity sequentially" + "filterlim (of_nat :: nat \ 'a::real_normed_algebra_1) at_infinity sequentially" proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) fix r :: real assume r: "r > 0" define n where "n = nat \r\" - from r have n: "\m\n. of_nat m \ r" unfolding n_def by linarith + from r have n: "\m\n. of_nat m \ r" + unfolding n_def by linarith from eventually_ge_at_top[of n] show "eventually (\m. norm (of_nat m :: 'a) \ r) sequentially" - by eventually_elim (insert n, simp_all) + by eventually_elim (use n in simp_all) qed 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"}. - + 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_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d::'a::real_normed_vector)" +lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d)" + for a d :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g="\x. x + d"]) - (auto intro!: tendsto_eq_intros filterlim_ident) - -lemma filtermap_nhds_minus: "filtermap (\x. - x) (nhds a) = nhds (- a::'a::real_normed_vector)" + (auto intro!: tendsto_eq_intros filterlim_ident) + +lemma filtermap_nhds_minus: "filtermap (\x. - x) (nhds a) = nhds (- a)" + for a :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g=uminus]) - (auto intro!: tendsto_eq_intros filterlim_ident) - -lemma filtermap_at_shift: "filtermap (\x. x - d) (at a) = at (a - d::'a::real_normed_vector)" + (auto intro!: tendsto_eq_intros filterlim_ident) + +lemma filtermap_at_shift: "filtermap (\x. x - d) (at a) = at (a - d)" + for 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)" +lemma filtermap_at_right_shift: "filtermap (\x. x - d) (at_right a) = at_right (a - d)" + for 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)" +lemma at_right_to_0: "at_right a = filtermap (\x. x + a) (at_right 0)" + for a :: real 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)" + "filterlim f F (at_right a) \ filterlim (\x. f (x + a)) F (at_right 0)" + for a :: real 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)" + "eventually P (at_right a) \ eventually (\x. P (x + a)) (at_right 0)" + for a :: real 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)" +lemma filtermap_at_minus: "filtermap (\x. - x) (at a) = at (- a)" + for 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))" +lemma at_left_minus: "at_left a = filtermap (\x. - x) (at_right (- a))" + for a :: real 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))" +lemma at_right_minus: "at_right a = filtermap (\x. - x) (at_left (- a))" + for a :: real 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))" + "filterlim f F (at_left a) \ filterlim (\x. f (- x)) F (at_right (-a))" + for a :: real 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))" + "eventually P (at_left a) \ eventually (\x. P (- x)) (at_right (-a))" + for a :: real unfolding at_left_minus[of a] by (simp add: eventually_filtermap) lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" @@ -1167,7 +1232,7 @@ 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] + and 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)" @@ -1176,7 +1241,8 @@ 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" + 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)" @@ -1188,41 +1254,56 @@ 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" + 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) + using * \0 < r\ + by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) qed qed lemma tendsto_add_filterlim_at_infinity: - assumes "(f \ (c :: 'b :: real_normed_vector)) (F :: 'a filter)" - assumes "filterlim g at_infinity F" - shows "filterlim (\x. f x + g x) at_infinity F" + fixes c :: "'b::real_normed_vector" + and F :: "'a filter" + assumes "(f \ c) F" + and "filterlim g at_infinity F" + shows "filterlim (\x. f x + g x) at_infinity F" proof (subst filterlim_at_infinity[OF order_refl], safe) - fix r :: real assume r: "r > 0" - from assms(1) have "((\x. norm (f x)) \ norm c) F" by (rule tendsto_norm) - hence "eventually (\x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all - moreover from r have "r + norm c + 1 > 0" by (intro add_pos_nonneg) simp_all + fix r :: real + assume r: "r > 0" + from assms(1) have "((\x. norm (f x)) \ norm c) F" + by (rule tendsto_norm) + then have "eventually (\x. norm (f x) < norm c + 1) F" + by (rule order_tendstoD) simp_all + moreover from r have "r + norm c + 1 > 0" + by (intro add_pos_nonneg) simp_all with assms(2) have "eventually (\x. norm (g x) \ r + norm c + 1) F" - unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all + unfolding filterlim_at_infinity[OF order_refl] + by (elim allE[of _ "r + norm c + 1"]) simp_all ultimately show "eventually (\x. norm (f x + g x) \ r) F" proof eventually_elim - fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \ norm (g x)" - from A B have "r \ norm (g x) - norm (f x)" by simp - also have "norm (g x) - norm (f x) \ norm (g x + f x)" by (rule norm_diff_ineq) - finally show "r \ norm (f x + g x)" by (simp add: add_ac) + fix x :: 'a + assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \ norm (g x)" + from A B have "r \ norm (g x) - norm (f x)" + by simp + also have "norm (g x) - norm (f x) \ norm (g x + f x)" + by (rule norm_diff_ineq) + finally show "r \ norm (f x + g x)" + by (simp add: add_ac) qed qed lemma tendsto_add_filterlim_at_infinity': + fixes c :: "'b::real_normed_vector" + and F :: "'a filter" assumes "filterlim f at_infinity F" - assumes "(g \ (c :: 'b :: real_normed_vector)) (F :: 'a filter)" - shows "filterlim (\x. f x + g x) at_infinity F" + and "(g \ c) F" + shows "filterlim (\x. f x + g x) at_infinity F" by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+ lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)" @@ -1272,7 +1353,8 @@ shows "filterlim inverse at_infinity (at (0::'a))" unfolding filterlim_at_infinity[OF order_refl] proof safe - fix r :: real assume "0 < r" + 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"]) @@ -1290,7 +1372,7 @@ 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) + 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" @@ -1301,36 +1383,40 @@ qed lemma tendsto_mult_filterlim_at_infinity: - assumes "F \ bot" "(f \ (c :: 'a :: real_normed_field)) F" "c \ 0" + fixes c :: "'a::real_normed_field" + assumes "F \ bot" "(f \ c) F" "c \ 0" assumes "filterlim g at_infinity F" - shows "filterlim (\x. f x * g x) at_infinity F" + shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ inverse c * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) - hence "filterlim (\x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" - unfolding filterlim_at using assms + then have "filterlim (\x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" + unfolding filterlim_at + using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) - thus ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all + then show ?thesis + by (subst filterlim_inverse_at_iff[symmetric]) simp_all 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) -lemma mult_nat_left_at_top: "c > 0 \ filterlim (\x. c * x :: nat) at_top sequentially" - by (rule filterlim_subseq) (auto simp: subseq_def) - -lemma mult_nat_right_at_top: "c > 0 \ filterlim (\x. x * c :: nat) at_top sequentially" +lemma mult_nat_left_at_top: "c > 0 \ filterlim (\x. c * x) at_top sequentially" + for c :: nat by (rule filterlim_subseq) (auto simp: subseq_def) -lemma at_to_infinity: - fixes x :: "'a :: {real_normed_field,field}" - shows "(at (0::'a)) = filtermap inverse at_infinity" +lemma mult_nat_right_at_top: "c > 0 \ filterlim (\x. x * c) at_top sequentially" + for c :: nat + by (rule filterlim_subseq) (auto simp: subseq_def) + +lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity" proof (rule antisym) have "(inverse \ (0::'a)) at_infinity" by (fact tendsto_inverse_0) then show "filtermap inverse at_infinity \ at (0::'a)" apply (simp add: le_principal eventually_filtermap eventually_at_infinity filterlim_def at_within_def) - apply (rule_tac x="1" in exI, auto) + apply (rule_tac x="1" in exI) + apply auto done next have "filtermap inverse (filtermap inverse (at (0::'a))) \ filtermap inverse at_infinity" @@ -1341,38 +1427,39 @@ qed lemma lim_at_infinity_0: - fixes l :: "'a :: {real_normed_field,field}" - shows "(f \ l) at_infinity \ ((f o inverse) \ l) (at (0::'a))" -by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) + fixes l :: "'a::{real_normed_field,field}" + shows "(f \ l) at_infinity \ ((f \ inverse) \ l) (at (0::'a))" + by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) lemma lim_zero_infinity: - fixes l :: "'a :: {real_normed_field,field}" + fixes l :: "'a::{real_normed_field,field}" shows "((\x. f(1 / x)) \ l) (at (0::'a)) \ (f \ l) at_infinity" -by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) + by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) 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}. - + 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" + assumes f: "(f \ c) F" + and c: "0 < c" + and 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" + 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_mono - simp: dist_real_def abs_real_def split: if_split_asm) + simp: dist_real_def abs_real_def split: if_split_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" + case (elim 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" @@ -1382,18 +1469,19 @@ 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" + and 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" + 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" + case (elim 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" @@ -1402,25 +1490,32 @@ qed lemma filterlim_tendsto_pos_mult_at_bot: - assumes "(f \ c) F" "0 < (c::real)" "filterlim g at_bot F" + fixes c :: real + assumes "(f \ c) F" "0 < c" "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_neg_mult_at_bot: - assumes c: "(f \ c) F" "(c::real) < 0" and g: "filterlim g at_top F" + fixes c :: real + assumes c: "(f \ c) F" "c < 0" and g: "filterlim g at_top F" shows "LIM x F. f x * g x :> at_bot" using c filterlim_tendsto_pos_mult_at_top[of "\x. - f x" "- c" F, OF _ _ g] unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp lemma filterlim_pow_at_top: fixes f :: "real \ real" - assumes "0 < n" and f: "LIM x F. f x :> at_top" + assumes "0 < n" + and f: "LIM x F. f x :> at_top" shows "LIM x F. (f x)^n :: real :> at_top" -using \0 < n\ proof (induct n) + using \0 < n\ +proof (induct n) + case 0 + then show ?case by simp +next case (Suc n) with f show ?case by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top) -qed simp +qed lemma filterlim_pow_at_bot_even: fixes f :: "real \ real" @@ -1434,11 +1529,12 @@ lemma filterlim_tendsto_add_at_top: assumes f: "(f \ c) F" - assumes g: "LIM x F. g x :> at_top" + and 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" + 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_mono simp: dist_real_def) moreover from g have "eventually (\x. Z - (c - 1) \ g x) F" @@ -1450,18 +1546,19 @@ 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" + and 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" + and 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" + 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" @@ -1473,34 +1570,43 @@ lemma tendsto_divide_0: fixes f :: "_ \ 'a::{real_normed_div_algebra, division_ring}" assumes f: "(f \ c) F" - assumes g: "LIM x F. g x :> at_infinity" + and 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) + 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 0 + then show ?case by simp +next case (Suc n) - have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" - by (simp add: field_simps of_nat_Suc x) + from x have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" + by (simp add: field_simps) also have "\ \ (x + 1)^Suc n" using Suc x by (simp add: mult_left_mono) finally show ?case . -qed simp +qed 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" + 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 + 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" @@ -1512,48 +1618,48 @@ subsection \Floor and Ceiling\ lemma eventually_floor_less: - fixes f::"'a \ 'b::{order_topology, floor_ceiling}" + fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" - assumes l: "l \ \" + and l: "l \ \" shows "\\<^sub>F x in F. of_int (floor l) < f x" by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l) lemma eventually_less_ceiling: - fixes f::"'a \ 'b::{order_topology, floor_ceiling}" + fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" - assumes l: "l \ \" + and l: "l \ \" shows "\\<^sub>F x in F. f x < of_int (ceiling l)" by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le) lemma eventually_floor_eq: - fixes f::"'a \ 'b::{order_topology, floor_ceiling}" + fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" - assumes l: "l \ \" + and l: "l \ \" shows "\\<^sub>F x in F. floor (f x) = floor l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma eventually_ceiling_eq: - fixes f::"'a \ 'b::{order_topology, floor_ceiling}" + fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" - assumes l: "l \ \" + and l: "l \ \" shows "\\<^sub>F x in F. ceiling (f x) = ceiling l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma tendsto_of_int_floor: - fixes f::"'a \ 'b::{order_topology, floor_ceiling}" + fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" - assumes "l \ \" - shows "((\x. of_int (floor (f x))::'c::{ring_1, topological_space}) \ of_int (floor l)) F" + and "l \ \" + shows "((\x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \ of_int (floor l)) F" using eventually_floor_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma tendsto_of_int_ceiling: - fixes f::"'a \ 'b::{order_topology, floor_ceiling}" + fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" - assumes "l \ \" - shows "((\x. of_int (ceiling (f x))::'c::{ring_1, topological_space}) \ of_int (ceiling l)) F" + and "l \ \" + shows "((\x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \ of_int (ceiling l)) F" using eventually_ceiling_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) @@ -1580,69 +1686,64 @@ shows "(X \ L) = (\r>0. \no. \n \ no. norm (X n - L) < r)" unfolding lim_sequentially 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" +lemma LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. norm (X n - L) < r) \ X \ L" + for L :: "'a::real_normed_vector" + by (simp add: LIMSEQ_iff) + +lemma LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. norm (X n - L) < r" + for L :: "'a::real_normed_vector" + 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 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))" +lemma Bseq_inverse_lemma: "r \ norm x \ 0 < r \ norm (inverse x) \ inverse r" + for x :: "'a::real_normed_div_algebra" + apply (subst nonzero_norm_inverse, clarsimp) + apply (erule (1) le_imp_inverse_le) + done + +lemma Bseq_inverse: "X \ a \ a \ 0 \ Bseq (\n. inverse (X n))" + for a :: "'a::real_normed_div_algebra" by (rule Bfun_inverse) -text\Transformation of limit.\ - -lemma Lim_transform: - fixes a b :: "'a::real_normed_vector" - shows "\(g \ a) F; ((\x. f x - g x) \ 0) F\ \ (f \ a) F" + +text \Transformation of limit.\ + +lemma Lim_transform: "(g \ a) F \ ((\x. f x - g x) \ 0) F \ (f \ a) F" + for a b :: "'a::real_normed_vector" using tendsto_add [of g a F "\x. f x - g x" 0] by simp -lemma Lim_transform2: - fixes a b :: "'a::real_normed_vector" - shows "\(f \ a) F; ((\x. f x - g x) \ 0) F\ \ (g \ a) F" +lemma Lim_transform2: "(f \ a) F \ ((\x. f x - g x) \ 0) F \ (g \ a) F" + for a b :: "'a::real_normed_vector" by (erule Lim_transform) (simp add: tendsto_minus_cancel) -proposition Lim_transform_eq: - fixes a :: "'a::real_normed_vector" - shows "((\x. f x - g x) \ 0) F \ (f \ a) F \ (g \ a) F" -using Lim_transform Lim_transform2 by blast +proposition Lim_transform_eq: "((\x. f x - g x) \ 0) F \ (f \ a) F \ (g \ a) F" + for a :: "'a::real_normed_vector" + using Lim_transform Lim_transform2 by blast lemma Lim_transform_eventually: "eventually (\x. f x = g x) net \ (f \ l) net \ (g \ l) net" apply (rule topological_tendstoI) apply (drule (2) topological_tendstoD) - apply (erule (1) eventually_elim2, simp) + apply (erule (1) eventually_elim2) + apply simp done lemma Lim_transform_within: assumes "(f \ l) (at x within S)" and "0 < d" - and "\x'. \x'\S; 0 < dist x' x; dist x' x < d\ \ f x' = g x'" + and "\x'. x'\S \ 0 < dist x' x \ dist x' x < d \ f x' = g x'" shows "(g \ l) (at x within S)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at x within S)" using assms by (auto simp: eventually_at) - show "(f \ l) (at x within S)" by fact + show "(f \ l) (at x within S)" + by fact qed -text\Common case assuming being away from some crucial point like 0.\ - +text \Common case assuming being away from some crucial point like 0.\ lemma Lim_transform_away_within: fixes a b :: "'a::t1_space" assumes "a \ b" @@ -1650,26 +1751,26 @@ and "(f \ l) (at a within S)" shows "(g \ l) (at a within S)" proof (rule Lim_transform_eventually) - show "(f \ l) (at a within S)" by fact + show "(f \ l) (at a within S)" + by fact show "eventually (\x. f x = g x) (at a within S)" unfolding eventually_at_topological - by (rule exI [where x="- {b}"], simp add: open_Compl assms) + by (rule exI [where x="- {b}"]) (simp add: open_Compl assms) qed lemma Lim_transform_away_at: fixes a b :: "'a::t1_space" - assumes ab: "a\b" + assumes ab: "a \ b" and fg: "\x. x \ a \ x \ b \ f x = g x" and fl: "(f \ l) (at a)" shows "(g \ l) (at a)" using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp -text\Alternatively, within an open set.\ - +text \Alternatively, within an open set.\ lemma Lim_transform_within_open: assumes "(f \ l) (at a within T)" and "open s" and "a \ s" - and "\x. \x\s; x \ a\ \ f x = g x" + and "\x. x\s \ x \ a \ f x = g x" shows "(g \ l) (at a within T)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at a within T)" @@ -1678,7 +1779,8 @@ show "(f \ l) (at a within T)" by fact qed -text\A congruence rule allowing us to transform limits assuming not at point.\ + +text \A congruence rule allowing us to transform limits assuming not at point.\ (* FIXME: Only one congruence rule for tendsto can be used at a time! *) @@ -1697,35 +1799,32 @@ shows "((\x. f x) \ x) (at a) \ ((g \ y) (at a))" unfolding tendsto_def eventually_at_topological using assms by 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" + +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" +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" + 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" +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" +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 @@ -1735,46 +1834,57 @@ lemma LIMSEQ_Suc_n_over_n: "(\n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially" - using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps) + using eventually_gt_at_top[of "0::nat"] + by eventually_elim (simp add: field_simps) have "(\n. 1 + inverse (of_nat n) :: 'a) \ 1 + 0" by (intro tendsto_add tendsto_const lim_inverse_n) - thus "(\n. 1 + inverse (of_nat n) :: 'a) \ 1" by simp + then show "(\n. 1 + inverse (of_nat n) :: 'a) \ 1" + by simp qed lemma LIMSEQ_n_over_Suc_n: "(\n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. inverse (of_nat (Suc n) / of_nat n :: 'a) = - of_nat n / of_nat (Suc n)) sequentially" + of_nat n / of_nat (Suc n)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps del: of_nat_Suc) have "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ inverse 1" by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all - thus "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ 1" by simp + then show "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ 1" + by simp qed + subsection \Convergence on sequences\ lemma convergent_cong: assumes "eventually (\x. f x = g x) sequentially" - shows "convergent f \ convergent g" - unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl) + shows "convergent f \ convergent g" + unfolding convergent_def + by (subst filterlim_cong[OF refl refl assms]) (rule refl) lemma convergent_Suc_iff: "convergent (\n. f (Suc n)) \ convergent f" by (auto simp: convergent_def LIMSEQ_Suc_iff) lemma convergent_ignore_initial_segment: "convergent (\n. f (n + m)) = convergent f" -proof (induction m arbitrary: f) +proof (induct m arbitrary: f) + case 0 + then show ?case by simp +next case (Suc m) - have "convergent (\n. f (n + Suc m)) \ convergent (\n. f (Suc n + m))" by simp - also have "\ \ convergent (\n. f (n + m))" by (rule convergent_Suc_iff) - also have "\ \ convergent f" by (rule Suc) + have "convergent (\n. f (n + Suc m)) \ convergent (\n. f (Suc n + m))" + by simp + also have "\ \ convergent (\n. f (n + m))" + by (rule convergent_Suc_iff) + also have "\ \ convergent f" + by (rule Suc) finally show ?case . -qed simp_all +qed lemma convergent_add: fixes X Y :: "nat \ 'a::real_normed_vector" assumes "convergent (\n. X n)" - assumes "convergent (\n. Y n)" + and "convergent (\n. Y n)" shows "convergent (\n. X n + Y n)" using assms unfolding convergent_def by (blast intro: tendsto_add) @@ -1783,9 +1893,14 @@ 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) + case True + then show ?thesis + using assms by (induct A set: finite) (simp_all add: convergent_const convergent_add) +next + case False + then show ?thesis + by (simp add: convergent_const) +qed lemma (in bounded_linear) convergent: assumes "convergent (\n. X n)" @@ -1793,17 +1908,18 @@ using assms unfolding convergent_def by (blast intro: tendsto) lemma (in bounded_bilinear) convergent: - assumes "convergent (\n. X n)" and "convergent (\n. Y n)" + assumes "convergent (\n. X n)" + and "convergent (\n. Y n)" shows "convergent (\n. X n ** Y n)" using assms unfolding convergent_def by (blast 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 +lemma convergent_minus_iff: "convergent X \ convergent (\n. - X n)" + for X :: "nat \ 'a::real_normed_vector" + apply (simp add: convergent_def) + apply (auto dest: tendsto_minus) + apply (drule tendsto_minus) + apply auto + done lemma convergent_diff: fixes X Y :: "nat \ 'a::real_normed_vector" @@ -1814,57 +1930,64 @@ lemma convergent_norm: assumes "convergent f" - shows "convergent (\n. norm (f n))" + shows "convergent (\n. norm (f n))" proof - - from assms have "f \ lim f" by (simp add: convergent_LIMSEQ_iff) - hence "(\n. norm (f n)) \ norm (lim f)" by (rule tendsto_norm) - thus ?thesis by (auto simp: convergent_def) + from assms have "f \ lim f" + by (simp add: convergent_LIMSEQ_iff) + then have "(\n. norm (f n)) \ norm (lim f)" + by (rule tendsto_norm) + then show ?thesis + by (auto simp: convergent_def) qed lemma convergent_of_real: - "convergent f \ convergent (\n. of_real (f n) :: 'a :: real_normed_algebra_1)" + "convergent f \ convergent (\n. of_real (f n) :: 'a::real_normed_algebra_1)" unfolding convergent_def by (blast intro!: tendsto_of_real) lemma convergent_add_const_iff: - "convergent (\n. c + f n :: 'a :: real_normed_vector) \ convergent f" + "convergent (\n. c + f n :: 'a::real_normed_vector) \ convergent f" proof assume "convergent (\n. c + f n)" - from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp + from convergent_diff[OF this convergent_const[of c]] show "convergent f" + by simp next assume "convergent f" - from convergent_add[OF convergent_const[of c] this] show "convergent (\n. c + f n)" by simp + from convergent_add[OF convergent_const[of c] this] show "convergent (\n. c + f n)" + by simp qed lemma convergent_add_const_right_iff: - "convergent (\n. f n + c :: 'a :: real_normed_vector) \ convergent f" + "convergent (\n. f n + c :: 'a::real_normed_vector) \ convergent f" using convergent_add_const_iff[of c f] by (simp add: add_ac) lemma convergent_diff_const_right_iff: - "convergent (\n. f n - c :: 'a :: real_normed_vector) \ convergent f" + "convergent (\n. f n - c :: 'a::real_normed_vector) \ convergent f" using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac) lemma convergent_mult: fixes X Y :: "nat \ 'a::real_normed_field" assumes "convergent (\n. X n)" - assumes "convergent (\n. Y n)" + and "convergent (\n. Y n)" shows "convergent (\n. X n * Y n)" using assms unfolding convergent_def by (blast intro: tendsto_mult) lemma convergent_mult_const_iff: assumes "c \ 0" - shows "convergent (\n. c * f n :: 'a :: real_normed_field) \ convergent f" + shows "convergent (\n. c * f n :: 'a::real_normed_field) \ convergent f" proof assume "convergent (\n. c * f n)" from assms convergent_mult[OF this convergent_const[of "inverse c"]] show "convergent f" by (simp add: field_simps) next assume "convergent f" - from convergent_mult[OF convergent_const[of c] this] show "convergent (\n. c * f n)" by simp + from convergent_mult[OF convergent_const[of c] this] show "convergent (\n. c * f n)" + by simp qed lemma convergent_mult_const_right_iff: + fixes c :: "'a::real_normed_field" assumes "c \ 0" - shows "convergent (\n. (f n :: 'a :: real_normed_field) * c) \ convergent f" + shows "convergent (\n. f n * c) \ convergent f" using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac) lemma convergent_imp_Bseq: "convergent f \ Bseq f" @@ -1874,60 +1997,66 @@ text \A monotone sequence converges to its least upper bound.\ lemma LIMSEQ_incseq_SUP: - fixes X :: "nat \ 'a::{conditionally_complete_linorder, linorder_topology}" + fixes X :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology}" assumes u: "bdd_above (range X)" - assumes X: "incseq X" + and X: "incseq X" shows "X \ (SUP i. X i)" by (rule order_tendstoI) - (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) + (auto simp: eventually_sequentially u less_cSUP_iff + intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) lemma LIMSEQ_decseq_INF: fixes X :: "nat \ 'a::{conditionally_complete_linorder, linorder_topology}" assumes u: "bdd_below (range X)" - assumes X: "decseq X" + and X: "decseq X" shows "X \ (INF i. X i)" by (rule order_tendstoI) - (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) - -text\Main monotonicity theorem\ - -lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent (X::nat\real)" - by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below) - -lemma Bseq_mono_convergent: "Bseq X \ (\m n. m \ n \ X m \ X n) \ convergent (X::nat\real)" + (auto simp: eventually_sequentially u cINF_less_iff + intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) + +text \Main monotonicity theorem.\ + +lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent X" + for X :: "nat \ real" + by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP + dest: Bseq_bdd_above Bseq_bdd_below) + +lemma Bseq_mono_convergent: "Bseq X \ (\m n. m \ n \ X m \ X n) \ convergent X" + for X :: "nat \ real" by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def) -lemma monoseq_imp_convergent_iff_Bseq: "monoseq (f :: nat \ real) \ convergent f \ Bseq f" +lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \ convergent f \ Bseq f" + for f :: "nat \ real" using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast lemma Bseq_monoseq_convergent'_inc: - "Bseq (\n. f (n + M) :: real) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" + fixes f :: "nat \ real" + shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Bseq_monoseq_convergent'_dec: - "Bseq (\n. f (n + M) :: real) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" + fixes f :: "nat \ real" + shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) - (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) - -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)" + (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) + +lemma Cauchy_iff: "Cauchy X \ (\e>0. \M. \m\M. \n\M. norm (X m - X n) < e)" + for X :: "nat \ 'a::real_normed_vector" 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 CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e) \ Cauchy X" + for X :: "nat \ 'a::real_normed_vector" + by (simp add: Cauchy_iff) + +lemma CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e" + for X :: "nat \ 'a::real_normed_vector" + by (simp add: Cauchy_iff) lemma incseq_convergent: fixes X :: "nat \ real" - assumes "incseq X" and "\i. X i \ B" + 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] @@ -1939,7 +2068,8 @@ lemma decseq_convergent: fixes X :: "nat \ real" - assumes "decseq X" and "\i. B \ X i" + 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] @@ -1949,69 +2079,85 @@ 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 +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 + apply (drule spec) + apply (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" +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 \ x \ 1 \ Bseq (\n. x ^ n)" + for x :: real + 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: "0 \ x \ x \ 1 \ monoseq (\n. x ^ n)" + for x :: real + apply (clarify intro!: mono_SucI2) + apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing) + apply auto + done + +lemma convergent_realpow: "0 \ x \ x \ 1 \ convergent (\n. x ^ n)" + for x :: real + by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) + +lemma LIMSEQ_inverse_realpow_zero: "1 < x \ (\n. inverse (x ^ n)) \ 0" + for x :: real 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" + fixes x :: real + assumes "0 \ x" "x < 1" + shows "(\n. x ^ n) \ 0" +proof (cases "x = 0") + case False + with \0 \ x\ have x0: "0 < x" by simp + then have "1 < inverse x" + using \x < 1\ by (rule one_less_inverse) + then have "(\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) - -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 + then show ?thesis by (simp add: power_inverse) +next + case True + show ?thesis + by (rule LIMSEQ_imp_Suc) (simp add: True) +qed + +lemma LIMSEQ_power_zero: "norm x < 1 \ (\n. x ^ n) \ 0" + for x :: "'a::real_normed_algebra_1" + 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"}\ +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]) @@ -2022,92 +2168,81 @@ 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_eq: "f \a\ L = (\r>0. \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r)" + for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" + 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" + "(\r. 0 < r \ \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r) \ f \a\ L" + for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" + by (simp add: LIM_eq) + +lemma LIM_D: "f \a\ L \ 0 < r \ \s>0.\x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" + for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" + by (simp add: LIM_eq) + +lemma LIM_offset: "f \a\ L \ (\x. f (x + k)) \(a - k)\ L" + for a :: "'a::real_normed_vector" + by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap) + +lemma LIM_offset_zero: "f \a\ L \ (\h. f (a + h)) \0\ L" + for a :: "'a::real_normed_vector" + by (drule LIM_offset [where k = a]) (simp add: add.commute) + +lemma LIM_offset_zero_cancel: "(\h. f (a + h)) \0\ L \ f \a\ L" + for a :: "'a::real_normed_vector" + by (drule LIM_offset [where k = "- a"]) simp + +lemma LIM_offset_zero_iff: "f \a\ L \ (\h. f (a + h)) \0\ L" + for f :: "'a :: real_normed_vector \ _" 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: "(f \ l) F \ ((\x. f x - l) \ 0) F" + for f :: "'a::topological_space \ 'b::real_normed_vector" + 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_zero_iff: "((\x. f x - l) \ 0) F = (f \ l) F" + for f :: "'a::metric_space \ 'b::real_normed_vector" + 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)" + and 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) + 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" + assumes "0 < R" + and "\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) + by (rule metric_LIM_equal2 [OF assms]) (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" + and g: "g \b\ c" + and 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]]) + 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" + and 1: "\x. x \ a \ 0 \ g x" + and 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) + fix x + assume x: "x \ a" + with 1 have "norm (g x - 0) = g x" by simp 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 @@ -2117,61 +2252,50 @@ 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 LIM_isCont_iff: "(f \a\ f a) = ((\h. f (a + h)) \0\ f a)" + for f :: "'a::real_normed_vector \ 'b::topological_space" + by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) + +lemma isCont_iff: "isCont f x = (\h. f (x + h)) \0\ f x" + for f :: "'a::real_normed_vector \ 'b::topological_space" + 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" + and g: "g \f a\ l" + and 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 (rule LIM_compose2 [OF f g inj]) + +lemma isCont_norm [simp]: "isCont f a \ isCont (\x. norm (f x)) a" + for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_norm) -lemma isCont_rabs [simp]: - fixes f :: "'a::t2_space \ real" - shows "isCont f a \ isCont (\x. \f x\) a" +lemma isCont_rabs [simp]: "isCont f a \ isCont (\x. \f x\) a" + for f :: "'a::t2_space \ real" by (fact continuous_rabs) -lemma isCont_add [simp]: - fixes f :: "'a::t2_space \ 'b::topological_monoid_add" - shows "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" +lemma isCont_add [simp]: "isCont f a \ isCont g a \ isCont (\x. f x + g x) a" + for f :: "'a::t2_space \ 'b::topological_monoid_add" 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" +lemma isCont_minus [simp]: "isCont f a \ isCont (\x. - f x) a" + for f :: "'a::t2_space \ 'b::real_normed_vector" 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" +lemma isCont_diff [simp]: "isCont f a \ isCont g a \ isCont (\x. f x - g x) a" + for f :: "'a::t2_space \ 'b::real_normed_vector" 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" +lemma isCont_mult [simp]: "isCont f a \ isCont g a \ isCont (\x. f x * g x) a" + for f g :: "'a::t2_space \ 'b::real_normed_algebra" by (fact continuous_mult) -lemma (in bounded_linear) isCont: - "isCont g a \ isCont (\x. f (g x)) a" +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" +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] = @@ -2180,16 +2304,15 @@ 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" +lemma isCont_power [simp]: "isCont f a \ isCont (\x. f x ^ n) a" + for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" by (fact continuous_power) -lemma isCont_setsum [simp]: - fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" - shows "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" +lemma isCont_setsum [simp]: "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" + for f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" by (auto intro: continuous_setsum) + subsection \Uniform Continuity\ lemma uniformly_continuous_on_def: @@ -2200,37 +2323,39 @@ uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric) -abbreviation isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where - "isUCont f \ uniformly_continuous_on UNIV f" - -lemma isUCont_def: "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" +abbreviation isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" + where "isUCont f \ uniformly_continuous_on UNIV f" + +lemma isUCont_def: "isUCont f \ (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" by (auto simp: uniformly_continuous_on_def dist_commute) -lemma isUCont_isCont: "isUCont f ==> isCont f x" +lemma isUCont_isCont: "isUCont f \ isCont f x" by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at) lemma uniformly_continuous_on_Cauchy: - fixes f::"'a::metric_space \ 'b::metric_space" + fixes f :: "'a::metric_space \ 'b::metric_space" assumes "uniformly_continuous_on S f" "Cauchy X" "\n. X n \ S" shows "Cauchy (\n. f (X n))" using assms - unfolding uniformly_continuous_on_def - apply - + apply (simp only: uniformly_continuous_on_def) apply (rule metric_CauchyI) - apply (drule_tac x=e in spec, safe) - apply (drule_tac e=d in metric_CauchyD, safe) - apply (rule_tac x=M in exI, simp) + apply (drule_tac x=e in spec) + apply safe + apply (drule_tac e=d in metric_CauchyD) + apply safe + apply (rule_tac x=M in exI) + apply simp done -lemma isUCont_Cauchy: - "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" +lemma isUCont_Cauchy: "isUCont f \ Cauchy X \ Cauchy (\n. f (X n))" by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all lemma (in bounded_linear) isUCont: "isUCont f" -unfolding isUCont_def dist_norm + 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" + fix r :: real + assume r: "0 < r" + obtain K where K: "0 < K" and norm_le: "norm (f x) \ norm x * K" for x using pos_bounded by blast show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" proof (rule exI, safe) @@ -2246,7 +2371,7 @@ qed lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" -by (rule isUCont [THEN isUCont_Cauchy]) + by (rule isUCont [THEN isUCont_Cauchy]) lemma LIM_less_bound: fixes f :: "real \ real" @@ -2268,16 +2393,21 @@ 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 } + have "f n \ g 0" for n + proof - + from \decseq g\ have "g n \ g 0" + by (rule decseqD) simp + with \\n. f n \ g n\[THEN spec, of n] show ?thesis + by auto + qed then obtain u where "f \ u" "\i. f i \ u" using incseq_convergent[OF \incseq f\] by auto - moreover - { fix n + moreover have "f 0 \ g n" for n + proof - 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 } + with \\n. f n \ g n\[THEN spec, of n] show ?thesis + by simp + qed 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\]] @@ -2287,8 +2417,8 @@ 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" + and trans: "\a b c. P a b \ P b c \ a \ b \ b \ c \ P a c" + and 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 - define bisect where "bisect = @@ -2298,57 +2428,73 @@ 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 [simp]: "l n \ u n" for n by (induct n) auto 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 + show "l n \ l (Suc n)" "u (Suc n) \ u n" for 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) + have "l n - u n = (a - b) / 2^n" for 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" + 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 \ x \ x \ b \ b - a < d \ P a b" for 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\) } + have "\ P (l n) (u n)" for n + proof (induct n) + case 0 + then show ?case + by (simp add: \\ P a b\) + next + case (Suc n) + with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case + by auto + qed moreover - { have "eventually (\n. x - d / 2 < l n) sequentially" + { + 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" + case (elim n) 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 } + 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" + fix C + assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" define T where "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: *) + then have *: "{a..c} = {a..b} \ {b..c}" + by auto + with trans obtain C1 C2 + where "C1\C" "finite C1" "{a..b} \ \C1" "C2\C" "finite C2" "{b..c} \ \C2" + by auto 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 + with C have "x \ \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_dist dist_real_def subset_eq Ball_def abs_less_iff) with \c \ C\ show ?case @@ -2378,17 +2524,18 @@ qed lemma open_Collect_positive: - fixes f :: "'a::t2_space \ real" - assumes f: "continuous_on s f" - shows "\A. open A \ A \ s = {x\s. 0 < f x}" - using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"] - by (auto simp: Int_def field_simps) + fixes f :: "'a::t2_space \ real" + assumes f: "continuous_on s f" + shows "\A. open A \ A \ s = {x\s. 0 < f x}" + using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"] + by (auto simp: Int_def field_simps) lemma open_Collect_less_Int: - fixes f g :: "'a::t2_space \ real" - assumes f: "continuous_on s f" and g: "continuous_on s g" - shows "\A. open A \ A \ s = {x\s. f x < g x}" - using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) + fixes f g :: "'a::t2_space \ real" + assumes f: "continuous_on s f" + and g: "continuous_on s g" + shows "\A. open A \ A \ s = {x\s. f x < g x}" + using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) subsection \Boundedness of continuous functions\ @@ -2399,14 +2546,14 @@ 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] + 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] + using continuous_attains_inf[of "{a..b}" f] by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_bounded: @@ -2421,21 +2568,23 @@ 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)" +lemma IVT_objl: + "(f a \ y \ y \ f b \ a \ b \ (\x. a \ x \ x \ b \ isCont f x)) \ + (\x. a \ x \ x \ b \ f x = y)" + for a y :: real 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)" +lemma IVT2_objl: + "(f b \ y \ y \ f a \ a \ b \ (\x. a \ x \ x \ b \ isCont f x)) \ + (\x. a \ x \ x \ b \ f x = y)" + for b y :: real 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)))" + (\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 @@ -2446,22 +2595,26 @@ 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) + apply simp + apply (metis order_trans) + apply simp + apply (metis order_trans) done qed -text\Continuity of inverse function\ +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" + 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}" + let ?A = "f (x - d)" + let ?B = "f (x + d)" + let ?D = "{x - d..x + d}" have f: "continuous_on ?D f" using cont by (intro continuous_at_imp_continuous_on ballI) auto @@ -2483,45 +2636,42 @@ 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 + 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)" + 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: "f \c\ l \ 0 < l \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ 0 < f x)" + for f :: "real \ real" + apply (drule (1) LIM_D) + apply clarify + apply (rule_tac x = s in exI) + apply (simp add: abs_less_iff) + done + +lemma LIM_fun_less_zero: "f \c\ l \ l < 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x < 0)" + for f :: "real \ real" + apply (drule LIM_D [where r="-l"]) + apply simp + apply clarify + apply (rule_tac x = s in exI) + apply (simp add: abs_less_iff) + done + +lemma LIM_fun_not_zero: "f \c\ l \ l \ 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x \ 0)" + for f :: "real \ real" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff) end