diff -r 4ddc49205f5d -r 68ca05a7f159 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Feb 23 13:27:19 2018 +0000 +++ b/src/HOL/Limits.thy Fri Feb 23 14:56:32 2018 +0000 @@ -52,6 +52,9 @@ for f :: "_ \ real" by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) +lemma filterlim_real_at_infinity_sequentially: "filterlim real at_infinity sequentially" + by (simp add: filterlim_at_top_imp_at_infinity 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) @@ -1256,6 +1259,20 @@ for a :: real unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) +lemma at_to_0: "at a = filtermap (\x. x + a) (at 0)" + for a :: "'a::real_normed_vector" + using filtermap_at_shift[of "-a" 0] by simp + +lemma filterlim_at_to_0: + "filterlim f F (at a) \ filterlim (\x. f (x + a)) F (at 0)" + for a :: "'a::real_normed_vector" + unfolding filterlim_def filtermap_filtermap at_to_0[of a] .. + +lemma eventually_at_to_0: + "eventually P (at a) \ eventually (\x. P (x + a)) (at 0)" + for a :: "'a::real_normed_vector" + unfolding at_to_0[of a] by (simp add: eventually_filtermap) + 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]) @@ -1268,6 +1285,7 @@ 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) \ filterlim (\x. f (- x)) F (at_right (-a))" for a :: real @@ -1486,6 +1504,36 @@ for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) +lemma filterlim_times_pos: + "LIM x F1. c * f x :> at_right l" + if "filterlim f (at_right p) F1" "0 < c" "l = c * p" + for c::"'a::{linordered_field, linorder_topology}" + unfolding filterlim_iff +proof safe + fix P + assume "\\<^sub>F x in at_right l. P x" + then obtain d where "c * p < d" "\y. y > c * p \ y < d \ P y" + unfolding \l = _ \ eventually_at_right_field + by auto + then have "\\<^sub>F a in at_right p. P (c * a)" + by (auto simp: eventually_at_right_field \0 < c\ field_simps intro!: exI[where x="d/c"]) + from that(1)[unfolded filterlim_iff, rule_format, OF this] + show "\\<^sub>F x in F1. P (c * f x)" . +qed + +lemma filtermap_nhds_times: "c \ 0 \ filtermap (times c) (nhds a) = nhds (c * a)" + for a c :: "'a::real_normed_field" + by (rule filtermap_fun_inverse[where g="\x. inverse c * x"]) + (auto intro!: tendsto_eq_intros filterlim_ident) + +lemma filtermap_times_pos_at_right: + fixes c::"'a::{linordered_field, linorder_topology}" + assumes "c > 0" + shows "filtermap (times c) (at_right p) = at_right (c * p)" + using assms + by (intro filtermap_fun_inverse[where g="\x. inverse c * x"]) + (auto intro!: filterlim_ident filterlim_times_pos) + lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity" proof (rule antisym) have "(inverse \ (0::'a)) at_infinity" @@ -1944,6 +1992,10 @@ lemma lim_inverse_n: "((\n. inverse(of_nat n)) \ (0::'a::real_normed_field)) sequentially" using lim_1_over_n by (simp add: inverse_eq_divide) +lemma lim_inverse_n': "((\n. 1 / n) \ 0) sequentially" + using lim_inverse_n + by (simp add: inverse_eq_divide) + 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"