# HG changeset patch # User hoelzl # Date 1354041071 -3600 # Node ID 491c5c81c2e804a7b3067dddce9ea44110729740 # Parent dea9363887a680c0f87c9869008c1c458aac9daf introduce filter_lim as a generatlization of tendsto diff -r dea9363887a6 -r 491c5c81c2e8 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Tue Nov 27 13:48:40 2012 +0100 +++ b/src/HOL/Limits.thy Tue Nov 27 19:31:11 2012 +0100 @@ -280,23 +280,70 @@ lemma filtermap_bot [simp]: "filtermap f bot = bot" by (simp add: filter_eq_iff eventually_filtermap) - -subsection {* Sequentially *} +subsection {* Order filters *} -definition sequentially :: "nat filter" - where "sequentially = Abs_filter (\P. \k. \n\k. P n)" +definition at_top :: "('a::order) filter" + where "at_top = Abs_filter (\P. \k. \n\k. P n)" -lemma eventually_sequentially: - "eventually P sequentially \ (\N. \n\N. P n)" -unfolding sequentially_def +lemma eventually_at_top_linorder: + fixes P :: "'a::linorder \ bool" shows "eventually P at_top \ (\N. \n\N. P n)" + unfolding at_top_def proof (rule eventually_Abs_filter, rule is_filter.intro) - fix P Q :: "nat \ bool" + fix P Q :: "'a \ bool" assume "\i. \n\i. P n" and "\j. \n\j. Q n" then obtain i j where "\n\i. P n" and "\n\j. Q n" by auto then have "\n\max i j. P n \ Q n" by simp then show "\k. \n\k. P n \ Q n" .. qed auto +lemma eventually_at_top_dense: + fixes P :: "'a::dense_linorder \ bool" shows "eventually P at_top \ (\N. \n>N. P n)" + unfolding eventually_at_top_linorder +proof safe + fix N assume "\n\N. P n" then show "\N. \n>N. P n" by (auto intro!: exI[of _ N]) +next + fix N assume "\n>N. P n" + moreover from gt_ex[of N] guess y .. + ultimately show "\N. \n\N. P n" by (auto intro!: exI[of _ y]) +qed + +definition at_bot :: "('a::order) filter" + where "at_bot = Abs_filter (\P. \k. \n\k. P n)" + +lemma eventually_at_bot_linorder: + fixes P :: "'a::linorder \ bool" shows "eventually P at_bot \ (\N. \n\N. P n)" + unfolding at_bot_def +proof (rule eventually_Abs_filter, rule is_filter.intro) + fix P Q :: "'a \ bool" + assume "\i. \n\i. P n" and "\j. \n\j. Q n" + then obtain i j where "\n\i. P n" and "\n\j. Q n" by auto + then have "\n\min i j. P n \ Q n" by simp + then show "\k. \n\k. P n \ Q n" .. +qed auto + +lemma eventually_at_bot_dense: + fixes P :: "'a::dense_linorder \ bool" shows "eventually P at_bot \ (\N. \nn\N. P n" then show "\N. \nnN. \n\N. P n" by (auto intro!: exI[of _ y]) +qed + +subsection {* Sequentially *} + +abbreviation sequentially :: "nat filter" + where "sequentially == at_top" + +lemma sequentially_def: "sequentially = Abs_filter (\P. \k. \n\k. P n)" + unfolding at_top_def by simp + +lemma eventually_sequentially: + "eventually P sequentially \ (\N. \n\N. P n)" + by (rule eventually_at_top_linorder) + lemma sequentially_bot [simp, intro]: "sequentially \ bot" unfolding filter_eq_iff eventually_sequentially by auto @@ -340,6 +387,9 @@ lemma within_empty [simp]: "F within {} = bot" unfolding filter_eq_iff eventually_within by simp +lemma within_le: "F within S \ F" + unfolding le_filter_def eventually_within by (auto elim: eventually_elim1) + lemma eventually_nhds: "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" unfolding nhds_def @@ -564,12 +614,24 @@ subsection {* Limits *} -definition (in topological_space) +definition filter_lim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where + "filter_lim f F2 F1 \ filtermap f F1 \ F2" + +syntax + "_LIM" :: "pttrns \ 'a \ 'b \ 'a \ bool" ("(3LIM (_)/ (_)./ (_) :> (_))" [1000, 10, 0, 10] 10) + +translations + "LIM x F1. f :> F2" == "CONST filter_lim (%x. f) F2 F1" + +lemma filter_limE: "(LIM x F1. f x :> F2) \ eventually P F2 \ eventually (\x. P (f x)) F1" + by (auto simp: filter_lim_def eventually_filtermap le_filter_def) + +lemma filter_limI: "(\P. eventually P F2 \ eventually (\x. P (f x)) F1) \ (LIM x F1. f x :> F2)" + by (auto simp: filter_lim_def eventually_filtermap le_filter_def) + +abbreviation (in topological_space) 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" + "(f ---> l) F \ filter_lim f (nhds l) F" ML {* structure Tendsto_Intros = Named_Thms @@ -581,6 +643,15 @@ setup Tendsto_Intros.setup +lemma tendsto_def: "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" + unfolding filter_lim_def +proof safe + fix S assume "open S" "l \ S" "filtermap f F \ nhds l" + then show "eventually (\x. f x \ S) F" + unfolding eventually_nhds eventually_filtermap le_filter_def + by (auto elim!: allE[of _ "\x. x \ S"] eventually_rev_mp) +qed (auto elim!: eventually_rev_mp simp: eventually_nhds eventually_filtermap le_filter_def) + lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" unfolding tendsto_def le_filter_def by fast @@ -706,15 +777,6 @@ 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]: @@ -1011,4 +1073,25 @@ shows "\(f ---> l) F; l \ 0\ \ ((\x. sgn (f x)) ---> sgn l) F" unfolding sgn_div_norm by (simp add: tendsto_intros) +subsection {* Limits to @{const at_top} and @{const at_bot} *} + +lemma filter_lim_at_top: + fixes f :: "'a \ ('b::dense_linorder)" + shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" + by (safe elim!: filter_limE intro!: filter_limI) + (auto simp: eventually_at_top_dense elim!: eventually_elim1) + +lemma filter_lim_at_bot: + fixes f :: "'a \ ('b::dense_linorder)" + shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x < Z) F)" + by (safe elim!: filter_limE intro!: filter_limI) + (auto simp: eventually_at_bot_dense elim!: eventually_elim1) + +lemma filter_lim_real_sequentially: "LIM x sequentially. real x :> at_top" + unfolding filter_lim_at_top + apply (intro allI) + apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI) + apply (auto simp: natceiling_le_eq) + done + end diff -r dea9363887a6 -r 491c5c81c2e8 src/HOL/Log.thy --- a/src/HOL/Log.thy Tue Nov 27 13:48:40 2012 +0100 +++ b/src/HOL/Log.thy Tue Nov 27 19:31:11 2012 +0100 @@ -359,12 +359,12 @@ qed lemma tendsto_neg_powr: - assumes "s < 0" and "real_tendsto_inf f F" + assumes "s < 0" and "LIM x F. f x :> at_top" 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) + from assms have "eventually (\x. Z < f x) F" by (simp add: filter_lim_at_top) moreover from assms have "\x. Z < x \ x powr s < Z powr s" by (auto simp: Z_def intro!: powr_less_mono2_neg)