# HG changeset patch # User wenzelm # Date 1323982012 -3600 # Node ID 36f3f0010b7d55a7bf0b456582b1e0a174922668 # Parent e7dbb27c13083ef00b65c58405d14189516faffe# Parent 629d123b7dbe63e5e2a7d11ba74484c9fad440c8 merged; diff -r 629d123b7dbe -r 36f3f0010b7d src/HOL/Limits.thy --- a/src/HOL/Limits.thy Thu Dec 15 19:53:28 2011 +0100 +++ b/src/HOL/Limits.thy Thu Dec 15 21:46:52 2011 +0100 @@ -101,6 +101,18 @@ shows "eventually (\i. R i) F" using assms by (auto elim!: eventually_rev_mp) +lemma eventually_subst: + assumes "eventually (\n. P n = Q n) F" + shows "eventually P F = eventually Q F" (is "?L = ?R") +proof - + from assms have "eventually (\x. P x \ Q x) F" + and "eventually (\x. Q x \ P x) F" + by (auto elim: eventually_elim1) + then show ?thesis by (auto elim: eventually_elim2) +qed + + + subsection {* Finer-than relation *} text {* @{term "F \ F'"} means that filter @{term F} is finer than @@ -280,6 +292,11 @@ unfolding le_filter_def eventually_sequentially by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) +lemma eventually_sequentiallyI: + assumes "\x. c \ x \ P x" + shows "eventually P sequentially" +using assms by (auto simp: eventually_sequentially) + subsection {* Standard filters *} @@ -537,6 +554,9 @@ tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "--->" 55) where "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" +definition real_tendsto_inf :: "('a \ real) \ 'a filter \ bool" where + "real_tendsto_inf f F \ \x. eventually (\y. x < f y) F" + ML {* structure Tendsto_Intros = Named_Thms ( @@ -673,6 +693,15 @@ using le_less_trans by (rule eventually_elim2) qed +lemma real_tendsto_inf_real: "real_tendsto_inf real sequentially" +proof (unfold real_tendsto_inf_def, rule allI) + fix x show "eventually (\y. x < real y) sequentially" + by (rule eventually_sequentiallyI[of "natceiling (x + 1)"]) + (simp add: natceiling_le_eq) +qed + + + subsubsection {* Distance and norms *} lemma tendsto_dist [tendsto_intros]: @@ -769,6 +798,22 @@ by (simp add: tendsto_const) qed +lemma real_tendsto_sandwich: + fixes f g h :: "'a \ real" + assumes ev: "eventually (\n. f n \ g n) net" "eventually (\n. g n \ h n) net" + assumes lim: "(f ---> c) net" "(h ---> c) net" + shows "(g ---> c) net" +proof - + have "((\n. g n - f n) ---> 0) net" + proof (rule metric_tendsto_imp_tendsto) + show "eventually (\n. dist (g n - f n) 0 \ dist (h n - f n) 0) net" + using ev by (rule eventually_elim2) (simp add: dist_real_def) + show "((\n. h n - f n) ---> 0) net" + using tendsto_diff[OF lim(2,1)] by simp + qed + from tendsto_add[OF this lim(1)] show ?thesis by simp +qed + subsubsection {* Linear operators and multiplication *} lemma (in bounded_linear) tendsto: diff -r 629d123b7dbe -r 36f3f0010b7d src/HOL/Log.thy --- a/src/HOL/Log.thy Thu Dec 15 19:53:28 2011 +0100 +++ b/src/HOL/Log.thy Thu Dec 15 21:46:52 2011 +0100 @@ -285,32 +285,40 @@ finally show ?thesis . qed -lemma LIMSEQ_neg_powr: - assumes s: "0 < s" - shows "(%x. (real x) powr - s) ----> 0" - apply (unfold LIMSEQ_iff) - apply clarsimp - apply (rule_tac x = "natfloor(r powr (1 / - s)) + 1" in exI) - apply clarify -proof - - fix r fix n - assume r: "0 < r" and n: "natfloor (r powr (1 / - s)) + 1 <= n" - have "r powr (1 / - s) < real(natfloor(r powr (1 / - s))) + 1" - by (rule real_natfloor_add_one_gt) - also have "... = real(natfloor(r powr (1 / -s)) + 1)" - by simp - also have "... <= real n" - apply (subst real_of_nat_le_iff) - apply (rule n) - done - finally have "r powr (1 / - s) < real n". - then have "real n powr (- s) < (r powr (1 / - s)) powr - s" - apply (intro powr_less_mono2_neg) - apply (auto simp add: s) - done - also have "... = r" - by (simp add: powr_powr s r less_imp_neq [THEN not_sym]) - finally show "real n powr - s < r" . +(* FIXME: generalize by replacing d by with g x and g ---> d? *) +lemma tendsto_zero_powrI: + assumes "eventually (\x. 0 < f x ) F" and "(f ---> 0) F" + assumes "0 < d" + shows "((\x. f x powr d) ---> 0) F" +proof (rule tendstoI) + fix e :: real assume "0 < e" + def Z \ "e powr (1 / d)" + with `0 < e` have "0 < Z" by simp + with assms have "eventually (\x. 0 < f x \ dist (f x) 0 < Z) F" + by (intro eventually_conj tendstoD) + moreover + from assms have "\x. 0 < x \ dist x 0 < Z \ x powr d < Z powr d" + by (intro powr_less_mono2) (auto simp: dist_real_def) + with assms `0 < e` have "\x. 0 < x \ dist x 0 < Z \ dist (x powr d) 0 < e" + unfolding dist_real_def Z_def by (auto simp: powr_powr) + ultimately + show "eventually (\x. dist (f x powr d) 0 < e) F" by (rule eventually_elim1) +qed + +lemma tendsto_neg_powr: + assumes "s < 0" and "real_tendsto_inf f F" + shows "((\x. f x powr s) ---> 0) F" +proof (rule tendstoI) + fix e :: real assume "0 < e" + def Z \ "e powr (1 / s)" + from assms have "eventually (\x. Z < f x) F" by (simp add: real_tendsto_inf_def) + moreover + from assms have "\x. Z < x \ x powr s < Z powr s" + by (auto simp: Z_def intro!: powr_less_mono2_neg) + with assms `0 < e` have "\x. Z < x \ dist (x powr s) 0 < e" + by (simp add: powr_powr Z_def dist_real_def) + ultimately + show "eventually (\x. dist (f x powr s) 0 < e) F" by (rule eventually_elim1) qed end diff -r 629d123b7dbe -r 36f3f0010b7d src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Thu Dec 15 19:53:28 2011 +0100 +++ b/src/HOL/Orderings.thy Thu Dec 15 21:46:52 2011 +0100 @@ -1057,14 +1057,24 @@ by (simp add: max_def) lemma min_leastR: "(\x\'a\order. least \ x) \ min x least = least" -apply (simp add: min_def) -apply (blast intro: antisym) -done +by (simp add: min_def) (blast intro: antisym) lemma max_leastR: "(\x\'a\order. least \ x) \ max x least = x" -apply (simp add: max_def) -apply (blast intro: antisym) -done +by (simp add: max_def) (blast intro: antisym) + +lemma min_greatestL: "(\x::'a::order. x \ greatest) \ min greatest x = x" +by (simp add: min_def) (blast intro: antisym) + +lemma max_greatestL: "(\x::'a::order. x \ greatest) \ max greatest x = greatest" +by (simp add: max_def) (blast intro: antisym) + +lemma min_greatestR: "(\x. x \ greatest) \ min x greatest = x" +by (simp add: min_def) + +lemma max_greatestR: "(\x. x \ greatest) \ max x greatest = greatest" +by (simp add: max_def) + + subsection {* (Unique) top and bottom elements *}