# HG changeset patch # User hoelzl # Date 1363945302 -3600 # Node ID cad22a3cc09c7c4cce65dea6cf8e7bf131718122 # Parent a981a5c8a505f895fdae2b9ae83a42532b730a18 move topological_space to its own theory diff -r a981a5c8a505 -r cad22a3cc09c src/HOL/Lim.thy --- a/src/HOL/Lim.thy Thu Mar 21 16:58:14 2013 +0100 +++ b/src/HOL/Lim.thy Fri Mar 22 10:41:42 2013 +0100 @@ -10,28 +10,12 @@ imports SEQ begin -text{*Standard Definitions*} - -abbreviation - LIM :: "['a::topological_space \ 'b::topological_space, 'a, 'b] \ bool" - ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where - "f -- a --> L \ (f ---> L) (at a)" - -definition - isCont :: "['a::topological_space \ 'b::topological_space, 'a] \ bool" where - "isCont f a = (f -- a --> (f a))" - definition isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where "isUCont f = (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" subsection {* Limits of Functions *} -lemma LIM_def: "f -- a --> L = - (\r > 0. \s > 0. \x. x \ a & dist x a < s - --> dist (f x) L < r)" -unfolding tendsto_iff eventually_at .. - lemma metric_LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) \ f -- a --> L" @@ -81,8 +65,6 @@ shows "(\h. f (a + h)) -- 0 --> L \ f -- a --> L" by (drule_tac k="- a" in LIM_offset, simp) -lemma LIM_cong_limit: "\ f -- x --> L ; K = L \ \ f -- x --> K" by simp - lemma LIM_zero: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "(f ---> l) F \ ((\x. f x - l) ---> 0) F" @@ -114,36 +96,6 @@ by (rule metric_LIM_imp_LIM [OF f], simp add: dist_norm le) -lemma LIM_const_not_eq: - fixes a :: "'a::perfect_space" - fixes k L :: "'b::t2_space" - shows "k \ L \ \ (\x. k) -- a --> L" - by (simp add: tendsto_const_iff) - -lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] - -lemma LIM_const_eq: - fixes a :: "'a::perfect_space" - fixes k L :: "'b::t2_space" - shows "(\x. k) -- a --> L \ k = L" - by (simp add: tendsto_const_iff) - -lemma LIM_unique: - fixes a :: "'a::perfect_space" - fixes L M :: "'b::t2_space" - shows "\f -- a --> L; f -- a --> M\ \ L = M" - using at_neq_bot by (rule tendsto_unique) - -text{*Limits are equal for functions equal except at limit point*} -lemma LIM_equal: - "[| \x. x \ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" -unfolding tendsto_def eventually_at_topological by simp - -lemma LIM_cong: - "\a = b; \x. x \ b \ f x = g x; l = m\ - \ ((\x. f x) -- a --> l) = ((\x. g x) -- b --> m)" -by (simp add: LIM_equal) - lemma metric_LIM_equal2: assumes 1: "0 < R" assumes 2: "\x. \x \ a; dist x a < R\ \ f x = g x" @@ -163,13 +115,6 @@ shows "g -- a --> l \ f -- a --> l" by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm) -lemma LIM_compose_eventually: - assumes f: "f -- a --> b" - assumes g: "g -- b --> c" - assumes inj: "eventually (\x. f x \ b) (at a)" - shows "(\x. g (f x)) -- a --> c" - using g f inj by (rule tendsto_compose_eventually) - lemma metric_LIM_compose2: assumes f: "f -- a --> b" assumes g: "g -- b --> c" @@ -186,16 +131,13 @@ shows "(\x. g (f x)) -- a --> c" by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) -lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" - unfolding o_def by (rule tendsto_compose) - lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" assumes f: "f -- a --> 0" assumes 1: "\x. x \ a \ 0 \ g x" assumes 2: "\x. x \ a \ g x \ f x" shows "g -- a --> 0" -proof (rule LIM_imp_LIM [OF f]) +proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) fix x assume x: "x \ a" have "norm (g x - 0) = g x" by (simp add: 1 x) also have "g x \ f x" by (rule 2 [OF x]) @@ -217,12 +159,6 @@ shows "isCont f x = (\h. f (x + h)) -- 0 --> f x" by (simp add: isCont_def LIM_isCont_iff) -lemma isCont_ident [simp]: "isCont (\x. x) a" - unfolding isCont_def by (rule tendsto_ident_at) - -lemma isCont_const [simp]: "isCont (\x. k) a" - unfolding isCont_def by (rule tendsto_const) - lemma isCont_norm [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "isCont f a \ isCont (\x. norm (f x)) a" @@ -263,10 +199,6 @@ shows "\isCont f a; isCont g a; g a \ 0\ \ isCont (\x. f x / g x) a" unfolding isCont_def by (rule tendsto_divide) -lemma isCont_tendsto_compose: - "\isCont g l; (f ---> l) F\ \ ((\x. g (f x)) ---> g l) F" - unfolding isCont_def by (rule tendsto_compose) - lemma metric_isCont_LIM_compose2: assumes f [unfolded isCont_def]: "isCont f a" assumes g: "g -- f a --> l" @@ -282,12 +214,6 @@ shows "(\x. g (f x)) -- a --> l" by (rule LIM_compose2 [OF f g inj]) -lemma isCont_o2: "\isCont f a; isCont g (f a)\ \ isCont (\x. g (f x)) a" - unfolding isCont_def by (rule tendsto_compose) - -lemma isCont_o: "\isCont f a; isCont g (f a)\ \ isCont (g o f) a" - unfolding o_def by (rule isCont_o2) - lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" unfolding isCont_def by (rule tendsto) @@ -372,7 +298,7 @@ def F \ "\n::nat. SOME x. x \ s \ x \ a \ dist x a < ?I n \ \ P x" assume "\ eventually P (at a within s)" hence P: "\d>0. \x. x \ s \ x \ a \ dist x a < d \ \ P x" - unfolding Limits.eventually_within Limits.eventually_at by fast + unfolding eventually_within eventually_at by fast hence "\n. \x. x \ s \ x \ a \ dist x a < ?I n \ \ P x" by simp hence F: "\n. F n \ s \ F n \ a \ dist (F n) a < ?I n \ \ P (F n)" unfolding F_def by (rule someI_ex) diff -r a981a5c8a505 -r cad22a3cc09c src/HOL/Limits.thy --- a/src/HOL/Limits.thy Thu Mar 21 16:58:14 2013 +0100 +++ b/src/HOL/Limits.thy Fri Mar 22 10:41:42 2013 +0100 @@ -8,457 +8,9 @@ imports RealVector begin -subsection {* Filters *} - -text {* - This definition also allows non-proper filters. -*} - -locale is_filter = - fixes F :: "('a \ bool) \ bool" - assumes True: "F (\x. True)" - assumes conj: "F (\x. P x) \ F (\x. Q x) \ F (\x. P x \ Q x)" - assumes mono: "\x. P x \ Q x \ F (\x. P x) \ F (\x. Q x)" - -typedef 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" -proof - show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) -qed - -lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" - using Rep_filter [of F] by simp - -lemma Abs_filter_inverse': - assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" - using assms by (simp add: Abs_filter_inverse) - - -subsection {* Eventually *} - -definition eventually :: "('a \ bool) \ 'a filter \ bool" - where "eventually P F \ Rep_filter F P" - -lemma eventually_Abs_filter: - assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" - unfolding eventually_def using assms by (simp add: Abs_filter_inverse) - -lemma filter_eq_iff: - shows "F = F' \ (\P. eventually P F = eventually P F')" - unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. - -lemma eventually_True [simp]: "eventually (\x. True) F" - unfolding eventually_def - by (rule is_filter.True [OF is_filter_Rep_filter]) - -lemma always_eventually: "\x. P x \ eventually P F" -proof - - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) - thus "eventually P F" by simp -qed - -lemma eventually_mono: - "(\x. P x \ Q x) \ eventually P F \ eventually Q F" - unfolding eventually_def - by (rule is_filter.mono [OF is_filter_Rep_filter]) - -lemma eventually_conj: - assumes P: "eventually (\x. P x) F" - assumes Q: "eventually (\x. Q x) F" - shows "eventually (\x. P x \ Q x) F" - using assms unfolding eventually_def - by (rule is_filter.conj [OF is_filter_Rep_filter]) - -lemma eventually_Ball_finite: - assumes "finite A" and "\y\A. eventually (\x. P x y) net" - shows "eventually (\x. \y\A. P x y) net" -using assms by (induct set: finite, simp, simp add: eventually_conj) - -lemma eventually_all_finite: - fixes P :: "'a \ 'b::finite \ bool" - assumes "\y. eventually (\x. P x y) net" - shows "eventually (\x. \y. P x y) net" -using eventually_Ball_finite [of UNIV P] assms by simp - -lemma eventually_mp: - assumes "eventually (\x. P x \ Q x) F" - assumes "eventually (\x. P x) F" - shows "eventually (\x. Q x) F" -proof (rule eventually_mono) - show "\x. (P x \ Q x) \ P x \ Q x" by simp - show "eventually (\x. (P x \ Q x) \ P x) F" - using assms by (rule eventually_conj) -qed - -lemma eventually_rev_mp: - assumes "eventually (\x. P x) F" - assumes "eventually (\x. P x \ Q x) F" - shows "eventually (\x. Q x) F" -using assms(2) assms(1) by (rule eventually_mp) - -lemma eventually_conj_iff: - "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" - by (auto intro: eventually_conj elim: eventually_rev_mp) - -lemma eventually_elim1: - assumes "eventually (\i. P i) F" - assumes "\i. P i \ Q i" - shows "eventually (\i. Q i) F" - using assms by (auto elim!: eventually_rev_mp) - -lemma eventually_elim2: - assumes "eventually (\i. P i) F" - assumes "eventually (\i. Q i) F" - assumes "\i. P i \ Q i \ R i" - 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 - -ML {* - fun eventually_elim_tac ctxt thms thm = - let - val thy = Proof_Context.theory_of ctxt - val mp_thms = thms RL [@{thm eventually_rev_mp}] - val raw_elim_thm = - (@{thm allI} RS @{thm always_eventually}) - |> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms - |> fold (fn _ => fn thm => @{thm impI} RS thm) thms - val cases_prop = prop_of (raw_elim_thm RS thm) - val cases = (Rule_Cases.make_common (thy, cases_prop) [(("elim", []), [])]) - in - CASES cases (rtac raw_elim_thm 1) thm - end -*} - -method_setup eventually_elim = {* - Scan.succeed (fn ctxt => METHOD_CASES (eventually_elim_tac ctxt)) -*} "elimination of eventually quantifiers" - - -subsection {* Finer-than relation *} - -text {* @{term "F \ F'"} means that filter @{term F} is finer than -filter @{term F'}. *} - -instantiation filter :: (type) complete_lattice -begin - -definition le_filter_def: - "F \ F' \ (\P. eventually P F' \ eventually P F)" - -definition - "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" - -definition - "top = Abs_filter (\P. \x. P x)" - -definition - "bot = Abs_filter (\P. True)" - -definition - "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" - -definition - "inf F F' = Abs_filter - (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" - -definition - "Sup S = Abs_filter (\P. \F\S. eventually P F)" - -definition - "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" - -lemma eventually_top [simp]: "eventually P top \ (\x. P x)" - unfolding top_filter_def - by (rule eventually_Abs_filter, rule is_filter.intro, auto) - -lemma eventually_bot [simp]: "eventually P bot" - unfolding bot_filter_def - by (subst eventually_Abs_filter, rule is_filter.intro, auto) - -lemma eventually_sup: - "eventually P (sup F F') \ eventually P F \ eventually P F'" - unfolding sup_filter_def - by (rule eventually_Abs_filter, rule is_filter.intro) - (auto elim!: eventually_rev_mp) - -lemma eventually_inf: - "eventually P (inf F F') \ - (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" - unfolding inf_filter_def - apply (rule eventually_Abs_filter, rule is_filter.intro) - apply (fast intro: eventually_True) - apply clarify - apply (intro exI conjI) - apply (erule (1) eventually_conj) - apply (erule (1) eventually_conj) - apply simp - apply auto - done - -lemma eventually_Sup: - "eventually P (Sup S) \ (\F\S. eventually P F)" - unfolding Sup_filter_def - apply (rule eventually_Abs_filter, rule is_filter.intro) - apply (auto intro: eventually_conj elim!: eventually_rev_mp) - done - -instance proof - fix F F' F'' :: "'a filter" and S :: "'a filter set" - { show "F < F' \ F \ F' \ \ F' \ F" - by (rule less_filter_def) } - { show "F \ F" - unfolding le_filter_def by simp } - { assume "F \ F'" and "F' \ F''" thus "F \ F''" - unfolding le_filter_def by simp } - { assume "F \ F'" and "F' \ F" thus "F = F'" - unfolding le_filter_def filter_eq_iff by fast } - { show "F \ top" - unfolding le_filter_def eventually_top by (simp add: always_eventually) } - { show "bot \ F" - unfolding le_filter_def by simp } - { show "F \ sup F F'" and "F' \ sup F F'" - unfolding le_filter_def eventually_sup by simp_all } - { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" - unfolding le_filter_def eventually_sup by simp } - { show "inf F F' \ F" and "inf F F' \ F'" - unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } - { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" - unfolding le_filter_def eventually_inf - by (auto elim!: eventually_mono intro: eventually_conj) } - { assume "F \ S" thus "F \ Sup S" - unfolding le_filter_def eventually_Sup by simp } - { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" - unfolding le_filter_def eventually_Sup by simp } - { assume "F'' \ S" thus "Inf S \ F''" - unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } - { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" - unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } -qed - -end - -lemma filter_leD: - "F \ F' \ eventually P F' \ eventually P F" - unfolding le_filter_def by simp - -lemma filter_leI: - "(\P. eventually P F' \ eventually P F) \ F \ F'" - unfolding le_filter_def by simp - -lemma eventually_False: - "eventually (\x. False) F \ F = bot" - unfolding filter_eq_iff by (auto elim: eventually_rev_mp) - -abbreviation (input) trivial_limit :: "'a filter \ bool" - where "trivial_limit F \ F = bot" - -lemma trivial_limit_def: "trivial_limit F \ eventually (\x. False) F" - by (rule eventually_False [symmetric]) - -lemma eventually_const: "\ trivial_limit net \ eventually (\x. P) net \ P" - by (cases P) (simp_all add: eventually_False) - - -subsection {* Map function for filters *} - -definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" - where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" - -lemma eventually_filtermap: - "eventually P (filtermap f F) = eventually (\x. P (f x)) F" - unfolding filtermap_def - apply (rule eventually_Abs_filter) - apply (rule is_filter.intro) - apply (auto elim!: eventually_rev_mp) - done - -lemma filtermap_ident: "filtermap (\x. x) F = F" - by (simp add: filter_eq_iff eventually_filtermap) - -lemma filtermap_filtermap: - "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" - by (simp add: filter_eq_iff eventually_filtermap) - -lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" - unfolding le_filter_def eventually_filtermap by simp - -lemma filtermap_bot [simp]: "filtermap f bot = bot" - by (simp add: filter_eq_iff eventually_filtermap) - -lemma filtermap_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" - by (auto simp: filter_eq_iff eventually_filtermap eventually_sup) - -subsection {* Order filters *} - -definition at_top :: "('a::order) filter" - where "at_top = Abs_filter (\P. \k. \n\k. P n)" - -lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" - unfolding at_top_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\max i j. P n \ Q n" by simp - then show "\k. \n\k. P n \ Q n" .. -qed auto - -lemma eventually_ge_at_top: - "eventually (\x. (c::_::linorder) \ x) at_top" - unfolding eventually_at_top_linorder by auto - -lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::dense_linorder. \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 - -lemma eventually_gt_at_top: - "eventually (\x. (c::_::dense_linorder) < x) at_top" - unfolding eventually_at_top_dense by auto - -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_le_at_bot: - "eventually (\x. x \ (c::_::linorder)) at_bot" - unfolding eventually_at_bot_linorder by 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 - -lemma eventually_gt_at_bot: - "eventually (\x. x < (c::_::dense_linorder)) at_bot" - unfolding eventually_at_bot_dense by auto - -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 - -lemmas trivial_limit_sequentially = sequentially_bot - -lemma eventually_False_sequentially [simp]: - "\ eventually (\n. False) sequentially" - by (simp add: eventually_False) - -lemma le_sequentially: - "F \ sequentially \ (\N. eventually (\n. N \ n) F)" - 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 *} - -definition within :: "'a filter \ 'a set \ 'a filter" (infixr "within" 70) - where "F within S = Abs_filter (\P. eventually (\x. x \ S \ P x) F)" - -definition (in topological_space) nhds :: "'a \ 'a filter" - where "nhds a = Abs_filter (\P. \S. open S \ a \ S \ (\x\S. P x))" - -definition (in topological_space) at :: "'a \ 'a filter" - where "at a = nhds a within - {a}" - -abbreviation at_right :: "'a\{topological_space, order} \ 'a filter" where - "at_right x \ at x within {x <..}" - -abbreviation at_left :: "'a\{topological_space, order} \ 'a filter" where - "at_left x \ at x within {..< x}" - definition at_infinity :: "'a::real_normed_vector filter" where "at_infinity = Abs_filter (\P. \r. \x. r \ norm x \ P x)" -lemma eventually_within: - "eventually P (F within S) = eventually (\x. x \ S \ P x) F" - unfolding within_def - by (rule eventually_Abs_filter, rule is_filter.intro) - (auto elim!: eventually_rev_mp) - -lemma within_UNIV [simp]: "F within UNIV = F" - unfolding filter_eq_iff eventually_within by simp - -lemma within_empty [simp]: "F within {} = bot" - unfolding filter_eq_iff eventually_within by simp - -lemma within_within_eq: "(F within S) within T = F within (S \ T)" - by (auto simp: filter_eq_iff eventually_within elim: eventually_elim1) - -lemma at_within_eq: "at x within T = nhds x within (T - {x})" - unfolding at_def within_within_eq by (simp add: ac_simps Diff_eq) - -lemma within_le: "F within S \ F" - unfolding le_filter_def eventually_within by (auto elim: eventually_elim1) - -lemma le_withinI: "F \ F' \ eventually (\x. x \ S) F \ F \ F' within S" - unfolding le_filter_def eventually_within by (auto elim: eventually_elim2) - -lemma le_within_iff: "eventually (\x. x \ S) F \ F \ F' within S \ F \ F'" - by (blast intro: within_le le_withinI order_trans) - -lemma eventually_nhds: - "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" -unfolding nhds_def -proof (rule eventually_Abs_filter, rule is_filter.intro) - have "open UNIV \ a \ UNIV \ (\x\UNIV. True)" by simp - thus "\S. open S \ a \ S \ (\x\S. True)" .. -next - fix P Q - assume "\S. open S \ a \ S \ (\x\S. P x)" - and "\T. open T \ a \ T \ (\x\T. Q x)" - then obtain S T where - "open S \ a \ S \ (\x\S. P x)" - "open T \ a \ T \ (\x\T. Q x)" by auto - hence "open (S \ T) \ a \ S \ T \ (\x\(S \ T). P x \ Q x)" - by (simp add: open_Int) - thus "\S. open S \ a \ S \ (\x\S. P x \ Q x)" .. -qed auto lemma eventually_nhds_metric: "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" @@ -472,18 +24,10 @@ apply (erule le_less_trans [OF dist_triangle]) done -lemma nhds_neq_bot [simp]: "nhds a \ bot" - unfolding trivial_limit_def eventually_nhds by simp - -lemma eventually_at_topological: - "eventually P (at a) \ (\S. open S \ a \ S \ (\x\S. x \ a \ P x))" -unfolding at_def eventually_within eventually_nhds by simp - lemma eventually_at: fixes a :: "'a::metric_space" shows "eventually P (at a) \ (\d>0. \x. x \ a \ dist x a < d \ P x)" unfolding at_def eventually_within eventually_nhds_metric by auto - lemma eventually_within_less: (* COPY FROM Topo/eventually_within *) "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" unfolding eventually_within eventually_at dist_nz by auto @@ -492,29 +36,6 @@ "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" unfolding eventually_within_less by auto (metis dense order_le_less_trans) -lemma at_eq_bot_iff: "at a = bot \ open {a}" - unfolding trivial_limit_def eventually_at_topological - by (safe, case_tac "S = {a}", simp, fast, fast) - -lemma at_neq_bot [simp]: "at (a::'a::perfect_space) \ bot" - by (simp add: at_eq_bot_iff not_open_singleton) - -lemma trivial_limit_at_left_real [simp]: (* maybe generalize type *) - "\ trivial_limit (at_left (x::real))" - unfolding trivial_limit_def eventually_within_le - apply clarsimp - apply (rule_tac x="x - d/2" in bexI) - apply (auto simp: dist_real_def) - done - -lemma trivial_limit_at_right_real [simp]: (* maybe generalize type *) - "\ trivial_limit (at_right (x::real))" - unfolding trivial_limit_def eventually_within_le - apply clarsimp - apply (rule_tac x="x + d/2" in bexI) - apply (auto simp: dist_real_def) - done - lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def @@ -722,214 +243,9 @@ lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] - -subsection {* Limits *} - -definition filterlim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where - "filterlim 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 filterlim (%x. f) F2 F1" - -lemma filterlim_iff: - "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" - unfolding filterlim_def le_filter_def eventually_filtermap .. - -lemma filterlim_compose: - "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" - unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) - -lemma filterlim_mono: - "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" - unfolding filterlim_def by (metis filtermap_mono order_trans) - -lemma filterlim_ident: "LIM x F. x :> F" - by (simp add: filterlim_def filtermap_ident) - -lemma filterlim_cong: - "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" - by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) - -lemma filterlim_within: - "(LIM x F1. f x :> F2 within S) \ (eventually (\x. f x \ S) F1 \ (LIM x F1. f x :> F2))" - unfolding filterlim_def -proof safe - assume "filtermap f F1 \ F2 within S" then show "eventually (\x. f x \ S) F1" - by (auto simp: le_filter_def eventually_filtermap eventually_within elim!: allE[of _ "\x. x \ S"]) -qed (auto intro: within_le order_trans simp: le_within_iff eventually_filtermap) - -lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\x. f (g x)) F1 F2" - unfolding filterlim_def filtermap_filtermap .. - -lemma filterlim_sup: - "filterlim f F F1 \ filterlim f F F2 \ filterlim f F (sup F1 F2)" - unfolding filterlim_def filtermap_sup by auto - -lemma filterlim_Suc: "filterlim Suc sequentially sequentially" - by (simp add: filterlim_iff eventually_sequentially) (metis le_Suc_eq) - -abbreviation (in topological_space) - tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "--->" 55) where - "(f ---> l) F \ filterlim f (nhds l) F" - -lemma tendsto_eq_rhs: "(f ---> x) F \ x = y \ (f ---> y) F" - by simp - -ML {* - -structure Tendsto_Intros = Named_Thms -( - val name = @{binding tendsto_intros} - val description = "introduction rules for tendsto" -) - -*} - -setup {* - Tendsto_Intros.setup #> - Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros}, - map (fn thm => @{thm tendsto_eq_rhs} OF [thm]) o Tendsto_Intros.get o Context.proof_of); -*} - -lemma tendsto_def: "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" - unfolding filterlim_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 filterlim_at: - "(LIM x F. f x :> at b) \ (eventually (\x. f x \ b) F \ (f ---> b) F)" - by (simp add: at_def filterlim_within) - -lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" - unfolding tendsto_def le_filter_def by fast - -lemma topological_tendstoI: - "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) - \ (f ---> l) F" - unfolding tendsto_def by auto - -lemma topological_tendstoD: - "(f ---> l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" - unfolding tendsto_def by auto - -lemma tendstoI: - assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) F" - shows "(f ---> l) F" - apply (rule topological_tendstoI) - apply (simp add: open_dist) - apply (drule (1) bspec, clarify) - apply (drule assms) - apply (erule eventually_elim1, simp) - done - -lemma tendstoD: - "(f ---> l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" - apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD) - apply (clarsimp simp add: open_dist) - apply (rule_tac x="e - dist x l" in exI, clarsimp) - apply (simp only: less_diff_eq) - apply (erule le_less_trans [OF dist_triangle]) - apply simp - apply simp - done - -lemma tendsto_iff: - "(f ---> l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" - using tendstoI tendstoD by fast - -lemma order_tendstoI: - fixes y :: "_ :: order_topology" - assumes "\a. a < y \ eventually (\x. a < f x) F" - assumes "\a. y < a \ eventually (\x. f x < a) F" - shows "(f ---> y) F" -proof (rule topological_tendstoI) - fix S assume "open S" "y \ S" - then show "eventually (\x. f x \ S) F" - unfolding open_generated_order - proof induct - case (UN K) - then obtain k where "y \ k" "k \ K" by auto - with UN(2)[of k] show ?case - by (auto elim: eventually_elim1) - qed (insert assms, auto elim: eventually_elim2) -qed - -lemma order_tendstoD: - fixes y :: "_ :: order_topology" - assumes y: "(f ---> y) F" - shows "a < y \ eventually (\x. a < f x) F" - and "y < a \ eventually (\x. f x < a) F" - using topological_tendstoD[OF y, of "{..< a}"] topological_tendstoD[OF y, of "{a <..}"] by auto - -lemma order_tendsto_iff: - fixes f :: "_ \ 'a :: order_topology" - shows "(f ---> x) F \(\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" - by (metis order_tendstoI order_tendstoD) - lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) -lemma tendsto_bot [simp]: "(f ---> a) bot" - unfolding tendsto_def by simp - -lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" - unfolding tendsto_def eventually_at_topological by auto - -lemma tendsto_ident_at_within [tendsto_intros]: - "((\x. x) ---> a) (at a within S)" - unfolding tendsto_def eventually_within eventually_at_topological by auto - -lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) F" - by (simp add: tendsto_def) - -lemma tendsto_unique: - fixes f :: "'a \ 'b::t2_space" - assumes "\ trivial_limit F" and "(f ---> a) F" and "(f ---> b) F" - shows "a = b" -proof (rule ccontr) - assume "a \ b" - obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" - using hausdorff [OF `a \ b`] by fast - have "eventually (\x. f x \ U) F" - using `(f ---> a) F` `open U` `a \ U` by (rule topological_tendstoD) - moreover - have "eventually (\x. f x \ V) F" - using `(f ---> b) F` `open V` `b \ V` by (rule topological_tendstoD) - ultimately - have "eventually (\x. False) F" - proof eventually_elim - case (elim x) - hence "f x \ U \ V" by simp - with `U \ V = {}` show ?case by simp - qed - with `\ trivial_limit F` show "False" - by (simp add: trivial_limit_def) -qed - -lemma tendsto_const_iff: - fixes a b :: "'a::t2_space" - assumes "\ trivial_limit F" shows "((\x. a) ---> b) F \ a = b" - by (safe intro!: tendsto_const tendsto_unique [OF assms tendsto_const]) - -lemma tendsto_at_iff_tendsto_nhds: - "(g ---> g l) (at l) \ (g ---> g l) (nhds l)" - unfolding tendsto_def at_def eventually_within - by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) - -lemma tendsto_compose: - "(g ---> g l) (at l) \ (f ---> l) F \ ((\x. g (f x)) ---> g l) F" - unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) - -lemma tendsto_compose_eventually: - "(g ---> m) (at l) \ (f ---> l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) ---> m) F" - by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) lemma metric_tendsto_imp_tendsto: assumes f: "(f ---> a) F" @@ -941,21 +257,6 @@ with le show "eventually (\x. dist (g x) b < e) F" using le_less_trans by (rule eventually_elim2) qed - -lemma increasing_tendsto: - fixes f :: "_ \ 'a::order_topology" - assumes bdd: "eventually (\n. f n \ l) F" - and en: "\x. x < l \ eventually (\n. x < f n) F" - shows "(f ---> l) F" - using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) - -lemma decreasing_tendsto: - fixes f :: "_ \ 'a::order_topology" - assumes bdd: "eventually (\n. l \ f n) F" - and en: "\x. l < x \ eventually (\n. f n < x) F" - shows "(f ---> l) F" - using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) - subsubsection {* Distance and norms *} lemma tendsto_dist [tendsto_intros]: @@ -1057,19 +358,6 @@ by (simp add: tendsto_const) qed -lemma tendsto_sandwich: - fixes f g h :: "'a \ 'b::order_topology" - 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 (rule order_tendstoI) - fix a show "a < c \ eventually (\x. a < g x) net" - using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) -next - fix a show "c < a \ eventually (\x. g x < a) net" - using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) -qed - lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real] subsubsection {* Linear operators and multiplication *} @@ -1136,31 +424,6 @@ by (simp add: tendsto_const) qed -lemma tendsto_le: - fixes f g :: "'a \ 'b::linorder_topology" - assumes F: "\ trivial_limit F" - assumes x: "(f ---> x) F" and y: "(g ---> y) F" - assumes ev: "eventually (\x. g x \ f x) F" - shows "y \ x" -proof (rule ccontr) - assume "\ y \ x" - with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" - by (auto simp: not_le) - then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" - using x y by (auto intro: order_tendstoD) - with ev have "eventually (\x. False) F" - by eventually_elim (insert xy, fastforce) - with F show False - by (simp add: eventually_False) -qed - -lemma tendsto_le_const: - fixes f :: "'a \ 'b::linorder_topology" - assumes F: "\ trivial_limit F" - assumes x: "(f ---> x) F" and a: "eventually (\x. a \ f x) F" - shows "a \ x" - using F x tendsto_const a by (rule tendsto_le) - subsubsection {* Inverse and division *} lemma (in bounded_bilinear) Zfun_prod_Bfun: @@ -1281,75 +544,6 @@ 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 filterlim_at_top: - fixes f :: "'a \ ('b::linorder)" - shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z \ f x) F)" - by (auto simp: filterlim_iff eventually_at_top_linorder elim!: eventually_elim1) - -lemma filterlim_at_top_dense: - fixes f :: "'a \ ('b::dense_linorder)" - shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" - by (metis eventually_elim1[of _ F] eventually_gt_at_top order_less_imp_le - filterlim_at_top[of f F] filterlim_iff[of f at_top F]) - -lemma filterlim_at_top_ge: - fixes f :: "'a \ ('b::linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_top) \ (\Z\c. eventually (\x. Z \ f x) F)" - unfolding filterlim_at_top -proof safe - fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" - with *[THEN spec, of "max Z c"] show "eventually (\x. Z \ f x) F" - by (auto elim!: eventually_elim1) -qed simp - -lemma filterlim_at_top_at_top: - fixes f :: "'a::linorder \ 'b::linorder" - assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" - assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" - assumes Q: "eventually Q at_top" - assumes P: "eventually P at_top" - shows "filterlim f at_top at_top" -proof - - from P obtain x where x: "\y. x \ y \ P y" - unfolding eventually_at_top_linorder by auto - show ?thesis - proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) - fix z assume "x \ z" - with x have "P z" by auto - have "eventually (\x. g z \ x) at_top" - by (rule eventually_ge_at_top) - with Q show "eventually (\x. z \ f x) at_top" - by eventually_elim (metis mono bij `P z`) - qed -qed - -lemma filterlim_at_top_gt: - fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_top) \ (\Z>c. eventually (\x. Z \ f x) F)" - by (metis filterlim_at_top order_less_le_trans gt_ex filterlim_at_top_ge) - -lemma filterlim_at_bot: - fixes f :: "'a \ ('b::linorder)" - shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x \ Z) F)" - by (auto simp: filterlim_iff eventually_at_bot_linorder elim!: eventually_elim1) - -lemma filterlim_at_bot_le: - fixes f :: "'a \ ('b::linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_bot) \ (\Z\c. eventually (\x. Z \ f x) F)" - unfolding filterlim_at_bot -proof safe - fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" - with *[THEN spec, of "min Z c"] show "eventually (\x. Z \ f x) F" - by (auto elim!: eventually_elim1) -qed simp - -lemma filterlim_at_bot_lt: - fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" - shows "(LIM x F. f x :> at_bot) \ (\Zx. Z \ f x) F)" - by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans) - lemma filterlim_at_bot_at_right: fixes f :: "real \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" @@ -1429,13 +623,7 @@ *} -lemma at_eq_sup_left_right: "at (x::real) = sup (at_left x) (at_right x)" - by (auto simp: eventually_within at_def filter_eq_iff eventually_sup - elim: eventually_elim2 eventually_elim1) - -lemma filterlim_split_at_real: - "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at (x::real))" - by (subst at_eq_sup_left_right) (rule filterlim_sup) +lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d::real)" unfolding filter_eq_iff eventually_filtermap eventually_nhds_metric @@ -1486,14 +674,6 @@ "eventually P (at_left (a::real)) \ eventually (\x. P (- x)) (at_right (-a))" unfolding at_left_minus[of a] by (simp add: eventually_filtermap) -lemma filterlim_at_split: - "filterlim f F (at (x::real)) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" - by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) - -lemma eventually_at_split: - "eventually P (at (x::real)) \ eventually P (at_left x) \ eventually P (at_right x)" - by (subst at_eq_sup_left_right) (simp add: eventually_sup) - lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: real filter)" unfolding filter_eq_iff eventually_filtermap eventually_at_top_linorder eventually_at_bot_linorder by (metis le_minus_iff minus_minus) @@ -1765,5 +945,10 @@ by (auto simp: norm_power) qed simp + +(* Unfortunately eventually_within was overwritten by Multivariate_Analysis. + Hence it was references as Limits.within, but now it is Basic_Topology.eventually_within *) +lemmas eventually_within = eventually_within + end diff -r a981a5c8a505 -r cad22a3cc09c src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Mar 21 16:58:14 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Mar 22 10:41:42 2013 +0100 @@ -1340,11 +1340,11 @@ text {* Some property holds "sufficiently close" to the limit point. *} -lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *) +lemma eventually_at: (* FIXME: this replaces Metric_Spaces.eventually_at *) "eventually P (at a) \ (\d>0. \x. 0 < dist x a \ dist x a < d \ P x)" unfolding eventually_at dist_nz by auto -lemma eventually_within: (* FIXME: this replaces Limits.eventually_within *) +lemma eventually_within: (* FIXME: this replaces Topological_Spaces.eventually_within *) "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" by (rule eventually_within_less) @@ -1448,12 +1448,12 @@ from assms obtain T where T: "open T" "x \ T" "T \ S" .. { assume "?lhs" then obtain A where "open A" "x \ A" "\y\A. y \ x \ y \ S \ P y" - unfolding Limits.eventually_within Limits.eventually_at_topological + unfolding Limits.eventually_within eventually_at_topological by auto with T have "open (A \ T)" "x \ A \ T" "\y\(A \ T). y \ x \ P y" by auto then have "?rhs" - unfolding Limits.eventually_at_topological by auto + unfolding eventually_at_topological by auto } moreover { assume "?rhs" hence "?lhs" unfolding Limits.eventually_within diff -r a981a5c8a505 -r cad22a3cc09c src/HOL/RealVector.thy --- a/src/HOL/RealVector.thy Thu Mar 21 16:58:14 2013 +0100 +++ b/src/HOL/RealVector.thy Fri Mar 22 10:41:42 2013 +0100 @@ -434,131 +434,6 @@ by (rule Reals_cases) auto -subsection {* Topological spaces *} - -class "open" = - fixes "open" :: "'a set \ bool" - -class topological_space = "open" + - assumes open_UNIV [simp, intro]: "open UNIV" - assumes open_Int [intro]: "open S \ open T \ open (S \ T)" - assumes open_Union [intro]: "\S\K. open S \ open (\ K)" -begin - -definition - closed :: "'a set \ bool" where - "closed S \ open (- S)" - -lemma open_empty [intro, simp]: "open {}" - using open_Union [of "{}"] by simp - -lemma open_Un [intro]: "open S \ open T \ open (S \ T)" - using open_Union [of "{S, T}"] by simp - -lemma open_UN [intro]: "\x\A. open (B x) \ open (\x\A. B x)" - unfolding SUP_def by (rule open_Union) auto - -lemma open_Inter [intro]: "finite S \ \T\S. open T \ open (\S)" - by (induct set: finite) auto - -lemma open_INT [intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" - unfolding INF_def by (rule open_Inter) auto - -lemma closed_empty [intro, simp]: "closed {}" - unfolding closed_def by simp - -lemma closed_Un [intro]: "closed S \ closed T \ closed (S \ T)" - unfolding closed_def by auto - -lemma closed_UNIV [intro, simp]: "closed UNIV" - unfolding closed_def by simp - -lemma closed_Int [intro]: "closed S \ closed T \ closed (S \ T)" - unfolding closed_def by auto - -lemma closed_INT [intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" - unfolding closed_def by auto - -lemma closed_Inter [intro]: "\S\K. closed S \ closed (\ K)" - unfolding closed_def uminus_Inf by auto - -lemma closed_Union [intro]: "finite S \ \T\S. closed T \ closed (\S)" - by (induct set: finite) auto - -lemma closed_UN [intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" - unfolding SUP_def by (rule closed_Union) auto - -lemma open_closed: "open S \ closed (- S)" - unfolding closed_def by simp - -lemma closed_open: "closed S \ open (- S)" - unfolding closed_def by simp - -lemma open_Diff [intro]: "open S \ closed T \ open (S - T)" - unfolding closed_open Diff_eq by (rule open_Int) - -lemma closed_Diff [intro]: "closed S \ open T \ closed (S - T)" - unfolding open_closed Diff_eq by (rule closed_Int) - -lemma open_Compl [intro]: "closed S \ open (- S)" - unfolding closed_open . - -lemma closed_Compl [intro]: "open S \ closed (- S)" - unfolding open_closed . - -end - -inductive generate_topology for S where - UNIV: "generate_topology S UNIV" -| Int: "generate_topology S a \ generate_topology S b \ generate_topology S (a \ b)" -| UN: "(\k. k \ K \ generate_topology S k) \ generate_topology S (\K)" -| Basis: "s \ S \ generate_topology S s" - -hide_fact (open) UNIV Int UN Basis - -lemma generate_topology_Union: - "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" - unfolding SUP_def by (intro generate_topology.UN) auto - -lemma topological_space_generate_topology: - "class.topological_space (generate_topology S)" - by default (auto intro: generate_topology.intros) - -class order_topology = order + "open" + - assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" -begin - -subclass topological_space - unfolding open_generated_order - by (rule topological_space_generate_topology) - -lemma open_greaterThan [simp]: "open {a <..}" - unfolding open_generated_order by (auto intro: generate_topology.Basis) - -lemma open_lessThan [simp]: "open {..< a}" - unfolding open_generated_order by (auto intro: generate_topology.Basis) - -lemma open_greaterThanLessThan [simp]: "open {a <..< b}" - unfolding greaterThanLessThan_eq by (simp add: open_Int) - -end - -class linorder_topology = linorder + order_topology - -lemma closed_atMost [simp]: "closed {.. a::'a::linorder_topology}" - by (simp add: closed_open) - -lemma closed_atLeast [simp]: "closed {a::'a::linorder_topology ..}" - by (simp add: closed_open) - -lemma closed_atLeastAtMost [simp]: "closed {a::'a::linorder_topology .. b}" -proof - - have "{a .. b} = {a ..} \ {.. b}" - by auto - then show ?thesis - by (simp add: closed_Int) -qed - subsection {* Metric spaces *} class dist = @@ -1182,78 +1057,6 @@ lemma bounded_linear_of_real: "bounded_linear (\r. of_real r)" unfolding of_real_def by (rule bounded_linear_scaleR_left) -subsection{* Hausdorff and other separation properties *} - -class t0_space = topological_space + - assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" - -class t1_space = topological_space + - assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" - -instance t1_space \ t0_space -proof qed (fast dest: t1_space) - -lemma separation_t1: - fixes x y :: "'a::t1_space" - shows "x \ y \ (\U. open U \ x \ U \ y \ U)" - using t1_space[of x y] by blast - -lemma closed_singleton: - fixes a :: "'a::t1_space" - shows "closed {a}" -proof - - let ?T = "\{S. open S \ a \ S}" - have "open ?T" by (simp add: open_Union) - also have "?T = - {a}" - by (simp add: set_eq_iff separation_t1, auto) - finally show "closed {a}" unfolding closed_def . -qed - -lemma closed_insert [simp]: - fixes a :: "'a::t1_space" - assumes "closed S" shows "closed (insert a S)" -proof - - from closed_singleton assms - have "closed ({a} \ S)" by (rule closed_Un) - thus "closed (insert a S)" by simp -qed - -lemma finite_imp_closed: - fixes S :: "'a::t1_space set" - shows "finite S \ closed S" -by (induct set: finite, simp_all) - -text {* T2 spaces are also known as Hausdorff spaces. *} - -class t2_space = topological_space + - assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - -instance t2_space \ t1_space -proof qed (fast dest: hausdorff) - -lemma (in linorder) less_separate: - assumes "x < y" - shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" -proof cases - assume "\z. x < z \ z < y" - then guess z .. - then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" - by auto - then show ?thesis by blast -next - assume "\ (\z. x < z \ z < y)" - with `x < y` have "x \ {..< y} \ y \ {x <..} \ {x <..} \ {..< y} = {}" - by auto - then show ?thesis by blast -qed - -instance linorder_topology \ t2_space -proof - fix x y :: 'a - from less_separate[of x y] less_separate[of y x] - show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - by (elim neqE) (metis open_lessThan open_greaterThan Int_commute)+ -qed instance metric_space \ t2_space proof @@ -1269,22 +1072,6 @@ then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by blast qed - -lemma separation_t2: - fixes x y :: "'a::t2_space" - shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" - using hausdorff[of x y] by blast - -lemma separation_t0: - fixes x y :: "'a::t0_space" - shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" - using t0_space[of x y] by blast - -text {* A perfect space is a topological space with no isolated points. *} - -class perfect_space = topological_space + - assumes not_open_singleton: "\ open {x}" - instance real_normed_algebra_1 \ perfect_space proof fix x::'a diff -r a981a5c8a505 -r cad22a3cc09c src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Thu Mar 21 16:58:14 2013 +0100 +++ b/src/HOL/SEQ.thy Fri Mar 22 10:41:42 2013 +0100 @@ -10,219 +10,11 @@ header {* Sequences and Convergence *} theory SEQ -imports Limits RComplete +imports Limits begin -subsection {* Monotone sequences and subsequences *} - -definition - monoseq :: "(nat \ 'a::order) \ bool" where - --{*Definition of monotonicity. - The use of disjunction here complicates proofs considerably. - One alternative is to add a Boolean argument to indicate the direction. - Another is to develop the notions of increasing and decreasing first.*} - "monoseq X = ((\m. \n\m. X m \ X n) | (\m. \n\m. X n \ X m))" - -definition - incseq :: "(nat \ 'a::order) \ bool" where - --{*Increasing sequence*} - "incseq X \ (\m. \n\m. X m \ X n)" - -definition - decseq :: "(nat \ 'a::order) \ bool" where - --{*Decreasing sequence*} - "decseq X \ (\m. \n\m. X n \ X m)" - -definition - subseq :: "(nat \ nat) \ bool" where - --{*Definition of subsequence*} - "subseq f \ (\m. \n>m. f m < f n)" - -lemma incseq_mono: "mono f \ incseq f" - unfolding mono_def incseq_def by auto - -lemma incseq_SucI: - "(\n. X n \ X (Suc n)) \ incseq X" - using lift_Suc_mono_le[of X] - by (auto simp: incseq_def) - -lemma incseqD: "\i j. incseq f \ i \ j \ f i \ f j" - by (auto simp: incseq_def) - -lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" - using incseqD[of A i "Suc i"] by auto - -lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" - by (auto intro: incseq_SucI dest: incseq_SucD) - -lemma incseq_const[simp, intro]: "incseq (\x. k)" - unfolding incseq_def by auto - -lemma decseq_SucI: - "(\n. X (Suc n) \ X n) \ decseq X" - using order.lift_Suc_mono_le[OF dual_order, of X] - by (auto simp: decseq_def) - -lemma decseqD: "\i j. decseq f \ i \ j \ f j \ f i" - by (auto simp: decseq_def) - -lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" - using decseqD[of A i "Suc i"] by auto - -lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" - by (auto intro: decseq_SucI dest: decseq_SucD) - -lemma decseq_const[simp, intro]: "decseq (\x. k)" - unfolding decseq_def by auto - -lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" - unfolding monoseq_def incseq_def decseq_def .. - -lemma monoseq_Suc: - "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" - unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. - -lemma monoI1: "\m. \ n \ m. X m \ X n ==> monoseq X" -by (simp add: monoseq_def) - -lemma monoI2: "\m. \ n \ m. X n \ X m ==> monoseq X" -by (simp add: monoseq_def) - -lemma mono_SucI1: "\n. X n \ X (Suc n) ==> monoseq X" -by (simp add: monoseq_Suc) - -lemma mono_SucI2: "\n. X (Suc n) \ X n ==> monoseq X" -by (simp add: monoseq_Suc) - -lemma monoseq_minus: - fixes a :: "nat \ 'a::ordered_ab_group_add" - assumes "monoseq a" - shows "monoseq (\ n. - a n)" -proof (cases "\ m. \ n \ m. a m \ a n") - case True - hence "\ m. \ n \ m. - a n \ - a m" by auto - thus ?thesis by (rule monoI2) -next - case False - hence "\ m. \ n \ m. - a m \ - a n" using `monoseq a`[unfolded monoseq_def] by auto - thus ?thesis by (rule monoI1) -qed - -text{*Subsequence (alternative definition, (e.g. Hoskins)*} - -lemma subseq_Suc_iff: "subseq f = (\n. (f n) < (f (Suc n)))" -apply (simp add: subseq_def) -apply (auto dest!: less_imp_Suc_add) -apply (induct_tac k) -apply (auto intro: less_trans) -done - -text{* for any sequence, there is a monotonic subsequence *} -lemma seq_monosub: - fixes s :: "nat => 'a::linorder" - shows "\f. subseq f \ monoseq (\ n. (s (f n)))" -proof cases - let "?P p n" = "p > n \ (\m\p. s m \ s p)" - assume *: "\n. \p. ?P p n" - def f \ "nat_rec (SOME p. ?P p 0) (\_ n. SOME p. ?P p n)" - have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp - have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. - have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto - have P_Suc: "\i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto - then have "subseq f" unfolding subseq_Suc_iff by auto - moreover have "monoseq (\n. s (f n))" unfolding monoseq_Suc - proof (intro disjI2 allI) - fix n show "s (f (Suc n)) \ s (f n)" - proof (cases n) - case 0 with P_Suc[of 0] P_0 show ?thesis by auto - next - case (Suc m) - from P_Suc[of n] Suc have "f (Suc m) \ f (Suc (Suc m))" by simp - with P_Suc Suc show ?thesis by simp - qed - qed - ultimately show ?thesis by auto -next - let "?P p m" = "m < p \ s m < s p" - assume "\ (\n. \p>n. (\m\p. s m \ s p))" - then obtain N where N: "\p. p > N \ \m>p. s p < s m" by (force simp: not_le le_less) - def f \ "nat_rec (SOME p. ?P p (Suc N)) (\_ n. SOME p. ?P p n)" - have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp - have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. - have P_0: "?P (f 0) (Suc N)" - unfolding f_0 some_eq_ex[of "\p. ?P p (Suc N)"] using N[of "Suc N"] by auto - { fix i have "N < f i \ ?P (f (Suc i)) (f i)" - unfolding f_Suc some_eq_ex[of "\p. ?P p (f i)"] using N[of "f i"] . } - note P' = this - { fix i have "N < f i \ ?P (f (Suc i)) (f i)" - by (induct i) (insert P_0 P', auto) } - then have "subseq f" "monoseq (\x. s (f x))" - unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le) - then show ?thesis by auto -qed - -lemma seq_suble: assumes sf: "subseq f" shows "n \ f n" -proof(induct n) - case 0 thus ?case by simp -next - case (Suc n) - from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps - have "n < f (Suc n)" by arith - thus ?case by arith -qed - -lemma eventually_subseq: - "subseq r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" - unfolding eventually_sequentially by (metis seq_suble le_trans) - -lemma filterlim_subseq: "subseq f \ filterlim f sequentially sequentially" - unfolding filterlim_iff by (metis eventually_subseq) - -lemma subseq_o: "subseq r \ subseq s \ subseq (r \ s)" - unfolding subseq_def by simp - -lemma subseq_mono: assumes "subseq r" "m < n" shows "r m < r n" - using assms by (auto simp: subseq_def) - -lemma incseq_imp_monoseq: "incseq X \ monoseq X" - by (simp add: incseq_def monoseq_def) - -lemma decseq_imp_monoseq: "decseq X \ monoseq X" - by (simp add: decseq_def monoseq_def) - -lemma decseq_eq_incseq: - fixes X :: "nat \ 'a::ordered_ab_group_add" shows "decseq X = incseq (\n. - X n)" - by (simp add: decseq_def incseq_def) - -lemma INT_decseq_offset: - assumes "decseq F" - shows "(\i. F i) = (\i\{n..}. F i)" -proof safe - fix x i assume x: "x \ (\i\{n..}. F i)" - show "x \ F i" - proof cases - from x have "x \ F n" by auto - also assume "i \ n" with `decseq F` have "F n \ F i" - unfolding decseq_def by simp - finally show ?thesis . - qed (insert x, simp) -qed auto - subsection {* Defintions of limits *} -abbreviation (in topological_space) - LIMSEQ :: "[nat \ 'a, 'a] \ bool" - ("((_)/ ----> (_))" [60, 60] 60) where - "X ----> L \ (X ---> L) sequentially" - -definition - lim :: "(nat \ 'a::t2_space) \ 'a" where - --{*Standard definition of limit using choice operator*} - "lim X = (THE L. X ----> L)" - -definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where - "convergent X = (\L. X ----> L)" - definition Bseq :: "(nat => 'a::real_normed_vector) => bool" where --{*Standard definition for bounded sequence*} @@ -317,78 +109,10 @@ shows "\X ----> L; 0 < r\ \ \no. \n\no. norm (X n - L) < r" by (simp add: LIMSEQ_iff) -lemma LIMSEQ_const_iff: - fixes k l :: "'a::t2_space" - shows "(\n. k) ----> l \ k = l" - using trivial_limit_sequentially by (rule tendsto_const_iff) - -lemma LIMSEQ_SUP: - "incseq X \ X ----> (SUP i. X i :: 'a :: {complete_linorder, linorder_topology})" - by (intro increasing_tendsto) - (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) - -lemma LIMSEQ_INF: - "decseq X \ X ----> (INF i. X i :: 'a :: {complete_linorder, linorder_topology})" - by (intro decreasing_tendsto) - (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) - -lemma LIMSEQ_ignore_initial_segment: - "f ----> a \ (\n. f (n + k)) ----> a" -apply (rule topological_tendstoI) -apply (drule (2) topological_tendstoD) -apply (simp only: eventually_sequentially) -apply (erule exE, rename_tac N) -apply (rule_tac x=N in exI) -apply simp -done - -lemma LIMSEQ_offset: - "(\n. f (n + k)) ----> a \ f ----> a" -apply (rule topological_tendstoI) -apply (drule (2) topological_tendstoD) -apply (simp only: eventually_sequentially) -apply (erule exE, rename_tac N) -apply (rule_tac x="N + k" in exI) -apply clarify -apply (drule_tac x="n - k" in spec) -apply (simp add: le_diff_conv2) -done - -lemma LIMSEQ_Suc: "f ----> l \ (\n. f (Suc n)) ----> l" -by (drule_tac k="Suc 0" in LIMSEQ_ignore_initial_segment, simp) - -lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) ----> l \ f ----> l" -by (rule_tac k="Suc 0" in LIMSEQ_offset, simp) - -lemma LIMSEQ_Suc_iff: "(\n. f (Suc n)) ----> l = f ----> l" -by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) - lemma LIMSEQ_linear: "\ X ----> x ; l > 0 \ \ (\ n. X (n * l)) ----> x" unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute) -lemma LIMSEQ_unique: - fixes a b :: "'a::t2_space" - shows "\X ----> a; X ----> b\ \ a = b" - using trivial_limit_sequentially by (rule tendsto_unique) - -lemma increasing_LIMSEQ: - fixes f :: "nat \ real" - assumes inc: "\n. f n \ f (Suc n)" - and bdd: "\n. f n \ l" - and en: "\e. 0 < e \ \n. l \ f n + e" - shows "f ----> l" -proof (rule increasing_tendsto) - fix x assume "x < l" - with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" - by auto - from en[OF `0 < e`] obtain n where "l - e \ f n" - by (auto simp: field_simps) - with `e < l - x` `0 < e` have "x < f n" by simp - with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" - by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) -qed (insert bdd, auto) - lemma Bseq_inverse_lemma: fixes x :: "'a::real_normed_div_algebra" shows "\r \ norm x; 0 < r\ \ norm (inverse x) \ inverse r" @@ -443,37 +167,8 @@ using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] by auto -lemma LIMSEQ_le_const: - "\X ----> (x::'a::linorder_topology); \N. \n\N. a \ X n\ \ a \ x" - using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially) - -lemma LIMSEQ_le: - "\X ----> x; Y ----> y; \N. \n\N. X n \ Y n\ \ x \ (y::'a::linorder_topology)" - using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) - -lemma LIMSEQ_le_const2: - "\X ----> (x::'a::linorder_topology); \N. \n\N. X n \ a\ \ x \ a" - by (rule LIMSEQ_le[of X x "\n. a"]) (auto simp: tendsto_const) - subsection {* Convergence *} -lemma limI: "X ----> L ==> lim X = L" -apply (simp add: lim_def) -apply (blast intro: LIMSEQ_unique) -done - -lemma convergentD: "convergent X ==> \L. (X ----> L)" -by (simp add: convergent_def) - -lemma convergentI: "(X ----> L) ==> convergent X" -by (auto simp add: convergent_def) - -lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)" -by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) - -lemma convergent_const: "convergent (\n. c)" - by (rule convergentI, rule tendsto_const) - lemma convergent_add: fixes X Y :: "nat \ 'a::real_normed_vector" assumes "convergent (\n. X n)" @@ -508,22 +203,6 @@ apply (drule tendsto_minus, auto) done -lemma lim_le: "convergent f \ (\n. f n \ (x::real)) \ lim f \ x" - using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) - -lemma monoseq_le: - "monoseq a \ a ----> (x::real) \ - ((\ n. a n \ x) \ (\m. \n\m. a m \ a n)) \ ((\ n. x \ a n) \ (\m. \n\m. a n \ a m))" - by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) - -lemma LIMSEQ_subseq_LIMSEQ: - "\ X ----> L; subseq f \ \ (X o f) ----> L" - unfolding comp_def by (rule filterlim_compose[of X, OF _ filterlim_subseq]) - -lemma convergent_subseq_convergent: - "\convergent X; subseq f\ \ convergent (X o f)" - unfolding convergent_def by (auto intro: LIMSEQ_subseq_LIMSEQ) - subsection {* Bounded Monotonic Sequences *} @@ -665,14 +344,6 @@ by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff Bseq_mono_convergent) -subsubsection{*Increasing and Decreasing Series*} - -lemma incseq_le: "incseq X \ X ----> L \ X n \ (L::real)" - by (metis incseq_def LIMSEQ_le_const) - -lemma decseq_le: "decseq X \ X ----> L \ (L::real) \ X n" - by (metis decseq_def LIMSEQ_le_const2) - subsection {* Cauchy Sequences *} lemma metric_CauchyI: diff -r a981a5c8a505 -r cad22a3cc09c src/HOL/Topological_Spaces.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Topological_Spaces.thy Fri Mar 22 10:41:42 2013 +0100 @@ -0,0 +1,1511 @@ +(* Title: HOL/Basic_Topology.thy + Author: Brian Huffman + Author: Johannes Hölzl +*) + +header {* Topological Spaces *} + +theory Topological_Spaces +imports Main +begin + +subsection {* Topological space *} + +class "open" = + fixes "open" :: "'a set \ bool" + +class topological_space = "open" + + assumes open_UNIV [simp, intro]: "open UNIV" + assumes open_Int [intro]: "open S \ open T \ open (S \ T)" + assumes open_Union [intro]: "\S\K. open S \ open (\ K)" +begin + +definition + closed :: "'a set \ bool" where + "closed S \ open (- S)" + +lemma open_empty [intro, simp]: "open {}" + using open_Union [of "{}"] by simp + +lemma open_Un [intro]: "open S \ open T \ open (S \ T)" + using open_Union [of "{S, T}"] by simp + +lemma open_UN [intro]: "\x\A. open (B x) \ open (\x\A. B x)" + unfolding SUP_def by (rule open_Union) auto + +lemma open_Inter [intro]: "finite S \ \T\S. open T \ open (\S)" + by (induct set: finite) auto + +lemma open_INT [intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" + unfolding INF_def by (rule open_Inter) auto + +lemma closed_empty [intro, simp]: "closed {}" + unfolding closed_def by simp + +lemma closed_Un [intro]: "closed S \ closed T \ closed (S \ T)" + unfolding closed_def by auto + +lemma closed_UNIV [intro, simp]: "closed UNIV" + unfolding closed_def by simp + +lemma closed_Int [intro]: "closed S \ closed T \ closed (S \ T)" + unfolding closed_def by auto + +lemma closed_INT [intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" + unfolding closed_def by auto + +lemma closed_Inter [intro]: "\S\K. closed S \ closed (\ K)" + unfolding closed_def uminus_Inf by auto + +lemma closed_Union [intro]: "finite S \ \T\S. closed T \ closed (\S)" + by (induct set: finite) auto + +lemma closed_UN [intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" + unfolding SUP_def by (rule closed_Union) auto + +lemma open_closed: "open S \ closed (- S)" + unfolding closed_def by simp + +lemma closed_open: "closed S \ open (- S)" + unfolding closed_def by simp + +lemma open_Diff [intro]: "open S \ closed T \ open (S - T)" + unfolding closed_open Diff_eq by (rule open_Int) + +lemma closed_Diff [intro]: "closed S \ open T \ closed (S - T)" + unfolding open_closed Diff_eq by (rule closed_Int) + +lemma open_Compl [intro]: "closed S \ open (- S)" + unfolding closed_open . + +lemma closed_Compl [intro]: "open S \ closed (- S)" + unfolding open_closed . + +end + +subsection{* Hausdorff and other separation properties *} + +class t0_space = topological_space + + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" + +class t1_space = topological_space + + assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" + +instance t1_space \ t0_space +proof qed (fast dest: t1_space) + +lemma separation_t1: + fixes x y :: "'a::t1_space" + shows "x \ y \ (\U. open U \ x \ U \ y \ U)" + using t1_space[of x y] by blast + +lemma closed_singleton: + fixes a :: "'a::t1_space" + shows "closed {a}" +proof - + let ?T = "\{S. open S \ a \ S}" + have "open ?T" by (simp add: open_Union) + also have "?T = - {a}" + by (simp add: set_eq_iff separation_t1, auto) + finally show "closed {a}" unfolding closed_def . +qed + +lemma closed_insert [simp]: + fixes a :: "'a::t1_space" + assumes "closed S" shows "closed (insert a S)" +proof - + from closed_singleton assms + have "closed ({a} \ S)" by (rule closed_Un) + thus "closed (insert a S)" by simp +qed + +lemma finite_imp_closed: + fixes S :: "'a::t1_space set" + shows "finite S \ closed S" +by (induct set: finite, simp_all) + +text {* T2 spaces are also known as Hausdorff spaces. *} + +class t2_space = topological_space + + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + +instance t2_space \ t1_space +proof qed (fast dest: hausdorff) + +lemma separation_t2: + fixes x y :: "'a::t2_space" + shows "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" + using hausdorff[of x y] by blast + +lemma separation_t0: + fixes x y :: "'a::t0_space" + shows "x \ y \ (\U. open U \ ~(x\U \ y\U))" + using t0_space[of x y] by blast + +text {* A perfect space is a topological space with no isolated points. *} + +class perfect_space = topological_space + + assumes not_open_singleton: "\ open {x}" + + +subsection {* Generators for toplogies *} + +inductive generate_topology for S where + UNIV: "generate_topology S UNIV" +| Int: "generate_topology S a \ generate_topology S b \ generate_topology S (a \ b)" +| UN: "(\k. k \ K \ generate_topology S k) \ generate_topology S (\K)" +| Basis: "s \ S \ generate_topology S s" + +hide_fact (open) UNIV Int UN Basis + +lemma generate_topology_Union: + "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" + unfolding SUP_def by (intro generate_topology.UN) auto + +lemma topological_space_generate_topology: + "class.topological_space (generate_topology S)" + by default (auto intro: generate_topology.intros) + +subsection {* Order topologies *} + +class order_topology = order + "open" + + assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" +begin + +subclass topological_space + unfolding open_generated_order + by (rule topological_space_generate_topology) + +lemma open_greaterThan [simp]: "open {a <..}" + unfolding open_generated_order by (auto intro: generate_topology.Basis) + +lemma open_lessThan [simp]: "open {..< a}" + unfolding open_generated_order by (auto intro: generate_topology.Basis) + +lemma open_greaterThanLessThan [simp]: "open {a <..< b}" + unfolding greaterThanLessThan_eq by (simp add: open_Int) + +end + +class linorder_topology = linorder + order_topology + +lemma closed_atMost [simp]: "closed {.. a::'a::linorder_topology}" + by (simp add: closed_open) + +lemma closed_atLeast [simp]: "closed {a::'a::linorder_topology ..}" + by (simp add: closed_open) + +lemma closed_atLeastAtMost [simp]: "closed {a::'a::linorder_topology .. b}" +proof - + have "{a .. b} = {a ..} \ {.. b}" + by auto + then show ?thesis + by (simp add: closed_Int) +qed + +lemma (in linorder) less_separate: + assumes "x < y" + shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" +proof cases + assume "\z. x < z \ z < y" + then guess z .. + then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" + by auto + then show ?thesis by blast +next + assume "\ (\z. x < z \ z < y)" + with `x < y` have "x \ {..< y} \ y \ {x <..} \ {x <..} \ {..< y} = {}" + by auto + then show ?thesis by blast +qed + +instance linorder_topology \ t2_space +proof + fix x y :: 'a + from less_separate[of x y] less_separate[of y x] + show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" + by (elim neqE) (metis open_lessThan open_greaterThan Int_commute)+ +qed + +lemma open_right: + fixes S :: "'a :: {no_top, linorder_topology} set" + assumes "open S" "x \ S" shows "\b>x. {x ..< b} \ S" + using assms unfolding open_generated_order +proof induction + case (Int A B) + then obtain a b where "a > x" "{x ..< a} \ A" "b > x" "{x ..< b} \ B" by auto + then show ?case by (auto intro!: exI[of _ "min a b"]) +next + case (Basis S) + moreover from gt_ex[of x] guess b .. + ultimately show ?case by (fastforce intro: exI[of _ b]) +qed (fastforce intro: gt_ex)+ + +lemma open_left: + fixes S :: "'a :: {no_bot, linorder_topology} set" + assumes "open S" "x \ S" shows "\b S" + using assms unfolding open_generated_order +proof induction + case (Int A B) + then obtain a b where "a < x" "{a <.. x} \ A" "b < x" "{b <.. x} \ B" by auto + then show ?case by (auto intro!: exI[of _ "max a b"]) +next + case (Basis S) + moreover from lt_ex[of x] guess b .. + ultimately show ?case by (fastforce intro: exI[of _ b]) +next + case UN then show ?case by blast +qed (fastforce intro: lt_ex) + +subsection {* Filters *} + +text {* + This definition also allows non-proper filters. +*} + +locale is_filter = + fixes F :: "('a \ bool) \ bool" + assumes True: "F (\x. True)" + assumes conj: "F (\x. P x) \ F (\x. Q x) \ F (\x. P x \ Q x)" + assumes mono: "\x. P x \ Q x \ F (\x. P x) \ F (\x. Q x)" + +typedef 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" +proof + show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) +qed + +lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" + using Rep_filter [of F] by simp + +lemma Abs_filter_inverse': + assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" + using assms by (simp add: Abs_filter_inverse) + + +subsubsection {* Eventually *} + +definition eventually :: "('a \ bool) \ 'a filter \ bool" + where "eventually P F \ Rep_filter F P" + +lemma eventually_Abs_filter: + assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" + unfolding eventually_def using assms by (simp add: Abs_filter_inverse) + +lemma filter_eq_iff: + shows "F = F' \ (\P. eventually P F = eventually P F')" + unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. + +lemma eventually_True [simp]: "eventually (\x. True) F" + unfolding eventually_def + by (rule is_filter.True [OF is_filter_Rep_filter]) + +lemma always_eventually: "\x. P x \ eventually P F" +proof - + assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) + thus "eventually P F" by simp +qed + +lemma eventually_mono: + "(\x. P x \ Q x) \ eventually P F \ eventually Q F" + unfolding eventually_def + by (rule is_filter.mono [OF is_filter_Rep_filter]) + +lemma eventually_conj: + assumes P: "eventually (\x. P x) F" + assumes Q: "eventually (\x. Q x) F" + shows "eventually (\x. P x \ Q x) F" + using assms unfolding eventually_def + by (rule is_filter.conj [OF is_filter_Rep_filter]) + +lemma eventually_Ball_finite: + assumes "finite A" and "\y\A. eventually (\x. P x y) net" + shows "eventually (\x. \y\A. P x y) net" +using assms by (induct set: finite, simp, simp add: eventually_conj) + +lemma eventually_all_finite: + fixes P :: "'a \ 'b::finite \ bool" + assumes "\y. eventually (\x. P x y) net" + shows "eventually (\x. \y. P x y) net" +using eventually_Ball_finite [of UNIV P] assms by simp + +lemma eventually_mp: + assumes "eventually (\x. P x \ Q x) F" + assumes "eventually (\x. P x) F" + shows "eventually (\x. Q x) F" +proof (rule eventually_mono) + show "\x. (P x \ Q x) \ P x \ Q x" by simp + show "eventually (\x. (P x \ Q x) \ P x) F" + using assms by (rule eventually_conj) +qed + +lemma eventually_rev_mp: + assumes "eventually (\x. P x) F" + assumes "eventually (\x. P x \ Q x) F" + shows "eventually (\x. Q x) F" +using assms(2) assms(1) by (rule eventually_mp) + +lemma eventually_conj_iff: + "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" + by (auto intro: eventually_conj elim: eventually_rev_mp) + +lemma eventually_elim1: + assumes "eventually (\i. P i) F" + assumes "\i. P i \ Q i" + shows "eventually (\i. Q i) F" + using assms by (auto elim!: eventually_rev_mp) + +lemma eventually_elim2: + assumes "eventually (\i. P i) F" + assumes "eventually (\i. Q i) F" + assumes "\i. P i \ Q i \ R i" + 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 + +ML {* + fun eventually_elim_tac ctxt thms thm = + let + val thy = Proof_Context.theory_of ctxt + val mp_thms = thms RL [@{thm eventually_rev_mp}] + val raw_elim_thm = + (@{thm allI} RS @{thm always_eventually}) + |> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms + |> fold (fn _ => fn thm => @{thm impI} RS thm) thms + val cases_prop = prop_of (raw_elim_thm RS thm) + val cases = (Rule_Cases.make_common (thy, cases_prop) [(("elim", []), [])]) + in + CASES cases (rtac raw_elim_thm 1) thm + end +*} + +method_setup eventually_elim = {* + Scan.succeed (fn ctxt => METHOD_CASES (eventually_elim_tac ctxt)) +*} "elimination of eventually quantifiers" + + +subsubsection {* Finer-than relation *} + +text {* @{term "F \ F'"} means that filter @{term F} is finer than +filter @{term F'}. *} + +instantiation filter :: (type) complete_lattice +begin + +definition le_filter_def: + "F \ F' \ (\P. eventually P F' \ eventually P F)" + +definition + "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" + +definition + "top = Abs_filter (\P. \x. P x)" + +definition + "bot = Abs_filter (\P. True)" + +definition + "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" + +definition + "inf F F' = Abs_filter + (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" + +definition + "Sup S = Abs_filter (\P. \F\S. eventually P F)" + +definition + "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" + +lemma eventually_top [simp]: "eventually P top \ (\x. P x)" + unfolding top_filter_def + by (rule eventually_Abs_filter, rule is_filter.intro, auto) + +lemma eventually_bot [simp]: "eventually P bot" + unfolding bot_filter_def + by (subst eventually_Abs_filter, rule is_filter.intro, auto) + +lemma eventually_sup: + "eventually P (sup F F') \ eventually P F \ eventually P F'" + unfolding sup_filter_def + by (rule eventually_Abs_filter, rule is_filter.intro) + (auto elim!: eventually_rev_mp) + +lemma eventually_inf: + "eventually P (inf F F') \ + (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" + unfolding inf_filter_def + apply (rule eventually_Abs_filter, rule is_filter.intro) + apply (fast intro: eventually_True) + apply clarify + apply (intro exI conjI) + apply (erule (1) eventually_conj) + apply (erule (1) eventually_conj) + apply simp + apply auto + done + +lemma eventually_Sup: + "eventually P (Sup S) \ (\F\S. eventually P F)" + unfolding Sup_filter_def + apply (rule eventually_Abs_filter, rule is_filter.intro) + apply (auto intro: eventually_conj elim!: eventually_rev_mp) + done + +instance proof + fix F F' F'' :: "'a filter" and S :: "'a filter set" + { show "F < F' \ F \ F' \ \ F' \ F" + by (rule less_filter_def) } + { show "F \ F" + unfolding le_filter_def by simp } + { assume "F \ F'" and "F' \ F''" thus "F \ F''" + unfolding le_filter_def by simp } + { assume "F \ F'" and "F' \ F" thus "F = F'" + unfolding le_filter_def filter_eq_iff by fast } + { show "F \ top" + unfolding le_filter_def eventually_top by (simp add: always_eventually) } + { show "bot \ F" + unfolding le_filter_def by simp } + { show "F \ sup F F'" and "F' \ sup F F'" + unfolding le_filter_def eventually_sup by simp_all } + { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" + unfolding le_filter_def eventually_sup by simp } + { show "inf F F' \ F" and "inf F F' \ F'" + unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } + { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" + unfolding le_filter_def eventually_inf + by (auto elim!: eventually_mono intro: eventually_conj) } + { assume "F \ S" thus "F \ Sup S" + unfolding le_filter_def eventually_Sup by simp } + { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" + unfolding le_filter_def eventually_Sup by simp } + { assume "F'' \ S" thus "Inf S \ F''" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } + { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } +qed + +end + +lemma filter_leD: + "F \ F' \ eventually P F' \ eventually P F" + unfolding le_filter_def by simp + +lemma filter_leI: + "(\P. eventually P F' \ eventually P F) \ F \ F'" + unfolding le_filter_def by simp + +lemma eventually_False: + "eventually (\x. False) F \ F = bot" + unfolding filter_eq_iff by (auto elim: eventually_rev_mp) + +abbreviation (input) trivial_limit :: "'a filter \ bool" + where "trivial_limit F \ F = bot" + +lemma trivial_limit_def: "trivial_limit F \ eventually (\x. False) F" + by (rule eventually_False [symmetric]) + +lemma eventually_const: "\ trivial_limit net \ eventually (\x. P) net \ P" + by (cases P) (simp_all add: eventually_False) + + +subsubsection {* Map function for filters *} + +definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" + where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" + +lemma eventually_filtermap: + "eventually P (filtermap f F) = eventually (\x. P (f x)) F" + unfolding filtermap_def + apply (rule eventually_Abs_filter) + apply (rule is_filter.intro) + apply (auto elim!: eventually_rev_mp) + done + +lemma filtermap_ident: "filtermap (\x. x) F = F" + by (simp add: filter_eq_iff eventually_filtermap) + +lemma filtermap_filtermap: + "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" + by (simp add: filter_eq_iff eventually_filtermap) + +lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" + unfolding le_filter_def eventually_filtermap by simp + +lemma filtermap_bot [simp]: "filtermap f bot = bot" + by (simp add: filter_eq_iff eventually_filtermap) + +lemma filtermap_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" + by (auto simp: filter_eq_iff eventually_filtermap eventually_sup) + +subsubsection {* Order filters *} + +definition at_top :: "('a::order) filter" + where "at_top = Abs_filter (\P. \k. \n\k. P n)" + +lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" + unfolding at_top_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\max i j. P n \ Q n" by simp + then show "\k. \n\k. P n \ Q n" .. +qed auto + +lemma eventually_ge_at_top: + "eventually (\x. (c::_::linorder) \ x) at_top" + unfolding eventually_at_top_linorder by auto + +lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::dense_linorder. \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 + +lemma eventually_gt_at_top: + "eventually (\x. (c::_::dense_linorder) < x) at_top" + unfolding eventually_at_top_dense by auto + +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_le_at_bot: + "eventually (\x. x \ (c::_::linorder)) at_bot" + unfolding eventually_at_bot_linorder by 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 + +lemma eventually_gt_at_bot: + "eventually (\x. x < (c::_::dense_linorder)) at_bot" + unfolding eventually_at_bot_dense by auto + +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 + +lemmas trivial_limit_sequentially = sequentially_bot + +lemma eventually_False_sequentially [simp]: + "\ eventually (\n. False) sequentially" + by (simp add: eventually_False) + +lemma le_sequentially: + "F \ sequentially \ (\N. eventually (\n. N \ n) F)" + 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) + + +subsubsection {* Standard filters *} + +definition within :: "'a filter \ 'a set \ 'a filter" (infixr "within" 70) + where "F within S = Abs_filter (\P. eventually (\x. x \ S \ P x) F)" + +lemma eventually_within: + "eventually P (F within S) = eventually (\x. x \ S \ P x) F" + unfolding within_def + by (rule eventually_Abs_filter, rule is_filter.intro) + (auto elim!: eventually_rev_mp) + +lemma within_UNIV [simp]: "F within UNIV = F" + unfolding filter_eq_iff eventually_within by simp + +lemma within_empty [simp]: "F within {} = bot" + unfolding filter_eq_iff eventually_within by simp + +lemma within_within_eq: "(F within S) within T = F within (S \ T)" + by (auto simp: filter_eq_iff eventually_within elim: eventually_elim1) + +lemma within_le: "F within S \ F" + unfolding le_filter_def eventually_within by (auto elim: eventually_elim1) + +lemma le_withinI: "F \ F' \ eventually (\x. x \ S) F \ F \ F' within S" + unfolding le_filter_def eventually_within by (auto elim: eventually_elim2) + +lemma le_within_iff: "eventually (\x. x \ S) F \ F \ F' within S \ F \ F'" + by (blast intro: within_le le_withinI order_trans) + +subsubsection {* Topological filters *} + +definition (in topological_space) nhds :: "'a \ 'a filter" + where "nhds a = Abs_filter (\P. \S. open S \ a \ S \ (\x\S. P x))" + +definition (in topological_space) at :: "'a \ 'a filter" + where "at a = nhds a within - {a}" + +abbreviation at_right :: "'a\{topological_space, order} \ 'a filter" where + "at_right x \ at x within {x <..}" + +abbreviation at_left :: "'a\{topological_space, order} \ 'a filter" where + "at_left x \ at x within {..< x}" + +lemma eventually_nhds: + "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" + unfolding nhds_def +proof (rule eventually_Abs_filter, rule is_filter.intro) + have "open (UNIV :: 'a :: topological_space set) \ a \ UNIV \ (\x\UNIV. True)" by simp + thus "\S. open S \ a \ S \ (\x\S. True)" .. +next + fix P Q + assume "\S. open S \ a \ S \ (\x\S. P x)" + and "\T. open T \ a \ T \ (\x\T. Q x)" + then obtain S T where + "open S \ a \ S \ (\x\S. P x)" + "open T \ a \ T \ (\x\T. Q x)" by auto + hence "open (S \ T) \ a \ S \ T \ (\x\(S \ T). P x \ Q x)" + by (simp add: open_Int) + thus "\S. open S \ a \ S \ (\x\S. P x \ Q x)" .. +qed auto + +lemma nhds_neq_bot [simp]: "nhds a \ bot" + unfolding trivial_limit_def eventually_nhds by simp + +lemma eventually_at_topological: + "eventually P (at a) \ (\S. open S \ a \ S \ (\x\S. x \ a \ P x))" +unfolding at_def eventually_within eventually_nhds by simp + +lemma at_eq_bot_iff: "at a = bot \ open {a}" + unfolding trivial_limit_def eventually_at_topological + by (safe, case_tac "S = {a}", simp, fast, fast) + +lemma at_neq_bot [simp]: "at (a::'a::perfect_space) \ bot" + by (simp add: at_eq_bot_iff not_open_singleton) + +lemma eventually_at_right: + fixes x :: "'a :: {no_top, linorder_topology}" + shows "eventually P (at_right x) \ (\b. x < b \ (\z. x < z \ z < b \ P z))" + unfolding eventually_nhds eventually_within at_def +proof safe + fix S assume "open S" "x \ S" + note open_right[OF this] + moreover assume "\s\S. s \ - {x} \ s \ {x<..} \ P s" + ultimately show "\b>x. \z. x < z \ z < b \ P z" + by (auto simp: subset_eq Ball_def) +next + fix b assume "x < b" "\z. x < z \ z < b \ P z" + then show "\S. open S \ x \ S \ (\xa\S. xa \ - {x} \ xa \ {x<..} \ P xa)" + by (intro exI[of _ "{..< b}"]) auto +qed + +lemma eventually_at_left: + fixes x :: "'a :: {no_bot, linorder_topology}" + shows "eventually P (at_left x) \ (\b. x > b \ (\z. b < z \ z < x \ P z))" + unfolding eventually_nhds eventually_within at_def +proof safe + fix S assume "open S" "x \ S" + note open_left[OF this] + moreover assume "\s\S. s \ - {x} \ s \ {.. P s" + ultimately show "\bz. b < z \ z < x \ P z" + by (auto simp: subset_eq Ball_def) +next + fix b assume "b < x" "\z. b < z \ z < x \ P z" + then show "\S. open S \ x \ S \ (\xa\S. xa \ - {x} \ xa \ {.. P xa)" + by (intro exI[of _ "{b <..}"]) auto +qed + +lemma trivial_limit_at_left_real [simp]: + "\ trivial_limit (at_left (x::'a::{no_bot, dense_linorder, linorder_topology}))" + unfolding trivial_limit_def eventually_at_left by (auto dest: dense) + +lemma trivial_limit_at_right_real [simp]: + "\ trivial_limit (at_right (x::'a::{no_top, dense_linorder, linorder_topology}))" + unfolding trivial_limit_def eventually_at_right by (auto dest: dense) + +lemma at_within_eq: "at x within T = nhds x within (T - {x})" + unfolding at_def within_within_eq by (simp add: ac_simps Diff_eq) + +lemma at_eq_sup_left_right: "at (x::'a::linorder_topology) = sup (at_left x) (at_right x)" + by (auto simp: eventually_within at_def filter_eq_iff eventually_sup + elim: eventually_elim2 eventually_elim1) + +lemma eventually_at_split: + "eventually P (at (x::'a::linorder_topology)) \ eventually P (at_left x) \ eventually P (at_right x)" + by (subst at_eq_sup_left_right) (simp add: eventually_sup) + +subsection {* Limits *} + +definition filterlim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where + "filterlim 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 filterlim (%x. f) F2 F1" + +lemma filterlim_iff: + "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" + unfolding filterlim_def le_filter_def eventually_filtermap .. + +lemma filterlim_compose: + "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" + unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) + +lemma filterlim_mono: + "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" + unfolding filterlim_def by (metis filtermap_mono order_trans) + +lemma filterlim_ident: "LIM x F. x :> F" + by (simp add: filterlim_def filtermap_ident) + +lemma filterlim_cong: + "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" + by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) + +lemma filterlim_within: + "(LIM x F1. f x :> F2 within S) \ (eventually (\x. f x \ S) F1 \ (LIM x F1. f x :> F2))" + unfolding filterlim_def +proof safe + assume "filtermap f F1 \ F2 within S" then show "eventually (\x. f x \ S) F1" + by (auto simp: le_filter_def eventually_filtermap eventually_within elim!: allE[of _ "\x. x \ S"]) +qed (auto intro: within_le order_trans simp: le_within_iff eventually_filtermap) + +lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\x. f (g x)) F1 F2" + unfolding filterlim_def filtermap_filtermap .. + +lemma filterlim_sup: + "filterlim f F F1 \ filterlim f F F2 \ filterlim f F (sup F1 F2)" + unfolding filterlim_def filtermap_sup by auto + +lemma filterlim_Suc: "filterlim Suc sequentially sequentially" + by (simp add: filterlim_iff eventually_sequentially) (metis le_Suc_eq) + +subsubsection {* Tendsto *} + +abbreviation (in topological_space) + tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "--->" 55) where + "(f ---> l) F \ filterlim f (nhds l) F" + +lemma tendsto_eq_rhs: "(f ---> x) F \ x = y \ (f ---> y) F" + by simp + +ML {* + +structure Tendsto_Intros = Named_Thms +( + val name = @{binding tendsto_intros} + val description = "introduction rules for tendsto" +) + +*} + +setup {* + Tendsto_Intros.setup #> + Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros}, + map (fn thm => @{thm tendsto_eq_rhs} OF [thm]) o Tendsto_Intros.get o Context.proof_of); +*} + +lemma tendsto_def: "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" + unfolding filterlim_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 filterlim_at: + "(LIM x F. f x :> at b) \ (eventually (\x. f x \ b) F \ (f ---> b) F)" + by (simp add: at_def filterlim_within) + +lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" + unfolding tendsto_def le_filter_def by fast + +lemma topological_tendstoI: + "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) + \ (f ---> l) F" + unfolding tendsto_def by auto + +lemma topological_tendstoD: + "(f ---> l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" + unfolding tendsto_def by auto + +lemma order_tendstoI: + fixes y :: "_ :: order_topology" + assumes "\a. a < y \ eventually (\x. a < f x) F" + assumes "\a. y < a \ eventually (\x. f x < a) F" + shows "(f ---> y) F" +proof (rule topological_tendstoI) + fix S assume "open S" "y \ S" + then show "eventually (\x. f x \ S) F" + unfolding open_generated_order + proof induct + case (UN K) + then obtain k where "y \ k" "k \ K" by auto + with UN(2)[of k] show ?case + by (auto elim: eventually_elim1) + qed (insert assms, auto elim: eventually_elim2) +qed + +lemma order_tendstoD: + fixes y :: "_ :: order_topology" + assumes y: "(f ---> y) F" + shows "a < y \ eventually (\x. a < f x) F" + and "y < a \ eventually (\x. f x < a) F" + using topological_tendstoD[OF y, of "{..< a}"] topological_tendstoD[OF y, of "{a <..}"] by auto + +lemma order_tendsto_iff: + fixes f :: "_ \ 'a :: order_topology" + shows "(f ---> x) F \(\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" + by (metis order_tendstoI order_tendstoD) + +lemma tendsto_bot [simp]: "(f ---> a) bot" + unfolding tendsto_def by simp + +lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" + unfolding tendsto_def eventually_at_topological by auto + +lemma tendsto_ident_at_within [tendsto_intros]: + "((\x. x) ---> a) (at a within S)" + unfolding tendsto_def eventually_within eventually_at_topological by auto + +lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) F" + by (simp add: tendsto_def) + +lemma tendsto_unique: + fixes f :: "'a \ 'b::t2_space" + assumes "\ trivial_limit F" and "(f ---> a) F" and "(f ---> b) F" + shows "a = b" +proof (rule ccontr) + assume "a \ b" + obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" + using hausdorff [OF `a \ b`] by fast + have "eventually (\x. f x \ U) F" + using `(f ---> a) F` `open U` `a \ U` by (rule topological_tendstoD) + moreover + have "eventually (\x. f x \ V) F" + using `(f ---> b) F` `open V` `b \ V` by (rule topological_tendstoD) + ultimately + have "eventually (\x. False) F" + proof eventually_elim + case (elim x) + hence "f x \ U \ V" by simp + with `U \ V = {}` show ?case by simp + qed + with `\ trivial_limit F` show "False" + by (simp add: trivial_limit_def) +qed + +lemma tendsto_const_iff: + fixes a b :: "'a::t2_space" + assumes "\ trivial_limit F" shows "((\x. a) ---> b) F \ a = b" + by (safe intro!: tendsto_const tendsto_unique [OF assms tendsto_const]) + +lemma increasing_tendsto: + fixes f :: "_ \ 'a::order_topology" + assumes bdd: "eventually (\n. f n \ l) F" + and en: "\x. x < l \ eventually (\n. x < f n) F" + shows "(f ---> l) F" + using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) + +lemma decreasing_tendsto: + fixes f :: "_ \ 'a::order_topology" + assumes bdd: "eventually (\n. l \ f n) F" + and en: "\x. l < x \ eventually (\n. f n < x) F" + shows "(f ---> l) F" + using assms by (intro order_tendstoI) (auto elim!: eventually_elim1) + +lemma tendsto_sandwich: + fixes f g h :: "'a \ 'b::order_topology" + 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 (rule order_tendstoI) + fix a show "a < c \ eventually (\x. a < g x) net" + using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) +next + fix a show "c < a \ eventually (\x. g x < a) net" + using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) +qed + +lemma tendsto_le: + fixes f g :: "'a \ 'b::linorder_topology" + assumes F: "\ trivial_limit F" + assumes x: "(f ---> x) F" and y: "(g ---> y) F" + assumes ev: "eventually (\x. g x \ f x) F" + shows "y \ x" +proof (rule ccontr) + assume "\ y \ x" + with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" + by (auto simp: not_le) + then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" + using x y by (auto intro: order_tendstoD) + with ev have "eventually (\x. False) F" + by eventually_elim (insert xy, fastforce) + with F show False + by (simp add: eventually_False) +qed + +lemma tendsto_le_const: + fixes f :: "'a \ 'b::linorder_topology" + assumes F: "\ trivial_limit F" + assumes x: "(f ---> x) F" and a: "eventually (\x. a \ f x) F" + shows "a \ x" + using F x tendsto_const a by (rule tendsto_le) + +subsection {* Limits to @{const at_top} and @{const at_bot} *} + +lemma filterlim_at_top: + fixes f :: "'a \ ('b::linorder)" + shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z \ f x) F)" + by (auto simp: filterlim_iff eventually_at_top_linorder elim!: eventually_elim1) + +lemma filterlim_at_top_dense: + fixes f :: "'a \ ('b::dense_linorder)" + shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" + by (metis eventually_elim1[of _ F] eventually_gt_at_top order_less_imp_le + filterlim_at_top[of f F] filterlim_iff[of f at_top F]) + +lemma filterlim_at_top_ge: + fixes f :: "'a \ ('b::linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_top) \ (\Z\c. eventually (\x. Z \ f x) F)" + unfolding filterlim_at_top +proof safe + fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" + with *[THEN spec, of "max Z c"] show "eventually (\x. Z \ f x) F" + by (auto elim!: eventually_elim1) +qed simp + +lemma filterlim_at_top_at_top: + fixes f :: "'a::linorder \ 'b::linorder" + assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" + assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" + assumes Q: "eventually Q at_top" + assumes P: "eventually P at_top" + shows "filterlim f at_top at_top" +proof - + from P obtain x where x: "\y. x \ y \ P y" + unfolding eventually_at_top_linorder by auto + show ?thesis + proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) + fix z assume "x \ z" + with x have "P z" by auto + have "eventually (\x. g z \ x) at_top" + by (rule eventually_ge_at_top) + with Q show "eventually (\x. z \ f x) at_top" + by eventually_elim (metis mono bij `P z`) + qed +qed + +lemma filterlim_at_top_gt: + fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_top) \ (\Z>c. eventually (\x. Z \ f x) F)" + by (metis filterlim_at_top order_less_le_trans gt_ex filterlim_at_top_ge) + +lemma filterlim_at_bot: + fixes f :: "'a \ ('b::linorder)" + shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x \ Z) F)" + by (auto simp: filterlim_iff eventually_at_bot_linorder elim!: eventually_elim1) + +lemma filterlim_at_bot_le: + fixes f :: "'a \ ('b::linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_bot) \ (\Z\c. eventually (\x. Z \ f x) F)" + unfolding filterlim_at_bot +proof safe + fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" + with *[THEN spec, of "min Z c"] show "eventually (\x. Z \ f x) F" + by (auto elim!: eventually_elim1) +qed simp + +lemma filterlim_at_bot_lt: + fixes f :: "'a \ ('b::dense_linorder)" and c :: "'b" + shows "(LIM x F. f x :> at_bot) \ (\Zx. Z \ f x) F)" + by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans) + +lemma filterlim_at_bot_at_right: + fixes f :: "'a::{no_top, linorder_topology} \ 'b::linorder" + assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" + assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" + assumes Q: "eventually Q (at_right a)" and bound: "\b. Q b \ a < b" + assumes P: "eventually P at_bot" + shows "filterlim f at_bot (at_right a)" +proof - + from P obtain x where x: "\y. y \ x \ P y" + unfolding eventually_at_bot_linorder by auto + show ?thesis + proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) + fix z assume "z \ x" + with x have "P z" by auto + have "eventually (\x. x \ g z) (at_right a)" + using bound[OF bij(2)[OF `P z`]] + unfolding eventually_at_right by (auto intro!: exI[of _ "g z"]) + with Q show "eventually (\x. f x \ z) (at_right a)" + by eventually_elim (metis bij `P z` mono) + qed +qed + +lemma filterlim_at_top_at_left: + fixes f :: "'a::{no_bot, linorder_topology} \ 'b::linorder" + assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" + assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" + assumes Q: "eventually Q (at_left a)" and bound: "\b. Q b \ b < a" + assumes P: "eventually P at_top" + shows "filterlim f at_top (at_left a)" +proof - + from P obtain x where x: "\y. x \ y \ P y" + unfolding eventually_at_top_linorder by auto + show ?thesis + proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) + fix z assume "x \ z" + with x have "P z" by auto + have "eventually (\x. g z \ x) (at_left a)" + using bound[OF bij(2)[OF `P z`]] + unfolding eventually_at_left by (auto intro!: exI[of _ "g z"]) + with Q show "eventually (\x. z \ f x) (at_left a)" + by eventually_elim (metis bij `P z` mono) + qed +qed + +lemma filterlim_split_at: + "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at (x::'a::linorder_topology))" + by (subst at_eq_sup_left_right) (rule filterlim_sup) + +lemma filterlim_at_split: + "filterlim f F (at (x::'a::linorder_topology)) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" + by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) + + +subsection {* Limits on sequences *} + +abbreviation (in topological_space) + LIMSEQ :: "[nat \ 'a, 'a] \ bool" + ("((_)/ ----> (_))" [60, 60] 60) where + "X ----> L \ (X ---> L) sequentially" + +definition + lim :: "(nat \ 'a::t2_space) \ 'a" where + --{*Standard definition of limit using choice operator*} + "lim X = (THE L. X ----> L)" + +definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where + "convergent X = (\L. X ----> L)" + +subsubsection {* Monotone sequences and subsequences *} + +definition + monoseq :: "(nat \ 'a::order) \ bool" where + --{*Definition of monotonicity. + The use of disjunction here complicates proofs considerably. + One alternative is to add a Boolean argument to indicate the direction. + Another is to develop the notions of increasing and decreasing first.*} + "monoseq X = ((\m. \n\m. X m \ X n) | (\m. \n\m. X n \ X m))" + +definition + incseq :: "(nat \ 'a::order) \ bool" where + --{*Increasing sequence*} + "incseq X \ (\m. \n\m. X m \ X n)" + +definition + decseq :: "(nat \ 'a::order) \ bool" where + --{*Decreasing sequence*} + "decseq X \ (\m. \n\m. X n \ X m)" + +definition + subseq :: "(nat \ nat) \ bool" where + --{*Definition of subsequence*} + "subseq f \ (\m. \n>m. f m < f n)" + +lemma incseq_mono: "mono f \ incseq f" + unfolding mono_def incseq_def by auto + +lemma incseq_SucI: + "(\n. X n \ X (Suc n)) \ incseq X" + using lift_Suc_mono_le[of X] + by (auto simp: incseq_def) + +lemma incseqD: "\i j. incseq f \ i \ j \ f i \ f j" + by (auto simp: incseq_def) + +lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" + using incseqD[of A i "Suc i"] by auto + +lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" + by (auto intro: incseq_SucI dest: incseq_SucD) + +lemma incseq_const[simp, intro]: "incseq (\x. k)" + unfolding incseq_def by auto + +lemma decseq_SucI: + "(\n. X (Suc n) \ X n) \ decseq X" + using order.lift_Suc_mono_le[OF dual_order, of X] + by (auto simp: decseq_def) + +lemma decseqD: "\i j. decseq f \ i \ j \ f j \ f i" + by (auto simp: decseq_def) + +lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" + using decseqD[of A i "Suc i"] by auto + +lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" + by (auto intro: decseq_SucI dest: decseq_SucD) + +lemma decseq_const[simp, intro]: "decseq (\x. k)" + unfolding decseq_def by auto + +lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" + unfolding monoseq_def incseq_def decseq_def .. + +lemma monoseq_Suc: + "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" + unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. + +lemma monoI1: "\m. \ n \ m. X m \ X n ==> monoseq X" +by (simp add: monoseq_def) + +lemma monoI2: "\m. \ n \ m. X n \ X m ==> monoseq X" +by (simp add: monoseq_def) + +lemma mono_SucI1: "\n. X n \ X (Suc n) ==> monoseq X" +by (simp add: monoseq_Suc) + +lemma mono_SucI2: "\n. X (Suc n) \ X n ==> monoseq X" +by (simp add: monoseq_Suc) + +lemma monoseq_minus: + fixes a :: "nat \ 'a::ordered_ab_group_add" + assumes "monoseq a" + shows "monoseq (\ n. - a n)" +proof (cases "\ m. \ n \ m. a m \ a n") + case True + hence "\ m. \ n \ m. - a n \ - a m" by auto + thus ?thesis by (rule monoI2) +next + case False + hence "\ m. \ n \ m. - a m \ - a n" using `monoseq a`[unfolded monoseq_def] by auto + thus ?thesis by (rule monoI1) +qed + +text{*Subsequence (alternative definition, (e.g. Hoskins)*} + +lemma subseq_Suc_iff: "subseq f = (\n. (f n) < (f (Suc n)))" +apply (simp add: subseq_def) +apply (auto dest!: less_imp_Suc_add) +apply (induct_tac k) +apply (auto intro: less_trans) +done + +text{* for any sequence, there is a monotonic subsequence *} +lemma seq_monosub: + fixes s :: "nat => 'a::linorder" + shows "\f. subseq f \ monoseq (\ n. (s (f n)))" +proof cases + let "?P p n" = "p > n \ (\m\p. s m \ s p)" + assume *: "\n. \p. ?P p n" + def f \ "nat_rec (SOME p. ?P p 0) (\_ n. SOME p. ?P p n)" + have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp + have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. + have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto + have P_Suc: "\i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto + then have "subseq f" unfolding subseq_Suc_iff by auto + moreover have "monoseq (\n. s (f n))" unfolding monoseq_Suc + proof (intro disjI2 allI) + fix n show "s (f (Suc n)) \ s (f n)" + proof (cases n) + case 0 with P_Suc[of 0] P_0 show ?thesis by auto + next + case (Suc m) + from P_Suc[of n] Suc have "f (Suc m) \ f (Suc (Suc m))" by simp + with P_Suc Suc show ?thesis by simp + qed + qed + ultimately show ?thesis by auto +next + let "?P p m" = "m < p \ s m < s p" + assume "\ (\n. \p>n. (\m\p. s m \ s p))" + then obtain N where N: "\p. p > N \ \m>p. s p < s m" by (force simp: not_le le_less) + def f \ "nat_rec (SOME p. ?P p (Suc N)) (\_ n. SOME p. ?P p n)" + have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp + have f_Suc: "\i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc .. + have P_0: "?P (f 0) (Suc N)" + unfolding f_0 some_eq_ex[of "\p. ?P p (Suc N)"] using N[of "Suc N"] by auto + { fix i have "N < f i \ ?P (f (Suc i)) (f i)" + unfolding f_Suc some_eq_ex[of "\p. ?P p (f i)"] using N[of "f i"] . } + note P' = this + { fix i have "N < f i \ ?P (f (Suc i)) (f i)" + by (induct i) (insert P_0 P', auto) } + then have "subseq f" "monoseq (\x. s (f x))" + unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le) + then show ?thesis by auto +qed + +lemma seq_suble: assumes sf: "subseq f" shows "n \ f n" +proof(induct n) + case 0 thus ?case by simp +next + case (Suc n) + from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps + have "n < f (Suc n)" by arith + thus ?case by arith +qed + +lemma eventually_subseq: + "subseq r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" + unfolding eventually_sequentially by (metis seq_suble le_trans) + +lemma filterlim_subseq: "subseq f \ filterlim f sequentially sequentially" + unfolding filterlim_iff by (metis eventually_subseq) + +lemma subseq_o: "subseq r \ subseq s \ subseq (r \ s)" + unfolding subseq_def by simp + +lemma subseq_mono: assumes "subseq r" "m < n" shows "r m < r n" + using assms by (auto simp: subseq_def) + +lemma incseq_imp_monoseq: "incseq X \ monoseq X" + by (simp add: incseq_def monoseq_def) + +lemma decseq_imp_monoseq: "decseq X \ monoseq X" + by (simp add: decseq_def monoseq_def) + +lemma decseq_eq_incseq: + fixes X :: "nat \ 'a::ordered_ab_group_add" shows "decseq X = incseq (\n. - X n)" + by (simp add: decseq_def incseq_def) + +lemma INT_decseq_offset: + assumes "decseq F" + shows "(\i. F i) = (\i\{n..}. F i)" +proof safe + fix x i assume x: "x \ (\i\{n..}. F i)" + show "x \ F i" + proof cases + from x have "x \ F n" by auto + also assume "i \ n" with `decseq F` have "F n \ F i" + unfolding decseq_def by simp + finally show ?thesis . + qed (insert x, simp) +qed auto + +lemma LIMSEQ_const_iff: + fixes k l :: "'a::t2_space" + shows "(\n. k) ----> l \ k = l" + using trivial_limit_sequentially by (rule tendsto_const_iff) + +lemma LIMSEQ_SUP: + "incseq X \ X ----> (SUP i. X i :: 'a :: {complete_linorder, linorder_topology})" + by (intro increasing_tendsto) + (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) + +lemma LIMSEQ_INF: + "decseq X \ X ----> (INF i. X i :: 'a :: {complete_linorder, linorder_topology})" + by (intro decreasing_tendsto) + (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) + +lemma LIMSEQ_ignore_initial_segment: + "f ----> a \ (\n. f (n + k)) ----> a" +apply (rule topological_tendstoI) +apply (drule (2) topological_tendstoD) +apply (simp only: eventually_sequentially) +apply (erule exE, rename_tac N) +apply (rule_tac x=N in exI) +apply simp +done + +lemma LIMSEQ_offset: + "(\n. f (n + k)) ----> a \ f ----> a" +apply (rule topological_tendstoI) +apply (drule (2) topological_tendstoD) +apply (simp only: eventually_sequentially) +apply (erule exE, rename_tac N) +apply (rule_tac x="N + k" in exI) +apply clarify +apply (drule_tac x="n - k" in spec) +apply (simp add: le_diff_conv2) +done + +lemma LIMSEQ_Suc: "f ----> l \ (\n. f (Suc n)) ----> l" +by (drule_tac k="Suc 0" in LIMSEQ_ignore_initial_segment, simp) + +lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) ----> l \ f ----> l" +by (rule_tac k="Suc 0" in LIMSEQ_offset, simp) + +lemma LIMSEQ_Suc_iff: "(\n. f (Suc n)) ----> l = f ----> l" +by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) + +lemma LIMSEQ_unique: + fixes a b :: "'a::t2_space" + shows "\X ----> a; X ----> b\ \ a = b" + using trivial_limit_sequentially by (rule tendsto_unique) + +lemma LIMSEQ_le_const: + "\X ----> (x::'a::linorder_topology); \N. \n\N. a \ X n\ \ a \ x" + using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially) + +lemma LIMSEQ_le: + "\X ----> x; Y ----> y; \N. \n\N. X n \ Y n\ \ x \ (y::'a::linorder_topology)" + using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) + +lemma LIMSEQ_le_const2: + "\X ----> (x::'a::linorder_topology); \N. \n\N. X n \ a\ \ x \ a" + by (rule LIMSEQ_le[of X x "\n. a"]) (auto simp: tendsto_const) + +lemma convergentD: "convergent X ==> \L. (X ----> L)" +by (simp add: convergent_def) + +lemma convergentI: "(X ----> L) ==> convergent X" +by (auto simp add: convergent_def) + +lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)" +by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) + +lemma convergent_const: "convergent (\n. c)" + by (rule convergentI, rule tendsto_const) + +lemma monoseq_le: + "monoseq a \ a ----> (x::'a::linorder_topology) \ + ((\ n. a n \ x) \ (\m. \n\m. a m \ a n)) \ ((\ n. x \ a n) \ (\m. \n\m. a n \ a m))" + by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) + +lemma LIMSEQ_subseq_LIMSEQ: + "\ X ----> L; subseq f \ \ (X o f) ----> L" + unfolding comp_def by (rule filterlim_compose[of X, OF _ filterlim_subseq]) + +lemma convergent_subseq_convergent: + "\convergent X; subseq f\ \ convergent (X o f)" + unfolding convergent_def by (auto intro: LIMSEQ_subseq_LIMSEQ) + +lemma limI: "X ----> L ==> lim X = L" +apply (simp add: lim_def) +apply (blast intro: LIMSEQ_unique) +done + +lemma lim_le: "convergent f \ (\n. f n \ (x::'a::linorder_topology)) \ lim f \ x" + using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) + +subsubsection{*Increasing and Decreasing Series*} + +lemma incseq_le: "incseq X \ X ----> L \ X n \ (L::'a::linorder_topology)" + by (metis incseq_def LIMSEQ_le_const) + +lemma decseq_le: "decseq X \ X ----> L \ (L::'a::linorder_topology) \ X n" + by (metis decseq_def LIMSEQ_le_const2) + +subsection {* Function limit at a point *} + +abbreviation + LIM :: "('a::topological_space \ 'b::topological_space) \ 'a \ 'b \ bool" + ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where + "f -- a --> L \ (f ---> L) (at a)" + +lemma LIM_const_not_eq[tendsto_intros]: + fixes a :: "'a::perfect_space" + fixes k L :: "'b::t2_space" + shows "k \ L \ \ (\x. k) -- a --> L" + by (simp add: tendsto_const_iff) + +lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] + +lemma LIM_const_eq: + fixes a :: "'a::perfect_space" + fixes k L :: "'b::t2_space" + shows "(\x. k) -- a --> L \ k = L" + by (simp add: tendsto_const_iff) + +lemma LIM_unique: + fixes a :: "'a::perfect_space" and L M :: "'b::t2_space" + shows "f -- a --> L \ f -- a --> M \ L = M" + using at_neq_bot by (rule tendsto_unique) + +text {* Limits are equal for functions equal except at limit point *} + +lemma LIM_equal: "\x. x \ a --> (f x = g x) \ (f -- a --> l) \ (g -- a --> l)" + unfolding tendsto_def eventually_at_topological by simp + +lemma LIM_cong: "a = b \ (\x. x \ b \ f x = g x) \ l = m \ (f -- a --> l) \ (g -- b --> m)" + by (simp add: LIM_equal) + +lemma LIM_cong_limit: "f -- x --> L \ K = L \ f -- x --> K" + by simp + +lemma tendsto_at_iff_tendsto_nhds: + "g -- l --> g l \ (g ---> g l) (nhds l)" + unfolding tendsto_def at_def eventually_within + by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) + +lemma tendsto_compose: + "g -- l --> g l \ (f ---> l) F \ ((\x. g (f x)) ---> g l) F" + unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) + +lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" + unfolding o_def by (rule tendsto_compose) + +lemma tendsto_compose_eventually: + "g -- l --> m \ (f ---> l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) ---> m) F" + by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) + +lemma LIM_compose_eventually: + assumes f: "f -- a --> b" + assumes g: "g -- b --> c" + assumes inj: "eventually (\x. f x \ b) (at a)" + shows "(\x. g (f x)) -- a --> c" + using g f inj by (rule tendsto_compose_eventually) + +subsection {* Continuity *} + +definition isCont :: "('a::topological_space \ 'b::topological_space) \ 'a \ bool" where + "isCont f a \ f -- a --> f a" + +lemma isCont_ident [simp]: "isCont (\x. x) a" + unfolding isCont_def by (rule tendsto_ident_at) + +lemma isCont_const [simp]: "isCont (\x. k) a" + unfolding isCont_def by (rule tendsto_const) + +lemma isCont_tendsto_compose: "isCont g l \ (f ---> l) F \ ((\x. g (f x)) ---> g l) F" + unfolding isCont_def by (rule tendsto_compose) + +lemma isCont_o2: "isCont f a \ isCont g (f a) \ isCont (\x. g (f x)) a" + unfolding isCont_def by (rule tendsto_compose) + +lemma isCont_o: "isCont f a \ isCont g (f a) \ isCont (g o f) a" + unfolding o_def by (rule isCont_o2) + +end +