diff --git a/src/HOL/Analysis/Abstract_Metric_Spaces.thy b/src/HOL/Analysis/Abstract_Metric_Spaces.thy --- a/src/HOL/Analysis/Abstract_Metric_Spaces.thy +++ b/src/HOL/Analysis/Abstract_Metric_Spaces.thy @@ -1,3348 +1,4565 @@ section \Abstract Metric Spaces\ theory Abstract_Metric_Spaces imports Elementary_Metric_Spaces Abstract_Limits Abstract_Topological_Spaces begin (*Avoid a clash with the existing metric_space locale (from the type class)*) locale Metric_space = fixes M :: "'a set" and d :: "'a \ 'a \ real" assumes nonneg [simp]: "\x y. 0 \ d x y" assumes commute: "\x y. d x y = d y x" assumes zero [simp]: "\x y. \x \ M; y \ M\ \ d x y = 0 \ x=y" assumes triangle: "\x y z. \x \ M; y \ M; z \ M\ \ d x z \ d x y + d y z" text \Link with the type class version\ interpretation Met_TC: Metric_space UNIV dist by (simp add: dist_commute dist_triangle Metric_space.intro) context Metric_space begin lemma subspace: "M' \ M \ Metric_space M' d" by (simp add: commute in_mono Metric_space.intro triangle) lemma abs_mdist [simp] : "\d x y\ = d x y" by simp lemma mdist_pos_less: "\x \ y; x \ M; y \ M\ \ 0 < d x y" by (metis less_eq_real_def nonneg zero) lemma mdist_zero [simp]: "x \ M \ d x x = 0" by simp lemma mdist_pos_eq [simp]: "\x \ M; y \ M\ \ 0 < d x y \ x \ y" using mdist_pos_less zero by fastforce lemma triangle': "\x \ M; y \ M; z \ M\ \ d x z \ d x y + d z y" by (simp add: commute triangle) lemma triangle'': "\x \ M; y \ M; z \ M\ \ d x z \ d y x + d y z" by (simp add: commute triangle) lemma mdist_reverse_triangle: "\x \ M; y \ M; z \ M\ \ \d x y - d y z\ \ d x z" by (smt (verit) commute triangle) text\ Open and closed balls \ definition mball where "mball x r \ {y. x \ M \ y \ M \ d x y < r}" definition mcball where "mcball x r \ {y. x \ M \ y \ M \ d x y \ r}" lemma in_mball [simp]: "y \ mball x r \ x \ M \ y \ M \ d x y < r" by (simp add: local.Metric_space_axioms Metric_space.mball_def) lemma centre_in_mball_iff [iff]: "x \ mball x r \ x \ M \ 0 < r" using in_mball mdist_zero by force lemma mball_subset_mspace: "mball x r \ M" by auto lemma mball_eq_empty: "mball x r = {} \ (x \ M) \ r \ 0" by (smt (verit, best) Collect_empty_eq centre_in_mball_iff mball_def nonneg) lemma mball_subset: "\d x y + a \ b; y \ M\ \ mball x a \ mball y b" by (smt (verit) commute in_mball subsetI triangle) lemma disjoint_mball: "r + r' \ d x x' \ disjnt (mball x r) (mball x' r')" by (smt (verit) commute disjnt_iff in_mball triangle) lemma mball_subset_concentric: "r \ s \ mball x r \ mball x s" by auto lemma in_mcball [simp]: "y \ mcball x r \ x \ M \ y \ M \ d x y \ r" by (simp add: local.Metric_space_axioms Metric_space.mcball_def) lemma centre_in_mcball_iff [iff]: "x \ mcball x r \ x \ M \ 0 \ r" using mdist_zero by force lemma mcball_eq_empty: "mcball x r = {} \ (x \ M) \ r < 0" by (smt (verit, best) Collect_empty_eq centre_in_mcball_iff empty_iff mcball_def nonneg) lemma mcball_subset_mspace: "mcball x r \ M" by auto lemma mball_subset_mcball: "mball x r \ mcball x r" by auto lemma mcball_subset: "\d x y + a \ b; y \ M\ \ mcball x a \ mcball y b" by (smt (verit) in_mcball mdist_reverse_triangle subsetI) lemma mcball_subset_concentric: "r \ s \ mcball x r \ mcball x s" by force lemma mcball_subset_mball: "\d x y + a < b; y \ M\ \ mcball x a \ mball y b" by (smt (verit) commute in_mball in_mcball subsetI triangle) lemma mcball_subset_mball_concentric: "a < b \ mcball x a \ mball x b" by force end subsection \Metric topology \ context Metric_space begin definition mopen where "mopen U \ U \ M \ (\x. x \ U \ (\r>0. mball x r \ U))" definition mtopology :: "'a topology" where "mtopology \ topology mopen" lemma is_topology_metric_topology [iff]: "istopology mopen" proof - have "\S T. \mopen S; mopen T\ \ mopen (S \ T)" by (smt (verit, del_insts) Int_iff in_mball mopen_def subset_eq) moreover have "\\. (\K\\. mopen K) \ mopen (\\)" using mopen_def by fastforce ultimately show ?thesis by (simp add: istopology_def) qed lemma openin_mtopology: "openin mtopology U \ U \ M \ (\x. x \ U \ (\r>0. mball x r \ U))" by (simp add: mopen_def mtopology_def) lemma topspace_mtopology [simp]: "topspace mtopology = M" by (meson order.refl mball_subset_mspace openin_mtopology openin_subset openin_topspace subset_antisym zero_less_one) lemma subtopology_mspace [simp]: "subtopology mtopology M = mtopology" by (metis subtopology_topspace topspace_mtopology) lemma open_in_mspace [iff]: "openin mtopology M" by (metis openin_topspace topspace_mtopology) lemma closedin_mspace [iff]: "closedin mtopology M" by (metis closedin_topspace topspace_mtopology) lemma openin_mball [iff]: "openin mtopology (mball x r)" proof - have "\y. \x \ M; d x y < r\ \ \s>0. mball y s \ mball x r" by (metis add_diff_cancel_left' add_diff_eq commute less_add_same_cancel1 mball_subset order_refl) then show ?thesis by (auto simp: openin_mtopology) qed lemma mcball_eq_cball [simp]: "Met_TC.mcball = cball" by force lemma mball_eq_ball [simp]: "Met_TC.mball = ball" by force lemma mopen_eq_open [simp]: "Met_TC.mopen = open" by (force simp: open_contains_ball Met_TC.mopen_def) lemma limitin_iff_tendsto [iff]: "limitin Met_TC.mtopology \ x F = tendsto \ x F" by (simp add: Met_TC.mtopology_def) lemma mtopology_is_euclideanreal [simp]: "Met_TC.mtopology = euclideanreal" by (simp add: Met_TC.mtopology_def) (* lemma metric_injective_image: "\f m s. f ` s \ M \ (\x y. x \ s \ y \ s \ f x = f y \ x = y) \ (mspace(metric(s,\(x,y). d (f x) (f y))) = s) \ (d(metric(s,\(x,y). d (f x) (f y))) = \(x,y). d (f x) (f y))" oops REWRITE_TAC[\; FORALL_IN_IMAGE; INJECTIVE_ON_ALT] THEN REPEAT GEN_TAC THEN STRIP_TAC THEN REWRITE_TAC[mspace; d; GSYM PAIR_EQ] THEN REWRITE_TAC[GSYM(CONJUNCT2 metric_tybij); is_metric_space] THEN REWRITE_TAC[GSYM mspace; GSYM d] THEN ASM_SIMP_TAC[MDIST_POS_LE; MDIST_TRIANGLE; MDIST_0] THEN ASM_MESON_TAC[MDIST_SYM]);; *) lemma mtopology_base: "mtopology = topology(arbitrary union_of (\U. \x \ M. \r>0. U = mball x r))" proof - have "\S. \x r. x \ M \ 0 < r \ S = mball x r \ openin mtopology S" using openin_mball by blast moreover have "\U x. \openin mtopology U; x \ U\ \ \B. (\x r. x \ M \ 0 < r \ B = mball x r) \ x \ B \ B \ U" by (metis centre_in_mball_iff in_mono openin_mtopology) ultimately show ?thesis by (smt (verit) topology_base_unique) qed lemma closedin_metric: "closedin mtopology C \ C \ M \ (\x. x \ M - C \ (\r>0. disjnt C (mball x r)))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding closedin_def openin_mtopology by (metis Diff_disjoint disjnt_def disjnt_subset2 topspace_mtopology) show "?rhs \ ?lhs" unfolding closedin_def openin_mtopology disjnt_def by (metis Diff_subset Diff_triv Int_Diff Int_commute inf.absorb_iff2 mball_subset_mspace topspace_mtopology) qed lemma closedin_mcball [iff]: "closedin mtopology (mcball x r)" proof - have "\ra>0. disjnt (mcball x r) (mball y ra)" if "x \ M" for y by (metis disjnt_empty1 gt_ex mcball_eq_empty that) moreover have "disjnt (mcball x r) (mball y (d x y - r))" if "y \ M" "d x y > r" for y using that disjnt_iff in_mball in_mcball mdist_reverse_triangle by force ultimately show ?thesis using closedin_metric mcball_subset_mspace by fastforce qed lemma mball_iff_mcball: "(\r>0. mball x r \ U) = (\r>0. mcball x r \ U)" by (meson dense mball_subset_mcball mcball_subset_mball_concentric order_trans) lemma openin_mtopology_mcball: "openin mtopology U \ U \ M \ (\x. x \ U \ (\r. 0 < r \ mcball x r \ U))" using mball_iff_mcball openin_mtopology by presburger lemma metric_derived_set_of: "mtopology derived_set_of S = {x \ M. \r>0. \y\S. y\x \ y \ mball x r}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" unfolding openin_mtopology derived_set_of_def by clarsimp (metis in_mball openin_mball openin_mtopology zero) show "?rhs \ ?lhs" unfolding openin_mtopology derived_set_of_def by clarify (metis subsetD topspace_mtopology) qed lemma metric_closure_of: "mtopology closure_of S = {x \ M. \r>0. \y \ S. y \ mball x r}" proof - have "\x r. \0 < r; x \ mtopology closure_of S\ \ \y\S. y \ mball x r" by (metis centre_in_mball_iff in_closure_of openin_mball topspace_mtopology) moreover have "\x T. \x \ M; \r>0. \y\S. y \ mball x r\ \ x \ mtopology closure_of S" by (smt (verit) in_closure_of in_mball openin_mtopology subsetD topspace_mtopology) ultimately show ?thesis by (auto simp: in_closure_of) qed lemma metric_closure_of_alt: "mtopology closure_of S = {x \ M. \r>0. \y \ S. y \ mcball x r}" proof - have "\x r. \\r>0. x \ M \ (\y\S. y \ mcball x r); 0 < r\ \ \y\S. y \ M \ d x y < r" by (meson dense in_mcball le_less_trans) then show ?thesis by (fastforce simp: metric_closure_of in_closure_of) qed lemma metric_interior_of: "mtopology interior_of S = {x \ M. \\>0. mball x \ \ S}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" using interior_of_maximal_eq openin_mtopology by fastforce show "?rhs \ ?lhs" using interior_of_def openin_mball by fastforce qed lemma metric_interior_of_alt: "mtopology interior_of S = {x \ M. \\>0. mcball x \ \ S}" by (fastforce simp: mball_iff_mcball metric_interior_of) lemma in_interior_of_mball: "x \ mtopology interior_of S \ x \ M \ (\\>0. mball x \ \ S)" using metric_interior_of by force lemma in_interior_of_mcball: "x \ mtopology interior_of S \ x \ M \ (\\>0. mcball x \ \ S)" using metric_interior_of_alt by force lemma Hausdorff_space_mtopology: "Hausdorff_space mtopology" unfolding Hausdorff_space_def proof clarify fix x y assume x: "x \ topspace mtopology" and y: "y \ topspace mtopology" and "x \ y" then have gt0: "d x y / 2 > 0" by auto have "disjnt (mball x (d x y / 2)) (mball y (d x y / 2))" by (simp add: disjoint_mball) then show "\U V. openin mtopology U \ openin mtopology V \ x \ U \ y \ V \ disjnt U V" by (metis centre_in_mball_iff gt0 openin_mball topspace_mtopology x y) qed subsection\Bounded sets\ definition mbounded where "mbounded S \ (\x B. S \ mcball x B)" lemma mbounded_pos: "mbounded S \ (\x B. 0 < B \ S \ mcball x B)" proof - have "\x' r'. 0 < r' \ S \ mcball x' r'" if "S \ mcball x r" for x r by (metis gt_ex less_eq_real_def linorder_not_le mcball_subset_concentric order_trans that) then show ?thesis by (auto simp: mbounded_def) qed lemma mbounded_alt: "mbounded S \ S \ M \ (\B. \x \ S. \y \ S. d x y \ B)" proof - have "\x B. S \ mcball x B \ \x\S. \y\S. d x y \ 2 * B" by (smt (verit, best) commute in_mcball subsetD triangle) then show ?thesis apply (auto simp: mbounded_def subset_iff) apply blast+ done qed lemma mbounded_alt_pos: "mbounded S \ S \ M \ (\B>0. \x \ S. \y \ S. d x y \ B)" by (smt (verit, del_insts) gt_ex mbounded_alt) lemma mbounded_subset: "\mbounded T; S \ T\ \ mbounded S" by (meson mbounded_def order_trans) lemma mbounded_subset_mspace: "mbounded S \ S \ M" by (simp add: mbounded_alt) lemma mbounded: "mbounded S \ S = {} \ (\x \ S. x \ M) \ (\y B. y \ M \ (\x \ S. d y x \ B))" by (meson all_not_in_conv in_mcball mbounded_def subset_iff) lemma mbounded_empty [iff]: "mbounded {}" by (simp add: mbounded) lemma mbounded_mcball: "mbounded (mcball x r)" using mbounded_def by auto lemma mbounded_mball [iff]: "mbounded (mball x r)" by (meson mball_subset_mcball mbounded_def) lemma mbounded_insert: "mbounded (insert a S) \ a \ M \ mbounded S" proof - have "\y B. \y \ M; \x\S. d y x \ B\ \ \y. y \ M \ (\B \ d y a. \x\S. d y x \ B)" by (metis order.trans nle_le) then show ?thesis by (auto simp: mbounded) qed lemma mbounded_Int: "mbounded S \ mbounded (S \ T)" by (meson inf_le1 mbounded_subset) lemma mbounded_Un: "mbounded (S \ T) \ mbounded S \ mbounded T" (is "?lhs=?rhs") proof assume R: ?rhs show ?lhs proof (cases "S={} \ T={}") case True then show ?thesis using R by auto next case False obtain x y B C where "S \ mcball x B" "T \ mcball y C" "B > 0" "C > 0" "x \ M" "y \ M" using R mbounded_pos by (metis False mcball_eq_empty subset_empty) then have "S \ T \ mcball x (B + C + d x y)" by (smt (verit) commute dual_order.trans le_supI mcball_subset mdist_pos_eq) then show ?thesis using mbounded_def by blast qed next show "?lhs \ ?rhs" using mbounded_def by auto qed lemma mbounded_Union: "\finite \; \X. X \ \ \ mbounded X\ \ mbounded (\\)" by (induction \ rule: finite_induct) (auto simp: mbounded_Un) lemma mbounded_closure_of: "mbounded S \ mbounded (mtopology closure_of S)" by (meson closedin_mcball closure_of_minimal mbounded_def) lemma mbounded_closure_of_eq: "S \ M \ (mbounded (mtopology closure_of S) \ mbounded S)" by (metis closure_of_subset mbounded_closure_of mbounded_subset topspace_mtopology) lemma maxdist_thm: assumes "mbounded S" and "x \ S" and "y \ S" shows "d x y = (SUP z\S. \d x z - d z y\)" proof - have "\d x z - d z y\ \ d x y" if "z \ S" for z by (metis all_not_in_conv assms mbounded mdist_reverse_triangle that) moreover have "d x y \ r" if "\z. z \ S \ \d x z - d z y\ \ r" for r :: real using that assms mbounded_subset_mspace mdist_zero by fastforce ultimately show ?thesis by (intro cSup_eq [symmetric]) auto qed lemma metric_eq_thm: "\S \ M; x \ S; y \ S\ \ (x = y) = (\z\S. d x z = d y z)" by (metis commute subset_iff zero) lemma compactin_imp_mbounded: assumes "compactin mtopology S" shows "mbounded S" proof - have "S \ M" and com: "\\. \\U\\. openin mtopology U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using assms by (auto simp: compactin_def mbounded_def) show ?thesis proof (cases "S = {}") case False with \S \ M\ obtain a where "a \ S" "a \ M" by blast with \S \ M\ gt_ex have "S \ \(range (mball a))" by force moreover have "\U \ range (mball a). openin mtopology U" by (simp add: openin_mball) ultimately obtain \ where "finite \" "\ \ range (mball a)" "S \ \\" by (meson com) then show ?thesis using mbounded_Union mbounded_subset by fastforce qed auto qed end subsection\Subspace of a metric space\ locale Submetric = Metric_space + fixes A assumes subset: "A \ M" sublocale Submetric \ sub: Metric_space A d by (simp add: subset subspace) context Submetric begin lemma mball_submetric_eq: "sub.mball a r = (if a \ A then A \ mball a r else {})" and mcball_submetric_eq: "sub.mcball a r = (if a \ A then A \ mcball a r else {})" using subset by force+ lemma mtopology_submetric: "sub.mtopology = subtopology mtopology A" unfolding topology_eq proof (intro allI iffI) fix S assume "openin sub.mtopology S" then have "\T. openin (subtopology mtopology A) T \ x \ T \ T \ S" if "x \ S" for x by (metis mball_submetric_eq openin_mball openin_subtopology_Int2 sub.centre_in_mball_iff sub.openin_mtopology subsetD that) then show "openin (subtopology mtopology A) S" by (meson openin_subopen) next fix S assume "openin (subtopology mtopology A) S" then obtain T where "openin mtopology T" "S = T \ A" by (meson openin_subtopology) then have "mopen T" by (simp add: mopen_def openin_mtopology) then have "sub.mopen (T \ A)" unfolding sub.mopen_def mopen_def by (metis inf.coboundedI2 mball_submetric_eq Int_iff \S = T \ A\ inf.bounded_iff subsetI) then show "openin sub.mtopology S" using \S = T \ A\ sub.mopen_def sub.openin_mtopology by force qed lemma mbounded_submetric: "sub.mbounded T \ mbounded T \ T \ A" by (meson mbounded_alt sub.mbounded_alt subset subset_trans) end lemma (in Metric_space) submetric_empty [iff]: "Submetric M d {}" by (simp add: Metric_space_axioms Submetric.intro Submetric_axioms_def) subsection \Abstract type of metric spaces\ typedef 'a metric = "{(M::'a set,d). Metric_space M d}" morphisms "dest_metric" "metric" proof - have "Metric_space {} (\x y. 0)" by (auto simp: Metric_space_def) then show ?thesis by blast qed definition mspace where "mspace m \ fst (dest_metric m)" definition mdist where "mdist m \ snd (dest_metric m)" lemma Metric_space_mspace_mdist: "Metric_space (mspace m) (mdist m)" by (metis Product_Type.Collect_case_prodD dest_metric mdist_def mspace_def) lemma mdist_nonneg [simp]: "\x y. 0 \ mdist m x y" by (metis Metric_space_def Metric_space_mspace_mdist) lemma mdist_commute: "\x y. mdist m x y = mdist m y x" by (metis Metric_space_def Metric_space_mspace_mdist) lemma mdist_zero [simp]: "\x y. \x \ mspace m; y \ mspace m\ \ mdist m x y = 0 \ x=y" by (meson Metric_space.zero Metric_space_mspace_mdist) lemma mdist_triangle: "\x y z. \x \ mspace m; y \ mspace m; z \ mspace m\ \ mdist m x z \ mdist m x y + mdist m y z" by (meson Metric_space.triangle Metric_space_mspace_mdist) lemma (in Metric_space) mspace_metric[simp]: "mspace (metric (M,d)) = M" by (simp add: mspace_def Metric_space_axioms metric_inverse) lemma (in Metric_space) mdist_metric[simp]: "mdist (metric (M,d)) = d" by (simp add: mdist_def Metric_space_axioms metric_inverse) lemma metric_collapse [simp]: "metric (mspace m, mdist m) = m" by (simp add: dest_metric_inverse mdist_def mspace_def) definition mtopology_of :: "'a metric \ 'a topology" where "mtopology_of \ \m. Metric_space.mtopology (mspace m) (mdist m)" lemma topspace_mtopology_of [simp]: "topspace (mtopology_of m) = mspace m" by (simp add: Metric_space.topspace_mtopology Metric_space_mspace_mdist mtopology_of_def) lemma (in Metric_space) mtopology_of [simp]: "mtopology_of (metric (M,d)) = mtopology" by (simp add: mtopology_of_def) definition "mball_of \ \m. Metric_space.mball (mspace m) (mdist m)" lemma (in Metric_space) mball_of [simp]: "mball_of (metric (M,d)) = mball" by (simp add: mball_of_def) definition "mcball_of \ \m. Metric_space.mcball (mspace m) (mdist m)" lemma (in Metric_space) mcball_of [simp]: "mcball_of (metric (M,d)) = mcball" by (simp add: mcball_of_def) definition "euclidean_metric \ metric (UNIV,dist)" lemma mspace_euclidean_metric [simp]: "mspace euclidean_metric = UNIV" by (simp add: euclidean_metric_def) lemma mdist_euclidean_metric [simp]: "mdist euclidean_metric = dist" by (simp add: euclidean_metric_def) +lemma mtopology_of_euclidean [simp]: "mtopology_of euclidean_metric = euclidean" + by (simp add: Met_TC.mtopology_def mtopology_of_def) + text\ Subspace of a metric space\ definition submetric where "submetric \ \m S. metric (S \ mspace m, mdist m)" lemma mspace_submetric [simp]: "mspace (submetric m S) = S \ mspace m" unfolding submetric_def by (meson Metric_space.subspace inf_le2 Metric_space_mspace_mdist Metric_space.mspace_metric) lemma mdist_submetric [simp]: "mdist (submetric m S) = mdist m" unfolding submetric_def by (meson Metric_space.subspace inf_le2 Metric_space.mdist_metric Metric_space_mspace_mdist) lemma submetric_UNIV [simp]: "submetric m UNIV = m" by (simp add: submetric_def dest_metric_inverse mdist_def mspace_def) lemma submetric_submetric [simp]: "submetric (submetric m S) T = submetric m (S \ T)" by (metis submetric_def Int_assoc inf_commute mdist_submetric mspace_submetric) lemma submetric_mspace [simp]: "submetric m (mspace m) = m" by (simp add: submetric_def dest_metric_inverse mdist_def mspace_def) lemma submetric_restrict: "submetric m S = submetric m (mspace m \ S)" by (metis submetric_mspace submetric_submetric) +lemma mtopology_of_submetric: "mtopology_of (submetric m A) = subtopology (mtopology_of m) A" +proof - + interpret Submetric "mspace m" "mdist m" "A \ mspace m" + using Metric_space_mspace_mdist Submetric.intro Submetric_axioms.intro inf_le2 by blast + have "sub.mtopology = subtopology (mtopology_of m) A" + by (metis inf_commute mtopology_of_def mtopology_submetric subtopology_mspace subtopology_subtopology) + then show ?thesis + by (simp add: submetric_def) +qed + subsection\The discrete metric\ locale discrete_metric = fixes M :: "'a set" definition (in discrete_metric) dd :: "'a \ 'a \ real" where "dd \ \x y::'a. if x=y then 0 else 1" lemma metric_M_dd: "Metric_space M discrete_metric.dd" by (simp add: discrete_metric.dd_def Metric_space.intro) sublocale discrete_metric \ disc: Metric_space M dd by (simp add: metric_M_dd) lemma (in discrete_metric) mopen_singleton: assumes "x \ M" shows "disc.mopen {x}" proof - have "disc.mball x (1/2) \ {x}" by (smt (verit) dd_def disc.in_mball less_divide_eq_1_pos singleton_iff subsetI) with assms show ?thesis using disc.mopen_def half_gt_zero_iff zero_less_one by blast qed lemma (in discrete_metric) mtopology_discrete_metric: "disc.mtopology = discrete_topology M" proof - have "\x. x \ M \ openin disc.mtopology {x}" by (simp add: disc.mtopology_def mopen_singleton) then show ?thesis by (metis disc.topspace_mtopology discrete_topology_unique) qed lemma (in discrete_metric) discrete_ultrametric: "dd x z \ max (dd x y) (dd y z)" by (simp add: dd_def) lemma (in discrete_metric) dd_le1: "dd x y \ 1" by (simp add: dd_def) lemma (in discrete_metric) mbounded_discrete_metric: "disc.mbounded S \ S \ M" by (meson dd_le1 disc.mbounded_alt) subsection\Metrizable spaces\ definition metrizable_space where "metrizable_space X \ \M d. Metric_space M d \ X = Metric_space.mtopology M d" lemma (in Metric_space) metrizable_space_mtopology: "metrizable_space mtopology" using local.Metric_space_axioms metrizable_space_def by blast lemma openin_mtopology_eq_open [simp]: "openin Met_TC.mtopology = open" by (simp add: Met_TC.mtopology_def) lemma closedin_mtopology_eq_closed [simp]: "closedin Met_TC.mtopology = closed" proof - have "(euclidean::'a topology) = Met_TC.mtopology" by (simp add: Met_TC.mtopology_def) then show ?thesis using closed_closedin by fastforce qed lemma compactin_mtopology_eq_compact [simp]: "compactin Met_TC.mtopology = compact" by (simp add: compactin_def compact_eq_Heine_Borel fun_eq_iff) meson lemma metrizable_space_discrete_topology: "metrizable_space(discrete_topology U)" by (metis discrete_metric.mtopology_discrete_metric metric_M_dd metrizable_space_def) lemma empty_metrizable_space: "topspace X = {} \ metrizable_space X" by (metis metrizable_space_discrete_topology subtopology_eq_discrete_topology_empty) lemma metrizable_space_subtopology: assumes "metrizable_space X" shows "metrizable_space(subtopology X S)" proof - obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d" using assms metrizable_space_def by blast then interpret Submetric M d "M \ S" by (simp add: Submetric.intro Submetric_axioms_def) show ?thesis unfolding metrizable_space_def by (metis X mtopology_submetric sub.Metric_space_axioms subtopology_restrict topspace_mtopology) qed lemma homeomorphic_metrizable_space_aux: assumes "X homeomorphic_space Y" "metrizable_space X" shows "metrizable_space Y" proof - obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d" using assms by (auto simp: metrizable_space_def) then interpret m: Metric_space M d by simp obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" and fg: "(\x \ M. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce define d' where "d' x y \ d (g x) (g y)" for x y interpret m': Metric_space "topspace Y" "d'" unfolding d'_def proof show "(d (g x) (g y) = 0) = (x = y)" if "x \ topspace Y" "y \ topspace Y" for x y by (metis fg X hmg homeomorphic_imp_surjective_map imageI m.topspace_mtopology m.zero that) show "d (g x) (g z) \ d (g x) (g y) + d (g y) (g z)" if "x \ topspace Y" and "y \ topspace Y" and "z \ topspace Y" for x y z by (metis X that hmg homeomorphic_eq_everything_map imageI m.topspace_mtopology m.triangle) qed (auto simp: m.nonneg m.commute) have "Y = Metric_space.mtopology (topspace Y) d'" unfolding topology_eq proof (intro allI) fix S have "openin m'.mtopology S" if S: "S \ topspace Y" and "openin X (g ` S)" unfolding m'.openin_mtopology proof (intro conjI that strip) fix y assume "y \ S" then obtain r where "r>0" and r: "m.mball (g y) r \ g ` S" using X \openin X (g ` S)\ m.openin_mtopology using \y \ S\ by auto then have "g ` m'.mball y r \ m.mball (g y) r" using X d'_def hmg homeomorphic_imp_surjective_map by fastforce with S fg have "m'.mball y r \ S" by (smt (verit, del_insts) image_iff m'.in_mball r subset_iff) then show "\r>0. m'.mball y r \ S" using \0 < r\ by blast qed moreover have "openin X (g ` S)" if ope': "openin m'.mtopology S" proof - have "\r>0. m.mball (g y) r \ g ` S" if "y \ S" for y proof - have y: "y \ topspace Y" using m'.openin_mtopology ope' that by blast obtain r where "r > 0" and r: "m'.mball y r \ S" using ope' by (meson \y \ S\ m'.openin_mtopology) moreover have "\x. \x \ M; d (g y) x < r\ \ \u. u \ topspace Y \ d' y u < r \ x = g u" using fg X d'_def hmf homeomorphic_imp_surjective_map by fastforce ultimately have "m.mball (g y) r \ g ` m'.mball y r" using y by (force simp: m'.openin_mtopology) then show ?thesis using \0 < r\ r by blast qed then show ?thesis using X hmg homeomorphic_imp_surjective_map m.openin_mtopology ope' openin_subset by fastforce qed ultimately have "(S \ topspace Y \ openin X (g ` S)) = openin m'.mtopology S" using m'.topspace_mtopology openin_subset by blast then show "openin Y S = openin m'.mtopology S" by (simp add: m'.mopen_def homeomorphic_map_openness_eq [OF hmg]) qed then show ?thesis using m'.metrizable_space_mtopology by force qed lemma homeomorphic_metrizable_space: assumes "X homeomorphic_space Y" shows "metrizable_space X \ metrizable_space Y" using assms homeomorphic_metrizable_space_aux homeomorphic_space_sym by metis lemma metrizable_space_retraction_map_image: "retraction_map X Y r \ metrizable_space X \ metrizable_space Y" using hereditary_imp_retractive_property metrizable_space_subtopology homeomorphic_metrizable_space by blast lemma metrizable_imp_Hausdorff_space: "metrizable_space X \ Hausdorff_space X" by (metis Metric_space.Hausdorff_space_mtopology metrizable_space_def) (** lemma metrizable_imp_kc_space: "metrizable_space X \ kc_space X" oops MESON_TAC[METRIZABLE_IMP_HAUSDORFF_SPACE; HAUSDORFF_IMP_KC_SPACE]);; lemma kc_space_mtopology: "kc_space mtopology" oops REWRITE_TAC[GSYM FORALL_METRIZABLE_SPACE; METRIZABLE_IMP_KC_SPACE]);; **) lemma metrizable_imp_t1_space: "metrizable_space X \ t1_space X" by (simp add: Hausdorff_imp_t1_space metrizable_imp_Hausdorff_space) lemma closed_imp_gdelta_in: assumes X: "metrizable_space X" and S: "closedin X S" shows "gdelta_in X S" proof - obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d" using X metrizable_space_def by blast then interpret M: Metric_space M d by blast have "S \ M" using M.closedin_metric \X = M.mtopology\ S by blast show ?thesis proof (cases "S = {}") case True then show ?thesis by simp next case False have "\y\S. d x y < inverse (1 + real n)" if "x \ S" for x n using \S \ M\ M.mdist_zero [of x] that by force moreover have "x \ S" if "x \ M" and \
: "\n. \y\S. d x y < inverse(Suc n)" for x proof - have *: "\y\S. d x y < \" if "\ > 0" for \ by (metis \
that not0_implies_Suc order_less_le order_less_le_trans real_arch_inverse) have "closedin M.mtopology S" using S by (simp add: Xeq) then show ?thesis apply (simp add: M.closedin_metric) by (metis * \x \ M\ M.in_mball disjnt_insert1 insert_absorb subsetD) qed ultimately have Seq: "S = \(range (\n. {x\M. \y\S. d x y < inverse(Suc n)}))" using \S \ M\ by force have "openin M.mtopology {xa \ M. \y\S. d xa y < inverse (1 + real n)}" for n proof (clarsimp simp: M.openin_mtopology) fix x y assume "x \ M" "y \ S" and dxy: "d x y < inverse (1 + real n)" then have "\z. \z \ M; d x z < inverse (1 + real n) - d x y\ \ \y\S. d z y < inverse (1 + real n)" by (smt (verit) M.commute M.triangle \S \ M\ in_mono) with dxy show "\r>0. M.mball x r \ {z \ M. \y\S. d z y < inverse (1 + real n)}" by (rule_tac x="inverse(Suc n) - d x y" in exI) auto qed then show ?thesis apply (subst Seq) apply (force simp: Xeq intro: gdelta_in_Inter open_imp_gdelta_in) done qed qed lemma open_imp_fsigma_in: "\metrizable_space X; openin X S\ \ fsigma_in X S" by (meson closed_imp_gdelta_in fsigma_in_gdelta_in openin_closedin openin_subset) -(*NEEDS first_countable -lemma first_countable_mtopology: - "first_countable mtopology" -oops - GEN_TAC THEN REWRITE_TAC[first_countable; TOPSPACE_MTOPOLOGY] THEN - X_GEN_TAC `x::A` THEN DISCH_TAC THEN - EXISTS_TAC `{ mball m (x::A,r) | rational r \ 0 < r}` THEN - REWRITE_TAC[FORALL_IN_GSPEC; OPEN_IN_MBALL; EXISTS_IN_GSPEC] THEN - ONCE_REWRITE_TAC[SET_RULE - `{f x | S x \ Q x} = f ` {x \ S. Q x}`] THEN - SIMP_TAC[COUNTABLE_IMAGE; COUNTABLE_RATIONAL; COUNTABLE_RESTRICT] THEN - REWRITE_TAC[OPEN_IN_MTOPOLOGY] THEN - X_GEN_TAC `U::A=>bool` THEN STRIP_TAC THEN - FIRST_X_ASSUM(MP_TAC \ SPEC `x::A`) THEN - ASM_REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN - X_GEN_TAC `r::real` THEN STRIP_TAC THEN FIRST_ASSUM - (MP_TAC \ SPEC `r::real` \ MATCH_MP RATIONAL_APPROXIMATION_BELOW) THEN - MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `q::real` THEN - REWRITE_TAC[REAL_SUB_REFL] THEN STRIP_TAC THEN - ASM_SIMP_TAC[CENTRE_IN_MBALL] THEN - TRANS_TAC SUBSET_TRANS `mball m (x::A,r)` THEN - ASM_SIMP_TAC[MBALL_SUBSET_CONCENTRIC; REAL_LT_IMP_LE]);; - -lemma metrizable_imp_first_countable: - "metrizable_space X \ first_countable X" -oops - REWRITE_TAC[FORALL_METRIZABLE_SPACE; FIRST_COUNTABLE_MTOPOLOGY]);; -*) - lemma mball_eq_ball [simp]: "Met_TC.mball = ball" by force lemma mopen_eq_open [simp]: "Met_TC.mopen = open" by (force simp: open_contains_ball Met_TC.mopen_def) lemma metrizable_space_euclidean: "metrizable_space (euclidean :: 'a::metric_space topology)" unfolding metrizable_space_def by (metis Met_TC.Metric_space_axioms Met_TC.mtopology_def mopen_eq_open) lemma (in Metric_space) regular_space_mtopology: "regular_space mtopology" unfolding regular_space_def proof clarify fix C a assume C: "closedin mtopology C" and a: "a \ topspace mtopology" and "a \ C" have "openin mtopology (topspace mtopology - C)" by (simp add: C openin_diff) then obtain r where "r>0" and r: "mball a r \ topspace mtopology - C" unfolding openin_mtopology using \a \ C\ a by auto show "\U V. openin mtopology U \ openin mtopology V \ a \ U \ C \ V \ disjnt U V" proof (intro exI conjI) show "a \ mball a (r/2)" using \0 < r\ a by force show "C \ topspace mtopology - mcball a (r/2)" using C \0 < r\ r by (fastforce simp: closedin_metric) qed (auto simp: openin_mball closedin_mcball openin_diff disjnt_iff) qed lemma metrizable_imp_regular_space: "metrizable_space X \ regular_space X" by (metis Metric_space.regular_space_mtopology metrizable_space_def) lemma regular_space_euclidean: "regular_space (euclidean :: 'a::metric_space topology)" by (simp add: metrizable_imp_regular_space metrizable_space_euclidean) subsection\Limits at a point in a topological space\ lemma (in Metric_space) eventually_atin_metric: "eventually P (atin mtopology a) \ (a \ M \ (\\>0. \x. x \ M \ 0 < d x a \ d x a < \ \ P x))" (is "?lhs=?rhs") proof (cases "a \ M") case True show ?thesis proof assume L: ?lhs with True obtain U where "openin mtopology U" "a \ U" and U: "\x\U - {a}. P x" by (auto simp: eventually_atin) then obtain r where "r>0" and "mball a r \ U" by (meson openin_mtopology) with U show ?rhs by (smt (verit, ccfv_SIG) commute in_mball insert_Diff_single insert_iff subset_iff) next assume ?rhs then obtain \ where "\>0" and \: "\x. x \ M \ 0 < d x a \ d x a < \ \ P x" using True by blast then have "\x \ mball a \ - {a}. P x" by (simp add: commute) then show ?lhs unfolding eventually_atin openin_mtopology by (metis True \0 < \\ centre_in_mball_iff openin_mball openin_mtopology) qed qed auto subsection \Normal spaces and metric spaces\ lemma (in Metric_space) normal_space_mtopology: "normal_space mtopology" unfolding normal_space_def proof clarify fix S T assume "closedin mtopology S" then have "\x. x \ M - S \ (\r>0. mball x r \ M - S)" by (simp add: closedin_def openin_mtopology) then obtain \ where d0: "\x. x \ M - S \ \ x > 0 \ mball x (\ x) \ M - S" by metis assume "closedin mtopology T" then have "\x. x \ M - T \ (\r>0. mball x r \ M - T)" by (simp add: closedin_def openin_mtopology) then obtain \ where e: "\x. x \ M - T \ \ x > 0 \ mball x (\ x) \ M - T" by metis assume "disjnt S T" have "S \ M" "T \ M" using \closedin mtopology S\ \closedin mtopology T\ closedin_metric by blast+ have \: "\x. x \ T \ \ x > 0 \ mball x (\ x) \ M - S" by (meson DiffI \T \ M\ \disjnt S T\ d0 disjnt_iff subsetD) have \: "\x. x \ S \ \ x > 0 \ mball x (\ x) \ M - T" by (meson Diff_iff \S \ M\ \disjnt S T\ disjnt_iff e subsetD) show "\U V. openin mtopology U \ openin mtopology V \ S \ U \ T \ V \ disjnt U V" proof (intro exI conjI) show "openin mtopology (\x\S. mball x (\ x / 2))" "openin mtopology (\x\T. mball x (\ x / 2))" by force+ show "S \ (\x\S. mball x (\ x / 2))" using \ \S \ M\ by force show "T \ (\x\T. mball x (\ x / 2))" using \ \T \ M\ by force show "disjnt (\x\S. mball x (\ x / 2)) (\x\T. mball x (\ x / 2))" using \ \ apply (clarsimp simp: disjnt_iff subset_iff) by (smt (verit, ccfv_SIG) field_sum_of_halves triangle') qed qed lemma metrizable_imp_normal_space: "metrizable_space X \ normal_space X" by (metis Metric_space.normal_space_mtopology metrizable_space_def) subsection\Topological limitin in metric spaces\ lemma (in Metric_space) limitin_mspace: "limitin mtopology f l F \ l \ M" using limitin_topspace by fastforce lemma (in Metric_space) limitin_metric_unique: "\limitin mtopology f l1 F; limitin mtopology f l2 F; F \ bot\ \ l1 = l2" by (meson Hausdorff_space_mtopology limitin_Hausdorff_unique) lemma (in Metric_space) limitin_metric: "limitin mtopology f l F \ l \ M \ (\\>0. eventually (\x. f x \ M \ d (f x) l < \) F)" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding limitin_def proof (intro conjI strip) show "l \ M" using L limitin_mspace by blast fix \::real assume "\>0" then have "\\<^sub>F x in F. f x \ mball l \" using L openin_mball by (fastforce simp: limitin_def) then show "\\<^sub>F x in F. f x \ M \ d (f x) l < \" using commute eventually_mono by fastforce qed next assume R: ?rhs then show ?lhs by (force simp: limitin_def commute openin_mtopology subset_eq elim: eventually_mono) qed lemma (in Metric_space) limit_metric_sequentially: "limitin mtopology f l sequentially \ l \ M \ (\\>0. \N. \n\N. f n \ M \ d (f n) l < \)" by (auto simp: limitin_metric eventually_sequentially) lemma (in Submetric) limitin_submetric_iff: "limitin sub.mtopology f l F \ l \ A \ eventually (\x. f x \ A) F \ limitin mtopology f l F" (is "?lhs=?rhs") by (simp add: limitin_subtopology mtopology_submetric) lemma (in Metric_space) metric_closedin_iff_sequentially_closed: "closedin mtopology S \ S \ M \ (\\ l. range \ \ S \ limitin mtopology \ l sequentially \ l \ S)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (force simp: closedin_metric limitin_closedin range_subsetD) next assume R: ?rhs show ?lhs unfolding closedin_metric proof (intro conjI strip) show "S \ M" using R by blast fix x assume "x \ M - S" have False if "\r>0. \y. y \ M \ y \ S \ d x y < r" proof - have "\n. \y. y \ M \ y \ S \ d x y < inverse(Suc n)" using that by auto then obtain \ where \: "\n. \ n \ M \ \ n \ S \ d x (\ n) < inverse(Suc n)" by metis then have "range \ \ M" by blast have "\N. \n\N. d x (\ n) < \" if "\>0" for \ proof - have "real (Suc (nat \inverse \\)) \ inverse \" by linarith then have "\n \ nat \inverse \\. d x (\ n) < \" by (metis \ inverse_inverse_eq inverse_le_imp_le nat_ceiling_le_eq nle_le not_less_eq_eq order.strict_trans2 that) then show ?thesis .. qed with \ have "limitin mtopology \ x sequentially" using \x \ M - S\ commute limit_metric_sequentially by auto then show ?thesis by (metis R DiffD2 \ image_subset_iff \x \ M - S\) qed then show "\r>0. disjnt S (mball x r)" by (meson disjnt_iff in_mball) qed qed lemma (in Metric_space) limit_atin_metric: "limitin X f y (atin mtopology x) \ y \ topspace X \ (x \ M \ (\V. openin X V \ y \ V \ (\\>0. \x'. x' \ M \ 0 < d x' x \ d x' x < \ \ f x' \ V)))" by (force simp: limitin_def eventually_atin_metric) lemma (in Metric_space) limitin_metric_dist_null: "limitin mtopology f l F \ l \ M \ eventually (\x. f x \ M) F \ ((\x. d (f x) l) \ 0) F" by (simp add: limitin_metric tendsto_iff eventually_conj_iff all_conj_distrib imp_conjR gt_ex) subsection\Cauchy sequences and complete metric spaces\ context Metric_space begin definition MCauchy :: "(nat \ 'a) \ bool" where "MCauchy \ \ range \ \ M \ (\\>0. \N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \)" definition mcomplete where "mcomplete \ (\\. MCauchy \ \ (\x. limitin mtopology \ x sequentially))" lemma mcomplete_empty [iff]: "Metric_space.mcomplete {} d" by (simp add: Metric_space.MCauchy_def Metric_space.mcomplete_def subspace) lemma MCauchy_imp_MCauchy_suffix: "MCauchy \ \ MCauchy (\ \ (+)n)" unfolding MCauchy_def image_subset_iff comp_apply by (metis UNIV_I add.commute trans_le_add1) lemma mcomplete: "mcomplete \ (\\. (\\<^sub>F n in sequentially. \ n \ M) \ (\\>0. \N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \) \ (\x. limitin mtopology \ x sequentially))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof clarify fix \ assume "\\<^sub>F n in sequentially. \ n \ M" and \: "\\>0. \N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" then obtain N where "\n. n\N \ \ n \ M" by (auto simp: eventually_sequentially) with \ have "MCauchy (\ \ (+)N)" unfolding MCauchy_def image_subset_iff comp_apply by (meson le_add1 trans_le_add2) then obtain x where "limitin mtopology (\ \ (+)N) x sequentially" using L MCauchy_imp_MCauchy_suffix mcomplete_def by blast then have "limitin mtopology \ x sequentially" unfolding o_def by (auto simp: add.commute limitin_sequentially_offset_rev) then show "\x. limitin mtopology \ x sequentially" .. qed qed (simp add: mcomplete_def MCauchy_def image_subset_iff) lemma mcomplete_empty_mspace: "M = {} \ mcomplete" using MCauchy_def mcomplete_def by blast lemma MCauchy_const [simp]: "MCauchy (\n. a) \ a \ M" using MCauchy_def mdist_zero by auto lemma convergent_imp_MCauchy: assumes "range \ \ M" and lim: "limitin mtopology \ l sequentially" shows "MCauchy \" unfolding MCauchy_def image_subset_iff proof (intro conjI strip) fix \::real assume "\ > 0" then have "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) l < \/2" using half_gt_zero lim limitin_metric by blast then obtain N where "\n. n\N \ \ n \ M \ d (\ n) l < \/2" by (force simp: eventually_sequentially) then show "\N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" by (smt (verit) Metric_space.limitin_mspace Metric_space.mdist_reverse_triangle Metric_space_axioms field_sum_of_halves lim) qed (use assms in blast) lemma mcomplete_alt: "mcomplete \ (\\. MCauchy \ \ range \ \ M \ (\x. limitin mtopology \ x sequentially))" using MCauchy_def convergent_imp_MCauchy mcomplete_def by blast lemma MCauchy_subsequence: assumes "strict_mono r" "MCauchy \" shows "MCauchy (\ \ r)" proof - have "d (\ (r n)) (\ (r n')) < \" if "N \ n" "N \ n'" "strict_mono r" "\n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" for \ N n n' using that by (meson le_trans strict_mono_imp_increasing) moreover have "range (\x. \ (r x)) \ M" using MCauchy_def assms by blast ultimately show ?thesis using assms by (simp add: MCauchy_def) metis qed lemma MCauchy_offset: assumes cau: "MCauchy (\ \ (+)k)" and \: "\n. n < k \ \ n \ M" shows "MCauchy \" unfolding MCauchy_def image_subset_iff proof (intro conjI strip) fix n show "\ n \ M" using assms unfolding MCauchy_def image_subset_iff by (metis UNIV_I comp_apply le_iff_add linorder_not_le) next fix \ :: real assume "\ > 0" obtain N where "\n n'. N \ n \ N \ n' \ d ((\ \ (+)k) n) ((\ \ (+)k) n') < \" using cau \\ > 0\ by (fastforce simp: MCauchy_def) then show "\N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" unfolding o_def apply (rule_tac x="k+N" in exI) by (smt (verit, del_insts) add.assoc le_add1 less_eqE) qed lemma MCauchy_convergent_subsequence: assumes cau: "MCauchy \" and "strict_mono r" and lim: "limitin mtopology (\ \ r) a sequentially" shows "limitin mtopology \ a sequentially" unfolding limitin_metric proof (intro conjI strip) show "a \ M" by (meson assms limitin_mspace) fix \ :: real assume "\ > 0" then obtain N1 where N1: "\n n'. \n\N1; n'\N1\ \ d (\ n) (\ n') < \/2" using cau unfolding MCauchy_def by (meson half_gt_zero) obtain N2 where N2: "\n. n \ N2 \ (\ \ r) n \ M \ d ((\ \ r) n) a < \/2" by (metis (no_types, lifting) lim \\ > 0\ half_gt_zero limit_metric_sequentially) have "\ n \ M \ d (\ n) a < \" if "n \ max N1 N2" for n proof (intro conjI) show "\ n \ M" using MCauchy_def cau by blast have "N1 \ r n" by (meson \strict_mono r\ le_trans max.cobounded1 strict_mono_imp_increasing that) then show "d (\ n) a < \" using N1[of n "r n"] N2[of n] \\ n \ M\ \a \ M\ triangle that by fastforce qed then show "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) a < \" using eventually_sequentially by blast qed lemma MCauchy_interleaving_gen: "MCauchy (\n. if even n then x(n div 2) else y(n div 2)) \ (MCauchy x \ MCauchy y \ (\n. d (x n) (y n)) \ 0)" (is "?lhs=?rhs") proof assume L: ?lhs have evens: "strict_mono (\n::nat. 2 * n)" and odds: "strict_mono (\n::nat. Suc (2 * n))" by (auto simp: strict_mono_def) show ?rhs proof (intro conjI) show "MCauchy x" "MCauchy y" using MCauchy_subsequence [OF evens L] MCauchy_subsequence [OF odds L] by (auto simp: o_def) show "(\n. d (x n) (y n)) \ 0" unfolding LIMSEQ_iff proof (intro strip) fix \ :: real assume "\ > 0" then obtain N where N: "\n n'. \n\N; n'\N\ \ d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \" using L MCauchy_def by fastforce have "d (x n) (y n) < \" if "n\N" for n using N [of "2*n" "Suc(2*n)"] that by auto then show "\N. \n\N. norm (d (x n) (y n) - 0) < \" by auto qed qed next assume R: ?rhs show ?lhs unfolding MCauchy_def proof (intro conjI strip) show "range (\n. if even n then x (n div 2) else y (n div 2)) \ M" using R by (auto simp: MCauchy_def) fix \ :: real assume "\ > 0" obtain Nx where Nx: "\n n'. \n\Nx; n'\Nx\ \ d (x n) (x n') < \/2" by (meson half_gt_zero MCauchy_def R \\ > 0\) obtain Ny where Ny: "\n n'. \n\Ny; n'\Ny\ \ d (y n) (y n') < \/2" by (meson half_gt_zero MCauchy_def R \\ > 0\) obtain Nxy where Nxy: "\n. n\Nxy \ d (x n) (y n) < \/2" using R \\ > 0\ half_gt_zero unfolding LIMSEQ_iff by (metis abs_mdist diff_zero real_norm_def) define N where "N \ 2 * Max{Nx,Ny,Nxy}" show "\N. \n n'. N \ n \ N \ n' \ d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \" proof (intro exI strip) fix n n' assume "N \ n" and "N \ n'" then have "n div 2 \ Nx" "n div 2 \ Ny" "n div 2 \ Nxy" "n' div 2 \ Nx" "n' div 2 \ Ny" by (auto simp: N_def) then have dxyn: "d (x (n div 2)) (y (n div 2)) < \/2" and dxnn': "d (x (n div 2)) (x (n' div 2)) < \/2" and dynn': "d (y (n div 2)) (y (n' div 2)) < \/2" using Nx Ny Nxy by blast+ have inM: "x (n div 2) \ M" "x (n' div 2) \ M""y (n div 2) \ M" "y (n' div 2) \ M" using Metric_space.MCauchy_def Metric_space_axioms R by blast+ show "d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \" proof (cases "even n") case nt: True show ?thesis proof (cases "even n'") case True with \\ > 0\ nt dxnn' show ?thesis by auto next case False with nt dxyn dynn' inM triangle show ?thesis by fastforce qed next case nf: False show ?thesis proof (cases "even n'") case True then show ?thesis by (smt (verit) \\ > 0\ dxyn dxnn' triangle commute inM field_sum_of_halves) next case False with \\ > 0\ nf dynn' show ?thesis by auto qed qed qed qed qed lemma MCauchy_interleaving: "MCauchy (\n. if even n then \(n div 2) else a) \ range \ \ M \ limitin mtopology \ a sequentially" (is "?lhs=?rhs") proof - have "?lhs \ (MCauchy \ \ a \ M \ (\n. d (\ n) a) \ 0)" by (simp add: MCauchy_interleaving_gen [where y = "\n. a"]) also have "... = ?rhs" by (metis MCauchy_def always_eventually convergent_imp_MCauchy limitin_metric_dist_null range_subsetD) finally show ?thesis . qed lemma mcomplete_nest: "mcomplete \ (\C::nat \'a set. (\n. closedin mtopology (C n)) \ (\n. C n \ {}) \ decseq C \ (\\>0. \n a. C n \ mcball a \) \ \ (range C) \ {})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding imp_conjL proof (intro strip) fix C :: "nat \ 'a set" assume clo: "\n. closedin mtopology (C n)" and ne: "\n. C n \ ({}::'a set)" and dec: "decseq C" and cover [rule_format]: "\\>0. \n a. C n \ mcball a \" obtain \ where \: "\n. \ n \ C n" by (meson ne empty_iff set_eq_iff) have "MCauchy \" unfolding MCauchy_def proof (intro conjI strip) show "range \ \ M" using \ clo metric_closedin_iff_sequentially_closed by auto fix \ :: real assume "\ > 0" then obtain N a where N: "C N \ mcball a (\/3)" using cover by fastforce have "d (\ m) (\ n) < \" if "N \ m" "N \ n" for m n proof - have "d a (\ m) \ \/3" "d a (\ n) \ \/3" using dec N \ that by (fastforce simp: decseq_def)+ then have "d (\ m) (\ n) \ \/3 + \/3" using triangle \ commute dec decseq_def subsetD that N by (smt (verit, ccfv_threshold) in_mcball) also have "... < \" using \\ > 0\ by auto finally show ?thesis . qed then show "\N. \m n. N \ m \ N \ n \ d (\ m) (\ n) < \" by blast qed then obtain x where x: "limitin mtopology \ x sequentially" using L mcomplete_def by blast have "x \ C n" for n proof (rule limitin_closedin [OF x]) show "closedin mtopology (C n)" by (simp add: clo) show "\\<^sub>F x in sequentially. \ x \ C n" by (metis \ dec decseq_def eventually_sequentiallyI subsetD) qed auto then show "\ (range C) \ {}" by blast qed next assume R: ?rhs show ?lhs unfolding mcomplete_def proof (intro strip) fix \ assume "MCauchy \" then have "range \ \ M" using MCauchy_def by blast define C where "C \ \n. mtopology closure_of (\ ` {n..})" have "\n. closedin mtopology (C n)" by (auto simp: C_def) moreover have ne: "\n. C n \ {}" using \MCauchy \\ by (auto simp: C_def MCauchy_def disjnt_iff closure_of_eq_empty_gen) moreover have dec: "decseq C" unfolding monotone_on_def proof (intro strip) fix m n::nat assume "m \ n" then have "{n..} \ {m..}" by auto then show "C n \ C m" unfolding C_def by (meson closure_of_mono image_mono) qed moreover have C: "\N u. C N \ mcball u \" if "\>0" for \ proof - obtain N where "\m n. N \ m \ N \ n \ d (\ m) (\ n) < \" by (meson MCauchy_def \0 < \\ \MCauchy \\) then have "\ ` {N..} \ mcball (\ N) \" using MCauchy_def \MCauchy \\ by (force simp: less_eq_real_def) then have "C N \ mcball (\ N) \" by (simp add: C_def closure_of_minimal) then show ?thesis by blast qed ultimately obtain l where x: "l \ \ (range C)" by (metis R ex_in_conv) then have *: "\\ N. 0 < \ \ \n'. N \ n' \ l \ M \ \ n' \ M \ d l (\ n') < \" by (force simp: C_def metric_closure_of) then have "l \ M" using gt_ex by blast show "\l. limitin mtopology \ l sequentially" unfolding limitin_metric proof (intro conjI strip exI) show "l \ M" using \\n. closedin mtopology (C n)\ closedin_subset x by fastforce fix \::real assume "\ > 0" obtain N where N: "\m n. N \ m \ N \ n \ d (\ m) (\ n) < \/2" by (meson MCauchy_def \0 < \\ \MCauchy \\ half_gt_zero) with * [of "\/2" N] have "\n\N. \ n \ M \ d (\ n) l < \" by (smt (verit) \range \ \ M\ commute field_sum_of_halves range_subsetD triangle) then show "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) l < \" using eventually_sequentially by blast qed qed qed lemma mcomplete_nest_sing: "mcomplete \ (\C. (\n. closedin mtopology (C n)) \ (\n. C n \ {}) \ decseq C \ (\e>0. \n a. C n \ mcball a e) \ (\l. l \ M \ \ (range C) = {l}))" proof - have *: False if clo: "\n. closedin mtopology (C n)" and cover: "\\>0. \n a. C n \ mcball a \" and no_sing: "\y. y \ M \ \ (range C) \ {y}" and l: "\n. l \ C n" for C :: "nat \ 'a set" and l proof - have inM: "\x. x \ \ (range C) \ x \ M" using closedin_metric clo by fastforce then have "l \ M" by (simp add: l) have False if l': "l' \ \ (range C)" and "l' \ l" for l' proof - have "l' \ M" using inM l' by blast obtain n a where na: "C n \ mcball a (d l l' / 3)" using inM \l \ M\ l' \l' \ l\ cover by force then have "d a l \ (d l l' / 3)" "d a l' \ (d l l' / 3)" "a \ M" using l l' na in_mcball by auto then have "d l l' \ (d l l' / 3) + (d l l' / 3)" using \l \ M\ \l' \ M\ mdist_reverse_triangle by fastforce then show False using nonneg [of l l'] \l' \ l\ \l \ M\ \l' \ M\ zero by force qed then show False by (metis l \l \ M\ no_sing INT_I empty_iff insertI1 is_singletonE is_singletonI') qed show ?thesis unfolding mcomplete_nest imp_conjL apply (intro all_cong1 imp_cong refl) using * by (smt (verit) Inter_iff ex_in_conv range_constant range_eqI) qed lemma mcomplete_fip: "mcomplete \ (\\. (\C \ \. closedin mtopology C) \ (\e>0. \C a. C \ \ \ C \ mcball a e) \ (\\. finite \ \ \ \ \ \ \ \ \ {}) \ \ \ \ {})" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs unfolding mcomplete_nest_sing imp_conjL proof (intro strip) fix \ :: "'a set set" assume clo: "\C\\. closedin mtopology C" and cover: "\e>0. \C a. C \ \ \ C \ mcball a e" and fip: "\\. finite \ \ \ \ \ \ \ \ \ {}" then have "\n. \C. C \ \ \ (\a. C \ mcball a (inverse (Suc n)))" by simp then obtain C where C: "\n. C n \ \" and coverC: "\n. \a. C n \ mcball a (inverse (Suc n))" by metis define D where "D \ \n. \ (C ` {..n})" have cloD: "closedin mtopology (D n)" for n unfolding D_def using clo C by blast have neD: "D n \ {}" for n using fip C by (simp add: D_def image_subset_iff) have decD: "decseq D" by (force simp: D_def decseq_def) have coverD: "\n a. D n \ mcball a \" if " \ >0" for \ proof - obtain n where "inverse (Suc n) < \" using \0 < \\ reals_Archimedean by blast then obtain a where "C n \ mcball a \" by (meson coverC less_eq_real_def mcball_subset_concentric order_trans) then show ?thesis unfolding D_def by blast qed have *: "a \ \\" if a: "\ (range D) = {a}" and "a \ M" for a proof - have aC: "a \ C n" for n using that by (auto simp: D_def) have eqa: "\u. (\n. u \ C n) \ a = u" using that by (auto simp: D_def) have "a \ T" if "T \ \" for T proof - have cloT: "closedin mtopology (T \ D n)" for n using clo cloD that by blast have "\ (insert T (C ` {..n})) \ {}" for n using that C by (intro fip [rule_format]) auto then have neT: "T \ D n \ {}" for n by (simp add: D_def) have decT: "decseq (\n. T \ D n)" by (force simp: D_def decseq_def) have coverT: "\n a. T \ D n \ mcball a \" if " \ >0" for \ by (meson coverD le_infI2 that) show ?thesis using L [unfolded mcomplete_nest_sing, rule_format, of "\n. T \ D n"] a by (force simp: cloT neT decT coverT) qed then show ?thesis by auto qed show "\ \ \ {}" by (metis L cloD neD decD coverD * empty_iff mcomplete_nest_sing) qed next assume R [rule_format]: ?rhs show ?lhs unfolding mcomplete_nest imp_conjL proof (intro strip) fix C :: "nat \ 'a set" assume clo: "\n. closedin mtopology (C n)" and ne: "\n. C n \ {}" and dec: "decseq C" and cover: "\\>0. \n a. C n \ mcball a \" have "\(C ` N) \ {}" if "finite N" for N proof - obtain k where "N \ {..k}" using \finite N\ finite_nat_iff_bounded_le by auto with dec have "C k \ \(C ` N)" by (auto simp: decseq_def) then show ?thesis using ne by force qed with clo cover R [of "range C"] show "\ (range C) \ {}" by (metis (no_types, opaque_lifting) finite_subset_image image_iff UNIV_I) qed qed lemma mcomplete_fip_sing: "mcomplete \ (\\. (\C\\. closedin mtopology C) \ (\e>0. \c a. c \ \ \ c \ mcball a e) \ (\\. finite \ \ \ \ \ \ \ \ \ {}) \ (\l. l \ M \ \ \ = {l}))" (is "?lhs = ?rhs") proof have *: "l \ M" "\ \ = {l}" if clo: "Ball \ (closedin mtopology)" and cover: "\e>0. \C a. C \ \ \ C \ mcball a e" and fin: "\\. finite \ \ \ \ \ \ \ \ \ {}" and l: "l \ \ \" for \ :: "'a set set" and l proof - show "l \ M" by (meson Inf_lower2 clo cover gt_ex metric_closedin_iff_sequentially_closed subsetD that(4)) show "\ \ = {l}" proof (cases "\ = {}") case True then show ?thesis using cover mbounded_pos by auto next case False have CM: "\a. a \ \\ \ a \ M" using False clo closedin_subset by fastforce have "l' \ \ \" if "l' \ l" for l' proof assume l': "l' \ \ \" with CM have "l' \ M" by blast with that \l \ M\ have gt0: "0 < d l l'" by simp then obtain C a where "C \ \" and C: "C \ mcball a (d l l' / 3)" using cover [rule_format, of "d l l' / 3"] by auto then have "d a l \ (d l l' / 3)" "d a l' \ (d l l' / 3)" "a \ M" using l l' in_mcball by auto then have "d l l' \ (d l l' / 3) + (d l l' / 3)" using \l \ M\ \l' \ M\ mdist_reverse_triangle by fastforce with gt0 show False by auto qed then show ?thesis using l by fastforce qed qed assume L: ?lhs with * show ?rhs unfolding mcomplete_fip imp_conjL ex_in_conv [symmetric] by (elim all_forward imp_forward2 asm_rl) (blast intro: elim: ) next assume ?rhs then show ?lhs unfolding mcomplete_fip by (force elim!: all_forward) qed end lemma MCauchy_iff_Cauchy [iff]: "Met_TC.MCauchy = Cauchy" by (force simp: Cauchy_def Met_TC.MCauchy_def) lemma mcomplete_iff_complete [iff]: "Met_TC.mcomplete (Pure.type ::'a::metric_space itself) \ complete (UNIV::'a set)" by (auto simp: Met_TC.mcomplete_def complete_def) lemma euclidean_metric: "Met_TC.mcomplete (Pure.type ::'a::euclidean_space itself)" using complete_UNIV mcomplete_iff_complete by blast context Submetric begin lemma MCauchy_submetric: "sub.MCauchy \ \ range \ \ A \ MCauchy \" using MCauchy_def sub.MCauchy_def subset by force lemma closedin_mcomplete_imp_mcomplete: assumes clo: "closedin mtopology A" and "mcomplete" shows "sub.mcomplete" unfolding sub.mcomplete_def proof (intro strip) fix \ assume "sub.MCauchy \" then have \: "MCauchy \" "range \ \ A" using MCauchy_submetric by blast+ then obtain x where x: "limitin mtopology \ x sequentially" using \mcomplete\ unfolding mcomplete_def by blast then have "x \ A" using \ clo metric_closedin_iff_sequentially_closed by force with \ x show "\x. limitin sub.mtopology \ x sequentially" using limitin_submetric_iff range_subsetD by fastforce qed lemma sequentially_closedin_mcomplete_imp_mcomplete: assumes "mcomplete" and "\\ l. range \ \ A \ limitin mtopology \ l sequentially \ l \ A" shows "sub.mcomplete" using assms closedin_mcomplete_imp_mcomplete metric_closedin_iff_sequentially_closed subset by blast end context Metric_space begin lemma mcomplete_Un: assumes A: "Submetric M d A" "Metric_space.mcomplete A d" and B: "Submetric M d B" "Metric_space.mcomplete B d" shows "Submetric M d (A \ B)" "Metric_space.mcomplete (A \ B) d" proof - show "Submetric M d (A \ B)" by (meson assms le_sup_iff Submetric_axioms_def Submetric_def) then interpret MAB: Metric_space "A \ B" d by (meson Submetric.subset subspace) interpret MA: Metric_space A d by (meson A Submetric.subset subspace) interpret MB: Metric_space B d by (meson B Submetric.subset subspace) show "Metric_space.mcomplete (A \ B) d" unfolding MAB.mcomplete_def proof (intro strip) fix \ assume "MAB.MCauchy \" then have "range \ \ A \ B" using MAB.MCauchy_def by blast then have "UNIV \ \ -` A \ \ -` B" by blast then consider "infinite (\ -` A)" | "infinite (\ -` B)" using finite_subset by auto then show "\x. limitin MAB.mtopology \ x sequentially" proof cases case 1 then obtain r where "strict_mono r" and r: "\n::nat. r n \ \ -` A" using infinite_enumerate by blast then have "MA.MCauchy (\ \ r)" using MA.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence \MAB.MCauchy \\ by auto with A obtain x where "limitin MA.mtopology (\ \ r) x sequentially" using MA.mcomplete_def by blast then have "limitin MAB.mtopology (\ \ r) x sequentially" by (metis MA.limit_metric_sequentially MAB.limit_metric_sequentially UnCI) then show ?thesis using MAB.MCauchy_convergent_subsequence \MAB.MCauchy \\ \strict_mono r\ by blast next case 2 then obtain r where "strict_mono r" and r: "\n::nat. r n \ \ -` B" using infinite_enumerate by blast then have "MB.MCauchy (\ \ r)" using MB.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence \MAB.MCauchy \\ by auto with B obtain x where "limitin MB.mtopology (\ \ r) x sequentially" using MB.mcomplete_def by blast then have "limitin MAB.mtopology (\ \ r) x sequentially" by (metis MB.limit_metric_sequentially MAB.limit_metric_sequentially UnCI) then show ?thesis using MAB.MCauchy_convergent_subsequence \MAB.MCauchy \\ \strict_mono r\ by blast qed qed qed lemma mcomplete_Union: assumes "finite \" and "\A. A \ \ \ Submetric M d A" "\A. A \ \ \ Metric_space.mcomplete A d" shows "Submetric M d (\\)" "Metric_space.mcomplete (\\) d" using assms by (induction rule: finite_induct) (auto simp: mcomplete_Un) lemma mcomplete_Inter: assumes "finite \" "\ \ {}" and sub: "\A. A \ \ \ Submetric M d A" and comp: "\A. A \ \ \ Metric_space.mcomplete A d" shows "Submetric M d (\\)" "Metric_space.mcomplete (\\) d" proof - show "Submetric M d (\\)" using assms unfolding Submetric_def Submetric_axioms_def by (metis Inter_lower equals0I inf.orderE le_inf_iff) then interpret MS: Submetric M d "\\" by (meson Submetric.subset subspace) show "Metric_space.mcomplete (\\) d" unfolding MS.sub.mcomplete_def proof (intro strip) fix \ assume "MS.sub.MCauchy \" then have "range \ \ \\" using MS.MCauchy_submetric by blast obtain A where "A \ \" and A: "Metric_space.mcomplete A d" using assms by blast then have "range \ \ A" using \range \ \ \\\ by blast interpret SA: Submetric M d A by (meson \A \ \\ sub Submetric.subset subspace) have "MCauchy \" using MS.MCauchy_submetric \MS.sub.MCauchy \\ by blast then obtain x where x: "limitin SA.sub.mtopology \ x sequentially" by (metis A SA.sub.MCauchy_def SA.sub.mcomplete_alt MCauchy_def \range \ \ A\) show "\x. limitin MS.sub.mtopology \ x sequentially" apply (rule_tac x="x" in exI) unfolding MS.limitin_submetric_iff proof (intro conjI) show "x \ \ \" proof clarsimp fix U assume "U \ \" interpret SU: Submetric M d U by (meson \U \ \\ sub Submetric.subset subspace) have "range \ \ U" using \U \ \\ \range \ \ \ \\ by blast moreover have "Metric_space.mcomplete U d" by (simp add: \U \ \\ comp) ultimately obtain x' where x': "limitin SU.sub.mtopology \ x' sequentially" using MCauchy_def SU.sub.MCauchy_def SU.sub.mcomplete_alt \MCauchy \\ by meson have "x' = x" proof (intro limitin_metric_unique) show "limitin mtopology \ x' sequentially" by (meson SU.Submetric_axioms Submetric.limitin_submetric_iff x') show "limitin mtopology \ x sequentially" by (meson SA.Submetric_axioms Submetric.limitin_submetric_iff x) qed auto then show "x \ U" using SU.sub.limitin_mspace x' by blast qed show "\\<^sub>F n in sequentially. \ n \ \\" by (meson \range \ \ \ \\ always_eventually range_subsetD) show "limitin mtopology \ x sequentially" by (meson SA.Submetric_axioms Submetric.limitin_submetric_iff x) qed qed qed lemma mcomplete_Int: assumes A: "Submetric M d A" "Metric_space.mcomplete A d" and B: "Submetric M d B" "Metric_space.mcomplete B d" shows "Submetric M d (A \ B)" "Metric_space.mcomplete (A \ B) d" using mcomplete_Inter [of "{A,B}"] assms by force+ subsection\Totally bounded subsets of metric spaces\ definition mtotally_bounded where "mtotally_bounded S \ \\>0. \K. finite K \ K \ S \ S \ (\x\K. mball x \)" lemma mtotally_bounded_empty [iff]: "mtotally_bounded {}" by (simp add: mtotally_bounded_def) lemma finite_imp_mtotally_bounded: "\finite S; S \ M\ \ mtotally_bounded S" by (auto simp: mtotally_bounded_def) lemma mtotally_bounded_imp_subset: "mtotally_bounded S \ S \ M" by (force simp: mtotally_bounded_def intro!: zero_less_one) lemma mtotally_bounded_sing [simp]: "mtotally_bounded {x} \ x \ M" by (meson empty_subsetI finite.simps finite_imp_mtotally_bounded insert_subset mtotally_bounded_imp_subset) lemma mtotally_bounded_Un: assumes "mtotally_bounded S" "mtotally_bounded T" shows "mtotally_bounded (S \ T)" proof - have "\K. finite K \ K \ S \ T \ S \ T \ (\x\K. mball x e)" if "e>0" and K: "finite K \ K \ S \ S \ (\x\K. mball x e)" and L: "finite L \ L \ T \ T \ (\x\L. mball x e)" for K L e using that by (rule_tac x="K \ L" in exI) auto with assms show ?thesis unfolding mtotally_bounded_def by presburger qed lemma mtotally_bounded_Union: assumes "finite f" "\S. S \ f \ mtotally_bounded S" shows "mtotally_bounded (\f)" using assms by (induction f) (auto simp: mtotally_bounded_Un) lemma mtotally_bounded_imp_mbounded: assumes "mtotally_bounded S" shows "mbounded S" proof - obtain K where "finite K \ K \ S \ S \ (\x\K. mball x 1)" using assms by (force simp: mtotally_bounded_def) then show ?thesis by (smt (verit) finite_imageI image_iff mbounded_Union mbounded_mball mbounded_subset) qed lemma mtotally_bounded_sequentially: "mtotally_bounded S \ S \ M \ (\\::nat \ 'a. range \ \ S \ (\r. strict_mono r \ MCauchy (\ \ r)))" (is "_ \ _ \ ?rhs") proof (cases "S \ M") case True show ?thesis proof - { fix \ :: "nat \ 'a" assume L: "mtotally_bounded S" and \: "range \ \ S" have "\j > i. d (\ i) (\ j) < 3*\/2 \ infinite (\ -` mball (\ j) (\/2))" if inf: "infinite (\ -` mball (\ i) \)" and "\ > 0" for i \ proof - obtain K where "finite K" "K \ S" and K: "S \ (\x\K. mball x (\/4))" by (metis L mtotally_bounded_def \\ > 0\ zero_less_divide_iff zero_less_numeral) then have K_imp_ex: "\y. y \ S \ \x\K. d x y < \/4" by fastforce have False if "\x\K. d x (\ i) < \ + \/4 \ finite (\ -` mball x (\/4))" proof - have "\w. w \ K \ d w (\ i) < 5 * \/4 \ d w (\ j) < \/4" if "d (\ i) (\ j) < \" for j proof - obtain w where w: "d w (\ j) < \/4" "w \ K" using K_imp_ex \ by blast then have "d w (\ i) < \ + \/4" by (smt (verit, ccfv_SIG) True \K \ S\ \ rangeI subset_eq that triangle') with w show ?thesis using in_mball by auto qed then have "(\ -` mball (\ i) \) \ (\x\K. if d x (\ i) < \ + \/4 then \ -` mball x (\/4) else {})" using True \K \ S\ by force then show False using finite_subset inf \finite K\ that by fastforce qed then obtain x where "x \ K" and dxi: "d x (\ i) < \ + \/4" and infx: "infinite (\ -` mball x (\/4))" by blast then obtain j where "j \ (\ -` mball x (\/4)) - {..i}" using bounded_nat_set_is_finite by (meson Diff_infinite_finite finite_atMost) then have "j > i" and dxj: "d x (\ j) < \/4" by auto have "(\ -` mball x (\/4)) \ (\ -` mball y (\/2))" if "d x y < \/4" "y \ M" for y using that by (simp add: mball_subset Metric_space_axioms vimage_mono) then have infj: "infinite (\ -` mball (\ j) (\/2))" by (meson True \d x (\ j) < \/4\ \ in_mono infx rangeI finite_subset) have "\ i \ M" "\ j \ M" "x \ M" using True \K \ S\ \x \ K\ \ by force+ then have "d (\ i) (\ j) \ d x (\ i) + d x (\ j)" using triangle'' by blast also have "\ < 3*\/2" using dxi dxj by auto finally have "d (\ i) (\ j) < 3*\/2" . with \i < j\ infj show ?thesis by blast qed then obtain nxt where nxt: "\i \. \\ > 0; infinite (\ -` mball (\ i) \)\ \ nxt i \ > i \ d (\ i) (\ (nxt i \)) < 3*\/2 \ infinite (\ -` mball (\ (nxt i \)) (\/2))" by metis have "mbounded S" using L by (simp add: mtotally_bounded_imp_mbounded) then obtain B where B: "\y \ S. d (\ 0) y \ B" and "B > 0" by (meson \ mbounded_alt_pos range_subsetD) define eps where "eps \ \n. (B+1) / 2^n" have [simp]: "eps (Suc n) = eps n / 2" "eps n > 0" for n using \B > 0\ by (auto simp: eps_def) have "UNIV \ \ -` mball (\ 0) (B+1)" using B True \ unfolding image_iff subset_iff by (smt (verit, best) UNIV_I in_mball vimageI) then have inf0: "infinite (\ -` mball (\ 0) (eps 0))" using finite_subset by (auto simp: eps_def) define r where "r \ rec_nat 0 (\n rec. nxt rec (eps n))" have [simp]: "r 0 = 0" "r (Suc n) = nxt (r n) (eps n)" for n by (auto simp: r_def) have \rM[simp]: "\ (r n) \ M" for n using True \ by blast have inf: "infinite (\ -` mball (\ (r n)) (eps n))" for n proof (induction n) case 0 then show ?case by (simp add: inf0) next case (Suc n) then show ?case using nxt [of "eps n" "r n"] by simp qed then have "r (Suc n) > r n" for n by (simp add: nxt) then have "strict_mono r" by (simp add: strict_mono_Suc_iff) have d_less: "d (\ (r n)) (\ (r (Suc n))) < 3 * eps n / 2" for n using nxt [OF _ inf] by simp have eps_plus: "eps (k + n) = eps n * (1/2)^k" for k n by (simp add: eps_def power_add field_simps) have *: "d (\ (r n)) (\ (r (k + n))) < 3 * eps n" for n k proof - have "d (\ (r n)) (\ (r (k+n))) \ 3/2 * eps n * (\i (r n)) (\ (r (Suc k + n))) \ d (\ (r n)) (\ (r (k + n))) + d (\ (r (k + n))) (\ (r (Suc (k + n))))" by (metis \rM add.commute add_Suc_right triangle) with d_less[of "k+n"] Suc show ?case by (simp add: algebra_simps eps_plus) qed also have "\ < 3/2 * eps n * 2" using geometric_sum [of "1/2::real" k] by simp finally show ?thesis by simp qed have "\N. \n\N. \n'\N. d (\ (r n)) (\ (r n')) < \" if "\ > 0" for \ proof - define N where "N \ nat \(log 2 (6*(B+1) / \))\" have \
: "b \ 2 ^ nat \log 2 b\" for b by (smt (verit) less_log_of_power real_nat_ceiling_ge) have N: "6 * eps N \ \" using \
[of "(6*(B+1) / \)"] that by (auto simp: N_def eps_def field_simps) have "d (\ (r N)) (\ (r n)) < 3 * eps N" if "n \ N" for n by (metis * add.commute nat_le_iff_add that) then have "\n\N. \n'\N. d (\ (r n)) (\ (r n')) < 3 * eps N + 3 * eps N" by (smt (verit, best) \rM triangle'') with N show ?thesis by fastforce qed then have "MCauchy (\ \ r)" unfolding MCauchy_def using True \ by auto then have "\r. strict_mono r \ MCauchy (\ \ r)" using \strict_mono r\ by blast } moreover { assume R: ?rhs have "mtotally_bounded S" unfolding mtotally_bounded_def proof (intro strip) fix \ :: real assume "\ > 0" have False if \
: "\K. \finite K; K \ S\ \ \s\S. s \ (\x\K. mball x \)" proof - obtain f where f: "\K. \finite K; K \ S\ \ f K \ S \ f K \ (\x\K. mball x \)" using \
by metis define \ where "\ \ wfrec less_than (\seq n. f (seq ` {.._eq: "\ n = f (\ ` {.._def]) have [simp]: "\ n \ S" for n using wf_less_than proof (induction n rule: wf_induct_rule) case (less n) with f show ?case by (auto simp: \_eq [of n]) qed then have "range \ \ S" by blast have \: "p < n \ \ \ d (\ p) (\ n)" for n p using f[of "\ ` {.._eq [of n] Ball_def) then obtain r where "strict_mono r" "MCauchy (\ \ r)" by (meson R \range \ \ S\) with \0 < \\ obtain N where N: "\n n'. \n\N; n'\N\ \ d (\ (r n)) (\ (r n')) < \" by (force simp: MCauchy_def) show ?thesis using N [of N "Suc (r N)"] \strict_mono r\ by (smt (verit) Suc_le_eq \ le_SucI order_refl strict_mono_imp_increasing) qed then show "\K. finite K \ K \ S \ S \ (\x\K. mball x \)" by blast qed } ultimately show ?thesis using True by blast qed qed (use mtotally_bounded_imp_subset in auto) lemma mtotally_bounded_subset: "\mtotally_bounded S; T \ S\ \ mtotally_bounded T" by (meson mtotally_bounded_sequentially order_trans) lemma mtotally_bounded_submetric: assumes "mtotally_bounded S" "S \ T" "T \ M" shows "Metric_space.mtotally_bounded T d S" proof - interpret Submetric M d T by (simp add: Metric_space_axioms assms Submetric.intro Submetric_axioms.intro) show ?thesis using assms unfolding sub.mtotally_bounded_def mtotally_bounded_def by (force simp: subset_iff elim!: all_forward ex_forward) qed lemma mtotally_bounded_absolute: "mtotally_bounded S \ S \ M \ Metric_space.mtotally_bounded S d S " proof - have "mtotally_bounded S" if "S \ M" "Metric_space.mtotally_bounded S d S" proof - interpret Submetric M d S by (simp add: Metric_space_axioms Submetric_axioms.intro Submetric_def \S \ M\) show ?thesis using that by (metis MCauchy_submetric Metric_space.mtotally_bounded_sequentially Metric_space_axioms subspace) qed moreover have "mtotally_bounded S \ Metric_space.mtotally_bounded S d S" by (simp add: mtotally_bounded_imp_subset mtotally_bounded_submetric) ultimately show ?thesis using mtotally_bounded_imp_subset by blast qed lemma mtotally_bounded_closure_of: assumes "mtotally_bounded S" shows "mtotally_bounded (mtopology closure_of S)" proof - have "S \ M" by (simp add: assms mtotally_bounded_imp_subset) have "mtotally_bounded(mtopology closure_of S)" unfolding mtotally_bounded_def proof (intro strip) fix \::real assume "\ > 0" then obtain K where "finite K" "K \ S" and K: "S \ (\x\K. mball x (\/2))" by (metis assms mtotally_bounded_def half_gt_zero) have "mtopology closure_of S \ (\x\K. mball x \)" unfolding metric_closure_of proof clarsimp fix x assume "x \ M" and x: "\r>0. \y\S. y \ M \ d x y < r" then obtain y where "y \ S" and y: "d x y < \/2" using \0 < \\ half_gt_zero by blast then obtain x' where "x' \ K" "y \ mball x' (\/2)" using K by auto then have "d x' x < \/2 + \/2" using triangle y \x \ M\ commute by fastforce then show "\x'\K. x' \ M \ d x' x < \" using \K \ S\ \S \ M\ \x' \ K\ by force qed then show "\K. finite K \ K \ mtopology closure_of S \ mtopology closure_of S \ (\x\K. mball x \)" using closure_of_subset_Int \K \ S\ \finite K\ K by fastforce qed then show ?thesis by (simp add: assms inf.absorb2 mtotally_bounded_imp_subset) qed lemma mtotally_bounded_closure_of_eq: "S \ M \ mtotally_bounded (mtopology closure_of S) \ mtotally_bounded S" by (metis closure_of_subset mtotally_bounded_closure_of mtotally_bounded_subset topspace_mtopology) lemma mtotally_bounded_cauchy_sequence: assumes "MCauchy \" shows "mtotally_bounded (range \)" unfolding MCauchy_def mtotally_bounded_def proof (intro strip) fix \::real assume "\ > 0" then obtain N where "\n. N \ n \ d (\ N) (\ n) < \" using assms by (force simp: MCauchy_def) then have "\m. \n\N. \ n \ M \ \ m \ M \ d (\ n) (\ m) < \" by (metis MCauchy_def assms mdist_zero nle_le range_subsetD) then show "\K. finite K \ K \ range \ \ range \ \ (\x\K. mball x \)" by (rule_tac x="\ ` {0..N}" in exI) force qed lemma MCauchy_imp_mbounded: "MCauchy \ \ mbounded (range \)" by (simp add: mtotally_bounded_cauchy_sequence mtotally_bounded_imp_mbounded) subsection\Compactness in metric spaces\ lemma Bolzano_Weierstrass_property: assumes "S \ U" "S \ M" shows "(\\::nat\'a. range \ \ S \ (\l r. l \ U \ strict_mono r \ limitin mtopology (\ \ r) l sequentially)) \ (\T. T \ S \ infinite T \ U \ mtopology derived_set_of T \ {})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof clarify fix T assume "T \ S" and "infinite T" and T: "U \ mtopology derived_set_of T = {}" then obtain \ :: "nat\'a" where "inj \" "range \ \ T" by (meson infinite_countable_subset) with L obtain l r where "l \ U" "strict_mono r" and lr: "limitin mtopology (\ \ r) l sequentially" by (meson \T \ S\ subset_trans) then obtain \ where "\ > 0" and \: "\y. y \ T \ y = l \ \ d l y < \" using T \T \ S\ \S \ M\ by (force simp: metric_derived_set_of limitin_metric disjoint_iff) with lr have "\\<^sub>F n in sequentially. \ (r n) \ M \ d (\ (r n)) l < \" by (auto simp: limitin_metric) then obtain N where N: "d (\ (r N)) l < \" "d (\ (r (Suc N))) l < \" using less_le_not_le by (auto simp: eventually_sequentially) moreover have "\ (r N) \ l \ \ (r (Suc N)) \ l" by (meson \inj \\ \strict_mono r\ injD n_not_Suc_n strict_mono_eq) ultimately show False using \ \range \ \ T\ commute by fastforce qed next assume R: ?rhs show ?lhs proof (intro strip) fix \ :: "nat \ 'a" assume "range \ \ S" show "\l r. l \ U \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" proof (cases "finite (range \)") case True then obtain m where "infinite (\ -` {\ m})" by (metis image_iff inf_img_fin_dom nat_not_finite) then obtain r where [iff]: "strict_mono r" and r: "\n::nat. r n \ \ -` {\ m}" using infinite_enumerate by blast have [iff]: "\ m \ U" "\ m \ M" using \range \ \ S\ assms by blast+ show ?thesis proof (intro conjI exI) show "limitin mtopology (\ \ r) (\ m) sequentially" using r by (simp add: limitin_metric) qed auto next case False then obtain l where "l \ U" and l: "l \ mtopology derived_set_of (range \)" by (meson R \range \ \ S\ disjoint_iff) then obtain g where g: "\\. \>0 \ \ (g \) \ l \ d l (\ (g \)) < \" by (simp add: metric_derived_set_of) metis have "range \ \ M" using \range \ \ S\ assms by auto have "l \ M" using l metric_derived_set_of by auto define E where \\a construction to ensure monotonicity\ "E \ \rec n. insert (inverse (Suc n)) ((\i. d l (\ i)) ` (\k wfrec less_than (\rec n. g (Min (E rec n)))" have "(\kk (r n)) > 0" for n using wf_less_than proof (induction n rule: wf_induct_rule) case (less n) then have *: "Min (E r n) > 0" using \l \ M\ \range \ \ M\ by (auto simp: E_def image_subset_iff) show ?case using g [OF *] r_eq [of n] by (metis \l \ M\ \range \ \ M\ mdist_pos_less range_subsetD) qed then have non_l: "\ (r n) \ l" for n using \range \ \ M\ mdist_pos_eq by blast have Min_pos: "Min (E r n) > 0" for n using dl_pos \l \ M\ \range \ \ M\ by (auto simp: E_def image_subset_iff) have d_small: "d (\(r n)) l < inverse(Suc n)" for n proof - have "d (\(r n)) l < Min (E r n)" by (simp add: \0 < Min (E r n)\ commute g r_eq) also have "... \ inverse(Suc n)" by (simp add: E_def) finally show ?thesis . qed have d_lt_d: "d l (\ (r n)) < d l (\ i)" if \
: "p < n" "i \ r p" "\ i \ l" for i p n proof - have 1: "d l (\ i) \ E r n" using \
\l \ M\ \range \ \ M\ by (force simp: E_def image_subset_iff image_iff) have "d l (\ (g (Min (E r n)))) < Min (E r n)" by (rule conjunct2 [OF g [OF Min_pos]]) also have "Min (E r n) \ d l (\ i)" using 1 unfolding E_def by (force intro!: Min.coboundedI) finally show ?thesis by (simp add: r_eq) qed have r: "r p < r n" if "p < n" for p n using d_lt_d [OF that] non_l by (meson linorder_not_le order_less_irrefl) show ?thesis proof (intro exI conjI) show "strict_mono r" by (simp add: r strict_monoI) show "limitin mtopology (\ \ r) l sequentially" unfolding limitin_metric proof (intro conjI strip \l \ M\) fix \ :: real assume "\ > 0" then have "\\<^sub>F n in sequentially. inverse(Suc n) < \" using Archimedean_eventually_inverse by auto then show "\\<^sub>F n in sequentially. (\ \ r) n \ M \ d ((\ \ r) n) l < \" by (smt (verit) \range \ \ M\ commute comp_apply d_small eventually_mono range_subsetD) qed qed (use \l \ U\ in auto) qed qed qed subsubsection \More on Bolzano Weierstrass\ lemma Bolzano_Weierstrass_A: assumes "compactin mtopology S" "T \ S" "infinite T" shows "S \ mtopology derived_set_of T \ {}" by (simp add: assms compactin_imp_Bolzano_Weierstrass) lemma Bolzano_Weierstrass_B: fixes \ :: "nat \ 'a" assumes "S \ M" "range \ \ S" and "\T. \T \ S \ infinite T\ \ S \ mtopology derived_set_of T \ {}" shows "\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" using Bolzano_Weierstrass_property assms by blast lemma Bolzano_Weierstrass_C: assumes "S \ M" assumes "\\:: nat \ 'a. range \ \ S \ (\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially)" shows "mtotally_bounded S" unfolding mtotally_bounded_sequentially by (metis convergent_imp_MCauchy assms image_comp image_mono subset_UNIV subset_trans) lemma Bolzano_Weierstrass_D: assumes "S \ M" "S \ \\" and opeU: "\U. U \ \ \ openin mtopology U" assumes \
: "(\\::nat\'a. range \ \ S \ (\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially))" shows "\\>0. \x \ S. \U \ \. mball x \ \ U" proof (rule ccontr) assume "\ (\\>0. \x \ S. \U \ \. mball x \ \ U)" then have "\n. \x\S. \U\\. \ mball x (inverse (Suc n)) \ U" by simp then obtain \ where "\n. \ n \ S" and \: "\n U. U \ \ \ \ mball (\ n) (inverse (Suc n)) \ U" by metis then obtain l r where "l \ S" "strict_mono r" and lr: "limitin mtopology (\ \ r) l sequentially" by (meson \
image_subsetI) with \S \ \\\ obtain B where "l \ B" "B \ \" by auto then obtain \ where "\ > 0" and \: "\z. \z \ M; d z l < \\ \ z \ B" by (metis opeU [OF \B \ \\] commute in_mball openin_mtopology subset_iff) then have "\\<^sub>F n in sequentially. \ (r n) \ M \ d (\ (r n)) l < \/2" using lr half_gt_zero unfolding limitin_metric o_def by blast moreover have "\\<^sub>F n in sequentially. inverse (real (Suc n)) < \/2" using Archimedean_eventually_inverse \0 < \\ half_gt_zero by blast ultimately obtain n where n: "d (\ (r n)) l < \/2" "inverse (real (Suc n)) < \/2" by (smt (verit, del_insts) eventually_sequentially le_add1 le_add2) have "x \ B" if "d (\ (r n)) x < inverse (Suc(r n))" "x \ M" for x proof - have rle: "inverse (real (Suc (r n))) \ inverse (real (Suc n))" using \strict_mono r\ strict_mono_imp_increasing by auto have "d x l \ d (\ (r n)) x + d (\ (r n)) l" using that by (metis triangle \\n. \ n \ S\ \l \ S\ \S \ M\ commute subsetD) also have "... < \" using that n rle by linarith finally show ?thesis by (simp add: \ that) qed then show False using \ [of B "r n"] by (simp add: \B \ \\ subset_iff) qed lemma Bolzano_Weierstrass_E: assumes "mtotally_bounded S" "S \ M" and S: "\\. \\U. U \ \ \ openin mtopology U; S \ \\\ \ \\>0. \x \ S. \U \ \. mball x \ \ U" shows "compactin mtopology S" proof (clarsimp simp: compactin_def assms) fix \ :: "'a set set" assume \: "\x\\. openin mtopology x" and "S \ \\" then obtain \ where "\>0" and \: "\x. x \ S \ \U \ \. mball x \ \ U" by (metis S) then obtain f where f: "\x. x \ S \ f x \ \ \ mball x \ \ f x" by metis then obtain K where "finite K" "K \ S" and K: "S \ (\x\K. mball x \)" by (metis \0 < \\ \mtotally_bounded S\ mtotally_bounded_def) show "\\. finite \ \ \ \ \ \ S \ \\" proof (intro conjI exI) show "finite (f ` K)" by (simp add: \finite K\) show "f ` K \ \" using \K \ S\ f by blast show "S \ \(f ` K)" using K \K \ S\ by (force dest: f) qed qed lemma compactin_eq_Bolzano_Weierstrass: "compactin mtopology S \ S \ M \ (\T. T \ S \ infinite T \ S \ mtopology derived_set_of T \ {})" using Bolzano_Weierstrass_C Bolzano_Weierstrass_D Bolzano_Weierstrass_E by (smt (verit, del_insts) Bolzano_Weierstrass_property compactin_imp_Bolzano_Weierstrass compactin_subspace subset_refl topspace_mtopology) lemma compactin_sequentially: shows "compactin mtopology S \ S \ M \ ((\\::nat\'a. range \ \ S \ (\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially)))" by (metis Bolzano_Weierstrass_property compactin_eq_Bolzano_Weierstrass subset_refl) lemma compactin_imp_mtotally_bounded: "compactin mtopology S \ mtotally_bounded S" by (simp add: Bolzano_Weierstrass_C compactin_sequentially) lemma lebesgue_number: "\compactin mtopology S; S \ \\; \U. U \ \ \ openin mtopology U\ \ \\>0. \x \ S. \U \ \. mball x \ \ U" by (simp add: Bolzano_Weierstrass_D compactin_sequentially) lemma compact_space_sequentially: "compact_space mtopology \ (\\::nat\'a. range \ \ M \ (\l r. l \ M \ strict_mono r \ limitin mtopology (\ \ r) l sequentially))" by (simp add: compact_space_def compactin_sequentially) lemma compact_space_eq_Bolzano_Weierstrass: "compact_space mtopology \ (\S. S \ M \ infinite S \ mtopology derived_set_of S \ {})" using Int_absorb1 [OF derived_set_of_subset_topspace [of mtopology]] by (force simp: compact_space_def compactin_eq_Bolzano_Weierstrass) lemma compact_space_nest: "compact_space mtopology \ (\C. (\n::nat. closedin mtopology (C n)) \ (\n. C n \ {}) \ decseq C \ \(range C) \ {})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof clarify fix C :: "nat \ 'a set" assume "\n. closedin mtopology (C n)" and "\n. C n \ {}" and "decseq C" and "\ (range C) = {}" then obtain K where K: "finite K" "\(C ` K) = {}" by (metis L compact_space_imp_nest) then obtain k where "K \ {..k}" using finite_nat_iff_bounded_le by auto then have "C k \ \(C ` K)" using \decseq C\ by (auto simp:decseq_def) then show False by (simp add: K \\n. C n \ {}\) qed next assume R [rule_format]: ?rhs show ?lhs unfolding compact_space_sequentially proof (intro strip) fix \ :: "nat \ 'a" assume \: "range \ \ M" have "mtopology closure_of \ ` {n..} \ {}" for n using \range \ \ M\ by (auto simp: closure_of_eq_empty image_subset_iff) moreover have "decseq (\n. mtopology closure_of \ ` {n..})" using closure_of_mono image_mono by (smt (verit) atLeast_subset_iff decseq_def) ultimately obtain l where l: "\n. l \ mtopology closure_of \ ` {n..}" using R [of "\n. mtopology closure_of (\ ` {n..})"] by auto then have "l \ M" and "\n. \r>0. \k\n. \ k \ M \ d l (\ k) < r" using metric_closure_of by fastforce+ then obtain f where f: "\n r. r>0 \ f n r \ n \ \ (f n r) \ M \ d l (\ (f n r)) < r" by metis define r where "r = rec_nat (f 0 1) (\n rec. (f (Suc rec) (inverse (Suc (Suc n)))))" have r: "d l (\(r n)) < inverse(Suc n)" for n by (induction n) (auto simp: rec_nat_0_imp [OF r_def] rec_nat_Suc_imp [OF r_def] f) have "r n < r(Suc n)" for n by (simp add: Suc_le_lessD f r_def) then have "strict_mono r" by (simp add: strict_mono_Suc_iff) moreover have "limitin mtopology (\ \ r) l sequentially" proof (clarsimp simp: limitin_metric \l \ M\) fix \ :: real assume "\ > 0" then have "(\\<^sub>F n in sequentially. inverse (real (Suc n)) < \)" using Archimedean_eventually_inverse by blast then show "\\<^sub>F n in sequentially. \ (r n) \ M \ d (\ (r n)) l < \" by eventually_elim (metis commute \range \ \ M\ order_less_trans r range_subsetD) qed ultimately show "\l r. l \ M \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" using \l \ M\ by blast qed qed lemma (in discrete_metric) mcomplete_discrete_metric: "disc.mcomplete" proof (clarsimp simp: disc.mcomplete_def) fix \ :: "nat \ 'a" assume "disc.MCauchy \" then obtain N where "\n. N \ n \ \ N = \ n" unfolding disc.MCauchy_def by (metis dd_def dual_order.refl order_less_irrefl zero_less_one) moreover have "range \ \ M" using \disc.MCauchy \\ disc.MCauchy_def by blast ultimately have "limitin disc.mtopology \ (\ N) sequentially" by (metis disc.limit_metric_sequentially disc.zero range_subsetD) then show "\x. limitin disc.mtopology \ x sequentially" .. qed lemma compact_space_imp_mcomplete: "compact_space mtopology \ mcomplete" by (simp add: compact_space_nest mcomplete_nest) lemma (in Submetric) compactin_imp_mcomplete: "compactin mtopology A \ sub.mcomplete" by (simp add: compactin_subspace mtopology_submetric sub.compact_space_imp_mcomplete) lemma (in Submetric) mcomplete_imp_closedin: assumes "sub.mcomplete" shows "closedin mtopology A" proof - have "l \ A" if "range \ \ A" and l: "limitin mtopology \ l sequentially" for \ :: "nat \ 'a" and l proof - have "sub.MCauchy \" using convergent_imp_MCauchy subset that by (force simp: MCauchy_submetric) then have "limitin sub.mtopology \ l sequentially" using assms unfolding sub.mcomplete_def using l limitin_metric_unique limitin_submetric_iff trivial_limit_sequentially by blast then show ?thesis using limitin_submetric_iff by blast qed then show ?thesis using metric_closedin_iff_sequentially_closed subset by auto qed lemma (in Submetric) closedin_eq_mcomplete: "mcomplete \ (closedin mtopology A \ sub.mcomplete)" using closedin_mcomplete_imp_mcomplete mcomplete_imp_closedin by blast lemma compact_space_eq_mcomplete_mtotally_bounded: "compact_space mtopology \ mcomplete \ mtotally_bounded M" by (meson Bolzano_Weierstrass_C compact_space_imp_mcomplete compact_space_sequentially limitin_mspace mcomplete_alt mtotally_bounded_sequentially subset_refl) lemma compact_closure_of_imp_mtotally_bounded: "\compactin mtopology (mtopology closure_of S); S \ M\ \ mtotally_bounded S" using compactin_imp_mtotally_bounded mtotally_bounded_closure_of_eq by blast lemma mtotally_bounded_eq_compact_closure_of: assumes "mcomplete" shows "mtotally_bounded S \ S \ M \ compactin mtopology (mtopology closure_of S)" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding compactin_subspace proof (intro conjI) show "S \ M" using L by (simp add: mtotally_bounded_imp_subset) show "mtopology closure_of S \ topspace mtopology" by (simp add: \S \ M\ closure_of_minimal) then have MSM: "mtopology closure_of S \ M" by auto interpret S: Submetric M d "mtopology closure_of S" proof qed (use MSM in auto) have "S.sub.mtotally_bounded (mtopology closure_of S)" using L mtotally_bounded_absolute mtotally_bounded_closure_of by blast then show "compact_space (subtopology mtopology (mtopology closure_of S))" using S.closedin_mcomplete_imp_mcomplete S.mtopology_submetric S.sub.compact_space_eq_mcomplete_mtotally_bounded assms by force qed qed (auto simp: compact_closure_of_imp_mtotally_bounded) lemma compact_closure_of_eq_Bolzano_Weierstrass: "compactin mtopology (mtopology closure_of S) \ (\T. infinite T \ T \ S \ T \ M \ mtopology derived_set_of T \ {})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro strip) fix T assume T: "infinite T \ T \ S \ T \ M" show "mtopology derived_set_of T \ {}" proof (intro compact_closure_of_imp_Bolzano_Weierstrass) show "compactin mtopology (mtopology closure_of S)" by (simp add: L) qed (use T in auto) qed next have "compactin mtopology (mtopology closure_of S)" if \
: "\T. \infinite T; T \ S\ \ mtopology derived_set_of T \ {}" and "S \ M" for S unfolding compactin_sequentially proof (intro conjI strip) show MSM: "mtopology closure_of S \ M" using closure_of_subset_topspace by fastforce fix \ :: "nat \ 'a" assume \: "range \ \ mtopology closure_of S" then have "\y \ S. d (\ n) y < inverse(Suc n)" for n by (simp add: metric_closure_of image_subset_iff) (metis inverse_Suc of_nat_Suc) then obtain \ where \: "\n. \ n \ S \ d (\ n) (\ n) < inverse(Suc n)" by metis then have "range \ \ S" by blast moreover have *: "\T. T \ S \ infinite T \ mtopology closure_of S \ mtopology derived_set_of T \ {}" using "\
"(1) derived_set_of_mono derived_set_of_subset_closure_of by fastforce moreover have "S \ mtopology closure_of S" by (simp add: \S \ M\ closure_of_subset) ultimately obtain l r where lr: "l \ mtopology closure_of S" "strict_mono r" "limitin mtopology (\ \ r) l sequentially" using Bolzano_Weierstrass_property \S \ M\ by metis then have "l \ M" using limitin_mspace by blast have dr_less: "d ((\ \ r) n) ((\ \ r) n) < inverse(Suc n)" for n proof - have "d ((\ \ r) n) ((\ \ r) n) < inverse(Suc (r n))" using \ by auto also have "... \ inverse(Suc n)" using lr strict_mono_imp_increasing by auto finally show ?thesis . qed have "limitin mtopology (\ \ r) l sequentially" unfolding limitin_metric proof (intro conjI strip) show "l \ M" using limitin_mspace lr by blast fix \ :: real assume "\ > 0" then have "\\<^sub>F n in sequentially. (\ \ r) n \ M \ d ((\ \ r) n) l < \/2" using lr half_gt_zero limitin_metric by blast moreover have "\\<^sub>F n in sequentially. inverse (real (Suc n)) < \/2" using Archimedean_eventually_inverse \0 < \\ half_gt_zero by blast then have "\\<^sub>F n in sequentially. d ((\ \ r) n) ((\ \ r) n) < \/2" by eventually_elim (smt (verit, del_insts) dr_less) ultimately have "\\<^sub>F n in sequentially. d ((\ \ r) n) l < \/2 + \/2" by eventually_elim (smt (verit) triangle \l \ M\ MSM \ comp_apply order_trans range_subsetD) then show "\\<^sub>F n in sequentially. (\ \ r) n \ M \ d ((\ \ r) n) l < \" apply eventually_elim using \mtopology closure_of S \ M\ \ by auto qed with lr show "\l r. l \ mtopology closure_of S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" by blast qed then show "?rhs \ ?lhs" by (metis Int_subset_iff closure_of_restrict inf_le1 topspace_mtopology) qed end lemma (in discrete_metric) mtotally_bounded_discrete_metric: "disc.mtotally_bounded S \ finite S \ S \ M" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof show "finite S" by (metis (no_types) L closure_of_subset_Int compactin_discrete_topology disc.mtotally_bounded_eq_compact_closure_of disc.topspace_mtopology discrete_metric.mcomplete_discrete_metric inf.absorb_iff2 mtopology_discrete_metric finite_subset) show "S \ M" by (simp add: L disc.mtotally_bounded_imp_subset) qed qed (simp add: disc.finite_imp_mtotally_bounded) context Metric_space begin lemma derived_set_of_infinite_openin_metric: "mtopology derived_set_of S = {x \ M. \U. x \ U \ openin mtopology U \ infinite(S \ U)}" by (simp add: derived_set_of_infinite_openin Hausdorff_space_mtopology) lemma derived_set_of_infinite_1: assumes "infinite (S \ mball x \)" shows "infinite (S \ mcball x \)" by (meson Int_mono assms finite_subset mball_subset_mcball subset_refl) lemma derived_set_of_infinite_2: assumes "openin mtopology U" "\\. 0 < \ \ infinite (S \ mcball x \)" and "x \ U" shows "infinite (S \ U)" by (metis assms openin_mtopology_mcball finite_Int inf.absorb_iff2 inf_assoc) lemma derived_set_of_infinite_mball: "mtopology derived_set_of S = {x \ M. \e>0. infinite(S \ mball x e)}" unfolding derived_set_of_infinite_openin_metric by (meson centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2) lemma derived_set_of_infinite_mcball: "mtopology derived_set_of S = {x \ M. \e>0. infinite(S \ mcball x e)}" unfolding derived_set_of_infinite_openin_metric by (meson centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2) end subsection\Continuous functions on metric spaces\ context Metric_space begin lemma continuous_map_to_metric: "continuous_map X mtopology f \ (\x \ topspace X. \\>0. \U. openin X U \ x \ U \ (\y\U. f y \ mball (f x) \))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" unfolding continuous_map_eq_topcontinuous_at topcontinuous_at_def by (metis centre_in_mball_iff openin_mball topspace_mtopology) next assume R: ?rhs then have "\x\topspace X. f x \ M" by (meson gt_ex in_mball) moreover have "\x V. \x \ topspace X; openin mtopology V; f x \ V\ \ \U. openin X U \ x \ U \ (\y\U. f y \ V)" unfolding openin_mtopology by (metis Int_iff R inf.orderE) ultimately show ?lhs by (simp add: continuous_map_eq_topcontinuous_at topcontinuous_at_def) qed lemma continuous_map_from_metric: "continuous_map mtopology X f \ f ` M \ topspace X \ (\a \ M. \U. openin X U \ f a \ U \ (\r>0. \x. x \ M \ d a x < r \ f x \ U))" proof (cases "f ` M \ topspace X") case True then show ?thesis by (fastforce simp: continuous_map openin_mtopology subset_eq) next case False then show ?thesis by (simp add: continuous_map_eq_image_closure_subset_gen) qed text \An abstract formulation, since the limits do not have to be sequential\ lemma continuous_map_uniform_limit: assumes contf: "\\<^sub>F \ in F. continuous_map X mtopology (f \)" and dfg: "\\. 0 < \ \ \\<^sub>F \ in F. \x \ topspace X. g x \ M \ d (f \ x) (g x) < \" and nontriv: "\ trivial_limit F" shows "continuous_map X mtopology g" unfolding continuous_map_to_metric proof (intro strip) fix x and \::real assume "x \ topspace X" and "\ > 0" then obtain \ where k: "continuous_map X mtopology (f \)" and gM: "\x \ topspace X. g x \ M" and third: "\x \ topspace X. d (f \ x) (g x) < \/3" using eventually_conj [OF contf] contf dfg [of "\/3"] eventually_happens' [OF nontriv] by (smt (verit, ccfv_SIG) zero_less_divide_iff) then obtain U where U: "openin X U" "x \ U" and Uthird: "\y\U. d (f \ y) (f \ x) < \/3" unfolding continuous_map_to_metric by (metis \0 < \\ \x \ topspace X\ commute divide_pos_pos in_mball zero_less_numeral) have f_inM: "f \ y \ M" if "y\U" for y using U k openin_subset that by (fastforce simp: continuous_map_def) have "d (g y) (g x) < \" if "y\U" for y proof - have "g y \ M" using U gM openin_subset that by blast have "d (g y) (g x) \ d (g y) (f \ x) + d (f \ x) (g x)" by (simp add: U \g y \ M\ \x \ topspace X\ f_inM gM triangle) also have "\ \ d (g y) (f \ y) + d (f \ y) (f \ x) + d (f \ x) (g x)" by (simp add: U \g y \ M\ commute f_inM that triangle') also have "\ < \/3 + \/3 + \/3" by (smt (verit) U(1) Uthird \x \ topspace X\ commute openin_subset subsetD that third) finally show ?thesis by simp qed with U gM show "\U. openin X U \ x \ U \ (\y\U. g y \ mball (g x) \)" by (metis commute in_mball in_mono openin_subset) qed lemma continuous_map_uniform_limit_alt: assumes contf: "\\<^sub>F \ in F. continuous_map X mtopology (f \)" and gim: "g ` (topspace X) \ M" and dfg: "\\. 0 < \ \ \\<^sub>F \ in F. \x \ topspace X. d (f \ x) (g x) < \" and nontriv: "\ trivial_limit F" shows "continuous_map X mtopology g" proof (rule continuous_map_uniform_limit [OF contf]) fix \ :: real assume "\ > 0" with gim dfg show "\\<^sub>F \ in F. \x\topspace X. g x \ M \ d (f \ x) (g x) < \" by (simp add: image_subset_iff) qed (use nontriv in auto) lemma continuous_map_uniformly_Cauchy_limit: assumes "mcomplete" assumes contf: "\\<^sub>F n in sequentially. continuous_map X mtopology (f n)" and Cauchy': "\\. \ > 0 \ \N. \m n x. N \ m \ N \ n \ x \ topspace X \ d (f m x) (f n x) < \" obtains g where "continuous_map X mtopology g" "\\. 0 < \ \ \\<^sub>F n in sequentially. \x\topspace X. d (f n x) (g x) < \" proof - have "\x. x \ topspace X \ \l. limitin mtopology (\n. f n x) l sequentially" using \mcomplete\ [unfolded mcomplete, rule_format] assms by (smt (verit) contf continuous_map_def eventually_mono topspace_mtopology) then obtain g where g: "\x. x \ topspace X \ limitin mtopology (\n. f n x) (g x) sequentially" by metis show thesis proof show "\\<^sub>F n in sequentially. \x\topspace X. d (f n x) (g x) < \" if "\ > 0" for \ :: real proof - obtain N where N: "\m n x. \N \ m; N \ n; x \ topspace X\ \ d (f m x) (f n x) < \/2" by (meson Cauchy' \0 < \\ half_gt_zero) obtain P where P: "\n x. \n \ P; x \ topspace X\ \ f n x \ M" using contf by (auto simp: eventually_sequentially continuous_map_def) show ?thesis proof (intro eventually_sequentiallyI strip) fix n x assume "max N P \ n" and x: "x \ topspace X" obtain L where "g x \ M" and L: "\n\L. f n x \ M \ d (f n x) (g x) < \/2" using g [OF x] \\ > 0\ unfolding limitin_metric by (metis (no_types, lifting) eventually_sequentially half_gt_zero) define n' where "n' \ Max{L,N,P}" have L': "\m \ n'. f m x \ M \ d (f m x) (g x) < \/2" using L by (simp add: n'_def) moreover have "d (f n x) (f n' x) < \/2" using N [of n n' x] \max N P \ n\ n'_def x by fastforce ultimately have "d (f n x) (g x) < \/2 + \/2" by (smt (verit, ccfv_SIG) P \g x \ M\ \max N P \ n\ le_refl max.bounded_iff mdist_zero triangle' x) then show "d (f n x) (g x) < \" by simp qed qed then show "continuous_map X mtopology g" by (smt (verit, del_insts) eventually_mono g limitin_mspace trivial_limit_sequentially continuous_map_uniform_limit [OF contf]) qed qed lemma metric_continuous_map: assumes "Metric_space M' d'" shows "continuous_map mtopology (Metric_space.mtopology M' d') f \ f ` M \ M' \ (\a \ M. \\>0. \\>0. (\x. x \ M \ d a x < \ \ d' (f a) (f x) < \))" (is "?lhs = ?rhs") proof - interpret M': Metric_space M' d' by (simp add: assms) show ?thesis proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` M \ M'" using L by (auto simp: continuous_map_def) fix a and \ :: real assume "a \ M" and "\ > 0" then have "openin mtopology {x \ M. f x \ M'.mball (f a) \}" "f a \ M'" using L unfolding continuous_map_def by fastforce+ then obtain \ where "\ > 0" "mball a \ \ {x \ M. f x \ M' \ d' (f a) (f x) < \}" using \0 < \\ \a \ M\ openin_mtopology by auto then show "\\>0. \x. x \ M \ d a x < \ \ d' (f a) (f x) < \" using \a \ M\ in_mball by blast qed next assume R: ?rhs show ?lhs unfolding continuous_map_def proof (intro conjI strip) fix U assume "openin M'.mtopology U" then show "openin mtopology {x \ topspace mtopology. f x \ U}" apply (simp add: continuous_map_def openin_mtopology M'.openin_mtopology subset_iff) by (metis R image_subset_iff) qed (use R in auto) qed qed end (*Metric_space*) subsection \Completely metrizable spaces\ text \These spaces are topologically complete\ definition completely_metrizable_space where "completely_metrizable_space X \ \M d. Metric_space M d \ Metric_space.mcomplete M d \ X = Metric_space.mtopology M d" lemma empty_completely_metrizable_space: "topspace X = {} \ completely_metrizable_space X" unfolding completely_metrizable_space_def subtopology_eq_discrete_topology_empty [symmetric] by (metis Metric_space.mcomplete_empty_mspace discrete_metric.mtopology_discrete_metric metric_M_dd) lemma completely_metrizable_imp_metrizable_space: "completely_metrizable_space X \ metrizable_space X" using completely_metrizable_space_def metrizable_space_def by auto lemma (in Metric_space) completely_metrizable_space_mtopology: "mcomplete \ completely_metrizable_space mtopology" using Metric_space_axioms completely_metrizable_space_def by blast lemma completely_metrizable_space_discrete_topology: "completely_metrizable_space (discrete_topology U)" unfolding completely_metrizable_space_def by (metis discrete_metric.mcomplete_discrete_metric discrete_metric.mtopology_discrete_metric metric_M_dd) lemma completely_metrizable_space_euclideanreal: "completely_metrizable_space euclideanreal" by (metis Met_TC.mtopology_is_euclideanreal Met_TC.completely_metrizable_space_mtopology euclidean_metric) lemma completely_metrizable_space_closedin: assumes X: "completely_metrizable_space X" and S: "closedin X S" shows "completely_metrizable_space(subtopology X S)" proof - obtain M d where "Metric_space M d" and comp: "Metric_space.mcomplete M d" and Xeq: "X = Metric_space.mtopology M d" using assms completely_metrizable_space_def by blast then interpret Metric_space M d by blast show ?thesis unfolding completely_metrizable_space_def proof (intro conjI exI) show "Metric_space S d" using S Xeq closedin_subset subspace by force have sub: "Submetric_axioms M S" by (metis S Xeq closedin_metric Submetric_axioms_def) then show "Metric_space.mcomplete S d" using S Submetric.closedin_mcomplete_imp_mcomplete Submetric_def Xeq comp by blast show "subtopology X S = Metric_space.mtopology S d" by (metis Metric_space_axioms Xeq sub Submetric.intro Submetric.mtopology_submetric) qed qed lemma homeomorphic_completely_metrizable_space_aux: assumes homXY: "X homeomorphic_space Y" and X: "completely_metrizable_space X" shows "completely_metrizable_space Y" proof - obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" and fg: "(\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" and fim: "f ` (topspace X) = topspace Y" and gim: "g ` (topspace Y) = topspace X" by (smt (verit, best) homXY homeomorphic_imp_surjective_map homeomorphic_maps_map homeomorphic_space_def) obtain M d where Md: "Metric_space M d" "Metric_space.mcomplete M d" and Xeq: "X = Metric_space.mtopology M d" using X by (auto simp: completely_metrizable_space_def) then interpret MX: Metric_space M d by metis define D where "D \ \x y. d (g x) (g y)" have "Metric_space (topspace Y) D" proof show "(D x y = 0) \ (x = y)" if "x \ topspace Y" "y \ topspace Y" for x y unfolding D_def by (metis that MX.topspace_mtopology MX.zero Xeq fg gim imageI) show "D x z \ D x y +D y z" if "x \ topspace Y" "y \ topspace Y" "z \ topspace Y" for x y z using that MX.triangle Xeq gim by (auto simp: D_def) qed (auto simp: D_def MX.commute) then interpret MY: Metric_space "topspace Y" "\x y. D x y" by metis show ?thesis unfolding completely_metrizable_space_def proof (intro exI conjI) show "Metric_space (topspace Y) D" using MY.Metric_space_axioms by blast have gball: "g ` MY.mball y r = MX.mball (g y) r" if "y \ topspace Y" for y r using that MX.topspace_mtopology Xeq gim unfolding MX.mball_def MY.mball_def by (auto simp: subset_iff image_iff D_def) have "\r>0. MY.mball y r \ S" if "openin Y S" and "y \ S" for S y proof - have "openin X (g`S)" using hmg homeomorphic_map_openness_eq that by auto then obtain r where "r>0" "MX.mball (g y) r \ g`S" using MX.openin_mtopology Xeq \y \ S\ by auto then show ?thesis by (smt (verit, ccfv_SIG) MY.in_mball gball fg image_iff in_mono openin_subset subsetI that(1)) qed moreover have "openin Y S" if "S \ topspace Y" and "\y. y \ S \ \r>0. MY.mball y r \ S" for S proof - have "\x. x \ g`S \ \r>0. MX.mball x r \ g`S" by (smt (verit) gball imageE image_mono subset_iff that) then have "openin X (g`S)" using MX.openin_mtopology Xeq gim that(1) by auto then show ?thesis using hmg homeomorphic_map_openness_eq that(1) by blast qed ultimately show Yeq: "Y = MY.mtopology" unfolding topology_eq MY.openin_mtopology by (metis openin_subset) show "MY.mcomplete" unfolding MY.mcomplete_def proof (intro strip) fix \ assume \: "MY.MCauchy \" have "MX.MCauchy (g \ \)" unfolding MX.MCauchy_def proof (intro conjI strip) show "range (g \ \) \ M" using MY.MCauchy_def Xeq \ gim by auto fix \ :: real assume "\ > 0" then obtain N where "\n n'. N \ n \ N \ n' \ D (\ n) (\ n') < \" using MY.MCauchy_def \ by presburger then show "\N. \n n'. N \ n \ N \ n' \ d ((g \ \) n) ((g \ \) n') < \" by (auto simp: o_def D_def) qed then obtain x where x: "limitin MX.mtopology (g \ \) x sequentially" "x \ topspace X" using MX.limitin_mspace MX.topspace_mtopology Md Xeq unfolding MX.mcomplete_def by blast with x have "limitin MY.mtopology (f \ (g \ \)) (f x) sequentially" by (metis Xeq Yeq continuous_map_limit hmf homeomorphic_imp_continuous_map) moreover have "f \ (g \ \) = \" using \MY.MCauchy \\ by (force simp add: fg MY.MCauchy_def subset_iff) ultimately have "limitin MY.mtopology \ (f x) sequentially" by simp then show "\y. limitin MY.mtopology \ y sequentially" by blast qed qed qed lemma homeomorphic_completely_metrizable_space: "X homeomorphic_space Y \ completely_metrizable_space X \ completely_metrizable_space Y" by (meson homeomorphic_completely_metrizable_space_aux homeomorphic_space_sym) lemma completely_metrizable_space_retraction_map_image: assumes r: "retraction_map X Y r" and X: "completely_metrizable_space X" shows "completely_metrizable_space Y" proof - obtain s where s: "retraction_maps X Y r s" using r retraction_map_def by blast then have "subtopology X (s ` topspace Y) homeomorphic_space Y" using retraction_maps_section_image2 by blast then show ?thesis by (metis X retract_of_space_imp_closedin retraction_maps_section_image1 homeomorphic_completely_metrizable_space completely_metrizable_space_closedin completely_metrizable_imp_metrizable_space metrizable_imp_Hausdorff_space s) qed subsection \Product metric\ text\For the nicest fit with the main Euclidean theories, we choose the Euclidean product, though other definitions of the product work.\ definition "prod_dist \ \d1 d2 (x,y) (x',y'). sqrt(d1 x x' ^ 2 + d2 y y' ^ 2)" locale Metric_space12 = M1: Metric_space M1 d1 + M2: Metric_space M2 d2 for M1 d1 M2 d2 lemma (in Metric_space12) prod_metric: "Metric_space (M1 \ M2) (prod_dist d1 d2)" proof fix x y z assume xyz: "x \ M1 \ M2" "y \ M1 \ M2" "z \ M1 \ M2" have "sqrt ((d1 x1 z1)\<^sup>2 + (d2 x2 z2)\<^sup>2) \ sqrt ((d1 x1 y1)\<^sup>2 + (d2 x2 y2)\<^sup>2) + sqrt ((d1 y1 z1)\<^sup>2 + (d2 y2 z2)\<^sup>2)" (is "sqrt ?L \ ?R") if "x = (x1, x2)" "y = (y1, y2)" "z = (z1, z2)" for x1 x2 y1 y2 z1 z2 proof - have tri: "d1 x1 z1 \ d1 x1 y1 + d1 y1 z1" "d2 x2 z2 \ d2 x2 y2 + d2 y2 z2" using that xyz M1.triangle [of x1 y1 z1] M2.triangle [of x2 y2 z2] by auto show ?thesis proof (rule real_le_lsqrt) have "?L \ (d1 x1 y1 + d1 y1 z1)\<^sup>2 + (d2 x2 y2 + d2 y2 z2)\<^sup>2" using tri by (smt (verit) M1.nonneg M2.nonneg power_mono) also have "... \ ?R\<^sup>2" by (metis real_sqrt_sum_squares_triangle_ineq sqrt_le_D) finally show "?L \ ?R\<^sup>2" . qed auto qed then show "prod_dist d1 d2 x z \ prod_dist d1 d2 x y + prod_dist d1 d2 y z" by (simp add: prod_dist_def case_prod_unfold) qed (auto simp: M1.commute M2.commute case_prod_unfold prod_dist_def) sublocale Metric_space12 \ Prod_metric: Metric_space "M1\M2" "prod_dist d1 d2" by (simp add: prod_metric) text \For easy reference to theorems outside of the locale\ lemma Metric_space12_mspace_mdist: "Metric_space12 (mspace m1) (mdist m1) (mspace m2) (mdist m2)" by (simp add: Metric_space12_def Metric_space_mspace_mdist) definition prod_metric where "prod_metric \ \m1 m2. metric (mspace m1 \ mspace m2, prod_dist (mdist m1) (mdist m2))" lemma submetric_prod_metric: "submetric (prod_metric m1 m2) (S \ T) = prod_metric (submetric m1 S) (submetric m2 T)" apply (simp add: prod_metric_def) by (simp add: submetric_def Metric_space.mspace_metric Metric_space.mdist_metric Metric_space12.prod_metric Metric_space12_def Metric_space_mspace_mdist Times_Int_Times) lemma mspace_prod_metric [simp]:" mspace (prod_metric m1 m2) = mspace m1 \ mspace m2" by (simp add: prod_metric_def Metric_space.mspace_metric Metric_space12.prod_metric Metric_space12_mspace_mdist) lemma mdist_prod_metric [simp]: "mdist (prod_metric m1 m2) = prod_dist (mdist m1) (mdist m2)" by (metis Metric_space.mdist_metric Metric_space12.prod_metric Metric_space12_mspace_mdist prod_metric_def) -context Metric_space12 +lemma prod_dist_dist [simp]: "prod_dist dist dist = dist" + by (simp add: prod_dist_def dist_prod_def fun_eq_iff) + +lemma prod_metric_euclidean [simp]: + "prod_metric euclidean_metric euclidean_metric = euclidean_metric" + by (simp add: prod_metric_def euclidean_metric_def) + +context Metric_space12 begin lemma component_le_prod_metric: "d1 x1 x2 \ prod_dist d1 d2 (x1,y1) (x2,y2)" "d2 y1 y2 \ prod_dist d1 d2 (x1,y1) (x2,y2)" by (auto simp: prod_dist_def) lemma prod_metric_le_components: "\x1 \ M1; y1 \ M1; x2 \ M2; y2 \ M2\ \ prod_dist d1 d2 (x1,x2) (y1,y2) \ d1 x1 y1 + d2 x2 y2" by (auto simp: prod_dist_def sqrt_sum_squares_le_sum) lemma mball_prod_metric_subset: "Prod_metric.mball (x,y) r \ M1.mball x r \ M2.mball y r" by clarsimp (smt (verit, best) component_le_prod_metric) lemma mcball_prod_metric_subset: "Prod_metric.mcball (x,y) r \ M1.mcball x r \ M2.mcball y r" by clarsimp (smt (verit, best) component_le_prod_metric) lemma mball_subset_prod_metric: "M1.mball x1 r1 \ M2.mball x2 r2 \ Prod_metric.mball (x1,x2) (r1 + r2)" using prod_metric_le_components by force lemma mcball_subset_prod_metric: "M1.mcball x1 r1 \ M2.mcball x2 r2 \ Prod_metric.mcball (x1,x2) (r1 + r2)" using prod_metric_le_components by force lemma mtopology_prod_metric: "Prod_metric.mtopology = prod_topology M1.mtopology M2.mtopology" unfolding prod_topology_def proof (rule topology_base_unique [symmetric]) fix U assume "U \ {S \ T |S T. openin M1.mtopology S \ openin M2.mtopology T}" then obtain S T where Ueq: "U = S \ T" and S: "openin M1.mtopology S" and T: "openin M2.mtopology T" by auto have "S \ M1" using M1.openin_mtopology S by auto have "T \ M2" using M2.openin_mtopology T by auto show "openin Prod_metric.mtopology U" unfolding Prod_metric.openin_mtopology proof (intro conjI strip) show "U \ M1 \ M2" using Ueq by (simp add: Sigma_mono \S \ M1\ \T \ M2\) fix z assume "z \ U" then obtain x1 x2 where "x1 \ S" "x2 \ T" and zeq: "z = (x1,x2)" using Ueq by blast obtain r1 where "r1>0" and r1: "M1.mball x1 r1 \ S" by (meson M1.openin_mtopology \openin M1.mtopology S\ \x1 \ S\) obtain r2 where "r2>0" and r2: "M2.mball x2 r2 \ T" by (meson M2.openin_mtopology \openin M2.mtopology T\ \x2 \ T\) have "Prod_metric.mball (x1,x2) (min r1 r2) \ U" proof (rule order_trans [OF mball_prod_metric_subset]) show "M1.mball x1 (min r1 r2) \ M2.mball x2 (min r1 r2) \ U" using Ueq r1 r2 by force qed then show "\r>0. Prod_metric.mball z r \ U" by (smt (verit, del_insts) zeq \0 < r1\ \0 < r2\) qed next fix U z assume "openin Prod_metric.mtopology U" and "z \ U" then have "U \ M1 \ M2" by (simp add: Prod_metric.openin_mtopology) then obtain x y where "x \ M1" "y \ M2" and zeq: "z = (x,y)" using \z \ U\ by blast obtain r where "r>0" and r: "Prod_metric.mball (x,y) r \ U" by (metis Prod_metric.openin_mtopology \openin Prod_metric.mtopology U\ \z \ U\ zeq) define B1 where "B1 \ M1.mball x (r/2)" define B2 where "B2 \ M2.mball y (r/2)" have "openin M1.mtopology B1" "openin M2.mtopology B2" by (simp_all add: B1_def B2_def) moreover have "(x,y) \ B1 \ B2" using \r > 0\ by (simp add: \x \ M1\ \y \ M2\ B1_def B2_def) moreover have "B1 \ B2 \ U" using r prod_metric_le_components by (force simp add: B1_def B2_def) ultimately show "\B. B \ {S \ T |S T. openin M1.mtopology S \ openin M2.mtopology T} \ z \ B \ B \ U" by (auto simp: zeq) qed lemma MCauchy_prod_metric: "Prod_metric.MCauchy \ \ M1.MCauchy (fst \ \) \ M2.MCauchy (snd \ \)" (is "?lhs \ ?rhs") proof safe assume L: ?lhs then have "range \ \ M1 \ M2" using Prod_metric.MCauchy_def by blast then have 1: "range (fst \ \) \ M1" and 2: "range (snd \ \) \ M2" by auto have N1: "\N. \n\N. \n'\N. d1 (fst (\ n)) (fst (\ n')) < \" and N2: "\N. \n\N. \n'\N. d2 (snd (\ n)) (snd (\ n')) < \" if "\>0" for \ :: real using that L unfolding Prod_metric.MCauchy_def by (smt (verit, del_insts) add.commute add_less_imp_less_left add_right_mono component_le_prod_metric prod.collapse)+ show "M1.MCauchy (fst \ \)" using 1 N1 M1.MCauchy_def by auto have "\N. \n\N. \n'\N. d2 (snd (\ n)) (snd (\ n')) < \" if "\>0" for \ :: real using that L unfolding Prod_metric.MCauchy_def by (smt (verit, del_insts) add.commute add_less_imp_less_left add_right_mono component_le_prod_metric prod.collapse) show "M2.MCauchy (snd \ \)" using 2 N2 M2.MCauchy_def by auto next assume M1: "M1.MCauchy (fst \ \)" and M2: "M2.MCauchy (snd \ \)" then have subM12: "range (fst \ \) \ M1" "range (snd \ \) \ M2" using M1.MCauchy_def M2.MCauchy_def by blast+ show ?lhs unfolding Prod_metric.MCauchy_def proof (intro conjI strip) show "range \ \ M1 \ M2" using subM12 by (smt (verit, best) SigmaI image_subset_iff o_apply prod.collapse) fix \ :: real assume "\ > 0" obtain N1 where N1: "\n n'. N1 \ n \ N1 \ n' \ d1 ((fst \ \) n) ((fst \ \) n') < \/2" by (meson M1.MCauchy_def \0 < \\ M1 zero_less_divide_iff zero_less_numeral) obtain N2 where N2: "\n n'. N2 \ n \ N2 \ n' \ d2 ((snd \ \) n) ((snd \ \) n') < \/2" by (meson M2.MCauchy_def \0 < \\ M2 zero_less_divide_iff zero_less_numeral) have "prod_dist d1 d2 (\ n) (\ n') < \" if "N1 \ n" and "N2 \ n" and "N1 \ n'" and "N2 \ n'" for n n' proof - obtain a b a' b' where \: "\ n = (a,b)" "\ n' = (a',b')" by fastforce+ have "prod_dist d1 d2 (a,b) (a',b') \ d1 a a' + d2 b b'" by (metis \range \ \ M1 \ M2\ \ mem_Sigma_iff prod_metric_le_components range_subsetD) also have "\ < \/2 + \/2" using N1 N2 \ that by fastforce finally show ?thesis by (simp add: \) qed then show "\N. \n n'. N \ n \ N \ n' \ prod_dist d1 d2 (\ n) (\ n') < \" by (metis order.trans linorder_le_cases) qed qed lemma mcomplete_prod_metric: "Prod_metric.mcomplete \ M1 = {} \ M2 = {} \ M1.mcomplete \ M2.mcomplete" (is "?lhs \ ?rhs") proof (cases "M1 = {} \ M2 = {}") case False then obtain x y where "x \ M1" "y \ M2" by blast have "M1.mcomplete \ M2.mcomplete \ Prod_metric.mcomplete" by (simp add: Prod_metric.mcomplete_def M1.mcomplete_def M2.mcomplete_def mtopology_prod_metric MCauchy_prod_metric limitin_pairwise) moreover { assume L: "Prod_metric.mcomplete" have "M1.mcomplete" unfolding M1.mcomplete_def proof (intro strip) fix \ assume "M1.MCauchy \" then have "Prod_metric.MCauchy (\n. (\ n, y))" using \y \ M2\ by (simp add: M1.MCauchy_def M2.MCauchy_def MCauchy_prod_metric) then obtain z where "limitin Prod_metric.mtopology (\n. (\ n, y)) z sequentially" using L Prod_metric.mcomplete_def by blast then show "\x. limitin M1.mtopology \ x sequentially" by (auto simp: Prod_metric.mcomplete_def M1.mcomplete_def mtopology_prod_metric limitin_pairwise o_def) qed } moreover { assume L: "Prod_metric.mcomplete" have "M2.mcomplete" unfolding M2.mcomplete_def proof (intro strip) fix \ assume "M2.MCauchy \" then have "Prod_metric.MCauchy (\n. (x, \ n))" using \x \ M1\ by (simp add: M2.MCauchy_def M1.MCauchy_def MCauchy_prod_metric) then obtain z where "limitin Prod_metric.mtopology (\n. (x, \ n)) z sequentially" using L Prod_metric.mcomplete_def by blast then show "\x. limitin M2.mtopology \ x sequentially" by (auto simp: Prod_metric.mcomplete_def M2.mcomplete_def mtopology_prod_metric limitin_pairwise o_def) qed } ultimately show ?thesis using False by blast qed auto lemma mbounded_prod_metric: "Prod_metric.mbounded U \ M1.mbounded (fst ` U) \ M2.mbounded (snd ` U)" proof - have "(\B. U \ Prod_metric.mcball (x,y) B) \ ((\B. (fst ` U) \ M1.mcball x B) \ (\B. (snd ` U) \ M2.mcball y B))" (is "?lhs \ ?rhs") for x y proof safe fix B assume "U \ Prod_metric.mcball (x, y) B" then have "(fst ` U) \ M1.mcball x B" "(snd ` U) \ M2.mcball y B" using mcball_prod_metric_subset by fastforce+ then show "\B. (fst ` U) \ M1.mcball x B" "\B. (snd ` U) \ M2.mcball y B" by auto next fix B1 B2 assume "(fst ` U) \ M1.mcball x B1" "(snd ` U) \ M2.mcball y B2" then have "fst ` U \ snd ` U \ M1.mcball x B1 \ M2.mcball y B2" by blast also have "\ \ Prod_metric.mcball (x, y) (B1+B2)" by (intro mcball_subset_prod_metric) finally show "\B. U \ Prod_metric.mcball (x, y) B" by (metis subsetD subsetI subset_fst_snd) qed then show ?thesis by (simp add: M1.mbounded_def M2.mbounded_def Prod_metric.mbounded_def) qed lemma mbounded_Times: "Prod_metric.mbounded (S \ T) \ S = {} \ T = {} \ M1.mbounded S \ M2.mbounded T" by (auto simp: mbounded_prod_metric) lemma mtotally_bounded_Times: "Prod_metric.mtotally_bounded (S \ T) \ S = {} \ T = {} \ M1.mtotally_bounded S \ M2.mtotally_bounded T" (is "?lhs \ _") proof (cases "S = {} \ T = {}") case False then obtain x y where "x \ S" "y \ T" by auto have "M1.mtotally_bounded S" if L: ?lhs unfolding M1.mtotally_bounded_sequentially proof (intro conjI strip) show "S \ M1" using Prod_metric.mtotally_bounded_imp_subset \y \ T\ that by blast fix \ :: "nat \ 'a" assume "range \ \ S" with L obtain r where "strict_mono r" "Prod_metric.MCauchy ((\n. (\ n,y)) \ r)" unfolding Prod_metric.mtotally_bounded_sequentially by (smt (verit) SigmaI \y \ T\ image_subset_iff) then have "M1.MCauchy (fst \ (\n. (\ n,y)) \ r)" by (simp add: MCauchy_prod_metric o_def) with \strict_mono r\ show "\r. strict_mono r \ M1.MCauchy (\ \ r)" by (auto simp add: o_def) qed moreover have "M2.mtotally_bounded T" if L: ?lhs unfolding M2.mtotally_bounded_sequentially proof (intro conjI strip) show "T \ M2" using Prod_metric.mtotally_bounded_imp_subset \x \ S\ that by blast fix \ :: "nat \ 'b" assume "range \ \ T" with L obtain r where "strict_mono r" "Prod_metric.MCauchy ((\n. (x,\ n)) \ r)" unfolding Prod_metric.mtotally_bounded_sequentially by (smt (verit) SigmaI \x \ S\ image_subset_iff) then have "M2.MCauchy (snd \ (\n. (x,\ n)) \ r)" by (simp add: MCauchy_prod_metric o_def) with \strict_mono r\ show "\r. strict_mono r \ M2.MCauchy (\ \ r)" by (auto simp add: o_def) qed moreover have ?lhs if 1: "M1.mtotally_bounded S" and 2: "M2.mtotally_bounded T" unfolding Prod_metric.mtotally_bounded_sequentially proof (intro conjI strip) show "S \ T \ M1 \ M2" using that by (auto simp: M1.mtotally_bounded_sequentially M2.mtotally_bounded_sequentially) fix \ :: "nat \ 'a \ 'b" assume \: "range \ \ S \ T" with 1 obtain r1 where r1: "strict_mono r1" "M1.MCauchy (fst \ \ \ r1)" apply (clarsimp simp: M1.mtotally_bounded_sequentially image_subset_iff) by (metis SigmaE comp_eq_dest_lhs fst_conv) from \ 2 obtain r2 where r2: "strict_mono r2" "M2.MCauchy (snd \ \ \ r1 \ r2)" apply (clarsimp simp: M2.mtotally_bounded_sequentially image_subset_iff) by (smt (verit, best) comp_apply mem_Sigma_iff prod.collapse) then have "M1.MCauchy (fst \ \ \ r1 \ r2)" by (simp add: M1.MCauchy_subsequence r1) with r2 have "Prod_metric.MCauchy (\ \ (r1 \ r2))" by (simp add: MCauchy_prod_metric o_def) then show "\r. strict_mono r \ Prod_metric.MCauchy (\ \ r)" using r1 r2 strict_mono_o by blast qed ultimately show ?thesis using False by blast qed auto lemma mtotally_bounded_prod_metric: "Prod_metric.mtotally_bounded U \ M1.mtotally_bounded (fst ` U) \ M2.mtotally_bounded (snd ` U)" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "U \ M1 \ M2" and *: "\\. range \ \ U \ \r::nat\nat. strict_mono r \ Prod_metric.MCauchy (\\r)" by (simp_all add: Prod_metric.mtotally_bounded_sequentially) show ?rhs unfolding M1.mtotally_bounded_sequentially M2.mtotally_bounded_sequentially proof (intro conjI strip) show "fst ` U \ M1" "snd ` U \ M2" using \U \ M1 \ M2\ by auto next fix \ :: "nat \ 'a" assume "range \ \ fst ` U" then obtain \ where \: "\n. \ n = fst (\ n) \ \ n \ U" unfolding image_subset_iff image_iff by (meson UNIV_I) then obtain r where "strict_mono r \ Prod_metric.MCauchy (\\r)" by (metis "*" image_subset_iff) with \ show "\r. strict_mono r \ M1.MCauchy (\ \ r)" by (auto simp: MCauchy_prod_metric o_def) next fix \:: "nat \ 'b" assume "range \ \ snd ` U" then obtain \ where \: "\n. \ n = snd (\ n) \ \ n \ U" unfolding image_subset_iff image_iff by (meson UNIV_I) then obtain r where "strict_mono r \ Prod_metric.MCauchy (\\r)" by (metis "*" image_subset_iff) with \ show "\r. strict_mono r \ M2.MCauchy (\ \ r)" by (auto simp: MCauchy_prod_metric o_def) qed next assume ?rhs then have "Prod_metric.mtotally_bounded ((fst ` U) \ (snd ` U))" by (simp add: mtotally_bounded_Times) then show ?lhs by (metis Prod_metric.mtotally_bounded_subset subset_fst_snd) qed end lemma metrizable_space_prod_topology: "metrizable_space (prod_topology X Y) \ topspace(prod_topology X Y) = {} \ metrizable_space X \ metrizable_space Y" (is "?lhs \ ?rhs") proof (cases "topspace(prod_topology X Y) = {}") case False then obtain x y where "x \ topspace X" "y \ topspace Y" by auto show ?thesis proof show "?rhs \ ?lhs" unfolding metrizable_space_def using Metric_space12.mtopology_prod_metric by (metis False Metric_space12.prod_metric Metric_space12_def) next assume L: ?lhs have "metrizable_space (subtopology (prod_topology X Y) (topspace X \ {y}))" "metrizable_space (subtopology (prod_topology X Y) ({x} \ topspace Y))" using L metrizable_space_subtopology by auto moreover have "(subtopology (prod_topology X Y) (topspace X \ {y})) homeomorphic_space X" by (metis \y \ topspace Y\ homeomorphic_space_prod_topology_sing1 homeomorphic_space_sym prod_topology_subtopology(2)) moreover have "(subtopology (prod_topology X Y) ({x} \ topspace Y)) homeomorphic_space Y" by (metis \x \ topspace X\ homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym prod_topology_subtopology(1)) ultimately show ?rhs by (simp add: homeomorphic_metrizable_space) qed qed (simp add: empty_metrizable_space) lemma completely_metrizable_space_prod_topology: "completely_metrizable_space (prod_topology X Y) \ topspace(prod_topology X Y) = {} \ completely_metrizable_space X \ completely_metrizable_space Y" (is "?lhs \ ?rhs") proof (cases "topspace(prod_topology X Y) = {}") case False then obtain x y where "x \ topspace X" "y \ topspace Y" by auto show ?thesis proof show "?rhs \ ?lhs" unfolding completely_metrizable_space_def by (metis False Metric_space12.mtopology_prod_metric Metric_space12.mcomplete_prod_metric Metric_space12.prod_metric Metric_space12_def) next assume L: ?lhs then have "Hausdorff_space (prod_topology X Y)" by (simp add: completely_metrizable_imp_metrizable_space metrizable_imp_Hausdorff_space) then have H: "Hausdorff_space X \ Hausdorff_space Y" using False Hausdorff_space_prod_topology by blast then have "closedin (prod_topology X Y) (topspace X \ {y}) \ closedin (prod_topology X Y) ({x} \ topspace Y)" using \x \ topspace X\ \y \ topspace Y\ by (auto simp: closedin_Hausdorff_sing_eq closedin_prod_Times_iff) with L have "completely_metrizable_space(subtopology (prod_topology X Y) (topspace X \ {y})) \ completely_metrizable_space(subtopology (prod_topology X Y) ({x} \ topspace Y))" by (simp add: completely_metrizable_space_closedin) moreover have "(subtopology (prod_topology X Y) (topspace X \ {y})) homeomorphic_space X" by (metis \y \ topspace Y\ homeomorphic_space_prod_topology_sing1 homeomorphic_space_sym prod_topology_subtopology(2)) moreover have "(subtopology (prod_topology X Y) ({x} \ topspace Y)) homeomorphic_space Y" by (metis \x \ topspace X\ homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym prod_topology_subtopology(1)) ultimately show ?rhs by (simp add: homeomorphic_completely_metrizable_space) qed qed (simp add: empty_completely_metrizable_space) +subsection \The "atin-within" filter for topologies\ + +(*FIXME: replace original definition of atin to be an abbreviation, like at / at_within + ("atin (_) (_)/ within (_)" [1000, 60] 60)*) +definition atin_within :: "['a topology, 'a, 'a set] \ 'a filter" + where "atin_within X a S = inf (nhdsin X a) (principal (topspace X \ S - {a}))" + +lemma atin_within_UNIV [simp]: + shows "atin_within X a UNIV = atin X a" + by (simp add: atin_def atin_within_def) + +lemma eventually_atin_subtopology: + assumes "a \ topspace X" + shows "eventually P (atin (subtopology X S) a) \ + (a \ S \ (\U. openin (subtopology X S) U \ a \ U \ (\x\U - {a}. P x)))" + using assms by (simp add: eventually_atin) + +lemma eventually_atin_within: + "eventually P (atin_within X a S) \ a \ topspace X \ (\T. openin X T \ a \ T \ (\x\T. x \ S \ x \ a \ P x))" +proof (cases "a \ topspace X") + case True + hence "eventually P (atin_within X a S) \ + (\T. openin X T \ a \ T \ + (\x\T. x \ topspace X \ x \ S \ x \ a \ P x))" + by (simp add: atin_within_def eventually_inf_principal eventually_nhdsin) + also have "\ \ (\T. openin X T \ a \ T \ (\x\T. x \ S \ x \ a \ P x))" + using openin_subset by (intro ex_cong) auto + finally show ?thesis by (simp add: True) +qed (simp add: atin_within_def) + +lemma atin_subtopology_within: + assumes "a \ S" + shows "atin (subtopology X S) a = atin_within X a S" +proof - + have "eventually P (atin (subtopology X S) a) \ eventually P (atin_within X a S)" for P + unfolding eventually_atin eventually_atin_within openin_subtopology + using assms by auto + then show ?thesis + by (meson le_filter_def order.eq_iff) +qed + +lemma limit_continuous_map_within: + "\continuous_map (subtopology X S) Y f; a \ S; a \ topspace X\ + \ limitin Y f (f a) (atin_within X a S)" + by (metis Int_iff atin_subtopology_within continuous_map_atin topspace_subtopology) + +lemma atin_subtopology_within_if: + shows "atin (subtopology X S) a = (if a \ S then atin_within X a S else bot)" + by (simp add: atin_subtopology_within) + +lemma trivial_limit_atpointof_within: + "trivial_limit(atin_within X a S) \ + (a \ X derived_set_of S)" + by (auto simp: trivial_limit_def eventually_atin_within in_derived_set_of) + +lemma derived_set_of_trivial_limit: + "a \ X derived_set_of S \ \ trivial_limit(atin_within X a S)" + by (simp add: trivial_limit_atpointof_within) + + +subsection\More sequential characterizations in a metric space\ + +context Metric_space +begin + +definition decreasing_dist :: "(nat \ 'a) \ 'a \ bool" + where "decreasing_dist \ x \ (\m n. m < n \ d (\ n) x < d (\ m) x)" + +lemma decreasing_dist_imp_inj: "decreasing_dist \ a \ inj \" + by (metis decreasing_dist_def dual_order.irrefl linorder_inj_onI') + +lemma eventually_atin_within_metric: + "eventually P (atin_within mtopology a S) \ + (a \ M \ (\\>0. \x. x \ M \ x \ S \ 0 < d x a \ d x a < \ \ P x))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs +unfolding eventually_atin_within openin_mtopology subset_iff + by (metis commute in_mball mdist_zero order_less_irrefl topspace_mtopology) +next + assume R: ?rhs + show ?lhs + proof (cases "a \ M") + case True + then obtain \ where "\ > 0" and \: "\x. \x \ M; x \ S; 0 < d x a; d x a < \\ \ P x" + using R by blast + then have "openin mtopology (mball a \) \ (\x \ mball a \. x \ S \ x \ a \ P x)" + by (simp add: commute openin_mball) + then show ?thesis + by (metis True \0 < \\ centre_in_mball_iff eventually_atin_within) + next + case False + with R show ?thesis + by (simp add: eventually_atin_within) + qed +qed + + +lemma eventually_atin_within_A: + assumes + "(\\. \range \ \ (S \ M) - {a}; decreasing_dist \ a; + inj \; limitin mtopology \ a sequentially\ + \ eventually (\n. P (\ n)) sequentially)" + shows "eventually P (atin_within mtopology a S)" +proof - + have False if SP: "\\. \>0 \ \x \ M-{a}. d x a < \ \ x \ S \ \ P x" and "a \ M" + proof - + define \ where "\ \ \n x. x \ M-{a} \ d x a < inverse (Suc n) \ x \ S \ \ P x" + obtain \ where \: "\n. \ n (\ n)" and dless: "\n. d (\(Suc n)) a < d (\ n) a" + proof - + obtain x0 where x0: "\ 0 x0" + using SP [OF zero_less_one] by (force simp: \_def) + have "\y. \ (Suc n) y \ d y a < d x a" if "\ n x" for n x + using SP [of "min (inverse (Suc (Suc n))) (d x a)"] \a \ M\ that + by (auto simp: \_def) + then obtain f where f: "\n x. \ n x \ \ (Suc n) (f n x) \ d (f n x) a < d x a" + by metis + show thesis + proof + show "\ n (rec_nat x0 f n)" for n + by (induction n) (auto simp: x0 dest: f) + with f show "d (rec_nat x0 f (Suc n)) a < d (rec_nat x0 f n) a" for n + by auto + qed + qed + have 1: "range \ \ (S \ M) - {a}" + using \ by (auto simp: \_def) + have "d (\(Suc (m+n))) a < d (\ n) a" for m n + by (induction m) (auto intro: order_less_trans dless) + then have 2: "decreasing_dist \ a" + unfolding decreasing_dist_def by (metis add.commute less_imp_Suc_add) + have "\\<^sub>F xa in sequentially. d (\ xa) a < \" if "\ > 0" for \ + proof - + obtain N where "inverse (Suc N) < \" + using \\ > 0\ reals_Archimedean by blast + with \ 2 show ?thesis + unfolding decreasing_dist_def by (smt (verit, best) \_def eventually_at_top_dense) + qed + then have 4: "limitin mtopology \ a sequentially" + using \ \a \ M\ by (simp add: \_def limitin_metric) + show False + using 2 assms [OF 1 _ decreasing_dist_imp_inj 4] \ by (force simp: \_def) + qed + then show ?thesis + by (fastforce simp: eventually_atin_within_metric) +qed + + +lemma eventually_atin_within_B: + assumes ev: "eventually P (atin_within mtopology a S)" + and ran: "range \ \ (S \ M) - {a}" + and lim: "limitin mtopology \ a sequentially" + shows "eventually (\n. P (\ n)) sequentially" +proof - + have "a \ M" + using lim limitin_mspace by auto + with ev obtain \ where "0 < \" + and \: "\\. \\ \ M; \ \ S; 0 < d \ a; d \ a < \\ \ P \" + by (auto simp: eventually_atin_within_metric) + then have *: "\n. \ n \ M \ d (\ n) a < \ \ P (\ n)" + using \a \ M\ ran by auto + have "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) a < \" + using lim \0 < \\ by (auto simp: limitin_metric) + then show ?thesis + by (simp add: "*" eventually_mono) +qed + +lemma eventually_atin_within_sequentially: + "eventually P (atin_within mtopology a S) \ + (\\. range \ \ (S \ M) - {a} \ + limitin mtopology \ a sequentially + \ eventually (\n. P(\ n)) sequentially)" + by (metis eventually_atin_within_A eventually_atin_within_B) + +lemma eventually_atin_within_sequentially_inj: + "eventually P (atin_within mtopology a S) \ + (\\. range \ \ (S \ M) - {a} \ inj \ \ + limitin mtopology \ a sequentially + \ eventually (\n. P(\ n)) sequentially)" + by (metis eventually_atin_within_A eventually_atin_within_B) + +lemma eventually_atin_within_sequentially_decreasing: + "eventually P (atin_within mtopology a S) \ + (\\. range \ \ (S \ M) - {a} \ decreasing_dist \ a \ + limitin mtopology \ a sequentially + \ eventually (\n. P(\ n)) sequentially)" + by (metis eventually_atin_within_A eventually_atin_within_B) + + +lemma eventually_atin_sequentially: + "eventually P (atin mtopology a) \ + (\\. range \ \ M - {a} \ limitin mtopology \ a sequentially + \ eventually (\n. P(\ n)) sequentially)" + using eventually_atin_within_sequentially [where S=UNIV] by simp + +lemma eventually_atin_sequentially_inj: + "eventually P (atin mtopology a) \ + (\\. range \ \ M - {a} \ inj \ \ + limitin mtopology \ a sequentially + \ eventually (\n. P(\ n)) sequentially)" + using eventually_atin_within_sequentially_inj [where S=UNIV] by simp + +lemma eventually_atin_sequentially_decreasing: + "eventually P (atin mtopology a) \ + (\\. range \ \ M - {a} \ decreasing_dist \ a \ + limitin mtopology \ a sequentially + \ eventually (\n. P(\ n)) sequentially)" + using eventually_atin_within_sequentially_decreasing [where S=UNIV] by simp end +context Metric_space12 +begin + +lemma limit_atin_sequentially_within: + "limitin M2.mtopology f l (atin_within M1.mtopology a S) \ + l \ M2 \ + (\\. range \ \ S \ M1 - {a} \ + limitin M1.mtopology \ a sequentially + \ limitin M2.mtopology (f \ \) l sequentially)" + by (auto simp: M1.eventually_atin_within_sequentially limitin_def) + +lemma limit_atin_sequentially_within_inj: + "limitin M2.mtopology f l (atin_within M1.mtopology a S) \ + l \ M2 \ + (\\. range \ \ S \ M1 - {a} \ inj \ \ + limitin M1.mtopology \ a sequentially + \ limitin M2.mtopology (f \ \) l sequentially)" + by (auto simp: M1.eventually_atin_within_sequentially_inj limitin_def) + +lemma limit_atin_sequentially_within_decreasing: + "limitin M2.mtopology f l (atin_within M1.mtopology a S) \ + l \ M2 \ + (\\. range \ \ S \ M1 - {a} \ M1.decreasing_dist \ a \ + limitin M1.mtopology \ a sequentially + \ limitin M2.mtopology (f \ \) l sequentially)" + by (auto simp: M1.eventually_atin_within_sequentially_decreasing limitin_def) + +lemma limit_atin_sequentially: + "limitin M2.mtopology f l (atin M1.mtopology a) \ + l \ M2 \ + (\\. range \ \ M1 - {a} \ + limitin M1.mtopology \ a sequentially + \ limitin M2.mtopology (f \ \) l sequentially)" + using limit_atin_sequentially_within [where S=UNIV] by simp + +lemma limit_atin_sequentially_inj: + "limitin M2.mtopology f l (atin M1.mtopology a) \ + l \ M2 \ + (\\. range \ \ M1 - {a} \ inj \ \ + limitin M1.mtopology \ a sequentially + \ limitin M2.mtopology (f \ \) l sequentially)" + using limit_atin_sequentially_within_inj [where S=UNIV] by simp + +lemma limit_atin_sequentially_decreasing: + "limitin M2.mtopology f l (atin M1.mtopology a) \ + l \ M2 \ + (\\. range \ \ M1 - {a} \ M1.decreasing_dist \ a \ + limitin M1.mtopology \ a sequentially + \ limitin M2.mtopology (f \ \) l sequentially)" + using limit_atin_sequentially_within_decreasing [where S=UNIV] by simp + +end + +text \An experiment: same result as within the locale, but using metric space variables\ +lemma limit_atin_sequentially_within: + "limitin (mtopology_of m2) f l (atin_within (mtopology_of m1) a S) \ + l \ mspace m2 \ + (\\. range \ \ S \ mspace m1 - {a} \ + limitin (mtopology_of m1) \ a sequentially + \ limitin (mtopology_of m2) (f \ \) l sequentially)" + using Metric_space12.limit_atin_sequentially_within [OF Metric_space12_mspace_mdist] + by (metis mtopology_of_def) + + +context Metric_space +begin + +lemma atin_within_imp_M: + "atin_within mtopology x S \ bot \ x \ M" + by (metis derived_set_of_trivial_limit in_derived_set_of topspace_mtopology) + +lemma atin_within_sequentially_sequence: + assumes "atin_within mtopology x S \ bot" + obtains \ where "range \ \ S \ M - {x}" + "decreasing_dist \ x" "inj \" "limitin mtopology \ x sequentially" + by (metis eventually_atin_within_A eventually_False assms) + +lemma derived_set_of_sequentially: + "mtopology derived_set_of S = + {x \ M. \\. range \ \ S \ M - {x} \ limitin mtopology \ x sequentially}" +proof - + have False + if "range \ \ S \ M - {x}" + and "limitin mtopology \ x sequentially" + and "atin_within mtopology x S = bot" + for x \ + proof - + have "\\<^sub>F n in sequentially. P (\ n)" for P + using that by (metis eventually_atin_within_B eventually_bot) + then show False + by (meson eventually_False_sequentially eventually_mono) + qed + then show ?thesis + using derived_set_of_trivial_limit + by (fastforce elim!: atin_within_sequentially_sequence intro: atin_within_imp_M) +qed + +lemma derived_set_of_sequentially_alt: + "mtopology derived_set_of S = + {x. \\. range \ \ S - {x} \ limitin mtopology \ x sequentially}" +proof - + have *: "\\. range \ \ S \ M - {x} \ limitin mtopology \ x sequentially" + if \: "range \ \ S - {x}" and lim: "limitin mtopology \ x sequentially" for x \ + proof - + obtain N where "\n\N. \ n \ M \ d (\ n) x < 1" + using lim limit_metric_sequentially by fastforce + with \ obtain a where a:"a \ S \ M - {x}" by auto + show ?thesis + proof (intro conjI exI) + show "range (\n. if \ n \ M then \ n else a) \ S \ M - {x}" + using a \ by fastforce + show "limitin mtopology (\n. if \ n \ M then \ n else a) x sequentially" + using lim limit_metric_sequentially by fastforce + qed + qed + show ?thesis + by (auto simp: limitin_mspace derived_set_of_sequentially intro!: *) +qed + +lemma derived_set_of_sequentially_inj: + "mtopology derived_set_of S = + {x \ M. \\. range \ \ S \ M - {x} \ inj \ \ limitin mtopology \ x sequentially}" +proof - + have False + if "x \ M" and "range \ \ S \ M - {x}" + and "limitin mtopology \ x sequentially" + and "atin_within mtopology x S = bot" + for x \ + proof - + have "\\<^sub>F n in sequentially. P (\ n)" for P + using that derived_set_of_sequentially_alt derived_set_of_trivial_limit by fastforce + then show False + by (meson eventually_False_sequentially eventually_mono) + qed + then show ?thesis + using derived_set_of_trivial_limit + by (fastforce elim!: atin_within_sequentially_sequence intro: atin_within_imp_M) +qed + + +lemma derived_set_of_sequentially_inj_alt: + "mtopology derived_set_of S = + {x. \\. range \ \ S - {x} \ inj \ \ limitin mtopology \ x sequentially}" +proof - + have "\\. range \ \ S - {x} \ inj \ \ limitin mtopology \ x sequentially" + if "atin_within mtopology x S \ bot" for x + by (metis Diff_empty Int_subset_iff atin_within_sequentially_sequence subset_Diff_insert that) + moreover have False + if "range (\x. \ (x::nat)) \ S - {x}" + and "limitin mtopology \ x sequentially" + and "atin_within mtopology x S = bot" + for x \ + proof - + have "\\<^sub>F n in sequentially. P (\ n)" for P + using that derived_set_of_sequentially_alt derived_set_of_trivial_limit by fastforce + then show False + by (meson eventually_False_sequentially eventually_mono) + qed + ultimately show ?thesis + using derived_set_of_trivial_limit by (fastforce intro: atin_within_imp_M) +qed + +lemma derived_set_of_sequentially_decreasing: + "mtopology derived_set_of S = + {x \ M. \\. range \ \ S - {x} \ decreasing_dist \ x \ limitin mtopology \ x sequentially}" +proof - + have "\\. range \ \ S - {x} \ decreasing_dist \ x \ limitin mtopology \ x sequentially" + if "atin_within mtopology x S \ bot" for x + by (metis Diff_empty atin_within_sequentially_sequence le_infE subset_Diff_insert that) + moreover have False + if "x \ M" and "range \ \ S - {x}" + and "limitin mtopology \ x sequentially" + and "atin_within mtopology x S = bot" + for x \ + proof - + have "\\<^sub>F n in sequentially. P (\ n)" for P + using that derived_set_of_sequentially_alt derived_set_of_trivial_limit by fastforce + then show False + by (meson eventually_False_sequentially eventually_mono) + qed + ultimately show ?thesis + using derived_set_of_trivial_limit by (fastforce intro: atin_within_imp_M) +qed + +lemma derived_set_of_sequentially_decreasing_alt: + "mtopology derived_set_of S = + {x. \\. range \ \ S - {x} \ decreasing_dist \ x \ limitin mtopology \ x sequentially}" + using derived_set_of_sequentially_alt derived_set_of_sequentially_decreasing by auto + +lemma closure_of_sequentially: + "mtopology closure_of S = + {x \ M. \\. range \ \ S \ M \ limitin mtopology \ x sequentially}" + by (auto simp: closure_of derived_set_of_sequentially) + +end (*Metric_space*) + + +subsection \Three strong notions of continuity for metric spaces\ + +subsubsection \Lipschitz continuity\ + +definition Lipschitz_continuous_map + where "Lipschitz_continuous_map \ + \m1 m2 f. f ` mspace m1 \ mspace m2 \ + (\B. \x \ mspace m1. \y \ mspace m1. mdist m2 (f x) (f y) \ B * mdist m1 x y)" + +lemma Lipschitz_continuous_map_image: + "Lipschitz_continuous_map m1 m2 f \ f ` mspace m1 \ mspace m2" + by (simp add: Lipschitz_continuous_map_def) + +lemma Lipschitz_continuous_map_pos: + "Lipschitz_continuous_map m1 m2 f \ + f ` mspace m1 \ mspace m2 \ + (\B>0. \x \ mspace m1. \y \ mspace m1. mdist m2 (f x) (f y) \ B * mdist m1 x y)" +proof - + have "B * mdist m1 x y \ (\B\ + 1) * mdist m1 x y" "\B\ + 1 > 0" for x y B + by (auto simp add: mult_right_mono) + then show ?thesis + unfolding Lipschitz_continuous_map_def by (meson dual_order.trans) +qed + + +lemma Lipschitz_continuous_map_eq: + assumes "Lipschitz_continuous_map m1 m2 f" "\x. x \ mspace m1 \ f x = g x" + shows "Lipschitz_continuous_map m1 m2 g" + using Lipschitz_continuous_map_def assms + by (metis (no_types, opaque_lifting) image_subset_iff) + +lemma Lipschitz_continuous_map_from_submetric: + assumes "Lipschitz_continuous_map m1 m2 f" + shows "Lipschitz_continuous_map (submetric m1 S) m2 f" + unfolding Lipschitz_continuous_map_def +proof + show "f ` mspace (submetric m1 S) \ mspace m2" + using Lipschitz_continuous_map_pos assms by fastforce +qed (use assms in \fastforce simp: Lipschitz_continuous_map_def\) + +lemma Lipschitz_continuous_map_from_submetric_mono: + "\Lipschitz_continuous_map (submetric m1 T) m2 f; S \ T\ + \ Lipschitz_continuous_map (submetric m1 S) m2 f" + by (metis Lipschitz_continuous_map_from_submetric inf.absorb_iff2 submetric_submetric) + +lemma Lipschitz_continuous_map_into_submetric: + "Lipschitz_continuous_map m1 (submetric m2 S) f \ + f ` mspace m1 \ S \ Lipschitz_continuous_map m1 m2 f" + by (auto simp: Lipschitz_continuous_map_def) + +lemma Lipschitz_continuous_map_const: + "Lipschitz_continuous_map m1 m2 (\x. c) \ + mspace m1 = {} \ c \ mspace m2" + unfolding Lipschitz_continuous_map_def image_subset_iff + by (metis all_not_in_conv mdist_nonneg mdist_zero mult_1) + +lemma Lipschitz_continuous_map_id: + "Lipschitz_continuous_map m1 m1 (\x. x)" + by (metis Lipschitz_continuous_map_def image_ident mult_1 order_refl) + +lemma Lipschitz_continuous_map_compose: + assumes f: "Lipschitz_continuous_map m1 m2 f" and g: "Lipschitz_continuous_map m2 m3 g" + shows "Lipschitz_continuous_map m1 m3 (g \ f)" + unfolding Lipschitz_continuous_map_def image_subset_iff +proof + show "\x\mspace m1. (g \ f) x \ mspace m3" + by (metis Lipschitz_continuous_map_def assms comp_apply image_subset_iff) + obtain B where B: "\x\mspace m1. \y\mspace m1. mdist m2 (f x) (f y) \ B * mdist m1 x y" + using assms unfolding Lipschitz_continuous_map_def by presburger + obtain C where "C>0" and C: "\x\mspace m2. \y\mspace m2. mdist m3 (g x) (g y) \ C * mdist m2 x y" + using assms unfolding Lipschitz_continuous_map_pos by metis + show "\B. \x\mspace m1. \y\mspace m1. mdist m3 ((g \ f) x) ((g \ f) y) \ B * mdist m1 x y" + apply (rule_tac x="C*B" in exI) + proof clarify + fix x y + assume \
: "x \ mspace m1" "y \ mspace m1" + then have "mdist m3 ((g \ f) x) ((g \ f) y) \ C * mdist m2 (f x) (f y)" + by (metis C Lipschitz_continuous_map_def f comp_apply image_subset_iff) + also have "\ \ C * B * mdist m1 x y" + by (simp add: "\
" B \0 < C\) + finally show "mdist m3 ((g \ f) x) ((g \ f) y) \ C * B * mdist m1 x y" . + qed +qed + +subsubsection \Uniform continuity\ + +definition uniformly_continuous_map + where "uniformly_continuous_map \ + \m1 m2 f. f ` mspace m1 \ mspace m2 \ + (\\>0. \\>0. \x \ mspace m1. \y \ mspace m1. + mdist m1 y x < \ \ mdist m2 (f y) (f x) < \)" + +lemma uniformly_continuous_map_image: + "uniformly_continuous_map m1 m2 f \ f ` mspace m1 \ mspace m2" + by (simp add: uniformly_continuous_map_def) + +lemma ucmap_A: + assumes "uniformly_continuous_map m1 m2 f" + and "(\n. mdist m1 (\ n) (\ n)) \ 0" + and "range \ \ mspace m1" "range \ \ mspace m1" + shows "(\n. mdist m2 (f (\ n)) (f (\ n))) \ 0" + using assms + unfolding uniformly_continuous_map_def image_subset_iff tendsto_iff + apply clarsimp + by (metis (mono_tags, lifting) eventually_sequentially) + +lemma ucmap_B: + assumes \
: "\\ \. \range \ \ mspace m1; range \ \ mspace m1; + (\n. mdist m1 (\ n) (\ n)) \ 0\ + \ (\n. mdist m2 (f (\ n)) (f (\ n))) \ 0" + and "0 < \" + and \: "range \ \ mspace m1" + and \: "range \ \ mspace m1" + and "(\n. mdist m1 (\ n) (\ n)) \ 0" + shows "\n. mdist m2 (f (\ (n::nat))) (f (\ n)) < \" + using \
[OF \ \] assms by (meson LIMSEQ_le_const linorder_not_less) + +lemma ucmap_C: + assumes \
: "\\ \ \. \\ > 0; range \ \ mspace m1; range \ \ mspace m1; + ((\n. mdist m1 (\ n) (\ n)) \ 0)\ + \ \n. mdist m2 (f (\ n)) (f (\ n)) < \" + and fim: "f ` mspace m1 \ mspace m2" + shows "uniformly_continuous_map m1 m2 f" +proof - + {assume "\ (\\>0. \\>0. \x\mspace m1. \y\mspace m1. mdist m1 y x < \ \ mdist m2 (f y) (f x) < \)" + then obtain \ where "\ > 0" + and "\n. \x\mspace m1. \y\mspace m1. mdist m1 y x < inverse(Suc n) \ mdist m2 (f y) (f x) \ \" + by (meson inverse_Suc linorder_not_le) + then obtain \ \ where space: "range \ \ mspace m1" "range \ \ mspace m1" + and dist: "\n. mdist m1 (\ n) (\ n) < inverse(Suc n) \ mdist m2 (f(\ n)) (f(\ n)) \ \" + by (metis image_subset_iff) + have False + using \
[OF \\ > 0\ space] dist Lim_null_comparison + by (smt (verit) LIMSEQ_norm_0 inverse_eq_divide mdist_commute mdist_nonneg real_norm_def) + } + moreover + have "t \ mspace m2" if "t \ f ` mspace m1" for t + using fim that by blast + ultimately show ?thesis + by (fastforce simp: uniformly_continuous_map_def) +qed + +lemma uniformly_continuous_map_sequentially: + "uniformly_continuous_map m1 m2 f \ + f ` mspace m1 \ mspace m2 \ + (\\ \. range \ \ mspace m1 \ range \ \ mspace m1 \ (\n. mdist m1 (\ n) (\ n)) \ 0 + \ (\n. mdist m2 (f (\ n)) (f (\ n))) \ 0)" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: ucmap_A uniformly_continuous_map_image) + show "?rhs \ ?lhs" + by (intro ucmap_B ucmap_C) auto +qed + + +lemma uniformly_continuous_map_sequentially_alt: + "uniformly_continuous_map m1 m2 f \ + f ` mspace m1 \ mspace m2 \ + (\\>0. \\ \. range \ \ mspace m1 \ range \ \ mspace m1 \ + ((\n. mdist m1 (\ n) (\ n)) \ 0) + \ (\n. mdist m2 (f (\ n)) (f (\ n)) < \))" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + using uniformly_continuous_map_image by (intro conjI strip ucmap_B | force simp: ucmap_A)+ + show "?rhs \ ?lhs" + by (intro ucmap_C) auto +qed + + +lemma uniformly_continuous_map_eq: + "\\x. x \ mspace m1 \ f x = g x; uniformly_continuous_map m1 m2 f\ + \ uniformly_continuous_map m1 m2 g" + by (simp add: uniformly_continuous_map_def) + +lemma uniformly_continuous_map_from_submetric: + assumes "uniformly_continuous_map m1 m2 f" + shows "uniformly_continuous_map (submetric m1 S) m2 f" + unfolding uniformly_continuous_map_def +proof + show "f ` mspace (submetric m1 S) \ mspace m2" + using assms by (auto simp: uniformly_continuous_map_def) +qed (use assms in \force simp: uniformly_continuous_map_def\) + +lemma uniformly_continuous_map_from_submetric_mono: + "\uniformly_continuous_map (submetric m1 T) m2 f; S \ T\ + \ uniformly_continuous_map (submetric m1 S) m2 f" + by (metis uniformly_continuous_map_from_submetric inf.absorb_iff2 submetric_submetric) + +lemma uniformly_continuous_map_into_submetric: + "uniformly_continuous_map m1 (submetric m2 S) f \ + f ` mspace m1 \ S \ uniformly_continuous_map m1 m2 f" + by (auto simp: uniformly_continuous_map_def) + +lemma uniformly_continuous_map_const: + "uniformly_continuous_map m1 m2 (\x. c) \ + mspace m1 = {} \ c \ mspace m2" + unfolding uniformly_continuous_map_def image_subset_iff + by (metis empty_iff equals0I mdist_zero) + +lemma uniformly_continuous_map_id [simp]: + "uniformly_continuous_map m1 m1 (\x. x)" + by (metis image_ident order_refl uniformly_continuous_map_def) + +lemma uniformly_continuous_map_compose: + assumes f: "uniformly_continuous_map m1 m2 f" and g: "uniformly_continuous_map m2 m3 g" + shows "uniformly_continuous_map m1 m3 (g \ f)" + by (smt (verit, ccfv_threshold) comp_apply f g image_subset_iff uniformly_continuous_map_def) + +lemma uniformly_continuous_map_real_const [simp]: + "uniformly_continuous_map m euclidean_metric (\x. c)" + by (simp add: euclidean_metric_def uniformly_continuous_map_const) + +text \Equivalence between "abstract" and "type class" notions\ +lemma uniformly_continuous_map_euclidean [simp]: + "uniformly_continuous_map (submetric euclidean_metric S) euclidean_metric f + = uniformly_continuous_on S f" + by (auto simp add: uniformly_continuous_map_def uniformly_continuous_on_def) + +subsubsection \Cauchy continuity\ + +definition Cauchy_continuous_map where + "Cauchy_continuous_map \ + \m1 m2 f. \\. Metric_space.MCauchy (mspace m1) (mdist m1) \ + \ Metric_space.MCauchy (mspace m2) (mdist m2) (f \ \)" + +lemma Cauchy_continuous_map_euclidean [simp]: + "Cauchy_continuous_map (submetric euclidean_metric S) euclidean_metric f + = Cauchy_continuous_on S f" + by (auto simp add: Cauchy_continuous_map_def Cauchy_continuous_on_def Cauchy_def Met_TC.subspace Metric_space.MCauchy_def) + +lemma Cauchy_continuous_map_image: + assumes "Cauchy_continuous_map m1 m2 f" + shows "f ` mspace m1 \ mspace m2" +proof clarsimp + fix x + assume "x \ mspace m1" + then have "Metric_space.MCauchy (mspace m1) (mdist m1) (\n. x)" + by (simp add: Metric_space.MCauchy_const Metric_space_mspace_mdist) + then have "Metric_space.MCauchy (mspace m2) (mdist m2) (f \ (\n. x))" + by (meson Cauchy_continuous_map_def assms) + then show "f x \ mspace m2" + by (simp add: Metric_space.MCauchy_def [OF Metric_space_mspace_mdist]) +qed + + +lemma Cauchy_continuous_map_eq: + "\\x. x \ mspace m1 \ f x = g x; Cauchy_continuous_map m1 m2 f\ + \ Cauchy_continuous_map m1 m2 g" + by (simp add: image_subset_iff Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist]) + +lemma Cauchy_continuous_map_from_submetric: + assumes "Cauchy_continuous_map m1 m2 f" + shows "Cauchy_continuous_map (submetric m1 S) m2 f" + using assms + by (simp add: image_subset_iff Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist]) + +lemma Cauchy_continuous_map_from_submetric_mono: + "\Cauchy_continuous_map (submetric m1 T) m2 f; S \ T\ + \ Cauchy_continuous_map (submetric m1 S) m2 f" + by (metis Cauchy_continuous_map_from_submetric inf.absorb_iff2 submetric_submetric) + +lemma Cauchy_continuous_map_into_submetric: + "Cauchy_continuous_map m1 (submetric m2 S) f \ + f ` mspace m1 \ S \ Cauchy_continuous_map m1 m2 f" (is "?lhs \ ?rhs") +proof - + have "?lhs \ Cauchy_continuous_map m1 m2 f" + by (simp add: Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist]) + moreover have "?rhs \ ?lhs" + by (bestsimp simp add: Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist]) + ultimately show ?thesis + by (metis Cauchy_continuous_map_image le_inf_iff mspace_submetric) +qed + +lemma Cauchy_continuous_map_const [simp]: + "Cauchy_continuous_map m1 m2 (\x. c) \ mspace m1 = {} \ c \ mspace m2" +proof - + have "mspace m1 = {} \ Cauchy_continuous_map m1 m2 (\x. c)" + by (simp add: Cauchy_continuous_map_def Metric_space.MCauchy_def Metric_space_mspace_mdist) + moreover have "c \ mspace m2 \ Cauchy_continuous_map m1 m2 (\x. c)" + by (simp add: Cauchy_continuous_map_def o_def Metric_space.MCauchy_const [OF Metric_space_mspace_mdist]) + ultimately show ?thesis + using Cauchy_continuous_map_image by blast +qed + +lemma Cauchy_continuous_map_id [simp]: + "Cauchy_continuous_map m1 m1 (\x. x)" + by (simp add: Cauchy_continuous_map_def o_def) + +lemma Cauchy_continuous_map_compose: + assumes f: "Cauchy_continuous_map m1 m2 f" and g: "Cauchy_continuous_map m2 m3 g" + shows "Cauchy_continuous_map m1 m3 (g \ f)" + by (metis (no_types, lifting) Cauchy_continuous_map_def f fun.map_comp g) + +lemma Lipschitz_imp_uniformly_continuous_map: + assumes "Lipschitz_continuous_map m1 m2 f" + shows "uniformly_continuous_map m1 m2 f" + proof - + have "f ` mspace m1 \ mspace m2" + by (simp add: Lipschitz_continuous_map_image assms) + moreover have "\\>0. \x\mspace m1. \y\mspace m1. mdist m1 y x < \ \ mdist m2 (f y) (f x) < \" + if "\ > 0" for \ + proof - + obtain B where "\x\mspace m1. \y\mspace m1. mdist m2 (f x) (f y) \ B * mdist m1 x y" + and "B>0" + using that assms by (force simp add: Lipschitz_continuous_map_pos) + then have "\x\mspace m1. \y\mspace m1. mdist m1 y x < \/B \ mdist m2 (f y) (f x) < \" + by (smt (verit, ccfv_SIG) less_divide_eq mdist_nonneg mult.commute that zero_less_divide_iff) + with \B>0\ show ?thesis + by (metis divide_pos_pos that) + qed + ultimately show ?thesis + by (auto simp: uniformly_continuous_map_def) +qed + +lemma uniformly_imp_Cauchy_continuous_map: + "uniformly_continuous_map m1 m2 f \ Cauchy_continuous_map m1 m2 f" + unfolding uniformly_continuous_map_def Cauchy_continuous_map_def + apply (simp add: image_subset_iff o_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist]) + by meson + +lemma locally_Cauchy_continuous_map: + assumes "\ > 0" + and \
: "\x. x \ mspace m1 \ Cauchy_continuous_map (submetric m1 (mball_of m1 x \)) m2 f" + shows "Cauchy_continuous_map m1 m2 f" + unfolding Cauchy_continuous_map_def +proof (intro strip) + interpret M1: Metric_space "mspace m1" "mdist m1" + by (simp add: Metric_space_mspace_mdist) + interpret M2: Metric_space "mspace m2" "mdist m2" + by (simp add: Metric_space_mspace_mdist) + fix \ + assume \: "M1.MCauchy \" + with \\ > 0\ obtain N where N: "\n n'. \n\N; n'\N\ \ mdist m1 (\ n) (\ n') < \" + using M1.MCauchy_def by fastforce + then have "M1.mball (\ N) \ \ mspace m1" + by (auto simp: image_subset_iff M1.mball_def) + then interpret MS1: Metric_space "mball_of m1 (\ N) \ \ mspace m1" "mdist m1" + by (simp add: M1.subspace) + show "M2.MCauchy (f \ \)" + proof (rule M2.MCauchy_offset) + have "M1.MCauchy (\ \ (+) N)" + by (simp add: M1.MCauchy_imp_MCauchy_suffix \) + moreover have "range (\ \ (+) N) \ M1.mball (\ N) \" + using N [OF order_refl] M1.MCauchy_def \ by fastforce + ultimately have "MS1.MCauchy (\ \ (+) N)" + unfolding M1.MCauchy_def MS1.MCauchy_def by (simp add: mball_of_def) + moreover have "\ N \ mspace m1" + using M1.MCauchy_def \ by auto + ultimately show "M2.MCauchy (f \ \ \ (+) N)" + unfolding comp_assoc + by (metis "\
" Cauchy_continuous_map_def mdist_submetric mspace_submetric) + next + fix n + have "\ n \ mspace m1" + by (meson Metric_space.MCauchy_def Metric_space_mspace_mdist \ range_subsetD) + then have "\ n \ mball_of m1 (\ n) \" + by (simp add: Metric_space.centre_in_mball_iff Metric_space_mspace_mdist assms(1) mball_of_def) + then show "(f \ \) n \ mspace m2" + using Cauchy_continuous_map_image [OF \
[of "\ n"]] \\ n \ mspace m1\ by auto + qed +qed + +context Metric_space12 +begin + +lemma Cauchy_continuous_imp_continuous_map: + assumes "Cauchy_continuous_map (metric (M1,d1)) (metric (M2,d2)) f" + shows "continuous_map M1.mtopology M2.mtopology f" +proof (clarsimp simp: continuous_map_atin) + fix x + assume "x \ M1" + show "limitin M2.mtopology f (f x) (atin M1.mtopology x)" + unfolding limit_atin_sequentially + proof (intro conjI strip) + show "f x \ M2" + using Cauchy_continuous_map_image \x \ M1\ assms by fastforce + fix \ + assume "range \ \ M1 - {x} \ limitin M1.mtopology \ x sequentially" + then have "M1.MCauchy (\n. if even n then \ (n div 2) else x)" + by (force simp add: M1.MCauchy_interleaving) + then have "M2.MCauchy (f \ (\n. if even n then \ (n div 2) else x))" + using assms by (simp add: Cauchy_continuous_map_def) + then show "limitin M2.mtopology (f \ \) (f x) sequentially" + using M2.MCauchy_interleaving [of "f \ \" "f x"] + by (simp add: o_def if_distrib cong: if_cong) + qed +qed + +lemma continuous_imp_Cauchy_continuous_map: + assumes "M1.mcomplete" + and f: "continuous_map M1.mtopology M2.mtopology f" + shows "Cauchy_continuous_map (metric (M1,d1)) (metric (M2,d2)) f" + unfolding Cauchy_continuous_map_def +proof clarsimp + fix \ + assume \: "M1.MCauchy \" + then obtain y where y: "limitin M1.mtopology \ y sequentially" + using M1.mcomplete_def assms by blast + have "range (f \ \) \ M2" + using \ f by (simp add: M2.subspace M1.MCauchy_def M1.metric_continuous_map image_subset_iff) + then show "M2.MCauchy (f \ \)" + using continuous_map_limit [OF f y] M2.convergent_imp_MCauchy + by blast +qed + +end + +text \The same outside the locale\ +lemma Cauchy_continuous_imp_continuous_map: + assumes "Cauchy_continuous_map m1 m2 f" + shows "continuous_map (mtopology_of m1) (mtopology_of m2) f" + using assms Metric_space12.Cauchy_continuous_imp_continuous_map [OF Metric_space12_mspace_mdist] + by (auto simp add: mtopology_of_def) + +lemma continuous_imp_Cauchy_continuous_map: + assumes "Metric_space.mcomplete (mspace m1) (mdist m1)" + and "continuous_map (mtopology_of m1) (mtopology_of m2) f" + shows "Cauchy_continuous_map m1 m2 f" + using assms Metric_space12.continuous_imp_Cauchy_continuous_map [OF Metric_space12_mspace_mdist] + by (auto simp add: mtopology_of_def) + +lemma uniformly_continuous_imp_continuous_map: + "uniformly_continuous_map m1 m2 f + \ continuous_map (mtopology_of m1) (mtopology_of m2) f" + by (simp add: Cauchy_continuous_imp_continuous_map uniformly_imp_Cauchy_continuous_map) + +lemma Lipschitz_continuous_imp_continuous_map: + "Lipschitz_continuous_map m1 m2 f + \ continuous_map (mtopology_of m1) (mtopology_of m2) f" + by (simp add: Lipschitz_imp_uniformly_continuous_map uniformly_continuous_imp_continuous_map) + +lemma Lipschitz_imp_Cauchy_continuous_map: + "Lipschitz_continuous_map m1 m2 f + \ Cauchy_continuous_map m1 m2 f" + by (simp add: Lipschitz_imp_uniformly_continuous_map uniformly_imp_Cauchy_continuous_map) + +lemma Cauchy_imp_uniformly_continuous_map: + assumes f: "Cauchy_continuous_map m1 m2 f" + and tbo: "Metric_space.mtotally_bounded (mspace m1) (mdist m1) (mspace m1)" + shows "uniformly_continuous_map m1 m2 f" + unfolding uniformly_continuous_map_sequentially_alt imp_conjL +proof (intro conjI strip) + show "f ` mspace m1 \ mspace m2" + by (simp add: Cauchy_continuous_map_image f) + interpret M1: Metric_space "mspace m1" "mdist m1" + by (simp add: Metric_space_mspace_mdist) + interpret M2: Metric_space "mspace m2" "mdist m2" + by (simp add: Metric_space_mspace_mdist) + fix \ :: real and \ \ + assume "\ > 0" + and \: "range \ \ mspace m1" + and \: "range \ \ mspace m1" + and 0: "(\n. mdist m1 (\ n) (\ n)) \ 0" + then obtain r1 where "strict_mono r1" and r1: "M1.MCauchy (\ \ r1)" + using M1.mtotally_bounded_sequentially tbo by meson + then obtain r2 where "strict_mono r2" and r2: "M1.MCauchy (\ \ r1 \ r2)" + by (metis M1.mtotally_bounded_sequentially tbo \ image_comp image_subset_iff range_subsetD) + define r where "r \ r1 \ r2" + have r: "strict_mono r" + by (simp add: r_def \strict_mono r1\ \strict_mono r2\ strict_mono_o) + define \ where "\ \ \n. if even n then (\ \ r) (n div 2) else (\ \ r) (n div 2)" + have "M1.MCauchy \" + unfolding \_def M1.MCauchy_interleaving_gen + proof (intro conjI) + show "M1.MCauchy (\ \ r)" + by (simp add: M1.MCauchy_subsequence \strict_mono r2\ fun.map_comp r1 r_def) + show "M1.MCauchy (\ \ r)" + by (simp add: fun.map_comp r2 r_def) + show "(\n. mdist m1 ((\ \ r) n) ((\ \ r) n)) \ 0" + using LIMSEQ_subseq_LIMSEQ [OF 0 r] by (simp add: o_def) + qed + then have "Metric_space.MCauchy (mspace m2) (mdist m2) (f \ \)" + by (meson Cauchy_continuous_map_def f) + then have "(\n. mdist m2 (f (\ (r n))) (f (\ (r n)))) \ 0" + using M2.MCauchy_interleaving_gen [of "f \ \ \ r" "f \ \ \ r"] + by (simp add: \_def o_def if_distrib cong: if_cong) + then show "\n. mdist m2 (f (\ n)) (f (\ n)) < \" + by (meson LIMSEQ_le_const \0 < \\ linorder_not_le) +qed + +lemma continuous_imp_uniformly_continuous_map: + "compact_space (mtopology_of m1) \ + continuous_map (mtopology_of m1) (mtopology_of m2) f + \ uniformly_continuous_map m1 m2 f" + by (simp add: Cauchy_imp_uniformly_continuous_map continuous_imp_Cauchy_continuous_map + Metric_space.compact_space_eq_mcomplete_mtotally_bounded + Metric_space_mspace_mdist mtopology_of_def) + +lemma continuous_eq_Cauchy_continuous_map: + "Metric_space.mcomplete (mspace m1) (mdist m1) \ + continuous_map (mtopology_of m1) (mtopology_of m2) f \ Cauchy_continuous_map m1 m2 f" + using Cauchy_continuous_imp_continuous_map continuous_imp_Cauchy_continuous_map by blast + +lemma continuous_eq_uniformly_continuous_map: + "compact_space (mtopology_of m1) + \ continuous_map (mtopology_of m1) (mtopology_of m2) f \ + uniformly_continuous_map m1 m2 f" + using continuous_imp_uniformly_continuous_map uniformly_continuous_imp_continuous_map by blast + +lemma Cauchy_eq_uniformly_continuous_map: + "Metric_space.mtotally_bounded (mspace m1) (mdist m1) (mspace m1) + \ Cauchy_continuous_map m1 m2 f \ uniformly_continuous_map m1 m2 f" + using Cauchy_imp_uniformly_continuous_map uniformly_imp_Cauchy_continuous_map by blast + +lemma Lipschitz_continuous_map_projections: + "Lipschitz_continuous_map (prod_metric m1 m2) m1 fst" + "Lipschitz_continuous_map (prod_metric m1 m2) m2 snd" + by (simp add: Lipschitz_continuous_map_def prod_dist_def; + metis mult_numeral_1 real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2)+ + +lemma Lipschitz_continuous_map_pairwise: + "Lipschitz_continuous_map m (prod_metric m1 m2) f \ + Lipschitz_continuous_map m m1 (fst \ f) \ Lipschitz_continuous_map m m2 (snd \ f)" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: Lipschitz_continuous_map_compose Lipschitz_continuous_map_projections) + have "Lipschitz_continuous_map m (prod_metric m1 m2) (\x. (f1 x, f2 x))" + if f1: "Lipschitz_continuous_map m m1 f1" and f2: "Lipschitz_continuous_map m m2 f2" for f1 f2 + proof - + obtain B1 where "B1 > 0" + and B1: "\x y. \x \ mspace m; y \ mspace m\ \ mdist m1 (f1 x) (f1 y) \ B1 * mdist m x y" + by (meson Lipschitz_continuous_map_pos f1) + obtain B2 where "B2 > 0" + and B2: "\x y. \x \ mspace m; y \ mspace m\ \ mdist m2 (f2 x) (f2 y) \ B2 * mdist m x y" + by (meson Lipschitz_continuous_map_pos f2) + show ?thesis + unfolding Lipschitz_continuous_map_pos + proof (intro exI conjI strip) + have f1im: "f1 ` mspace m \ mspace m1" + by (simp add: Lipschitz_continuous_map_image f1) + moreover have f2im: "f2 ` mspace m \ mspace m2" + by (simp add: Lipschitz_continuous_map_image f2) + ultimately show "(\x. (f1 x, f2 x)) ` mspace m \ mspace (prod_metric m1 m2)" + by auto + show "B1+B2 > 0" + using \0 < B1\ \0 < B2\ by linarith + fix x y + assume xy: "x \ mspace m" "y \ mspace m" + with f1im f2im have "mdist (prod_metric m1 m2) (f1 x, f2 x) (f1 y, f2 y) \ mdist m1 (f1 x) (f1 y) + mdist m2 (f2 x) (f2 y)" + unfolding mdist_prod_metric + by (intro Metric_space12.prod_metric_le_components [OF Metric_space12_mspace_mdist]) auto + also have "... \ (B1+B2) * mdist m x y" + using B1 [OF xy] B2 [OF xy] by (simp add: vector_space_over_itself.scale_left_distrib) + finally show "mdist (prod_metric m1 m2) (f1 x, f2 x) (f1 y, f2 y) \ (B1+B2) * mdist m x y" . + qed + qed + then show "?rhs \ ?lhs" + by force +qed + +lemma uniformly_continuous_map_pairwise: + "uniformly_continuous_map m (prod_metric m1 m2) f \ + uniformly_continuous_map m m1 (fst \ f) \ uniformly_continuous_map m m2 (snd \ f)" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: Lipschitz_continuous_map_projections Lipschitz_imp_uniformly_continuous_map uniformly_continuous_map_compose) + have "uniformly_continuous_map m (prod_metric m1 m2) (\x. (f1 x, f2 x))" + if f1: "uniformly_continuous_map m m1 f1" and f2: "uniformly_continuous_map m m2 f2" for f1 f2 + proof - + show ?thesis + unfolding uniformly_continuous_map_def + proof (intro conjI strip) + have f1im: "f1 ` mspace m \ mspace m1" + by (simp add: uniformly_continuous_map_image f1) + moreover have f2im: "f2 ` mspace m \ mspace m2" + by (simp add: uniformly_continuous_map_image f2) + ultimately show "(\x. (f1 x, f2 x)) ` mspace m \ mspace (prod_metric m1 m2)" + by auto + fix \:: real + assume "\ > 0" + obtain \1 where "\1>0" + and \1: "\x y. \x \ mspace m; y \ mspace m; mdist m y x < \1\ \ mdist m1 (f1 y) (f1 x) < \/2" + by (metis \0 < \\ f1 half_gt_zero uniformly_continuous_map_def) + obtain \2 where "\2>0" + and \2: "\x y. \x \ mspace m; y \ mspace m; mdist m y x < \2\ \ mdist m2 (f2 y) (f2 x) < \/2" + by (metis \0 < \\ f2 half_gt_zero uniformly_continuous_map_def) + show "\\>0. \x\mspace m. \y\mspace m. mdist m y x < \ \ mdist (prod_metric m1 m2) (f1 y, f2 y) (f1 x, f2 x) < \" + proof (intro exI conjI strip) + show "min \1 \2>0" + using \0 < \1\ \0 < \2\ by auto + fix x y + assume xy: "x \ mspace m" "y \ mspace m" and d: "mdist m y x < min \1 \2" + have *: "mdist m1 (f1 y) (f1 x) < \/2" "mdist m2 (f2 y) (f2 x) < \/2" + using \1 \2 d xy by auto + have "mdist (prod_metric m1 m2) (f1 y, f2 y) (f1 x, f2 x) \ mdist m1 (f1 y) (f1 x) + mdist m2 (f2 y) (f2 x)" + unfolding mdist_prod_metric using f1im f2im xy + by (intro Metric_space12.prod_metric_le_components [OF Metric_space12_mspace_mdist]) auto + also have "... < \/2 + \/2" + using * by simp + finally show "mdist (prod_metric m1 m2) (f1 y, f2 y) (f1 x, f2 x) < \" + by simp + qed + qed + qed + then show "?rhs \ ?lhs" + by force +qed + +lemma Cauchy_continuous_map_pairwise: + "Cauchy_continuous_map m (prod_metric m1 m2) f \ Cauchy_continuous_map m m1 (fst \ f) \ Cauchy_continuous_map m m2 (snd \ f)" + by (auto simp: Cauchy_continuous_map_def Metric_space12.MCauchy_prod_metric[OF Metric_space12_mspace_mdist] comp_assoc) + +lemma Lipschitz_continuous_map_paired: + "Lipschitz_continuous_map m (prod_metric m1 m2) (\x. (f x, g x)) \ + Lipschitz_continuous_map m m1 f \ Lipschitz_continuous_map m m2 g" + by (simp add: Lipschitz_continuous_map_pairwise o_def) + +lemma uniformly_continuous_map_paired: + "uniformly_continuous_map m (prod_metric m1 m2) (\x. (f x, g x)) \ + uniformly_continuous_map m m1 f \ uniformly_continuous_map m m2 g" + by (simp add: uniformly_continuous_map_pairwise o_def) + +lemma Cauchy_continuous_map_paired: + "Cauchy_continuous_map m (prod_metric m1 m2) (\x. (f x, g x)) \ + Cauchy_continuous_map m m1 f \ Cauchy_continuous_map m m2 g" + by (simp add: Cauchy_continuous_map_pairwise o_def) + +lemma mbounded_Lipschitz_continuous_image: + assumes f: "Lipschitz_continuous_map m1 m2 f" and S: "Metric_space.mbounded (mspace m1) (mdist m1) S" + shows "Metric_space.mbounded (mspace m2) (mdist m2) (f`S)" +proof - + obtain B where fim: "f ` mspace m1 \ mspace m2" + and "B>0" and B: "\x y. \x \ mspace m1; y \ mspace m1\ \ mdist m2 (f x) (f y) \ B * mdist m1 x y" + by (meson Lipschitz_continuous_map_pos f) + show ?thesis + unfolding Metric_space.mbounded_alt_pos [OF Metric_space_mspace_mdist] + proof + obtain C where "S \ mspace m1" and "C>0" and C: "\x y. \x \ S; y \ S\ \ mdist m1 x y \ C" + using S by (auto simp: Metric_space.mbounded_alt_pos [OF Metric_space_mspace_mdist]) + show "f ` S \ mspace m2" + using fim \S \ mspace m1\ by blast + have "\x y. \x \ S; y \ S\ \ mdist m2 (f x) (f y) \ B * C" + by (smt (verit) B C \0 < B\ \S \ mspace m1\ mdist_nonneg mult_mono subsetD) + moreover have "B*C > 0" + by (simp add: \0 < B\ \0 < C\) + ultimately show "\B>0. \x\f ` S. \y\f ` S. mdist m2 x y \ B" + by auto + qed +qed + +lemma mtotally_bounded_Cauchy_continuous_image: + assumes f: "Cauchy_continuous_map m1 m2 f" and S: "Metric_space.mtotally_bounded (mspace m1) (mdist m1) S" + shows "Metric_space.mtotally_bounded (mspace m2) (mdist m2) (f ` S)" + unfolding Metric_space.mtotally_bounded_sequentially[OF Metric_space_mspace_mdist] +proof (intro conjI strip) + have "S \ mspace m1" + using S by (simp add: Metric_space.mtotally_bounded_sequentially[OF Metric_space_mspace_mdist]) + then show "f ` S \ mspace m2" + using Cauchy_continuous_map_image f by blast + fix \ :: "nat \ 'b" + assume "range \ \ f ` S" + then have "\n. \x. \ n = f x \ x \ S" + by (meson imageE range_subsetD) + then obtain \ where \: "\n. \ n = f (\ n)" "range \ \ S" + by (metis image_subset_iff) + then have "\ = f \ \" + by fastforce + obtain r where "strict_mono r" "Metric_space.MCauchy (mspace m1) (mdist m1) (\ \ r)" + by (meson \ S Metric_space.mtotally_bounded_sequentially[OF Metric_space_mspace_mdist]) + then have "Metric_space.MCauchy (mspace m2) (mdist m2) (f \ \ \ r)" + using f unfolding Cauchy_continuous_map_def by (metis fun.map_comp) + then show "\r. strict_mono r \ Metric_space.MCauchy (mspace m2) (mdist m2) (\ \ r)" + using \\ = f \ \\ \strict_mono r\ by blast +qed + +lemma Lipschitz_coefficient_pos: + assumes "x \ mspace m" "y \ mspace m" "f x \ f y" + and "f ` mspace m \ mspace m2" + and "\x y. \x \ mspace m; y \ mspace m\ + \ mdist m2 (f x) (f y) \ k * mdist m x y" + shows "0 < k" + using assms by (smt (verit, best) image_subset_iff mdist_nonneg mdist_zero mult_nonpos_nonneg) + +lemma Lipschitz_continuous_map_metric: + "Lipschitz_continuous_map (prod_metric m m) euclidean_metric (\(x,y). mdist m x y)" +proof - + have "\x y x' y'. \x \ mspace m; y \ mspace m; x' \ mspace m; y' \ mspace m\ + \ \mdist m x y - mdist m x' y'\ \ 2 * sqrt ((mdist m x x')\<^sup>2 + (mdist m y y')\<^sup>2)" + by (smt (verit, del_insts) mdist_commute mdist_triangle real_sqrt_sum_squares_ge2) + then show ?thesis + by (fastforce simp add: Lipschitz_continuous_map_def prod_dist_def dist_real_def) +qed + +lemma Lipschitz_continuous_map_mdist: + assumes f: "Lipschitz_continuous_map m m' f" + and g: "Lipschitz_continuous_map m m' g" + shows "Lipschitz_continuous_map m euclidean_metric (\x. mdist m' (f x) (g x))" + (is "Lipschitz_continuous_map m _ ?h") +proof - + have eq: "?h = ((\(x,y). mdist m' x y) \ (\x. (f x,g x)))" + by force + show ?thesis + unfolding eq + proof (rule Lipschitz_continuous_map_compose) + show "Lipschitz_continuous_map m (prod_metric m' m') (\x. (f x, g x))" + by (simp add: Lipschitz_continuous_map_paired f g) + show "Lipschitz_continuous_map (prod_metric m' m') euclidean_metric (\(x,y). mdist m' x y)" + by (simp add: Lipschitz_continuous_map_metric) + qed +qed + +lemma uniformly_continuous_map_mdist: + assumes f: "uniformly_continuous_map m m' f" + and g: "uniformly_continuous_map m m' g" + shows "uniformly_continuous_map m euclidean_metric (\x. mdist m' (f x) (g x))" + (is "uniformly_continuous_map m _ ?h") +proof - + have eq: "?h = ((\(x,y). mdist m' x y) \ (\x. (f x,g x)))" + by force + show ?thesis + unfolding eq + proof (rule uniformly_continuous_map_compose) + show "uniformly_continuous_map m (prod_metric m' m') (\x. (f x, g x))" + by (simp add: uniformly_continuous_map_paired f g) + show "uniformly_continuous_map (prod_metric m' m') euclidean_metric (\(x,y). mdist m' x y)" + by (simp add: Lipschitz_continuous_map_metric Lipschitz_imp_uniformly_continuous_map) + qed +qed + +lemma Cauchy_continuous_map_mdist: + assumes f: "Cauchy_continuous_map m m' f" + and g: "Cauchy_continuous_map m m' g" + shows "Cauchy_continuous_map m euclidean_metric (\x. mdist m' (f x) (g x))" + (is "Cauchy_continuous_map m _ ?h") +proof - + have eq: "?h = ((\(x,y). mdist m' x y) \ (\x. (f x,g x)))" + by force + show ?thesis + unfolding eq + proof (rule Cauchy_continuous_map_compose) + show "Cauchy_continuous_map m (prod_metric m' m') (\x. (f x, g x))" + by (simp add: Cauchy_continuous_map_paired f g) + show "Cauchy_continuous_map (prod_metric m' m') euclidean_metric (\(x,y). mdist m' x y)" + by (simp add: Lipschitz_continuous_map_metric Lipschitz_imp_Cauchy_continuous_map) + qed +qed + +lemma mtopology_of_prod_metric [simp]: + "mtopology_of (prod_metric m1 m2) = prod_topology (mtopology_of m1) (mtopology_of m2)" + by (simp add: mtopology_of_def Metric_space12.mtopology_prod_metric[OF Metric_space12_mspace_mdist]) + +lemma continuous_map_metric: + "continuous_map (prod_topology (mtopology_of m) (mtopology_of m)) euclidean + (\(x,y). mdist m x y)" + using Lipschitz_continuous_imp_continuous_map [OF Lipschitz_continuous_map_metric] + by auto + +lemma continuous_map_mdist_alt: + assumes "continuous_map X (prod_topology (mtopology_of m) (mtopology_of m)) f" + shows "continuous_map X euclidean (\x. case_prod (mdist m) (f x))" + (is "continuous_map _ _ ?h") +proof - + have eq: "?h = case_prod (mdist m) \ f" + by force + show ?thesis + unfolding eq + using assms continuous_map_compose continuous_map_metric by blast +qed + +lemma continuous_map_mdist: + assumes f: "continuous_map X (mtopology_of m) f" + and g: "continuous_map X (mtopology_of m) g" + shows "continuous_map X euclidean (\x. mdist m (f x) (g x))" + (is "continuous_map X _ ?h") +proof - + have eq: "?h = ((\(x,y). mdist m x y) \ (\x. (f x,g x)))" + by force + show ?thesis + unfolding eq + proof (rule continuous_map_compose) + show "continuous_map X (prod_topology (mtopology_of m) (mtopology_of m)) (\x. (f x, g x))" + by (simp add: continuous_map_paired f g) + qed (simp add: continuous_map_metric) +qed + +lemma continuous_on_mdist: + "a \ mspace m \ continuous_map (mtopology_of m) euclidean (mdist m a)" + by (simp add: continuous_map_mdist) + +subsection \Isometries\ + +lemma (in Metric_space12) isometry_imp_embedding_map: + assumes fim: "f ` M1 \ M2" and d: "\x y. \x \ M1; y \ M1\ \ d2 (f x) (f y) = d1 x y" + shows "embedding_map M1.mtopology M2.mtopology f" +proof - + have "inj_on f M1" + by (metis M1.zero d inj_onI) + then obtain g where g: "\x. x \ M1 \ g (f x) = x" + by (metis inv_into_f_f) + have "homeomorphic_maps M1.mtopology (subtopology M2.mtopology (f ` topspace M1.mtopology)) f g" + unfolding homeomorphic_maps_def + proof (intro conjI; clarsimp) + show "continuous_map M1.mtopology (subtopology M2.mtopology (f ` M1)) f" + by (metis M1.metric_continuous_map M1.topspace_mtopology M2.Metric_space_axioms continuous_map_into_subtopology d fim order_refl) + have "Lipschitz_continuous_map (submetric (metric(M2,d2)) (f ` M1)) (metric(M1,d1)) g" + unfolding Lipschitz_continuous_map_def + proof (intro conjI exI strip; simp) + show "d1 (g x) (g y) \ 1 * d2 x y" if "x \ f ` M1 \ x \ M2" and "y \ f ` M1 \ y \ M2" for x y + using that d g by force + qed (use g in auto) + then have "continuous_map (mtopology_of (submetric (metric(M2,d2)) (f ` M1))) M1.mtopology g" + using Lipschitz_continuous_imp_continuous_map by force + moreover have "mtopology_of (submetric (metric(M2,d2)) (f ` M1)) = subtopology M2.mtopology (f ` M1)" + by (simp add: mtopology_of_submetric) + ultimately show "continuous_map (subtopology M2.mtopology (f ` M1)) M1.mtopology g" + by simp + qed (use g in auto) + then show ?thesis + by (auto simp: embedding_map_def homeomorphic_map_maps) +qed + +lemma (in Metric_space12) isometry_imp_homeomorphic_map: + assumes fim: "f ` M1 = M2" and d: "\x y. \x \ M1; y \ M1\ \ d2 (f x) (f y) = d1 x y" + shows "homeomorphic_map M1.mtopology M2.mtopology f" + by (metis M1.topspace_mtopology M2.subtopology_mspace d embedding_map_def fim isometry_imp_embedding_map subsetI) + +end + diff --git a/src/HOL/Analysis/Elementary_Metric_Spaces.thy b/src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy @@ -1,3212 +1,3271 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Elementary Metric Spaces\ theory Elementary_Metric_Spaces imports Abstract_Topology_2 Metric_Arith begin subsection \Open and closed balls\ definition\<^marker>\tag important\ ball :: "'a::metric_space \ real \ 'a set" where "ball x e = {y. dist x y < e}" definition\<^marker>\tag important\ cball :: "'a::metric_space \ real \ 'a set" where "cball x e = {y. dist x y \ e}" definition\<^marker>\tag important\ sphere :: "'a::metric_space \ real \ 'a set" where "sphere x e = {y. dist x y = e}" lemma mem_ball [simp, metric_unfold]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) lemma mem_cball [simp, metric_unfold]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) lemma mem_sphere [simp]: "y \ sphere x e \ dist x y = e" by (simp add: sphere_def) lemma ball_trivial [simp]: "ball x 0 = {}" by auto lemma cball_trivial [simp]: "cball x 0 = {x}" by auto lemma sphere_trivial [simp]: "sphere x 0 = {x}" by auto lemma disjoint_ballI: "dist x y \ r+s \ ball x r \ ball y s = {}" using dist_triangle_less_add not_le by fastforce lemma disjoint_cballI: "dist x y > r + s \ cball x r \ cball y s = {}" by (metis add_mono disjoint_iff_not_equal dist_triangle2 dual_order.trans leD mem_cball) lemma sphere_empty [simp]: "r < 0 \ sphere a r = {}" for a :: "'a::metric_space" by auto lemma centre_in_ball [simp]: "x \ ball x e \ 0 < e" by simp lemma centre_in_cball [simp]: "x \ cball x e \ 0 \ e" by simp lemma ball_subset_cball [simp, intro]: "ball x e \ cball x e" by (simp add: subset_eq) lemma mem_ball_imp_mem_cball: "x \ ball y e \ x \ cball y e" by auto lemma sphere_cball [simp,intro]: "sphere z r \ cball z r" by force lemma cball_diff_sphere: "cball a r - sphere a r = ball a r" by auto lemma subset_ball[intro]: "d \ e \ ball x d \ ball x e" by auto lemma subset_cball[intro]: "d \ e \ cball x d \ cball x e" by auto lemma mem_ball_leI: "x \ ball y e \ e \ f \ x \ ball y f" by auto lemma mem_cball_leI: "x \ cball y e \ e \ f \ x \ cball y f" by auto lemma cball_trans: "y \ cball z b \ x \ cball y a \ x \ cball z (b + a)" by metric lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" by auto lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" by auto lemma cball_max_Un: "cball a (max r s) = cball a r \ cball a s" by auto lemma cball_min_Int: "cball a (min r s) = cball a r \ cball a s" by auto lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r" by auto lemma open_ball [intro, simp]: "open (ball x e)" proof - have "open (dist x -` {.. (\x\S. \e>0. ball x e \ S)" by (simp add: open_dist subset_eq Ball_def dist_commute) lemma openI [intro?]: "(\x. x\S \ \e>0. ball x e \ S) \ open S" by (auto simp: open_contains_ball) lemma openE[elim?]: assumes "open S" "x\S" obtains e where "e>0" "ball x e \ S" using assms unfolding open_contains_ball by auto lemma open_contains_ball_eq: "open S \ x\S \ (\e>0. ball x e \ S)" by (metis open_contains_ball subset_eq centre_in_ball) lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" unfolding mem_ball set_eq_iff by (simp add: not_less) metric lemma ball_empty: "e \ 0 \ ball x e = {}" by simp lemma closed_cball [iff]: "closed (cball x e)" proof - have "closed (dist x -` {..e})" by (intro closed_vimage closed_atMost continuous_intros) also have "dist x -` {..e} = cball x e" by auto finally show ?thesis . qed lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" proof - { fix x and e::real assume "x\S" "e>0" "ball x e \ S" then have "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) } moreover { fix x and e::real assume "x\S" "e>0" "cball x e \ S" then have "\d>0. ball x d \ S" using mem_ball_imp_mem_cball by blast } ultimately show ?thesis unfolding open_contains_ball by auto qed lemma open_contains_cball_eq: "open S \ (\x. x \ S \ (\e>0. cball x e \ S))" by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball) lemma eventually_nhds_ball: "d > 0 \ eventually (\x. x \ ball z d) (nhds z)" by (rule eventually_nhds_in_open) simp_all lemma eventually_at_ball: "d > 0 \ eventually (\t. t \ ball z d \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma eventually_at_ball': "d > 0 \ eventually (\t. t \ ball z d \ t \ z \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma at_within_ball: "e > 0 \ dist x y < e \ at y within ball x e = at y" by (subst at_within_open) auto lemma atLeastAtMost_eq_cball: fixes a b::real shows "{a .. b} = cball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma cball_eq_atLeastAtMost: fixes a b::real shows "cball a b = {a - b .. a + b}" by (auto simp: dist_real_def) lemma greaterThanLessThan_eq_ball: fixes a b::real shows "{a <..< b} = ball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma ball_eq_greaterThanLessThan: fixes a b::real shows "ball a b = {a - b <..< a + b}" by (auto simp: dist_real_def) lemma interior_ball [simp]: "interior (ball x e) = ball x e" by (simp add: interior_open) lemma cball_eq_empty [simp]: "cball x e = {} \ e < 0" by (smt (verit, best) Diff_empty ball_eq_empty cball_diff_sphere centre_in_ball centre_in_cball sphere_empty) lemma cball_empty [simp]: "e < 0 \ cball x e = {}" by simp lemma cball_sing: fixes x :: "'a::metric_space" shows "e = 0 \ cball x e = {x}" by simp lemma ball_divide_subset: "d \ 1 \ ball x (e/d) \ ball x e" by (metis ball_eq_empty div_by_1 frac_le linear subset_ball zero_less_one) lemma ball_divide_subset_numeral: "ball x (e / numeral w) \ ball x e" using ball_divide_subset one_le_numeral by blast lemma cball_divide_subset: "d \ 1 \ cball x (e/d) \ cball x e" by (smt (verit, best) cball_empty div_by_1 frac_le subset_cball zero_le_divide_iff) lemma cball_divide_subset_numeral: "cball x (e / numeral w) \ cball x e" using cball_divide_subset one_le_numeral by blast lemma cball_scale: assumes "a \ 0" shows "(\x. a *\<^sub>R x) ` cball c r = cball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" proof - have 1: "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a proof safe fix x assume x: "x \ cball c r" have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" by (auto simp: dist_norm) also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" by (simp add: algebra_simps) finally show "a *\<^sub>R x \ cball (a *\<^sub>R c) (\a\ * r)" using that x by (auto simp: dist_norm) qed have "cball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` cball (a *\<^sub>R c) (\a\ * r)" unfolding image_image using assms by simp also have "\ \ (\x. a *\<^sub>R x) ` cball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" using assms by (intro image_mono 1) auto also have "\ = (\x. a *\<^sub>R x) ` cball c r" using assms by (simp add: algebra_simps) finally have "cball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` cball c r" . moreover from assms have "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" by (intro 1) auto ultimately show ?thesis by blast qed lemma ball_scale: assumes "a \ 0" shows "(\x. a *\<^sub>R x) ` ball c r = ball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" proof - have 1: "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a proof safe fix x assume x: "x \ ball c r" have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" by (auto simp: dist_norm) also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" by (simp add: algebra_simps) finally show "a *\<^sub>R x \ ball (a *\<^sub>R c) (\a\ * r)" using that x by (auto simp: dist_norm) qed have "ball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` ball (a *\<^sub>R c) (\a\ * r)" unfolding image_image using assms by simp also have "\ \ (\x. a *\<^sub>R x) ` ball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" using assms by (intro image_mono 1) auto also have "\ = (\x. a *\<^sub>R x) ` ball c r" using assms by (simp add: algebra_simps) finally have "ball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` ball c r" . moreover from assms have "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" by (intro 1) auto ultimately show ?thesis by blast qed lemma frequently_atE: fixes x :: "'a :: metric_space" assumes "frequently P (at x within s)" shows "\f. filterlim f (at x within s) sequentially \ (\n. P (f n))" proof - have "\y. y \ s \ (ball x (1 / real (Suc n)) - {x}) \ P y" for n proof - have "\z\s. z \ x \ dist z x < (1 / real (Suc n)) \ P z" by (metis assms divide_pos_pos frequently_at of_nat_0_less_iff zero_less_Suc zero_less_one) then show ?thesis by (auto simp: dist_commute conj_commute) qed then obtain f where f: "\n. f n \ s \ (ball x (1 / real (Suc n)) - {x}) \ P (f n)" by metis have "filterlim f (nhds x) sequentially" unfolding tendsto_iff proof clarify fix e :: real assume e: "e > 0" then obtain n where n: "Suc n > 1 / e" by (meson le_nat_floor lessI not_le) have "dist (f k) x < e" if "k \ n" for k proof - have "dist (f k) x < 1 / real (Suc k)" using f[of k] by (auto simp: dist_commute) also have "\ \ 1 / real (Suc n)" using that by (intro divide_left_mono) auto also have "\ < e" using n e by (simp add: field_simps) finally show ?thesis . qed thus "\\<^sub>F k in sequentially. dist (f k) x < e" unfolding eventually_at_top_linorder by blast qed moreover have "f n \ x" for n using f[of n] by auto ultimately have "filterlim f (at x within s) sequentially" using f by (auto simp: filterlim_at) with f show ?thesis by blast qed subsection \Limit Points\ lemma islimpt_approachable: fixes x :: "'a::metric_space" shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" unfolding islimpt_iff_eventually eventually_at by fast lemma islimpt_approachable_le: "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x \ e)" for x :: "'a::metric_space" unfolding islimpt_approachable using approachable_lt_le2 [where f="\y. dist y x" and P="\y. y \ S \ y = x" and Q="\x. True"] by auto lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \ x islimpt S" for x :: "'a::metric_space" by (metis islimpt_def islimpt_eq_acc_point mem_Collect_eq) lemma closed_limpts: "closed {x::'a::metric_space. x islimpt S}" using closed_limpt limpt_of_limpts by blast lemma limpt_of_closure: "x islimpt closure S \ x islimpt S" for x :: "'a::metric_space" by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts) lemma islimpt_eq_infinite_ball: "x islimpt S \ (\e>0. infinite(S \ ball x e))" unfolding islimpt_eq_acc_point by (metis open_ball Int_commute Int_mono finite_subset open_contains_ball_eq subset_eq) lemma islimpt_eq_infinite_cball: "x islimpt S \ (\e>0. infinite(S \ cball x e))" unfolding islimpt_eq_infinite_ball by (metis open_ball ball_subset_cball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq) subsection \Perfect Metric Spaces\ lemma perfect_choose_dist: "0 < r \ \a. a \ x \ dist a x < r" for x :: "'a::{perfect_space,metric_space}" using islimpt_UNIV [of x] by (simp add: islimpt_approachable) lemma cball_eq_sing: fixes x :: "'a::{metric_space,perfect_space}" shows "cball x e = {x} \ e = 0" by (smt (verit, best) open_ball ball_eq_empty ball_subset_cball cball_empty cball_trivial not_open_singleton subset_singleton_iff) subsection \Finite and discrete\ lemma finite_ball_include: fixes a :: "'a::metric_space" assumes "finite S" shows "\e>0. S \ ball a e" using assms proof induction case (insert x S) then obtain e0 where "e0>0" and e0:"S \ ball a e0" by auto define e where "e = max e0 (2 * dist a x)" have "e>0" unfolding e_def using \e0>0\ by auto moreover have "insert x S \ ball a e" using e0 \e>0\ unfolding e_def by auto ultimately show ?case by auto qed (auto intro: zero_less_one) lemma finite_set_avoid: fixes a :: "'a::metric_space" assumes "finite S" shows "\d>0. \x\S. x \ a \ d \ dist a x" using assms proof induction case (insert x S) then obtain d where "d > 0" and d: "\x\S. x \ a \ d \ dist a x" by blast show ?case by (smt (verit, del_insts) dist_pos_lt insert.IH insert_iff) qed (auto intro: zero_less_one) lemma discrete_imp_closed: fixes S :: "'a::metric_space set" assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" shows "closed S" proof - have False if C: "\e. e>0 \ \x'\S. x' \ x \ dist x' x < e" for x proof - from e have e2: "e/2 > 0" by arith from C[OF e2] obtain y where y: "y \ S" "y \ x" "dist y x < e/2" by blast from e2 y(2) have mp: "min (e/2) (dist x y) > 0" by simp from d y C[OF mp] show ?thesis by metric qed then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) qed lemma discrete_imp_not_islimpt: assumes e: "0 < e" and d: "\x y. x \ S \ y \ S \ dist y x < e \ y = x" shows "\ x islimpt S" proof assume "x islimpt S" hence "x islimpt S - {x}" by (meson islimpt_punctured) moreover from assms have "closed (S - {x})" by (intro discrete_imp_closed) auto ultimately show False unfolding closed_limpt by blast qed subsection \Interior\ lemma mem_interior: "x \ interior S \ (\e>0. ball x e \ S)" using open_contains_ball_eq [where S="interior S"] by (simp add: open_subset_interior) lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" by (meson ball_subset_cball interior_subset mem_interior open_contains_cball open_interior subset_trans) lemma ball_iff_cball: "(\r>0. ball x r \ U) \ (\r>0. cball x r \ U)" by (meson mem_interior mem_interior_cball) subsection \Frontier\ lemma frontier_straddle: fixes a :: "'a::metric_space" shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" unfolding frontier_def closure_interior by (auto simp: mem_interior subset_eq ball_def) subsection \Limits\ proposition Lim: "(f \ l) net \ trivial_limit net \ (\e>0. eventually (\x. dist (f x) l < e) net)" by (auto simp: tendsto_iff trivial_limit_eq) text \Show that they yield usual definitions in the various cases.\ proposition Lim_within_le: "(f \ l)(at a within S) \ (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a \ d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at_le) proposition Lim_within: "(f \ l) (at a within S) \ (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) corollary Lim_withinI [intro?]: assumes "\e. e > 0 \ \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l \ e" shows "(f \ l) (at a within S)" unfolding Lim_within by (smt (verit, ccfv_SIG) assms zero_less_dist_iff) proposition Lim_at: "(f \ l) (at a) \ (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) lemma Lim_transform_within_set: fixes a :: "'a::metric_space" and l :: "'b::metric_space" shows "\(f \ l) (at a within S); eventually (\x. x \ S \ x \ T) (at a)\ \ (f \ l) (at a within T)" by (simp add: eventually_at Lim_within) (smt (verit, best)) text \Another limit point characterization.\ lemma limpt_sequential_inj: fixes x :: "'a::metric_space" shows "x islimpt S \ (\f. (\n::nat. f n \ S - {x}) \ inj f \ (f \ x) sequentially)" (is "?lhs = ?rhs") proof assume ?lhs then have "\e>0. \x'\S. x' \ x \ dist x' x < e" by (force simp: islimpt_approachable) then obtain y where y: "\e. e>0 \ y e \ S \ y e \ x \ dist (y e) x < e" by metis define f where "f \ rec_nat (y 1) (\n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))" have [simp]: "f 0 = y 1" and fSuc: "f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n by (simp_all add: f_def) have f: "f n \ S \ (f n \ x) \ dist (f n) x < inverse(2 ^ n)" for n proof (induction n) case 0 show ?case by (simp add: y) next case (Suc n) then show ?case by (smt (verit, best) fSuc dist_pos_lt inverse_positive_iff_positive y zero_less_power) qed show ?rhs proof (intro exI conjI allI) show "\n. f n \ S - {x}" using f by blast have "dist (f n) x < dist (f m) x" if "m < n" for m n using that proof (induction n) case 0 then show ?case by simp next case (Suc n) then consider "m < n" | "m = n" using less_Suc_eq by blast then show ?case proof cases assume "m < n" have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x" by (simp add: fSuc) also have "\ < dist (f n) x" by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y) also have "\ < dist (f m) x" using Suc.IH \m < n\ by blast finally show ?thesis . next assume "m = n" then show ?case by (smt (verit, best) dist_pos_lt f fSuc y) qed qed then show "inj f" by (metis less_irrefl linorder_injI) have "\e n. \0 < e; nat \1 / e\ \ n\ \ dist (f n) x < e" apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]]) by (simp add: divide_simps order_le_less_trans) then show "f \ x" using lim_sequentially by blast qed next assume ?rhs then show ?lhs by (fastforce simp add: islimpt_approachable lim_sequentially) qed lemma Lim_dist_ubound: assumes "\(trivial_limit net)" and "(f \ l) net" and "eventually (\x. dist a (f x) \ e) net" shows "dist a l \ e" using assms by (fast intro: tendsto_le tendsto_intros) subsection \Continuity\ text\Derive the epsilon-delta forms, which we often use as "definitions"\ proposition continuous_within_eps_delta: "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" unfolding continuous_within and Lim_within by fastforce corollary continuous_at_eps_delta: "continuous (at x) f \ (\e > 0. \d > 0. \x'. dist x' x < d \ dist (f x') (f x) < e)" using continuous_within_eps_delta [of x UNIV f] by simp lemma continuous_at_right_real_increasing: fixes f :: "real \ real" assumes nondecF: "\x y. x \ y \ f x \ f y" shows "continuous (at_right a) f \ (\e>0. \d>0. f (a + d) - f a < e)" apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong) by (smt (verit, best) nondecF) lemma continuous_at_left_real_increasing: assumes nondecF: "\ x y. x \ y \ f x \ ((f y) :: real)" shows "(continuous (at_left (a :: real)) f) \ (\e > 0. \delta > 0. f a - f (a - delta) < e)" apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong) by (smt (verit) nondecF) text\Versions in terms of open balls.\ lemma continuous_within_ball: "continuous (at x within S) f \ (\e > 0. \d > 0. f ` (ball x d \ S) \ ball (f x) e)" (is "?lhs = ?rhs") proof assume ?lhs { fix e :: real assume "e > 0" then obtain d where "d>0" and d: "\y\S. 0 < dist y x \ dist y x < d \ dist (f y) (f x) < e" using \?lhs\[unfolded continuous_within Lim_within] by auto { fix y assume "y \ f ` (ball x d \ S)" then have "y \ ball (f x) e" using d \e > 0\ by (auto simp: dist_commute) } then have "\d>0. f ` (ball x d \ S) \ ball (f x) e" using \d > 0\ by blast } then show ?rhs by auto next assume ?rhs then show ?lhs apply (simp add: continuous_within Lim_within ball_def subset_eq) by (metis (mono_tags, lifting) Int_iff dist_commute mem_Collect_eq) qed lemma continuous_at_ball: "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" apply (simp add: continuous_at Lim_at subset_eq Ball_def Bex_def image_iff) by (smt (verit, ccfv_threshold) dist_commute dist_self) text\Define setwise continuity in terms of limits within the set.\ lemma continuous_on_iff: "continuous_on s f \ (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding continuous_on_def Lim_within by (metis dist_pos_lt dist_self) lemma continuous_within_E: assumes "continuous (at x within S) f" "e>0" obtains d where "d>0" "\x'. \x'\ S; dist x' x \ d\ \ dist (f x') (f x) < e" using assms unfolding continuous_within_eps_delta by (metis dense order_le_less_trans) lemma continuous_onI [intro?]: assumes "\x e. \e > 0; x \ S\ \ \d>0. \x'\S. dist x' x < d \ dist (f x') (f x) \ e" shows "continuous_on S f" apply (simp add: continuous_on_iff, clarify) apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) done text\Some simple consequential lemmas.\ lemma continuous_onE: assumes "continuous_on s f" "x\s" "e>0" obtains d where "d>0" "\x'. \x' \ s; dist x' x \ d\ \ dist (f x') (f x) < e" using assms unfolding continuous_on_iff by (metis dense order_le_less_trans) text\The usual transformation theorems.\ lemma continuous_transform_within: fixes f g :: "'a::metric_space \ 'b::topological_space" assumes "continuous (at x within s) f" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "continuous (at x within s) g" using assms unfolding continuous_within by (force intro: Lim_transform_within) subsection \Closure and Limit Characterization\ lemma closure_approachable: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" using dist_self by (force simp: closure_def islimpt_approachable) lemma closure_approachable_le: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x \ e)" unfolding closure_approachable using dense by force lemma closure_approachableD: assumes "x \ closure S" "e>0" shows "\y\S. dist x y < e" using assms unfolding closure_approachable by (auto simp: dist_commute) lemma closed_approachable: fixes S :: "'a::metric_space set" shows "closed S \ (\e>0. \y\S. dist y x < e) \ x \ S" by (metis closure_closed closure_approachable) lemma closure_contains_Inf: fixes S :: "real set" assumes "S \ {}" "bdd_below S" shows "Inf S \ closure S" proof - have *: "\x\S. Inf S \ x" using cInf_lower[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Inf S < Inf S + e" by simp with assms obtain x where "x \ S" "x < Inf S + e" using cInf_lessD by blast with * have "\x\S. dist x (Inf S) < e" using dist_real_def by force } then show ?thesis unfolding closure_approachable by auto qed lemma closure_contains_Sup: fixes S :: "real set" assumes "S \ {}" "bdd_above S" shows "Sup S \ closure S" proof - have *: "\x\S. x \ Sup S" using cSup_upper[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Sup S - e < Sup S" by simp with assms obtain x where "x \ S" "Sup S - e < x" using less_cSupE by blast with * have "\x\S. dist x (Sup S) < e" using dist_real_def by force } then show ?thesis unfolding closure_approachable by auto qed lemma not_trivial_limit_within_ball: "\ trivial_limit (at x within S) \ (\e>0. S \ ball x e - {x} \ {})" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S - {x}" and "dist y x < e" using \?lhs\ not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto then have "y \ S \ ball x e - {x}" unfolding ball_def by (simp add: dist_commute) then have "S \ ball x e - {x} \ {}" by blast } then show ?thesis by auto qed show ?lhs if ?rhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S \ ball x e - {x}" using \?rhs\ by blast then have "y \ S - {x}" and "dist y x < e" unfolding ball_def by (simp_all add: dist_commute) then have "\y \ S - {x}. dist y x < e" by auto } then show ?thesis using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto qed qed subsection \Boundedness\ (* FIXME: This has to be unified with BSEQ!! *) definition\<^marker>\tag important\ (in metric_space) bounded :: "'a set \ bool" where "bounded S \ (\x e. \y\S. dist x y \ e)" lemma bounded_subset_cball: "bounded S \ (\e x. S \ cball x e \ 0 \ e)" unfolding bounded_def subset_eq by auto (meson order_trans zero_le_dist) lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" unfolding bounded_def by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le) lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" unfolding bounded_any_center [where a=0] by (simp add: dist_norm) lemma bdd_above_norm: "bdd_above (norm ` X) \ bounded X" by (simp add: bounded_iff bdd_above_def) lemma bounded_norm_comp: "bounded ((\x. norm (f x)) ` S) = bounded (f ` S)" by (simp add: bounded_iff) lemma boundedI: assumes "\x. x \ S \ norm x \ B" shows "bounded S" using assms bounded_iff by blast lemma bounded_empty [simp]: "bounded {}" by (simp add: bounded_def) lemma bounded_subset: "bounded T \ S \ T \ bounded S" by (metis bounded_def subset_eq) lemma bounded_interior[intro]: "bounded S \ bounded(interior S)" by (metis bounded_subset interior_subset) lemma bounded_closure[intro]: assumes "bounded S" shows "bounded (closure S)" proof - from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto { fix y assume "y \ closure S" then obtain f where f: "\n. f n \ S" "(f \ y) sequentially" unfolding closure_sequential by auto have "\n. f n \ S \ dist x (f n) \ a" using a by simp then have "eventually (\n. dist x (f n) \ a) sequentially" by (simp add: f(1)) then have "dist x y \ a" using Lim_dist_ubound f(2) trivial_limit_sequentially by blast } then show ?thesis unfolding bounded_def by auto qed lemma bounded_closure_image: "bounded (f ` closure S) \ bounded (f ` S)" by (simp add: bounded_subset closure_subset image_mono) lemma bounded_cball[simp,intro]: "bounded (cball x e)" unfolding bounded_def using mem_cball by blast lemma bounded_ball[simp,intro]: "bounded (ball x e)" by (metis ball_subset_cball bounded_cball bounded_subset) lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" by (auto simp: bounded_def) (metis Un_iff bounded_any_center le_max_iff_disj) lemma bounded_Union[intro]: "finite F \ \S\F. bounded S \ bounded (\F)" by (induct rule: finite_induct[of F]) auto lemma bounded_UN [intro]: "finite A \ \x\A. bounded (B x) \ bounded (\x\A. B x)" by auto lemma bounded_insert [simp]: "bounded (insert x S) \ bounded S" proof - have "\y\{x}. dist x y \ 0" by simp then have "bounded {x}" unfolding bounded_def by fast then show ?thesis by (metis insert_is_Un bounded_Un) qed lemma bounded_subset_ballI: "S \ ball x r \ bounded S" by (meson bounded_ball bounded_subset) lemma bounded_subset_ballD: assumes "bounded S" shows "\r. 0 < r \ S \ ball x r" proof - obtain e::real and y where "S \ cball y e" "0 \ e" using assms by (auto simp: bounded_subset_cball) then show ?thesis by (intro exI[where x="dist x y + e + 1"]) metric qed lemma finite_imp_bounded [intro]: "finite S \ bounded S" by (induct set: finite) simp_all lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" by (metis Int_lower1 Int_lower2 bounded_subset) lemma bounded_diff[intro]: "bounded S \ bounded (S - T)" by (metis Diff_subset bounded_subset) lemma bounded_dist_comp: assumes "bounded (f ` S)" "bounded (g ` S)" shows "bounded ((\x. dist (f x) (g x)) ` S)" proof - from assms obtain M1 M2 where *: "dist (f x) undefined \ M1" "dist undefined (g x) \ M2" if "x \ S" for x by (auto simp: bounded_any_center[of _ undefined] dist_commute) have "dist (f x) (g x) \ M1 + M2" if "x \ S" for x using *[OF that] by metric then show ?thesis by (auto intro!: boundedI) qed lemma bounded_Times: assumes "bounded s" "bounded t" shows "bounded (s \ t)" proof - obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" using assms [unfolded bounded_def] by auto then have "\z\s \ t. dist (x, y) z \ sqrt (a\<^sup>2 + b\<^sup>2)" by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto qed subsection \Compactness\ lemma compact_imp_bounded: assumes "compact U" shows "bounded U" proof - have "compact U" "\x\U. open (ball x 1)" "U \ (\x\U. ball x 1)" using assms by auto then obtain D where D: "D \ U" "finite D" "U \ (\x\D. ball x 1)" by (metis compactE_image) from \finite D\ have "bounded (\x\D. ball x 1)" by (simp add: bounded_UN) then show "bounded U" using \U \ (\x\D. ball x 1)\ by (rule bounded_subset) qed lemma continuous_on_compact_bound: assumes "compact A" "continuous_on A f" obtains B where "B \ 0" "\x. x \ A \ norm (f x) \ B" proof - have "compact (f ` A)" by (metis assms compact_continuous_image) then obtain B where "\x\A. norm (f x) \ B" by (auto dest!: compact_imp_bounded simp: bounded_iff) hence "max B 0 \ 0" and "\x\A. norm (f x) \ max B 0" by auto thus ?thesis using that by blast qed lemma closure_Int_ball_not_empty: assumes "S \ closure T" "x \ S" "r > 0" shows "T \ ball x r \ {}" using assms centre_in_ball closure_iff_nhds_not_empty by blast lemma compact_sup_maxdistance: fixes S :: "'a::metric_space set" assumes "compact S" and "S \ {}" shows "\x\S. \y\S. \u\S. \v\S. dist u v \ dist x y" proof - have "compact (S \ S)" using \compact S\ by (intro compact_Times) moreover have "S \ S \ {}" using \S \ {}\ by auto moreover have "continuous_on (S \ S) (\x. dist (fst x) (snd x))" by (intro continuous_at_imp_continuous_on ballI continuous_intros) ultimately show ?thesis using continuous_attains_sup[of "S \ S" "\x. dist (fst x) (snd x)"] by auto qed text \ If \A\ is a compact subset of an open set \B\ in a metric space, then there exists an \\ > 0\ such that the Minkowski sum of \A\ with an open ball of radius \\\ is also a subset of \B\. \ lemma compact_subset_open_imp_ball_epsilon_subset: assumes "compact A" "open B" "A \ B" obtains e where "e > 0" "(\x\A. ball x e) \ B" proof - have "\x\A. \e. e > 0 \ ball x e \ B" using assms unfolding open_contains_ball by blast then obtain e where e: "\x. x \ A \ e x > 0" "\x. x \ A \ ball x (e x) \ B" by metis define C where "C = e ` A" obtain X where X: "X \ A" "finite X" "A \ (\c\X. ball c (e c / 2))" using assms(1) proof (rule compactE_image) show "open (ball x (e x / 2))" if "x \ A" for x by simp show "A \ (\c\A. ball c (e c / 2))" using e by auto qed auto define e' where "e' = Min (insert 1 ((\x. e x / 2) ` X))" have "e' > 0" unfolding e'_def using e X by (subst Min_gr_iff) auto have e': "e' \ e x / 2" if "x \ X" for x using that X unfolding e'_def by (intro Min.coboundedI) auto show ?thesis proof show "e' > 0" by fact next show "(\x\A. ball x e') \ B" proof clarify fix x y assume xy: "x \ A" "y \ ball x e'" from xy(1) X obtain z where z: "z \ X" "x \ ball z (e z / 2)" by auto have "dist y z \ dist x y + dist z x" by (metis dist_commute dist_triangle) also have "dist z x < e z / 2" using xy z by auto also have "dist x y < e'" using xy by auto also have "\ \ e z / 2" using z by (intro e') auto finally have "y \ ball z (e z)" by (simp add: dist_commute) also have "\ \ B" using z X by (intro e) auto finally show "y \ B" . qed qed qed lemma compact_subset_open_imp_cball_epsilon_subset: assumes "compact A" "open B" "A \ B" obtains e where "e > 0" "(\x\A. cball x e) \ B" proof - obtain e where "e > 0" and e: "(\x\A. ball x e) \ B" using compact_subset_open_imp_ball_epsilon_subset [OF assms] by blast then have "(\x\A. cball x (e / 2)) \ (\x\A. ball x e)" by auto with \0 < e\ that show ?thesis by (metis e half_gt_zero_iff order_trans) qed subsubsection\Totally bounded\ -lemma cauchy_def: "Cauchy S \ (\e>0. \N. \m n. m \ N \ n \ N \ dist (S m) (S n) < e)" - unfolding Cauchy_def by metis - proposition seq_compact_imp_totally_bounded: assumes "seq_compact S" shows "\e>0. \k. finite k \ k \ S \ S \ (\x\k. ball x e)" proof - { fix e::real assume "e > 0" assume *: "\k. finite k \ k \ S \ \ S \ (\x\k. ball x e)" let ?Q = "\x n r. r \ S \ (\m < (n::nat). \ (dist (x m) r < e))" have "\x. \n::nat. ?Q x n (x n)" proof (rule dependent_wellorder_choice) fix n x assume "\y. y < n \ ?Q x y (x y)" then have "\ S \ (\x\x ` {0..S" "z \ (\x\x ` {0..r. ?Q x n r" using z by auto qed simp then obtain x where "\n::nat. x n \ S" and x:"\n m. m < n \ \ (dist (x m) (x n) < e)" by blast then obtain l r where "l \ S" and r:"strict_mono r" and "((x \ r) \ l) sequentially" using assms by (metis seq_compact_def) then have "Cauchy (x \ r)" using LIMSEQ_imp_Cauchy by auto then obtain N::nat where "\m n. N \ m \ N \ n \ dist ((x \ r) m) ((x \ r) n) < e" - unfolding cauchy_def using \e > 0\ by blast + unfolding Cauchy_def using \e > 0\ by blast then have False using x[of "r N" "r (N+1)"] r by (auto simp: strict_mono_def) } then show ?thesis by metis qed subsubsection\Heine-Borel theorem\ proposition seq_compact_imp_Heine_Borel: fixes S :: "'a :: metric_space set" assumes "seq_compact S" shows "compact S" proof - from seq_compact_imp_totally_bounded[OF \seq_compact S\] obtain f where f: "\e>0. finite (f e) \ f e \ S \ S \ (\x\f e. ball x e)" unfolding choice_iff' .. define K where "K = (\(x, r). ball x r) ` ((\e \ \ \ {0 <..}. f e) \ \)" have "countably_compact S" using \seq_compact S\ by (rule seq_compact_imp_countably_compact) then show "compact S" proof (rule countably_compact_imp_compact) show "countable K" unfolding K_def using f by (auto intro: countable_finite countable_subset countable_rat intro!: countable_image countable_SIGMA countable_UN) show "\b\K. open b" by (auto simp: K_def) next fix T x assume T: "open T" "x \ T" and x: "x \ S" from openE[OF T] obtain e where "0 < e" "ball x e \ T" by auto then have "0 < e/2" "ball x (e/2) \ T" by auto from Rats_dense_in_real[OF \0 < e/2\] obtain r where "r \ \" "0 < r" "r < e/2" by auto from f[rule_format, of r] \0 < r\ \x \ S\ obtain k where "k \ f r" "x \ ball k r" by auto from \r \ \\ \0 < r\ \k \ f r\ have "ball k r \ K" by (auto simp: K_def) then show "\b\K. x \ b \ b \ S \ T" proof (rule bexI[rotated], safe) fix y assume "y \ ball k r" with \r < e/2\ \x \ ball k r\ have "dist x y < e" by (intro dist_triangle_half_r [of k _ e]) (auto simp: dist_commute) with \ball x e \ T\ show "y \ T" by auto next show "x \ ball k r" by fact qed qed qed proposition compact_eq_seq_compact_metric: "compact (S :: 'a::metric_space set) \ seq_compact S" using compact_imp_seq_compact seq_compact_imp_Heine_Borel by blast proposition compact_def: \ \this is the definition of compactness in HOL Light\ "compact (S :: 'a::metric_space set) \ (\f. (\n. f n \ S) \ (\l\S. \r::nat\nat. strict_mono r \ (f \ r) \ l))" unfolding compact_eq_seq_compact_metric seq_compact_def by auto subsubsection \Complete the chain of compactness variants\ proposition compact_eq_Bolzano_Weierstrass: fixes S :: "'a::metric_space set" shows "compact S \ (\T. infinite T \ T \ S \ (\x \ S. x islimpt T))" using Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass compact_eq_seq_compact_metric by blast proposition Bolzano_Weierstrass_imp_bounded: "(\T. \infinite T; T \ S\ \ (\x \ S. x islimpt T)) \ bounded S" using compact_imp_bounded unfolding compact_eq_Bolzano_Weierstrass by metis subsection \Banach fixed point theorem\ theorem banach_fix:\ \TODO: rename to \Banach_fix\\ assumes s: "complete s" "s \ {}" and c: "0 \ c" "c < 1" and f: "f ` s \ s" and lipschitz: "\x\s. \y\s. dist (f x) (f y) \ c * dist x y" shows "\!x\s. f x = x" proof - from c have "1 - c > 0" by simp from s(2) obtain z0 where z0: "z0 \ s" by blast define z where "z n = (f ^^ n) z0" for n with f z0 have z_in_s: "z n \ s" for n :: nat by (induct n) auto define d where "d = dist (z 0) (z 1)" have fzn: "f (z n) = z (Suc n)" for n by (simp add: z_def) have cf_z: "dist (z n) (z (Suc n)) \ (c ^ n) * d" for n :: nat proof (induct n) case 0 then show ?case by (simp add: d_def) next case (Suc m) with \0 \ c\ have "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp then show ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] by (simp add: fzn mult_le_cancel_left) qed have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \ (c ^ m) * d * (1 - c ^ n)" for n m :: nat proof (induct n) case 0 show ?case by simp next case (Suc k) from c have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" by (simp add: dist_triangle) also from c cf_z[of "m + k"] have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" by simp also from Suc have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" by (simp add: field_simps) also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" by (simp add: power_add field_simps) also from c have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" by (simp add: field_simps) finally show ?case by simp qed have "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" if "e > 0" for e proof (cases "d = 0") case True from \1 - c > 0\ have "(1 - c) * x \ 0 \ x \ 0" for x by (simp add: mult_le_0_iff) with c cf_z2[of 0] True have "z n = z0" for n by (simp add: z_def) with \e > 0\ show ?thesis by simp next case False with zero_le_dist[of "z 0" "z 1"] have "d > 0" by (metis d_def less_le) with \1 - c > 0\ \e > 0\ have "0 < e * (1 - c) / d" by simp with c obtain N where N: "c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto have *: "dist (z m) (z n) < e" if "m > n" and as: "m \ N" "n \ N" for m n :: nat proof - from c \n \ N\ have *: "c ^ n \ c ^ N" using power_decreasing[OF \n\N\, of c] by simp from c \m > n\ have "1 - c ^ (m - n) > 0" using power_strict_mono[of c 1 "m - n"] by simp with \d > 0\ \0 < 1 - c\ have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" by simp from cf_z2[of n "m - n"] \m > n\ have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" by (simp add: pos_le_divide_eq[OF \1 - c > 0\] mult.commute dist_commute) also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" using mult_right_mono[OF * order_less_imp_le[OF **]] by (simp add: mult.assoc) also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) also from c \d > 0\ \1 - c > 0\ have "\ = e * (1 - c ^ (m - n))" by simp also from c \1 - c ^ (m - n) > 0\ \e > 0\ have "\ \ e" using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto finally show ?thesis by simp qed have "dist (z n) (z m) < e" if "N \ m" "N \ n" for m n :: nat proof (cases "n = m") case True with \e > 0\ show ?thesis by simp next case False with *[of n m] *[of m n] and that show ?thesis by (auto simp: dist_commute nat_neq_iff) qed then show ?thesis by auto qed then have "Cauchy z" - by (simp add: cauchy_def) + by (metis metric_CauchyI) then obtain x where "x\s" and x:"(z \ x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto define e where "e = dist (f x) x" have "e = 0" proof (rule ccontr) assume "e \ 0" then have "e > 0" unfolding e_def using zero_le_dist[of "f x" x] by (metis dist_eq_0_iff dist_nz e_def) then obtain N where N:"\n\N. dist (z n) x < e/2" using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto then have N':"dist (z N) x < e/2" by auto have *: "c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 using zero_le_dist[of "z N" x] and c by (metis dist_eq_0_iff dist_nz order_less_asym less_le) have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] using z_in_s[of N] \x\s\ using c by auto also have "\ < e/2" using N' and c using * by auto finally show False unfolding fzn using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] unfolding e_def by auto qed then have "f x = x" by (auto simp: e_def) moreover have "y = x" if "f y = y" "y \ s" for y proof - from \x \ s\ \f x = x\ that have "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp with c and zero_le_dist[of x y] have "dist x y = 0" by (simp add: mult_le_cancel_right1) then show ?thesis by simp qed ultimately show ?thesis using \x\s\ by blast qed subsection \Edelstein fixed point theorem\ theorem Edelstein_fix: fixes S :: "'a::metric_space set" assumes S: "compact S" "S \ {}" and gs: "(g ` S) \ S" and dist: "\x\S. \y\S. x \ y \ dist (g x) (g y) < dist x y" shows "\!x\S. g x = x" proof - let ?D = "(\x. (x, x)) ` S" have D: "compact ?D" "?D \ {}" by (rule compact_continuous_image) (auto intro!: S continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) have "\x y e. x \ S \ y \ S \ 0 < e \ dist y x < e \ dist (g y) (g x) < e" using dist by fastforce then have "continuous_on S g" by (auto simp: continuous_on_iff) then have cont: "continuous_on ?D (\x. dist ((g \ fst) x) (snd x))" unfolding continuous_on_eq_continuous_within by (intro continuous_dist ballI continuous_within_compose) (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) obtain a where "a \ S" and le: "\x. x \ S \ dist (g a) a \ dist (g x) x" using continuous_attains_inf[OF D cont] by auto have "g a = a" proof (rule ccontr) assume "g a \ a" with \a \ S\ gs have "dist (g (g a)) (g a) < dist (g a) a" by (intro dist[rule_format]) auto moreover have "dist (g a) a \ dist (g (g a)) (g a)" using \a \ S\ gs by (intro le) auto ultimately show False by auto qed moreover have "\x. x \ S \ g x = x \ x = a" using dist[THEN bspec[where x=a]] \g a = a\ and \a\S\ by auto ultimately show "\!x\S. g x = x" using \a \ S\ by blast qed subsection \The diameter of a set\ definition\<^marker>\tag important\ diameter :: "'a::metric_space set \ real" where "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" lemma diameter_empty [simp]: "diameter{} = 0" by (auto simp: diameter_def) lemma diameter_singleton [simp]: "diameter{x} = 0" by (auto simp: diameter_def) lemma diameter_le: assumes "S \ {} \ 0 \ d" and no: "\x y. \x \ S; y \ S\ \ norm(x - y) \ d" shows "diameter S \ d" using assms by (auto simp: dist_norm diameter_def intro: cSUP_least) lemma diameter_bounded_bound: fixes S :: "'a :: metric_space set" assumes S: "bounded S" "x \ S" "y \ S" shows "dist x y \ diameter S" proof - from S obtain z d where z: "\x. x \ S \ dist z x \ d" unfolding bounded_def by auto have "bdd_above (case_prod dist ` (S\S))" proof (intro bdd_aboveI, safe) fix a b assume "a \ S" "b \ S" with z[of a] z[of b] dist_triangle[of a b z] show "dist a b \ 2 * d" by (simp add: dist_commute) qed moreover have "(x,y) \ S\S" using S by auto ultimately have "dist x y \ (SUP (x,y)\S\S. dist x y)" by (rule cSUP_upper2) simp with \x \ S\ show ?thesis by (auto simp: diameter_def) qed lemma diameter_lower_bounded: fixes S :: "'a :: metric_space set" assumes S: "bounded S" and d: "0 < d" "d < diameter S" shows "\x\S. \y\S. d < dist x y" proof (rule ccontr) assume contr: "\ ?thesis" moreover have "S \ {}" using d by (auto simp: diameter_def) ultimately have "diameter S \ d" by (auto simp: not_less diameter_def intro!: cSUP_least) with \d < diameter S\ show False by auto qed lemma diameter_bounded: assumes "bounded S" shows "\x\S. \y\S. dist x y \ diameter S" and "\d>0. d < diameter S \ (\x\S. \y\S. dist x y > d)" using diameter_bounded_bound[of S] diameter_lower_bounded[of S] assms by auto lemma bounded_two_points: "bounded S \ (\e. \x\S. \y\S. dist x y \ e)" by (meson bounded_def diameter_bounded(1)) lemma diameter_compact_attained: assumes "compact S" and "S \ {}" shows "\x\S. \y\S. dist x y = diameter S" proof - have b: "bounded S" using assms(1) by (rule compact_imp_bounded) then obtain x y where xys: "x\S" "y\S" and xy: "\u\S. \v\S. dist u v \ dist x y" using compact_sup_maxdistance[OF assms] by auto then have "diameter S \ dist x y" unfolding diameter_def by (force intro!: cSUP_least) then show ?thesis by (metis b diameter_bounded_bound order_antisym xys) qed lemma diameter_ge_0: assumes "bounded S" shows "0 \ diameter S" by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) lemma diameter_subset: assumes "S \ T" "bounded T" shows "diameter S \ diameter T" proof (cases "S = {} \ T = {}") case True with assms show ?thesis by (force simp: diameter_ge_0) next case False then have "bdd_above ((\x. case x of (x, xa) \ dist x xa) ` (T \ T))" using \bounded T\ diameter_bounded_bound by (force simp: bdd_above_def) with False \S \ T\ show ?thesis apply (simp add: diameter_def) apply (rule cSUP_subset_mono, auto) done qed lemma diameter_closure: assumes "bounded S" shows "diameter(closure S) = diameter S" proof (rule order_antisym) have "False" if d_less_d: "diameter S < diameter (closure S)" proof - define d where "d = diameter(closure S) - diameter(S)" have "d > 0" using that by (simp add: d_def) then have dle: "diameter(closure(S)) - d / 2 < diameter(closure(S))" by simp have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" by (simp add: d_def field_split_simps) have bocl: "bounded (closure S)" using assms by blast moreover have "0 \ diameter S" using assms diameter_ge_0 by blast ultimately obtain x y where "x \ closure S" "y \ closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" by (smt (verit) dle d_less_d d_def dd diameter_lower_bounded) then obtain x' y' where x'y': "x' \ S" "dist x' x < d/4" "y' \ S" "dist y' y < d/4" by (metis \0 < d\ zero_less_divide_iff zero_less_numeral closure_approachable) then have "dist x' y' \ diameter S" using assms diameter_bounded_bound by blast with x'y' have "dist x y \ d / 4 + diameter S + d / 4" by (meson add_mono dist_triangle dist_triangle3 less_eq_real_def order_trans) then show ?thesis using xy d_def by linarith qed then show "diameter (closure S) \ diameter S" by fastforce next show "diameter S \ diameter (closure S)" by (simp add: assms bounded_closure closure_subset diameter_subset) qed proposition Lebesgue_number_lemma: assumes "compact S" "\ \ {}" "S \ \\" and ope: "\B. B \ \ \ open B" obtains \ where "0 < \" "\T. \T \ S; diameter T < \\ \ \B \ \. T \ B" proof (cases "S = {}") case True then show ?thesis by (metis \\ \ {}\ zero_less_one empty_subsetI equals0I subset_trans that) next case False { fix x assume "x \ S" then obtain C where C: "x \ C" "C \ \" using \S \ \\\ by blast then obtain r where r: "r>0" "ball x (2*r) \ C" by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) then have "\r C. r > 0 \ ball x (2*r) \ C \ C \ \" using C by blast } then obtain r where r: "\x. x \ S \ r x > 0 \ (\C \ \. ball x (2*r x) \ C)" by metis then have "S \ (\x \ S. ball x (r x))" by auto then obtain \ where "finite \" "S \ \\" and \: "\ \ (\x. ball x (r x)) ` S" by (rule compactE [OF \compact S\]) auto then obtain S0 where "S0 \ S" "finite S0" and S0: "\ = (\x. ball x (r x)) ` S0" by (meson finite_subset_image) then have "S0 \ {}" using False \S \ \\\ by auto define \ where "\ = Inf (r ` S0)" have "\ > 0" using \finite S0\ \S0 \ S\ \S0 \ {}\ r by (auto simp: \_def finite_less_Inf_iff) show ?thesis proof show "0 < \" by (simp add: \0 < \\) show "\B \ \. T \ B" if "T \ S" and dia: "diameter T < \" for T proof (cases "T = {}") case True then show ?thesis using \\ \ {}\ by blast next case False then obtain y where "y \ T" by blast then have "y \ S" using \T \ S\ by auto then obtain x where "x \ S0" and x: "y \ ball x (r x)" using \S \ \\\ S0 that by blast have "ball y \ \ ball y (r x)" by (metis \_def \S0 \ {}\ \finite S0\ \x \ S0\ empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball) also have "... \ ball x (2*r x)" using x by metric finally obtain C where "C \ \" "ball y \ \ C" by (meson r \S0 \ S\ \x \ S0\ dual_order.trans subsetCE) have "bounded T" using \compact S\ bounded_subset compact_imp_bounded \T \ S\ by blast then have "T \ ball y \" using \y \ T\ dia diameter_bounded_bound by fastforce then show ?thesis by (meson \C \ \\ \ball y \ \ C\ subset_eq) qed qed qed subsection \Metric spaces with the Heine-Borel property\ text \ A metric space (or topological vector space) is said to have the Heine-Borel property if every closed and bounded subset is compact. \ class heine_borel = metric_space + assumes bounded_imp_convergent_subsequence: "bounded (range f) \ \l r. strict_mono (r::nat\nat) \ ((f \ r) \ l) sequentially" proposition bounded_closed_imp_seq_compact: fixes S::"'a::heine_borel set" assumes "bounded S" and "closed S" shows "seq_compact S" proof (unfold seq_compact_def, clarify) fix f :: "nat \ 'a" assume f: "\n. f n \ S" with \bounded S\ have "bounded (range f)" by (auto intro: bounded_subset) obtain l r where r: "strict_mono (r :: nat \ nat)" and l: "((f \ r) \ l) sequentially" using bounded_imp_convergent_subsequence [OF \bounded (range f)\] by auto from f have fr: "\n. (f \ r) n \ S" by simp have "l \ S" using \closed S\ fr l by (rule closed_sequentially) show "\l\S. \r. strict_mono r \ ((f \ r) \ l) sequentially" using \l \ S\ r l by blast qed lemma compact_eq_bounded_closed: fixes S :: "'a::heine_borel set" shows "compact S \ bounded S \ closed S" using bounded_closed_imp_seq_compact compact_eq_seq_compact_metric compact_imp_bounded compact_imp_closed by auto lemma bounded_infinite_imp_islimpt: fixes S :: "'a::heine_borel set" assumes "T \ S" "bounded S" "infinite T" obtains x where "x islimpt S" by (meson assms closed_limpt compact_eq_Bolzano_Weierstrass compact_eq_bounded_closed islimpt_subset) lemma compact_Inter: fixes \ :: "'a :: heine_borel set set" assumes com: "\S. S \ \ \ compact S" and "\ \ {}" shows "compact(\ \)" using assms by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed) lemma compact_closure [simp]: fixes S :: "'a::heine_borel set" shows "compact(closure S) \ bounded S" by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed) instance\<^marker>\tag important\ real :: heine_borel proof fix f :: "nat \ real" assume f: "bounded (range f)" obtain r :: "nat \ nat" where r: "strict_mono r" "monoseq (f \ r)" unfolding comp_def by (metis seq_monosub) then have "Bseq (f \ r)" unfolding Bseq_eq_bounded by (metis f BseqI' bounded_iff comp_apply rangeI) with r show "\l r. strict_mono r \ (f \ r) \ l" using Bseq_monoseq_convergent[of "f \ r"] by (auto simp: convergent_def) qed lemma compact_lemma_general: fixes f :: "nat \ 'a" fixes proj::"'a \ 'b \ 'c::heine_borel" (infixl "proj" 60) fixes unproj:: "('b \ 'c) \ 'a" assumes finite_basis: "finite basis" assumes bounded_proj: "\k. k \ basis \ bounded ((\x. x proj k) ` range f)" assumes proj_unproj: "\e k. k \ basis \ (unproj e) proj k = e k" assumes unproj_proj: "\x. unproj (\k. x proj k) = x" shows "\d\basis. \l::'a. \ r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof safe fix d :: "'b set" assume d: "d \ basis" with finite_basis have "finite d" by (blast intro: finite_subset) from this d show "\l::'a. \r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof (induct d) case empty then show ?case unfolding strict_mono_def by auto next case (insert k d) have k[intro]: "k \ basis" using insert by auto have s': "bounded ((\x. x proj k) ` range f)" using k by (rule bounded_proj) obtain l1::"'a" and r1 where r1: "strict_mono r1" and lr1: "\e > 0. \\<^sub>F n in sequentially. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e" using insert by auto have f': "\n. f (r1 n) proj k \ (\x. x proj k) ` range f" by simp have "bounded (range (\i. f (r1 i) proj k))" by (metis (lifting) bounded_subset f' image_subsetI s') then obtain l2 r2 where r2: "strict_mono r2" and lr2: "(\i. f (r1 (r2 i)) proj k) \ l2" using bounded_imp_convergent_subsequence[of "\i. f (r1 i) proj k"] by (auto simp: o_def) define r where "r = r1 \ r2" have r: "strict_mono r" using r1 and r2 unfolding r_def o_def strict_mono_def by auto moreover define l where "l = unproj (\i. if i = k then l2 else l1 proj i)" { fix e::real assume "e > 0" from lr1 \e > 0\ have N1: "\\<^sub>F n in sequentially. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e" by blast from lr2 \e > 0\ have N2: "\\<^sub>F n in sequentially. dist (f (r1 (r2 n)) proj k) l2 < e" by (rule tendstoD) from r2 N1 have N1': "\\<^sub>F n in sequentially. \i\d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e" by (rule eventually_subseq) have "\\<^sub>F n in sequentially. \i\insert k d. dist (f (r n) proj i) (l proj i) < e" using N1' N2 by eventually_elim (use insert.prems in \auto simp: l_def r_def o_def proj_unproj\) } ultimately show ?case by auto qed qed lemma bounded_fst: "bounded s \ bounded (fst ` s)" unfolding bounded_def by (metis (erased, opaque_lifting) dist_fst_le image_iff order_trans) lemma bounded_snd: "bounded s \ bounded (snd ` s)" unfolding bounded_def by (metis (no_types, opaque_lifting) dist_snd_le image_iff order.trans) instance\<^marker>\tag important\ prod :: (heine_borel, heine_borel) heine_borel proof fix f :: "nat \ 'a \ 'b" assume f: "bounded (range f)" then have "bounded (fst ` range f)" by (rule bounded_fst) then have s1: "bounded (range (fst \ f))" by (simp add: image_comp) obtain l1 r1 where r1: "strict_mono r1" and l1: "(\n. fst (f (r1 n))) \ l1" using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast from f have s2: "bounded (range (snd \ f \ r1))" by (auto simp: image_comp intro: bounded_snd bounded_subset) obtain l2 r2 where r2: "strict_mono r2" and l2: "(\n. snd (f (r1 (r2 n)))) \ l2" using bounded_imp_convergent_subsequence [OF s2] unfolding o_def by fast have l1': "((\n. fst (f (r1 (r2 n)))) \ l1) sequentially" using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def . have l: "((f \ (r1 \ r2)) \ (l1, l2)) sequentially" using tendsto_Pair [OF l1' l2] unfolding o_def by simp have r: "strict_mono (r1 \ r2)" using r1 r2 unfolding strict_mono_def by simp show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" using l r by fast qed subsection \Completeness\ proposition (in metric_space) completeI: assumes "\f. \n. f n \ s \ Cauchy f \ \l\s. f \ l" shows "complete s" using assms unfolding complete_def by fast proposition (in metric_space) completeE: assumes "complete s" and "\n. f n \ s" and "Cauchy f" obtains l where "l \ s" and "f \ l" using assms unfolding complete_def by fast (* TODO: generalize to uniform spaces *) lemma compact_imp_complete: fixes s :: "'a::metric_space set" assumes "compact s" shows "complete s" proof - { fix f assume as: "(\n::nat. f n \ s)" "Cauchy f" from as(1) obtain l r where lr: "l\s" "strict_mono r" "(f \ r) \ l" using assms unfolding compact_def by blast note lr' = seq_suble [OF lr(2)] { fix e :: real assume "e > 0" from as(2) obtain N where N:"\m n. N \ m \ N \ n \ dist (f m) (f n) < e/2" - unfolding cauchy_def using \e > 0\ by (meson half_gt_zero) + unfolding Cauchy_def using \e > 0\ by (meson half_gt_zero) then obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" by (metis dist_self lim_sequentially lr(3)) { fix n :: nat assume n: "n \ max N M" have "dist ((f \ r) n) l < e/2" using n M by auto moreover have "r n \ N" using lr'[of n] n by auto then have "dist (f n) ((f \ r) n) < e/2" using N and n by auto ultimately have "dist (f n) l < e" using n M by metric } then have "\N. \n\N. dist (f n) l < e" by blast } then have "\l\s. (f \ l) sequentially" using \l\s\ unfolding lim_sequentially by auto } then show ?thesis unfolding complete_def by auto qed proposition compact_eq_totally_bounded: "compact s \ complete s \ (\e>0. \k. finite k \ s \ (\x\k. ball x e))" (is "_ \ ?rhs") proof assume assms: "?rhs" then obtain k where k: "\e. 0 < e \ finite (k e)" "\e. 0 < e \ s \ (\x\k e. ball x e)" by (auto simp: choice_iff') show "compact s" proof cases assume "s = {}" then show "compact s" by (simp add: compact_def) next assume "s \ {}" show ?thesis unfolding compact_def proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ s" define e where "e n = 1 / (2 * Suc n)" for n then have [simp]: "\n. 0 < e n" by auto define B where "B n U = (SOME b. infinite {n. f n \ b} \ (\x. b \ ball x (e n) \ U))" for n U { fix n U assume "infinite {n. f n \ U}" then have "\b\k (e n). infinite {i\{n. f n \ U}. f i \ ball b (e n)}" using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq) then obtain a where "a \ k (e n)" "infinite {i \ {n. f n \ U}. f i \ ball a (e n)}" .. then have "\b. infinite {i. f i \ b} \ (\x. b \ ball x (e n) \ U)" by (intro exI[of _ "ball a (e n) \ U"] exI[of _ a]) (auto simp: ac_simps) from someI_ex[OF this] have "infinite {i. f i \ B n U}" "\x. B n U \ ball x (e n) \ U" unfolding B_def by auto } note B = this define F where "F = rec_nat (B 0 UNIV) B" { fix n have "infinite {i. f i \ F n}" by (induct n) (auto simp: F_def B) } then have F: "\n. \x. F (Suc n) \ ball x (e n) \ F n" using B by (simp add: F_def) then have F_dec: "\m n. m \ n \ F n \ F m" using decseq_SucI[of F] by (auto simp: decseq_def) obtain sel where sel: "\k i. i < sel k i" "\k i. f (sel k i) \ F k" proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI) fix k i have "infinite ({n. f n \ F k} - {.. i})" using \infinite {n. f n \ F k}\ by auto from infinite_imp_nonempty[OF this] show "\x>i. f x \ F k" by (simp add: set_eq_iff not_le conj_commute) qed define t where "t = rec_nat (sel 0 0) (\n i. sel (Suc n) i)" have "strict_mono t" unfolding strict_mono_Suc_iff by (simp add: t_def sel) moreover have "\i. (f \ t) i \ s" using f by auto moreover have t: "(f \ t) n \ F n" for n by (cases n) (simp_all add: t_def sel) have "Cauchy (f \ t)" proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE) fix r :: real and N n m assume "1 / Suc N < r" "Suc N \ n" "Suc N \ m" then have "(f \ t) n \ F (Suc N)" "(f \ t) m \ F (Suc N)" "2 * e N < r" using F_dec t by (auto simp: e_def field_simps) with F[of N] obtain x where "dist x ((f \ t) n) < e N" "dist x ((f \ t) m) < e N" by (auto simp: subset_eq) with \2 * e N < r\ show "dist ((f \ t) m) ((f \ t) n) < r" by metric qed ultimately show "\l\s. \r. strict_mono r \ (f \ r) \ l" using assms unfolding complete_def by blast qed qed qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded) lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)" proof - from assms obtain N :: nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" - unfolding cauchy_def by force + by (meson Cauchy_def zero_less_one) then have N:"\n. N \ n \ dist (s N) (s n) < 1" by auto moreover have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" unfolding bounded_any_center [where a="s N"] by auto ultimately show "?thesis" unfolding bounded_any_center [where a="s N"] apply (rule_tac x="max a 1" in exI, auto) apply (erule_tac x=y in allE) apply (erule_tac x=y in ballE, auto) done qed instance heine_borel < complete_space proof fix f :: "nat \ 'a" assume "Cauchy f" then show "convergent f" unfolding convergent_def using Cauchy_converges_subseq cauchy_imp_bounded bounded_imp_convergent_subsequence by blast qed lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)" proof (rule completeI) fix f :: "nat \ 'a" assume "Cauchy f" then show "\l\UNIV. f \ l" using Cauchy_convergent convergent_def by blast qed lemma complete_imp_closed: fixes S :: "'a::metric_space set" shows "complete S \ closed S" by (metis LIMSEQ_imp_Cauchy LIMSEQ_unique closed_sequential_limits completeE) lemma complete_Int_closed: fixes S :: "'a::metric_space set" assumes "complete S" and "closed t" shows "complete (S \ t)" by (metis Int_iff assms closed_sequentially completeE completeI) lemma complete_closed_subset: fixes S :: "'a::metric_space set" assumes "closed S" and "S \ t" and "complete t" shows "complete S" using assms complete_Int_closed [of t S] by (simp add: Int_absorb1) lemma complete_eq_closed: fixes S :: "('a::complete_space) set" shows "complete S \ closed S" proof assume "closed S" then show "complete S" using subset_UNIV complete_UNIV by (rule complete_closed_subset) next assume "complete S" then show "closed S" by (rule complete_imp_closed) qed lemma convergent_eq_Cauchy: fixes S :: "nat \ 'a::complete_space" shows "(\l. (S \ l) sequentially) \ Cauchy S" unfolding Cauchy_convergent_iff convergent_def .. lemma convergent_imp_bounded: fixes S :: "nat \ 'a::metric_space" shows "(S \ l) sequentially \ bounded (range S)" by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy) lemma frontier_subset_compact: fixes S :: "'a::heine_borel set" shows "compact S \ frontier S \ S" using frontier_subset_closed compact_eq_bounded_closed by blast -lemma continuous_closed_imp_Cauchy_continuous: - fixes S :: "('a::complete_space) set" - shows "\continuous_on S f; closed S; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" - by (meson LIMSEQ_imp_Cauchy completeE complete_eq_closed continuous_on_sequentially) - lemma banach_fix_type: fixes f::"'a::complete_space\'a" assumes c:"0 \ c" "c < 1" and lipschitz:"\x. \y. dist (f x) (f y) \ c * dist x y" shows "\!x. (f x = x)" using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] by auto +subsection \Cauchy continuity\ + +definition Cauchy_continuous_on where + "Cauchy_continuous_on \ \S f. \\. Cauchy \ \ range \ \ S \ Cauchy (f \ \)" + +lemma continuous_closed_imp_Cauchy_continuous: + fixes S :: "('a::complete_space) set" + shows "\continuous_on S f; closed S\ \ Cauchy_continuous_on S f" + unfolding Cauchy_continuous_on_def + by (metis LIMSEQ_imp_Cauchy completeE complete_eq_closed continuous_on_sequentially range_subsetD) + +lemma uniformly_continuous_imp_Cauchy_continuous: + fixes f :: "'a::metric_space \ 'b::metric_space" + shows "uniformly_continuous_on S f \ Cauchy_continuous_on S f" + by (simp add: uniformly_continuous_on_def Cauchy_continuous_on_def Cauchy_def image_subset_iff) metis + +lemma Cauchy_continuous_on_imp_continuous: + fixes f :: "'a::metric_space \ 'b::metric_space" + assumes "Cauchy_continuous_on S f" + shows "continuous_on S f" +proof - + have False if x: "\n. \x'\S. dist x' x < inverse(Suc n) \ \ dist (f x') (f x) < \" "\>0" "x \ S" for x and \::real + proof - + obtain \ where \: "\n. \ n \ S" and dx: "\n. dist (\ n) x < inverse(Suc n)" and dfx: "\n. \ dist (f (\ n)) (f x) < \" + using x by metis + define \ where "\ \ \n. if even n then \ n else x" + with \ \x \ S\ have "range \ \ S" + by auto + have "\ \ x" + unfolding tendsto_iff + proof (intro strip) + fix e :: real + assume "e>0" + then obtain N where "inverse (Suc N) < e" + using reals_Archimedean by blast + then have "\n. N \ n \ dist (\ n) x < e" + by (smt (verit, ccfv_SIG) dx inverse_Suc inverse_less_iff_less inverse_positive_iff_positive of_nat_Suc of_nat_mono) + with \e>0\ show "\\<^sub>F n in sequentially. dist (\ n) x < e" + by (auto simp add: eventually_sequentially \_def) + qed + then have "Cauchy \" + by (intro LIMSEQ_imp_Cauchy) + then have Cf: "Cauchy (f \ \)" + by (meson Cauchy_continuous_on_def \range \ \ S\ assms) + have "(f \ \) \ f x" + unfolding tendsto_iff + proof (intro strip) + fix e :: real + assume "e>0" + then obtain N where N: "\m\N. \n\N. dist ((f \ \) m) ((f \ \) n) < e" + using Cf unfolding Cauchy_def by presburger + moreover have "(f \ \) (Suc(N+N)) = f x" + by (simp add: \_def) + ultimately have "\n\N. dist ((f \ \) n) (f x) < e" + by (metis add_Suc le_add2) + then show "\\<^sub>F n in sequentially. dist ((f \ \) n) (f x) < e" + using eventually_sequentially by blast + qed + moreover have "\n. \ dist (f (\ (2*n))) (f x) < \" + using dfx by (simp add: \_def) + ultimately show False + using \\>0\ by (fastforce simp: mult_2 nat_le_iff_add tendsto_iff eventually_sequentially) + qed + then show ?thesis + unfolding continuous_on_iff by (meson inverse_Suc) +qed + subsection\<^marker>\tag unimportant\\ Finite intersection property\ text\Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\ lemma closed_imp_fip: fixes S :: "'a::heine_borel set" assumes "closed S" and T: "T \ \" "bounded T" and clof: "\T. T \ \ \ closed T" and "\\'. \finite \'; \' \ \\ \ S \ \\' \ {}" shows "S \ \\ \ {}" proof - have "compact (S \ T)" using \closed S\ clof compact_eq_bounded_closed T by blast then have "(S \ T) \ \\ \ {}" by (smt (verit, best) Inf_insert Int_assoc assms compact_imp_fip finite_insert insert_subset) then show ?thesis by blast qed lemma closed_imp_fip_compact: fixes S :: "'a::heine_borel set" shows "\closed S; \T. T \ \ \ compact T; \\'. \finite \'; \' \ \\ \ S \ \\' \ {}\ \ S \ \\ \ {}" by (metis closed_imp_fip compact_eq_bounded_closed equals0I finite.emptyI order.refl) lemma closed_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes "T \ \" "bounded T" and "\T. T \ \ \ closed T" and "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" using closed_imp_fip [OF closed_UNIV] assms by auto lemma compact_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes clof: "\T. T \ \ \ compact T" and none: "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" by (metis InterI clof closed_fip_Heine_Borel compact_eq_bounded_closed equals0D none) lemma compact_sequence_with_limit: fixes f :: "nat \ 'a::heine_borel" shows "f \ l \ compact (insert l (range f))" by (simp add: closed_limpt compact_eq_bounded_closed convergent_imp_bounded islimpt_insert sequence_unique_limpt) subsection \Properties of Balls and Spheres\ lemma compact_cball[simp]: fixes x :: "'a::heine_borel" shows "compact (cball x e)" using compact_eq_bounded_closed bounded_cball closed_cball by blast lemma compact_frontier_bounded[intro]: fixes S :: "'a::heine_borel set" shows "bounded S \ compact (frontier S)" unfolding frontier_def using compact_eq_bounded_closed by blast lemma compact_frontier[intro]: fixes S :: "'a::heine_borel set" shows "compact S \ compact (frontier S)" using compact_eq_bounded_closed compact_frontier_bounded by blast lemma no_limpt_imp_countable: assumes "\z. \z islimpt (X :: 'a :: {real_normed_vector, heine_borel} set)" shows "countable X" proof - have "countable (\n. cball 0 (real n) \ X)" proof (intro countable_UN[OF _ countable_finite]) fix n :: nat show "finite (cball 0 (real n) \ X)" using assms by (intro finite_not_islimpt_in_compact) auto qed auto also have "(\n. cball 0 (real n)) = (UNIV :: 'a set)" proof safe fix z :: 'a have "norm z \ 0" by simp hence "real (nat \norm z\) \ norm z" by linarith thus "z \ (\n. cball 0 (real n))" by auto qed auto hence "(\n. cball 0 (real n) \ X) = X" by blast finally show "countable X" . qed subsection \Distance from a Set\ lemma distance_attains_sup: assumes "compact s" "s \ {}" shows "\x\s. \y\s. dist a y \ dist a x" proof (rule continuous_attains_sup [OF assms]) { fix x assume "x\s" have "(dist a \ dist a x) (at x within s)" by (intro tendsto_dist tendsto_const tendsto_ident_at) } then show "continuous_on s (dist a)" unfolding continuous_on .. qed text \For \emph{minimal} distance, we only need closure, not compactness.\ lemma distance_attains_inf: fixes a :: "'a::heine_borel" assumes "closed s" and "s \ {}" obtains x where "x\s" "\y. y \ s \ dist a x \ dist a y" proof - from assms obtain b where "b \ s" by auto let ?B = "s \ cball a (dist b a)" have "?B \ {}" using \b \ s\ by (auto simp: dist_commute) moreover have "continuous_on ?B (dist a)" by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_ident continuous_const) moreover have "compact ?B" by (intro closed_Int_compact \closed s\ compact_cball) ultimately obtain x where "x \ ?B" "\y\?B. dist a x \ dist a y" by (metis continuous_attains_inf) with that show ?thesis by fastforce qed subsection \Infimum Distance\ definition\<^marker>\tag important\ "infdist x A = (if A = {} then 0 else INF a\A. dist x a)" lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)" by (auto intro!: zero_le_dist) lemma infdist_notempty: "A \ {} \ infdist x A = (INF a\A. dist x a)" by (simp add: infdist_def) lemma infdist_nonneg: "0 \ infdist x A" by (auto simp: infdist_def intro: cINF_greatest) lemma infdist_le: "a \ A \ infdist x A \ dist x a" by (auto intro: cINF_lower simp add: infdist_def) lemma infdist_le2: "a \ A \ dist x a \ d \ infdist x A \ d" by (auto intro!: cINF_lower2 simp add: infdist_def) lemma infdist_zero[simp]: "a \ A \ infdist a A = 0" by (auto intro!: antisym infdist_nonneg infdist_le2) lemma infdist_Un_min: assumes "A \ {}" "B \ {}" shows "infdist x (A \ B) = min (infdist x A) (infdist x B)" using assms by (simp add: infdist_def cINF_union inf_real_def) lemma infdist_triangle: "infdist x A \ infdist y A + dist x y" proof (cases "A = {}") case True then show ?thesis by (simp add: infdist_def) next case False then obtain a where "a \ A" by auto have "infdist x A \ Inf {dist x y + dist y a |a. a \ A}" proof (rule cInf_greatest) from \A \ {}\ show "{dist x y + dist y a |a. a \ A} \ {}" by simp fix d assume "d \ {dist x y + dist y a |a. a \ A}" then obtain a where d: "d = dist x y + dist y a" "a \ A" by auto show "infdist x A \ d" unfolding infdist_notempty[OF \A \ {}\] proof (rule cINF_lower2) show "a \ A" by fact show "dist x a \ d" unfolding d by (rule dist_triangle) qed simp qed also have "\ = dist x y + infdist y A" proof (rule cInf_eq, safe) fix a assume "a \ A" then show "dist x y + infdist y A \ dist x y + dist y a" by (auto intro: infdist_le) next fix i assume inf: "\d. d \ {dist x y + dist y a |a. a \ A} \ i \ d" then have "i - dist x y \ infdist y A" unfolding infdist_notempty[OF \A \ {}\] using \a \ A\ by (intro cINF_greatest) (auto simp: field_simps) then show "i \ dist x y + infdist y A" by simp qed finally show ?thesis by simp qed lemma infdist_triangle_abs: "\infdist x A - infdist y A\ \ dist x y" by (metis (full_types) abs_diff_le_iff diff_le_eq dist_commute infdist_triangle) lemma in_closure_iff_infdist_zero: assumes "A \ {}" shows "x \ closure A \ infdist x A = 0" proof assume "x \ closure A" show "infdist x A = 0" proof (rule ccontr) assume "infdist x A \ 0" with infdist_nonneg[of x A] have "infdist x A > 0" by auto then have "ball x (infdist x A) \ closure A = {}" by (smt (verit, best) \x \ closure A\ closure_approachableD infdist_le) then have "x \ closure A" by (metis \0 < infdist x A\ centre_in_ball disjoint_iff_not_equal) then show False using \x \ closure A\ by simp qed next assume x: "infdist x A = 0" then obtain a where "a \ A" by atomize_elim (metis all_not_in_conv assms) have False if "e > 0" "\ (\y\A. dist y x < e)" for e proof - have "infdist x A \ e" using \a \ A\ unfolding infdist_def using that by (force simp: dist_commute intro: cINF_greatest) with x \e > 0\ show False by auto qed then show "x \ closure A" using closure_approachable by blast qed lemma in_closed_iff_infdist_zero: assumes "closed A" "A \ {}" shows "x \ A \ infdist x A = 0" proof - have "x \ closure A \ infdist x A = 0" by (simp add: \A \ {}\ in_closure_iff_infdist_zero) with assms show ?thesis by simp qed lemma infdist_pos_not_in_closed: assumes "closed S" "S \ {}" "x \ S" shows "infdist x S > 0" by (smt (verit, best) assms in_closed_iff_infdist_zero infdist_nonneg) lemma infdist_attains_inf: fixes X::"'a::heine_borel set" assumes "closed X" assumes "X \ {}" obtains x where "x \ X" "infdist y X = dist y x" proof - have "bdd_below (dist y ` X)" by auto from distance_attains_inf[OF assms, of y] obtain x where "x \ X" "\z. z \ X \ dist y x \ dist y z" by auto then have "infdist y X = dist y x" by (metis antisym assms(2) cINF_greatest infdist_def infdist_le) with \x \ X\ show ?thesis .. qed text \Every metric space is a T4 space:\ instance metric_space \ t4_space proof fix S T::"'a set" assume H: "closed S" "closed T" "S \ T = {}" consider "S = {}" | "T = {}" | "S \ {} \ T \ {}" by auto then show "\U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" proof (cases) case 1 then show ?thesis by blast next case 2 then show ?thesis by blast next case 3 define U where "U = (\x\S. ball x ((infdist x T)/2))" have A: "open U" unfolding U_def by auto have "infdist x T > 0" if "x \ S" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have B: "S \ U" unfolding U_def by auto define V where "V = (\x\T. ball x ((infdist x S)/2))" have C: "open V" unfolding V_def by auto have "infdist x S > 0" if "x \ T" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have D: "T \ V" unfolding V_def by auto have "(ball x ((infdist x T)/2)) \ (ball y ((infdist y S)/2)) = {}" if "x \ S" "y \ T" for x y proof auto fix z assume H: "dist x z * 2 < infdist x T" "dist y z * 2 < infdist y S" have "2 * dist x y \ 2 * dist x z + 2 * dist y z" by metric also have "... < infdist x T + infdist y S" using H by auto finally have "dist x y < infdist x T \ dist x y < infdist y S" by auto then show False using infdist_le[OF \x \ S\, of y] infdist_le[OF \y \ T\, of x] by (auto simp add: dist_commute) qed then have E: "U \ V = {}" unfolding U_def V_def by auto show ?thesis using A B C D E by blast qed qed lemma tendsto_infdist [tendsto_intros]: assumes f: "(f \ l) F" shows "((\x. infdist (f x) A) \ infdist l A) F" proof (rule tendstoI) fix e ::real assume "e > 0" from tendstoD[OF f this] show "eventually (\x. dist (infdist (f x) A) (infdist l A) < e) F" proof (eventually_elim) fix x from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l] have "dist (infdist (f x) A) (infdist l A) \ dist (f x) l" by (simp add: dist_commute dist_real_def) also assume "dist (f x) l < e" finally show "dist (infdist (f x) A) (infdist l A) < e" . qed qed lemma continuous_infdist[continuous_intros]: assumes "continuous F f" shows "continuous F (\x. infdist (f x) A)" using assms unfolding continuous_def by (rule tendsto_infdist) lemma continuous_on_infdist [continuous_intros]: assumes "continuous_on S f" shows "continuous_on S (\x. infdist (f x) A)" using assms unfolding continuous_on by (auto intro: tendsto_infdist) lemma compact_infdist_le: fixes A::"'a::heine_borel set" assumes "A \ {}" assumes "compact A" assumes "e > 0" shows "compact {x. infdist x A \ e}" proof - from continuous_closed_vimage[of "{0..e}" "\x. infdist x A"] continuous_infdist[OF continuous_ident, of _ UNIV A] have "closed {x. infdist x A \ e}" by (auto simp: vimage_def infdist_nonneg) moreover from assms obtain x0 b where b: "\x. x \ A \ dist x0 x \ b" "closed A" by (auto simp: compact_eq_bounded_closed bounded_def) { fix y assume "infdist y A \ e" moreover from infdist_attains_inf[OF \closed A\ \A \ {}\, of y] obtain z where "z \ A" "infdist y A = dist y z" by blast ultimately have "dist x0 y \ b + e" using b by metric } then have "bounded {x. infdist x A \ e}" by (auto simp: bounded_any_center[where a=x0] intro!: exI[where x="b + e"]) ultimately show "compact {x. infdist x A \ e}" by (simp add: compact_eq_bounded_closed) qed subsection \Separation between Points and Sets\ proposition separate_point_closed: fixes S :: "'a::heine_borel set" assumes "closed S" and "a \ S" shows "\d>0. \x\S. d \ dist a x" by (metis assms distance_attains_inf empty_iff gt_ex zero_less_dist_iff) proposition separate_compact_closed: fixes S T :: "'a::heine_borel set" assumes "compact S" and T: "closed T" "S \ T = {}" shows "\d>0. \x\S. \y\T. d \ dist x y" proof cases assume "S \ {} \ T \ {}" then have "S \ {}" "T \ {}" by auto let ?inf = "\x. infdist x T" have "continuous_on S ?inf" by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) then obtain x where x: "x \ S" "\y\S. ?inf x \ ?inf y" using continuous_attains_inf[OF \compact S\ \S \ {}\] by auto then have "0 < ?inf x" using T \T \ {}\ in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg) moreover have "\x'\S. \y\T. ?inf x \ dist x' y" using x by (auto intro: order_trans infdist_le) ultimately show ?thesis by auto qed (auto intro!: exI[of _ 1]) proposition separate_closed_compact: fixes S T :: "'a::heine_borel set" assumes S: "closed S" and T: "compact T" and dis: "S \ T = {}" shows "\d>0. \x\S. \y\T. d \ dist x y" by (metis separate_compact_closed[OF T S] dis dist_commute inf_commute) proposition compact_in_open_separated: fixes A::"'a::heine_borel set" assumes A: "A \ {}" "compact A" assumes "open B" assumes "A \ B" obtains e where "e > 0" "{x. infdist x A \ e} \ B" proof atomize_elim have "closed (- B)" "compact A" "- B \ A = {}" using assms by (auto simp: open_Diff compact_eq_bounded_closed) from separate_closed_compact[OF this] obtain d'::real where d': "d'>0" "\x y. x \ B \ y \ A \ d' \ dist x y" by auto define d where "d = d' / 2" hence "d>0" "d < d'" using d' by auto with d' have d: "\x y. x \ B \ y \ A \ d < dist x y" by force show "\e>0. {x. infdist x A \ e} \ B" proof (rule ccontr) assume "\e. 0 < e \ {x. infdist x A \ e} \ B" with \d > 0\ obtain x where x: "infdist x A \ d" "x \ B" by auto then show False by (metis A compact_eq_bounded_closed infdist_attains_inf x d linorder_not_less) qed qed subsection \Uniform Continuity\ lemma uniformly_continuous_onE: assumes "uniformly_continuous_on s f" "0 < e" obtains d where "d>0" "\x x'. \x\s; x'\s; dist x' x < d\ \ dist (f x') (f x) < e" using assms by (auto simp: uniformly_continuous_on_def) lemma uniformly_continuous_on_sequentially: "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ (\n. dist (x n) (y n)) \ 0 \ (\n. dist (f(x n)) (f(y n))) \ 0)" (is "?lhs = ?rhs") proof assume ?lhs { fix x y assume x: "\n. x n \ s" and y: "\n. y n \ s" and xy: "((\n. dist (x n) (y n)) \ 0) sequentially" { fix e :: real assume "e > 0" then obtain d where "d > 0" and d: "\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" by (metis \?lhs\ uniformly_continuous_onE) obtain N where N: "\n\N. dist (x n) (y n) < d" using xy[unfolded lim_sequentially dist_norm] and \d>0\ by auto then have "\N. \n\N. dist (f (x n)) (f (y n)) < e" using d x y by blast } then have "((\n. dist (f(x n)) (f(y n))) \ 0) sequentially" unfolding lim_sequentially and dist_real_def by auto } then show ?rhs by auto next assume ?rhs { assume "\ ?lhs" then obtain e where "e > 0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto then obtain fa where fa: "\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] by (auto simp: Bex_def dist_commute) define x where "x n = fst (fa (inverse (real n + 1)))" for n define y where "y n = snd (fa (inverse (real n + 1)))" for n have xyn: "\n. x n \ s \ y n \ s" and xy0: "\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" unfolding x_def and y_def using fa by auto { fix e :: real assume "e > 0" then obtain N :: nat where "N \ 0" and N: "0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inverse[of e] by auto then have "\N. \n\N. dist (x n) (y n) < e" by (smt (verit, ccfv_SIG) inverse_le_imp_le nat_le_real_less of_nat_le_0_iff xy0) } then have "\e>0. \N. \n\N. dist (f (x n)) (f (y n)) < e" using \?rhs\[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding lim_sequentially dist_real_def by auto then have False using fxy and \e>0\ by auto } then show ?lhs unfolding uniformly_continuous_on_def by blast qed subsection \Continuity on a Compact Domain Implies Uniform Continuity\ text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ lemma Heine_Borel_lemma: assumes "compact S" and Ssub: "S \ \\" and opn: "\G. G \ \ \ open G" obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" proof - have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" proof - have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n using neg by simp then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" by metis then obtain l r where "l \ S" "strict_mono r" and to_l: "(f \ r) \ l" using \compact S\ compact_def that by metis then obtain G where "l \ G" "G \ \" using Ssub by auto then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" using opn open_dist by blast obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" using to_l apply (simp add: lim_sequentially) using \0 < e\ half_gt_zero that by blast obtain N2 where N2: "of_nat N2 > 2/e" using reals_Archimedean2 by blast obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" using fG [OF \G \ \\, of "r (max N1 N2)"] by blast then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" by simp also have "... \ 1 / real (Suc (max N1 N2))" by (meson Suc_le_mono \strict_mono r\ inverse_of_nat_le nat.discI seq_suble) also have "... \ 1 / real (Suc N2)" by (simp add: field_simps) also have "... < e/2" using N2 \0 < e\ by (simp add: field_simps) finally have "dist (f (r (max N1 N2))) x < e/2" . moreover have "dist (f (r (max N1 N2))) l < e/2" using N1 max.cobounded1 by blast ultimately have "dist x l < e" by metric then show ?thesis using e \x \ G\ by blast qed then show ?thesis by (meson that) qed lemma compact_uniformly_equicontinuous: assumes "compact S" and cont: "\x e. \x \ S; 0 < e\ \ \d. 0 < d \ (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" and "0 < e" obtains d where "0 < d" "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" proof - obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" using cont by metis let ?\ = "((\x. ball x (d x (e/2))) ` S)" have Ssub: "S \ \ ?\" using \0 < e\ d_pos by auto then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" by (rule Heine_Borel_lemma [OF \compact S\]) auto moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v proof - obtain G where "G \ ?\" "u \ G" "v \ G" using k that by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) then obtain w where w: "dist w u < d w (e/2)" "dist w v < d w (e/2)" "w \ S" by auto with that d_dist have "dist (f w) (f v) < e/2" by (metis \0 < e\ dist_commute half_gt_zero) moreover have "dist (f w) (f u) < e/2" using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) ultimately show ?thesis using dist_triangle_half_r by blast qed ultimately show ?thesis using that by blast qed corollary compact_uniformly_continuous: fixes f :: "'a :: metric_space \ 'b :: metric_space" assumes f: "continuous_on S f" and S: "compact S" shows "uniformly_continuous_on S f" using f unfolding continuous_on_iff uniformly_continuous_on_def by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) subsection\<^marker>\tag unimportant\\ Theorems relating continuity and uniform continuity to closures\ lemma continuous_on_closure: "continuous_on (closure S) f \ (\x e. x \ closure S \ 0 < e \ (\d. 0 < d \ (\y. y \ S \ dist y x < d \ dist (f y) (f x) < e)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding continuous_on_iff by (metis Un_iff closure_def) next assume R [rule_format]: ?rhs show ?lhs proof fix x and e::real assume "0 < e" and x: "x \ closure S" obtain \::real where "\ > 0" and \: "\y. \y \ S; dist y x < \\ \ dist (f y) (f x) < e/2" using R [of x "e/2"] \0 < e\ x by auto have "dist (f y) (f x) \ e" if y: "y \ closure S" and dyx: "dist y x < \/2" for y proof - obtain \'::real where "\' > 0" and \': "\z. \z \ S; dist z y < \'\ \ dist (f z) (f y) < e/2" using R [of y "e/2"] \0 < e\ y by auto obtain z where "z \ S" and z: "dist z y < min \' \ / 2" using closure_approachable y by (metis \0 < \'\ \0 < \\ divide_pos_pos min_less_iff_conj zero_less_numeral) have "dist (f z) (f y) < e/2" using \' [OF \z \ S\] z \0 < \'\ by metric moreover have "dist (f z) (f x) < e/2" using \[OF \z \ S\] z dyx by metric ultimately show ?thesis by metric qed then show "\d>0. \x'\closure S. dist x' x < d \ dist (f x') (f x) \ e" by (rule_tac x="\/2" in exI) (simp add: \\ > 0\) qed qed lemma continuous_on_closure_sequentially: fixes f :: "'a::metric_space \ 'b :: metric_space" shows "continuous_on (closure S) f \ (\x a. a \ closure S \ (\n. x n \ S) \ x \ a \ (f \ x) \ f a)" (is "?lhs = ?rhs") proof - have "continuous_on (closure S) f \ (\x \ closure S. continuous (at x within S) f)" by (force simp: continuous_on_closure continuous_within_eps_delta) also have "... = ?rhs" by (force simp: continuous_within_sequentially) finally show ?thesis . qed lemma uniformly_continuous_on_closure: fixes f :: "'a::metric_space \ 'b::metric_space" assumes ucont: "uniformly_continuous_on S f" and cont: "continuous_on (closure S) f" shows "uniformly_continuous_on (closure S) f" unfolding uniformly_continuous_on_def proof (intro allI impI) fix e::real assume "0 < e" then obtain d::real where "d>0" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (f x') (f x) < e/3" using ucont [unfolded uniformly_continuous_on_def, rule_format, of "e/3"] by auto show "\d>0. \x\closure S. \x'\closure S. dist x' x < d \ dist (f x') (f x) < e" proof (rule exI [where x="d/3"], clarsimp simp: \d > 0\) fix x y assume x: "x \ closure S" and y: "y \ closure S" and dyx: "dist y x * 3 < d" obtain d1::real where "d1 > 0" and d1: "\w. \w \ closure S; dist w x < d1\ \ dist (f w) (f x) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "x" "e/3"] \0 < e\ x by auto obtain x' where "x' \ S" and x': "dist x' x < min d1 (d / 3)" using closure_approachable [of x S] by (metis \0 < d1\ \0 < d\ divide_pos_pos min_less_iff_conj x zero_less_numeral) obtain d2::real where "d2 > 0" and d2: "\w \ closure S. dist w y < d2 \ dist (f w) (f y) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "y" "e/3"] \0 < e\ y by auto obtain y' where "y' \ S" and y': "dist y' y < min d2 (d / 3)" using closure_approachable [of y S] by (metis \0 < d2\ \0 < d\ divide_pos_pos min_less_iff_conj y zero_less_numeral) have "dist x' x < d/3" using x' by auto then have "dist x' y' < d" using dyx y' by metric then have "dist (f x') (f y') < e/3" by (rule d [OF \y' \ S\ \x' \ S\]) moreover have "dist (f x') (f x) < e/3" using \x' \ S\ closure_subset x' d1 by (simp add: closure_def) moreover have "dist (f y') (f y) < e/3" using \y' \ S\ closure_subset y' d2 by (simp add: closure_def) ultimately show "dist (f y) (f x) < e" by metric qed qed lemma uniformly_continuous_on_extension_at_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" assumes "x \ closure X" obtains l where "(f \ l) (at x within X)" proof - from assms obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs] obtain l where l: "(\n. f (xs n)) \ l" by atomize_elim (simp only: convergent_eq_Cauchy) have "(f \ l) (at x within X)" proof (safe intro!: Lim_within_LIMSEQ) fix xs' assume "\n. xs' n \ x \ xs' n \ X" and xs': "xs' \ x" then have "xs' n \ x" "xs' n \ X" for n by auto from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \xs' \ x\ \xs' _ \ X\] obtain l' where l': "(\n. f (xs' n)) \ l'" by atomize_elim (simp only: convergent_eq_Cauchy) show "(\n. f (xs' n)) \ l" proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' \ e/2" have "e' > 0" using \e > 0\ by (simp add: e'_def) have "\\<^sub>F n in sequentially. dist (f (xs n)) l < e'" by (simp add: \0 < e'\ l tendstoD) moreover from uc[unfolded uniformly_continuous_on_def, rule_format, OF \e' > 0\] obtain d where d: "d > 0" "\x x'. x \ X \ x' \ X \ dist x x' < d \ dist (f x) (f x') < e'" by auto have "\\<^sub>F n in sequentially. dist (xs n) (xs' n) < d" by (auto intro!: \0 < d\ order_tendstoD tendsto_eq_intros xs xs') ultimately show "\\<^sub>F n in sequentially. dist (f (xs' n)) l < e" proof eventually_elim case (elim n) have "dist (f (xs' n)) l \ dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l" by metric also have "dist (f (xs n)) (f (xs' n)) < e'" by (auto intro!: d xs \xs' _ \ _\ elim) also note \dist (f (xs n)) l < e'\ also have "e' + e' = e" by (simp add: e'_def) finally show ?case by simp qed qed qed thus ?thesis .. qed lemma uniformly_continuous_on_extension_on_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" obtains g where "uniformly_continuous_on (closure X) g" "\x. x \ X \ f x = g x" "\Y h x. X \ Y \ Y \ closure X \ continuous_on Y h \ (\x. x \ X \ f x = h x) \ x \ Y \ h x = g x" proof - from uc have cont_f: "continuous_on X f" by (simp add: uniformly_continuous_imp_continuous) obtain y where y: "(f \ y x) (at x within X)" if "x \ closure X" for x using uniformly_continuous_on_extension_at_closure[OF assms] by meson let ?g = "\x. if x \ X then f x else y x" have "uniformly_continuous_on (closure X) ?g" unfolding uniformly_continuous_on_def proof safe fix e::real assume "e > 0" define e' where "e' \ e / 3" have "e' > 0" using \e > 0\ by (simp add: e'_def) from uc[unfolded uniformly_continuous_on_def, rule_format, OF \0 < e'\] obtain d where "d > 0" and d: "\x x'. x \ X \ x' \ X \ dist x' x < d \ dist (f x') (f x) < e'" by auto define d' where "d' = d / 3" have "d' > 0" using \d > 0\ by (simp add: d'_def) show "\d>0. \x\closure X. \x'\closure X. dist x' x < d \ dist (?g x') (?g x) < e" proof (safe intro!: exI[where x=d'] \d' > 0\) fix x x' assume x: "x \ closure X" and x': "x' \ closure X" and dist: "dist x' x < d'" then obtain xs xs' where xs: "xs \ x" "\n. xs n \ X" and xs': "xs' \ x'" "\n. xs' n \ X" by (auto simp: closure_sequential) have "\\<^sub>F n in sequentially. dist (xs' n) x' < d'" and "\\<^sub>F n in sequentially. dist (xs n) x < d'" by (auto intro!: \0 < d'\ order_tendstoD tendsto_eq_intros xs xs') moreover have "(\x. f (xs x)) \ y x" if "x \ closure X" "x \ X" "xs \ x" "\n. xs n \ X" for xs x using that not_eventuallyD by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at) then have "(\x. f (xs' x)) \ ?g x'" "(\x. f (xs x)) \ ?g x" using x x' by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs) then have "\\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'" "\\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'" by (auto intro!: \0 < e'\ order_tendstoD tendsto_eq_intros) ultimately have "\\<^sub>F n in sequentially. dist (?g x') (?g x) < e" proof eventually_elim case (elim n) have "dist (?g x') (?g x) \ dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)" by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le) also from \dist (xs' n) x' < d'\ \dist x' x < d'\ \dist (xs n) x < d'\ have "dist (xs' n) (xs n) < d" unfolding d'_def by metric with \xs _ \ X\ \xs' _ \ X\ have "dist (f (xs' n)) (f (xs n)) < e'" by (rule d) also note \dist (f (xs' n)) (?g x') < e'\ also note \dist (f (xs n)) (?g x) < e'\ finally show ?case by (simp add: e'_def) qed then show "dist (?g x') (?g x) < e" by simp qed qed moreover have "f x = ?g x" if "x \ X" for x using that by simp moreover { fix Y h x assume Y: "x \ Y" "X \ Y" "Y \ closure X" and cont_h: "continuous_on Y h" and extension: "(\x. x \ X \ f x = h x)" { assume "x \ X" have "x \ closure X" using Y by auto then obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y have hx: "(\x. f (xs x)) \ h x" by (auto simp: subsetD extension) then have "(\x. f (xs x)) \ y x" using \x \ X\ not_eventuallyD xs(2) by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at xs) with hx have "h x = y x" by (rule LIMSEQ_unique) } then have "h x = ?g x" using extension by auto } ultimately show ?thesis .. qed lemma bounded_uniformly_continuous_image: fixes f :: "'a :: heine_borel \ 'b :: heine_borel" assumes "uniformly_continuous_on S f" "bounded S" shows "bounded(f ` S)" by (metis (no_types, lifting) assms bounded_closure_image compact_closure compact_continuous_image compact_eq_bounded_closed image_cong uniformly_continuous_imp_continuous uniformly_continuous_on_extension_on_closure) subsection \With Abstract Topology (TODO: move and remove dependency?)\ lemma openin_contains_ball: "openin (top_of_set T) S \ S \ T \ (\x \ S. \e. 0 < e \ ball x e \ T \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis IntD2 Int_commute Int_lower1 Int_mono inf.idem openE openin_open) next assume ?rhs then show ?lhs by (smt (verit) open_ball Int_commute Int_iff centre_in_ball in_mono openin_open_Int openin_subopen) qed lemma openin_contains_cball: "openin (top_of_set T) S \ S \ T \ (\x \ S. \e. 0 < e \ cball x e \ T \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp add: openin_contains_ball intro: exI [where x="_/2"]) next assume ?rhs then show ?lhs by (force simp add: openin_contains_ball) qed subsection \Closed Nest\ text \Bounded closed nest property (proof does not use Heine-Borel)\ lemma bounded_closed_nest: fixes S :: "nat \ ('a::heine_borel) set" assumes "\n. closed (S n)" and "\n. S n \ {}" and "\m n. m \ n \ S n \ S m" and "bounded (S 0)" obtains a where "\n. a \ S n" proof - from assms(2) obtain x where x: "\n. x n \ S n" using choice[of "\n x. x \ S n"] by auto from assms(4,1) have "seq_compact (S 0)" by (simp add: bounded_closed_imp_seq_compact) then obtain l r where lr: "l \ S 0" "strict_mono r" "(x \ r) \ l" using x and assms(3) unfolding seq_compact_def by blast have "\n. l \ S n" proof fix n :: nat have "closed (S n)" using assms(1) by simp moreover have "\i. (x \ r) i \ S i" using x and assms(3) and lr(2) [THEN seq_suble] by auto then have "\i. (x \ r) (i + n) \ S n" using assms(3) by (fast intro!: le_add2) moreover have "(\i. (x \ r) (i + n)) \ l" using lr(3) by (rule LIMSEQ_ignore_initial_segment) ultimately show "l \ S n" by (rule closed_sequentially) qed then show ?thesis using that by blast qed text \Decreasing case does not even need compactness, just completeness.\ lemma decreasing_closed_nest: fixes S :: "nat \ ('a::complete_space) set" assumes "\n. closed (S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x\S n. \y\S n. dist x y < e" obtains a where "\n. a \ S n" proof - obtain t where t: "\n. t n \ S n" by (meson assms(2) equals0I) { fix e :: real assume "e > 0" then obtain N where N: "\x\S N. \y\S N. dist x y < e" using assms(4) by blast { fix m n :: nat assume "N \ m \ N \ n" then have "t m \ S N" "t n \ S N" using assms(3) t unfolding subset_eq t by blast+ then have "dist (t m) (t n) < e" using N by auto } then have "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto } then have "Cauchy t" - unfolding cauchy_def by auto + by (metis metric_CauchyI) then obtain l where l:"(t \ l) sequentially" using complete_UNIV unfolding complete_def by auto { fix n :: nat { fix e :: real assume "e > 0" then obtain N :: nat where N: "\n\N. dist (t n) l < e" using l[unfolded lim_sequentially] by auto have "t (max n N) \ S n" by (meson assms(3) contra_subsetD max.cobounded1 t) then have "\y\S n. dist y l < e" using N max.cobounded2 by blast } then have "l \ S n" using closed_approachable[of "S n" l] assms(1) by auto } then show ?thesis using that by blast qed text \Strengthen it to the intersection actually being a singleton.\ lemma decreasing_closed_nest_sing: fixes S :: "nat \ 'a::complete_space set" assumes "\n. closed(S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x \ (S n). \ y\(S n). dist x y < e" shows "\a. \(range S) = {a}" proof - obtain a where a: "\n. a \ S n" using decreasing_closed_nest[of S] using assms by auto { fix b assume b: "b \ \(range S)" { fix e :: real assume "e > 0" then have "dist a b < e" using assms(4) and b and a by blast } then have "dist a b = 0" by (metis dist_eq_0_iff dist_nz less_le) } with a have "\(range S) = {a}" unfolding image_def by auto then show ?thesis .. qed subsection\<^marker>\tag unimportant\ \Making a continuous function avoid some value in a neighbourhood\ lemma continuous_within_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x within s) f" and "f x \ a" shows "\e>0. \y \ s. dist x y < e --> f y \ a" proof - obtain U where "open U" and "f x \ U" and "a \ U" using t1_space [OF \f x \ a\] by fast have "(f \ f x) (at x within s)" using assms(1) by (simp add: continuous_within) then have "eventually (\y. f y \ U) (at x within s)" using \open U\ and \f x \ U\ unfolding tendsto_def by fast then have "eventually (\y. f y \ a) (at x within s)" using \a \ U\ by (fast elim: eventually_mono) then show ?thesis using \f x \ a\ by (auto simp: dist_commute eventually_at) qed lemma continuous_at_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x) f" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms continuous_within_avoid[of x UNIV f a] by simp lemma continuous_on_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "x \ s" and "f x \ a" shows "\e>0. \y \ s. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] using assms(3) by auto lemma continuous_on_open_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "open s" and "x \ s" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] using continuous_at_avoid[of x f a] assms(4) by auto subsection \Consequences for Real Numbers\ lemma closed_contains_Inf: fixes S :: "real set" shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" by (metis closure_contains_Inf closure_closed) lemma closed_subset_contains_Inf: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_below A \ Inf A \ C" by (metis closure_contains_Inf closure_minimal subset_eq) lemma closed_contains_Sup: fixes S :: "real set" shows "S \ {} \ bdd_above S \ closed S \ Sup S \ S" by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup) lemma closed_subset_contains_Sup: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_above A \ Sup A \ C" by (metis closure_contains_Sup closure_minimal subset_eq) lemma atLeastAtMost_subset_contains_Inf: fixes A :: "real set" and a b :: real shows "A \ {} \ a \ b \ A \ {a..b} \ Inf A \ {a..b}" by (rule closed_subset_contains_Inf) (auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" by (simp add: bounded_iff) lemma bounded_imp_bdd_above: "bounded S \ bdd_above (S :: real set)" by (auto simp: bounded_def bdd_above_def dist_real_def) (metis abs_le_D1 abs_minus_commute diff_le_eq) lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" by (auto simp: bounded_def bdd_below_def dist_real_def) (metis abs_le_D1 add.commute diff_le_eq) lemma bounded_has_Sup: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Sup S" and "\b. (\x\S. x \ b) \ Sup S \ b" proof show "\b. (\x\S. x \ b) \ Sup S \ b" using assms by (metis cSup_least) qed (metis cSup_upper assms(1) bounded_imp_bdd_above) lemma Sup_insert: fixes S :: "real set" shows "bounded S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) lemma bounded_has_Inf: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Inf S" and "\b. (\x\S. x \ b) \ Inf S \ b" proof show "\b. (\x\S. x \ b) \ Inf S \ b" using assms by (metis cInf_greatest) qed (metis cInf_lower assms(1) bounded_imp_bdd_below) lemma Inf_insert: fixes S :: "real set" shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) lemma open_real: fixes s :: "real set" shows "open s \ (\x \ s. \e>0. \x'. \x' - x\ < e --> x' \ s)" unfolding open_dist dist_norm by simp lemma islimpt_approachable_real: fixes s :: "real set" shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ \x' - x\ < e)" unfolding islimpt_approachable dist_norm by simp lemma closed_real: fixes s :: "real set" shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ \x' - x\ < e) \ x \ s)" unfolding closed_limpt islimpt_approachable dist_norm by simp lemma continuous_at_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> \f x' - f x\ < e)" by (metis (mono_tags, opaque_lifting) LIM_eq continuous_within norm_eq_zero real_norm_def right_minus_eq) lemma continuous_on_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d \ \f x' - f x\ < e))" unfolding continuous_on_iff dist_norm by simp lemma continuous_on_closed_Collect_le: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" shows "closed {x \ s. f x \ g x}" proof - have "closed ((\x. g x - f x) -` {0..} \ s)" using closed_real_atLeast continuous_on_diff [OF g f] by (simp add: continuous_on_closed_vimage [OF s]) also have "((\x. g x - f x) -` {0..} \ s) = {x\s. f x \ g x}" by auto finally show ?thesis . qed lemma continuous_le_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. x \ a}" ] assms continuous_on_closed_Collect_le[of "UNIV" "\x. x" "\x. a"] by auto lemma continuous_ge_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. a \ x}"] assms continuous_on_closed_Collect_le[of "UNIV" "\x. a" "\x. x"] by auto subsection\The infimum of the distance between two sets\ definition\<^marker>\tag important\ setdist :: "'a::metric_space set \ 'a set \ real" where "setdist s t \ (if s = {} \ t = {} then 0 else Inf {dist x y| x y. x \ s \ y \ t})" lemma setdist_empty1 [simp]: "setdist {} t = 0" by (simp add: setdist_def) lemma setdist_empty2 [simp]: "setdist t {} = 0" by (simp add: setdist_def) lemma setdist_pos_le [simp]: "0 \ setdist s t" by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) lemma le_setdistI: assumes "s \ {}" "t \ {}" "\x y. \x \ s; y \ t\ \ d \ dist x y" shows "d \ setdist s t" using assms by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) lemma setdist_le_dist: "\x \ s; y \ t\ \ setdist s t \ dist x y" unfolding setdist_def by (auto intro!: bdd_belowI [where m=0] cInf_lower) lemma le_setdist_iff: "d \ setdist S T \ (\x \ S. \y \ T. d \ dist x y) \ (S = {} \ T = {} \ d \ 0)" by (smt (verit) le_setdistI setdist_def setdist_le_dist) lemma setdist_ltE: assumes "setdist S T < b" "S \ {}" "T \ {}" obtains x y where "x \ S" "y \ T" "dist x y < b" using assms by (auto simp: not_le [symmetric] le_setdist_iff) lemma setdist_refl: "setdist S S = 0" by (metis dist_eq_0_iff ex_in_conv order_antisym setdist_def setdist_le_dist setdist_pos_le) lemma setdist_sym: "setdist S T = setdist T S" by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) lemma setdist_triangle: "setdist S T \ setdist S {a} + setdist {a} T" proof (cases "S = {} \ T = {}") case True then show ?thesis using setdist_pos_le by fastforce next case False then have "\x. x \ S \ setdist S T - dist x a \ setdist {a} T" using dist_self dist_triangle3 empty_not_insert le_setdist_iff setdist_le_dist singleton_iff by (smt (verit, best) dist_self dist_triangle3 empty_not_insert le_setdist_iff setdist_le_dist singleton_iff) then have "setdist S T - setdist {a} T \ setdist S {a}" using False by (fastforce intro: le_setdistI) then show ?thesis by (simp add: algebra_simps) qed lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y" by (simp add: setdist_def) lemma setdist_Lipschitz: "\setdist {x} S - setdist {y} S\ \ dist x y" apply (subst setdist_singletons [symmetric]) by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\y. (setdist {y} S))" by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma continuous_on_setdist [continuous_intros]: "continuous_on T (\y. (setdist {y} S))" by (metis continuous_at_setdist continuous_at_imp_continuous_on) lemma uniformly_continuous_on_setdist: "uniformly_continuous_on T (\y. (setdist {y} S))" by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma setdist_subset_right: "\T \ {}; T \ u\ \ setdist S u \ setdist S T" by (smt (verit, best) in_mono le_setdist_iff) lemma setdist_subset_left: "\S \ {}; S \ T\ \ setdist T u \ setdist S u" by (metis setdist_subset_right setdist_sym) lemma setdist_closure_1 [simp]: "setdist (closure S) T = setdist S T" proof (cases "S = {} \ T = {}") case True then show ?thesis by force next case False { fix y assume "y \ T" have "continuous_on (closure S) (\a. dist a y)" by (auto simp: continuous_intros dist_norm) then have *: "\x. x \ closure S \ setdist S T \ dist x y" by (fast intro: setdist_le_dist \y \ T\ continuous_ge_on_closure) } then show ?thesis by (metis False antisym closure_eq_empty closure_subset le_setdist_iff setdist_subset_left) qed lemma setdist_closure_2 [simp]: "setdist T (closure S) = setdist T S" by (metis setdist_closure_1 setdist_sym) lemma setdist_eq_0I: "\x \ S; x \ T\ \ setdist S T = 0" by (metis antisym dist_self setdist_le_dist setdist_pos_le) lemma setdist_unique: "\a \ S; b \ T; \x y. x \ S \ y \ T ==> dist a b \ dist x y\ \ setdist S T = dist a b" by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) lemma setdist_le_sing: "x \ S ==> setdist S T \ setdist {x} T" using setdist_subset_left by auto lemma infdist_eq_setdist: "infdist x A = setdist {x} A" by (simp add: infdist_def setdist_def Setcompr_eq_image) lemma setdist_eq_infdist: "setdist A B = (if A = {} then 0 else INF a\A. infdist a B)" proof - have "Inf {dist x y |x y. x \ A \ y \ B} = (INF x\A. Inf (dist x ` B))" if "b \ B" "a \ A" for a b proof (rule order_antisym) have "Inf {dist x y |x y. x \ A \ y \ B} \ Inf (dist x ` B)" if "b \ B" "a \ A" "x \ A" for x proof - have "\b'. b' \ B \ Inf {dist x y |x y. x \ A \ y \ B} \ dist x b'" by (metis (mono_tags, lifting) ex_in_conv setdist_def setdist_le_dist that(3)) then show ?thesis by (smt (verit) cINF_greatest ex_in_conv \b \ B\) qed then show "Inf {dist x y |x y. x \ A \ y \ B} \ (INF x\A. Inf (dist x ` B))" by (metis (mono_tags, lifting) cINF_greatest emptyE that) next have *: "\x y. \b \ B; a \ A; x \ A; y \ B\ \ \a\A. Inf (dist a ` B) \ dist x y" by (meson bdd_below_image_dist cINF_lower) show "(INF x\A. Inf (dist x ` B)) \ Inf {dist x y |x y. x \ A \ y \ B}" proof (rule conditionally_complete_lattice_class.cInf_mono) show "bdd_below ((\x. Inf (dist x ` B)) ` A)" by (metis (no_types, lifting) bdd_belowI2 ex_in_conv infdist_def infdist_nonneg that(1)) qed (use that in \auto simp: *\) qed then show ?thesis by (auto simp: setdist_def infdist_def) qed lemma infdist_mono: assumes "A \ B" "A \ {}" shows "infdist x B \ infdist x A" by (simp add: assms infdist_eq_setdist setdist_subset_right) lemma infdist_singleton [simp]: "infdist x {y} = dist x y" by (simp add: infdist_eq_setdist) proposition setdist_attains_inf: assumes "compact B" "B \ {}" obtains y where "y \ B" "setdist A B = infdist y A" proof (cases "A = {}") case True then show thesis by (metis assms diameter_compact_attained infdist_def setdist_def that) next case False obtain y where "y \ B" and min: "\y'. y' \ B \ infdist y A \ infdist y' A" by (metis continuous_attains_inf [OF assms continuous_on_infdist] continuous_on_id) show thesis proof have "setdist A B = (INF y\B. infdist y A)" by (metis \B \ {}\ setdist_eq_infdist setdist_sym) also have "\ = infdist y A" proof (rule order_antisym) show "(INF y\B. infdist y A) \ infdist y A" by (meson \y \ B\ bdd_belowI2 cInf_lower image_eqI infdist_nonneg) next show "infdist y A \ (INF y\B. infdist y A)" by (simp add: \B \ {}\ cINF_greatest min) qed finally show "setdist A B = infdist y A" . qed (fact \y \ B\) qed end diff --git a/src/HOL/Analysis/Function_Metric.thy b/src/HOL/Analysis/Function_Metric.thy --- a/src/HOL/Analysis/Function_Metric.thy +++ b/src/HOL/Analysis/Function_Metric.thy @@ -1,386 +1,387 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP License: BSD *) section \Metrics on product spaces\ theory Function_Metric imports Function_Topology Elementary_Metric_Spaces begin text \In general, the product topology is not metrizable, unless the index set is countable. When the index set is countable, essentially any (convergent) combination of the metrics on the factors will do. We use below the simplest one, based on \L\<^sup>1\, but \L\<^sup>2\ would also work, for instance. What is not completely trivial is that the distance thus defined induces the same topology as the product topology. This is what we have to prove to show that we have an instance of \<^class>\metric_space\. The proofs below would work verbatim for general countable products of metric spaces. However, since distances are only implemented in terms of type classes, we only develop the theory for countable products of the same space.\ instantiation "fun" :: (countable, metric_space) metric_space begin definition\<^marker>\tag important\ dist_fun_def: "dist x y = (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" definition\<^marker>\tag important\ uniformity_fun_def: "(uniformity::(('a \ 'b) \ ('a \ 'b)) filter) = (INF e\{0<..}. principal {(x, y). dist (x::('a\'b)) y < e})" text \Except for the first one, the auxiliary lemmas below are only useful when proving the instance: once it is proved, they become trivial consequences of the general theory of metric spaces. It would thus be desirable to hide them once the instance is proved, but I do not know how to do this.\ lemma dist_fun_le_dist_first_terms: "dist x y \ 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N} + (1/2)^N" proof - have "(\n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) = (\n. (1 / 2) ^ (Suc N) * ((1/2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1))" by (rule suminf_cong, simp add: power_add) also have "... = (1/2)^(Suc N) * (\n. (1 / 2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)" apply (rule suminf_mult) by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... \ (1/2)^(Suc N) * (\n. (1 / 2) ^ n)" apply (simp, rule suminf_le, simp) by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... = (1/2)^(Suc N) * 2" using suminf_geometric[of "1/2"] by auto also have "... = (1/2)^N" by auto finally have *: "(\n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) \ (1/2)^N" by simp define M where "M = Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N}" have "dist (x (from_nat 0)) (y (from_nat 0)) \ M" unfolding M_def by (rule Max_ge, auto) then have [simp]: "M \ 0" by (meson dual_order.trans zero_le_dist) have "dist (x (from_nat n)) (y (from_nat n)) \ M" if "n\N" for n unfolding M_def apply (rule Max_ge) using that by auto then have i: "min (dist (x (from_nat n)) (y (from_nat n))) 1 \ M" if "n\N" for n using that by force have "(\n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \ (\n< Suc N. M * (1 / 2) ^ n)" by (rule sum_mono, simp add: i) also have "... = M * (\n M * (\n. (1 / 2) ^ n)" by (rule mult_left_mono, rule sum_le_suminf, auto simp add: summable_geometric_iff) also have "... = M * 2" using suminf_geometric[of "1/2"] by auto finally have **: "(\n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \ 2 * M" by simp have "dist x y = (\n. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" unfolding dist_fun_def by simp also have "... = (\n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) + (\nn. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... \ 2 * M + (1/2)^N" using * ** by auto finally show ?thesis unfolding M_def by simp qed lemma open_fun_contains_ball_aux: assumes "open (U::(('a \ 'b) set))" "x \ U" shows "\e>0. {y. dist x y < e} \ U" proof - have *: "openin (product_topology (\i. euclidean) UNIV) U" using open_fun_def assms by auto obtain X where H: "Pi\<^sub>E UNIV X \ U" "\i. openin euclidean (X i)" "finite {i. X i \ topspace euclidean}" "x \ Pi\<^sub>E UNIV X" using product_topology_open_contains_basis[OF * \x \ U\] by auto define I where "I = {i. X i \ topspace euclidean}" have "finite I" unfolding I_def using H(3) by auto { fix i have "x i \ X i" using \x \ U\ H by auto then have "\e. e>0 \ ball (x i) e \ X i" using \openin euclidean (X i)\ open_openin open_contains_ball by blast then obtain e where "e>0" "ball (x i) e \ X i" by blast define f where "f = min e (1/2)" have "f>0" "f<1" unfolding f_def using \e>0\ by auto moreover have "ball (x i) f \ X i" unfolding f_def using \ball (x i) e \ X i\ by auto ultimately have "\f. f > 0 \ f < 1 \ ball (x i) f \ X i" by auto } note * = this have "\e. \i. e i > 0 \ e i < 1 \ ball (x i) (e i) \ X i" by (rule choice, auto simp add: *) then obtain e where "\i. e i > 0" "\i. e i < 1" "\i. ball (x i) (e i) \ X i" by blast define m where "m = Min {(1/2)^(to_nat i) * e i|i. i \ I}" have "m > 0" if "I\{}" unfolding m_def Min_gr_iff using \finite I\ \I \ {}\ \\i. e i > 0\ by auto moreover have "{y. dist x y < m} \ U" proof (auto) fix y assume "dist x y < m" have "y i \ X i" if "i \ I" for i proof - have *: "summable (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) define n where "n = to_nat i" have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 < m" using \dist x y < m\ unfolding dist_fun_def using sum_le_suminf[OF *, of "{n}"] by auto then have "(1/2)^(to_nat i) * min (dist (x i) (y i)) 1 < m" using \n = to_nat i\ by auto also have "... \ (1/2)^(to_nat i) * e i" unfolding m_def apply (rule Min_le) using \finite I\ \i \ I\ by auto ultimately have "min (dist (x i) (y i)) 1 < e i" by (auto simp add: field_split_simps) then have "dist (x i) (y i) < e i" using \e i < 1\ by auto then show "y i \ X i" using \ball (x i) (e i) \ X i\ by auto qed then have "y \ Pi\<^sub>E UNIV X" using H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff) then show "y \ U" using \Pi\<^sub>E UNIV X \ U\ by auto qed ultimately have *: "\m>0. {y. dist x y < m} \ U" if "I \ {}" using that by auto have "U = UNIV" if "I = {}" using that H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff) then have "\m>0. {y. dist x y < m} \ U" if "I = {}" using that zero_less_one by blast then show "\m>0. {y. dist x y < m} \ U" using * by blast qed lemma ball_fun_contains_open_aux: fixes x::"('a \ 'b)" and e::real assumes "e>0" shows "\U. open U \ x \ U \ U \ {y. dist x y < e}" proof - have "\N::nat. 2^N > 8/e" by (simp add: real_arch_pow) then obtain N::nat where "2^N > 8/e" by auto define f where "f = e/4" have [simp]: "e>0" "f > 0" unfolding f_def using assms by auto define X::"('a \ 'b set)" where "X = (\i. if (\n\N. i = from_nat n) then {z. dist (x i) z < f} else UNIV)" have "{i. X i \ UNIV} \ from_nat`{0..N}" unfolding X_def by auto then have "finite {i. X i \ topspace euclidean}" unfolding topspace_euclidean using finite_surj by blast have "\i. open (X i)" unfolding X_def using metric_space_class.open_ball open_UNIV by auto then have "\i. openin euclidean (X i)" using open_openin by auto define U where "U = Pi\<^sub>E UNIV X" have "open U" unfolding open_fun_def product_topology_def apply (rule topology_generated_by_Basis) unfolding U_def using \\i. openin euclidean (X i)\ \finite {i. X i \ topspace euclidean}\ by auto moreover have "x \ U" unfolding U_def X_def by (auto simp add: PiE_iff) moreover have "dist x y < e" if "y \ U" for y proof - have *: "dist (x (from_nat n)) (y (from_nat n)) \ f" if "n \ N" for n using \y \ U\ unfolding U_def apply (auto simp add: PiE_iff) unfolding X_def using that by (metis less_imp_le mem_Collect_eq) have **: "Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N} \ f" apply (rule Max.boundedI) using * by auto have "dist x y \ 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N} + (1/2)^N" by (rule dist_fun_le_dist_first_terms) also have "... \ 2 * f + e / 8" apply (rule add_mono) using ** \2^N > 8/e\ by(auto simp add: field_split_simps) also have "... \ e/2 + e/8" unfolding f_def by auto also have "... < e" by auto finally show "dist x y < e" by simp qed ultimately show ?thesis by auto qed lemma fun_open_ball_aux: fixes U::"('a \ 'b) set" shows "open U \ (\x\U. \e>0. \y. dist x y < e \ y \ U)" proof (auto) assume "open U" fix x assume "x \ U" then show "\e>0. \y. dist x y < e \ y \ U" using open_fun_contains_ball_aux[OF \open U\ \x \ U\] by auto next assume H: "\x\U. \e>0. \y. dist x y < e \ y \ U" define K where "K = {V. open V \ V \ U}" { fix x assume "x \ U" then obtain e where e: "e>0" "{y. dist x y < e} \ U" using H by blast then obtain V where V: "open V" "x \ V" "V \ {y. dist x y < e}" using ball_fun_contains_open_aux[OF \e>0\, of x] by auto have "V \ K" unfolding K_def using e(2) V(1) V(3) by auto then have "x \ (\K)" using \x \ V\ by auto } then have "(\K) = U" unfolding K_def by auto moreover have "open (\K)" unfolding K_def by auto ultimately show "open U" by simp qed instance proof fix x y::"'a \ 'b" show "(dist x y = 0) = (x = y)" proof assume "x = y" then show "dist x y = 0" unfolding dist_fun_def using \x = y\ by auto next assume "dist x y = 0" have *: "summable (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 = 0" for n using \dist x y = 0\ unfolding dist_fun_def by (simp add: "*" suminf_eq_zero_iff) then have "dist (x (from_nat n)) (y (from_nat n)) = 0" for n by (metis div_0 min_def nonzero_mult_div_cancel_left power_eq_0_iff zero_eq_1_divide_iff zero_neq_numeral) then have "x (from_nat n) = y (from_nat n)" for n by auto then have "x i = y i" for i by (metis from_nat_to_nat) then show "x = y" by auto qed next text\The proof of the triangular inequality is trivial, modulo the fact that we are dealing with infinite series, hence we should justify the convergence at each step. In this respect, the following simplification rule is extremely handy.\ have [simp]: "summable (\n. (1/2)^n * min (dist (u (from_nat n)) (v (from_nat n))) 1)" for u v::"'a \ 'b" by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) fix x y z::"'a \ 'b" { fix n have *: "dist (x (from_nat n)) (y (from_nat n)) \ dist (x (from_nat n)) (z (from_nat n)) + dist (y (from_nat n)) (z (from_nat n))" by (simp add: dist_triangle2) have "0 \ dist (y (from_nat n)) (z (from_nat n))" using zero_le_dist by blast moreover { assume "min (dist (y (from_nat n)) (z (from_nat n))) 1 \ dist (y (from_nat n)) (z (from_nat n))" then have "1 \ min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1" by (metis (no_types) diff_le_eq diff_self min_def zero_le_dist zero_le_one) } ultimately have "min (dist (x (from_nat n)) (y (from_nat n))) 1 \ min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1" using * by linarith } note ineq = this have "dist x y = (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" unfolding dist_fun_def by simp also have "... \ (\n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1 + (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)" apply (rule suminf_le) using ineq apply (metis (no_types, opaque_lifting) add.right_neutral distrib_left le_divide_eq_numeral1(1) mult_2_right mult_left_mono zero_le_one zero_le_power) by (auto simp add: summable_add) also have "... = (\n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1) + (\n. (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)" by (rule suminf_add[symmetric], auto) also have "... = dist x z + dist y z" unfolding dist_fun_def by simp finally show "dist x y \ dist x z + dist y z" by simp next text\Finally, we show that the topology generated by the distance and the product topology coincide. This is essentially contained in Lemma \fun_open_ball_aux\, except that the condition to prove is expressed with filters. To deal with this, we copy some mumbo jumbo from Lemma \eventually_uniformity_metric\ in \<^file>\~~/src/HOL/Real_Vector_Spaces.thy\\ fix U::"('a \ 'b) set" have "eventually P uniformity \ (\e>0. \x (y::('a \ 'b)). dist x y < e \ P (x, y))" for P unfolding uniformity_fun_def apply (subst eventually_INF_base) by (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) then show "open U = (\x\U. \\<^sub>F (x', y) in uniformity. x' = x \ y \ U)" unfolding fun_open_ball_aux by simp qed (simp add: uniformity_fun_def) end text \Nice properties of spaces are preserved under countable products. In addition to first countability, second countability and metrizability, as we have seen above, completeness is also preserved, and therefore being Polish.\ instance "fun" :: (countable, complete_space) complete_space proof fix u::"nat \ ('a \ 'b)" assume "Cauchy u" have "Cauchy (\n. u n i)" for i - unfolding cauchy_def proof (intro impI allI) - fix e assume "e>(0::real)" + unfolding Cauchy_def + proof (intro strip) + fix e::real assume "e>0" obtain k where "i = from_nat k" using from_nat_to_nat[of i] by metis have "(1/2)^k * min e 1 > 0" using \e>0\ by auto then have "\N. \m n. N \ m \ N \ n \ dist (u m) (u n) < (1/2)^k * min e 1" - using \Cauchy u\ unfolding cauchy_def by blast + using \Cauchy u\ by (meson Cauchy_def) then obtain N::nat where C: "\m n. N \ m \ N \ n \ dist (u m) (u n) < (1/2)^k * min e 1" by blast have "\m n. N \ m \ N \ n \ dist (u m i) (u n i) < e" proof (auto) fix m n::nat assume "m \ N" "n \ N" have "(1/2)^k * min (dist (u m i) (u n i)) 1 = sum (\p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1) {k}" using \i = from_nat k\ by auto also have "... \ (\p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1)" apply (rule sum_le_suminf) by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... = dist (u m) (u n)" unfolding dist_fun_def by auto also have "... < (1/2)^k * min e 1" using C \m \ N\ \n \ N\ by auto finally have "min (dist (u m i) (u n i)) 1 < min e 1" by (auto simp add: field_split_simps) then show "dist (u m i) (u n i) < e" by auto qed - then show "\N. \m n. N \ m \ N \ n \ dist (u m i) (u n i) < e" + then show "\M. \m\M. \n\M. dist (u m i) (u n i) < e" by blast qed then have "\x. (\n. u n i) \ x" for i using Cauchy_convergent convergent_def by auto then have "\x. \i. (\n. u n i) \ x i" using choice by force then obtain x where *: "\i. (\n. u n i) \ x i" by blast have "u \ x" proof (rule metric_LIMSEQ_I) fix e assume [simp]: "e>(0::real)" have i: "\K. \n\K. dist (u n i) (x i) < e/4" for i by (rule metric_LIMSEQ_D, auto simp add: *) have "\K. \i. \n\K i. dist (u n i) (x i) < e/4" apply (rule choice) using i by auto then obtain K where K: "\i n. n \ K i \ dist (u n i) (x i) < e/4" by blast have "\N::nat. 2^N > 4/e" by (simp add: real_arch_pow) then obtain N::nat where "2^N > 4/e" by auto define L where "L = Max {K (from_nat n)|n. n \ N}" have "dist (u k) x < e" if "k \ L" for k proof - have *: "dist (u k (from_nat n)) (x (from_nat n)) \ e / 4" if "n \ N" for n proof - have "K (from_nat n) \ L" unfolding L_def apply (rule Max_ge) using \n \ N\ by auto then have "k \ K (from_nat n)" using \k \ L\ by auto then show ?thesis using K less_imp_le by auto qed have **: "Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \ N} \ e/4" apply (rule Max.boundedI) using * by auto have "dist (u k) x \ 2 * Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \ N} + (1/2)^N" using dist_fun_le_dist_first_terms by auto also have "... \ 2 * e/4 + e/4" apply (rule add_mono) using ** \2^N > 4/e\ less_imp_le by (auto simp add: field_split_simps) also have "... < e" by auto finally show ?thesis by simp qed then show "\L. \k\L. dist (u k) x < e" by blast qed then show "convergent u" using convergent_def by blast qed instance "fun" :: (countable, polish_space) polish_space by standard end \ No newline at end of file diff --git a/src/HOL/Analysis/Further_Topology.thy b/src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy +++ b/src/HOL/Analysis/Further_Topology.thy @@ -1,5516 +1,5517 @@ section \Extending Continous Maps, Invariance of Domain, etc\ (*FIX rename? *) text\Ported from HOL Light (moretop.ml) by L C Paulson\ theory Further_Topology imports Weierstrass_Theorems Polytope Complex_Transcendental Equivalence_Lebesgue_Henstock_Integration Retracts begin subsection\A map from a sphere to a higher dimensional sphere is nullhomotopic\ lemma spheremap_lemma1: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "subspace S" "subspace T" and dimST: "dim S < dim T" and "S \ T" and diff_f: "f differentiable_on sphere 0 1 \ S" shows "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" proof assume fim: "f ` (sphere 0 1 \ S) = sphere 0 1 \ T" have inS: "\x. \x \ S; x \ 0\ \ (x /\<^sub>R norm x) \ S" using subspace_mul \subspace S\ by blast have subS01: "(\x. x /\<^sub>R norm x) ` (S - {0}) \ sphere 0 1 \ S" using \subspace S\ subspace_mul by fastforce then have diff_f': "f differentiable_on (\x. x /\<^sub>R norm x) ` (S - {0})" by (rule differentiable_on_subset [OF diff_f]) define g where "g \ \x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)" have gdiff: "g differentiable_on S - {0}" unfolding g_def by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+ have geq: "g ` (S - {0}) = T - {0}" proof have "\u. \u \ S; norm u *\<^sub>R f (u /\<^sub>R norm u) \ T\ \ u = 0" by (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \subspace T\] fim image_subset_iff inf_le2 singletonD) then have "g ` (S - {0}) \ T" using g_def by blast moreover have "g ` (S - {0}) \ UNIV - {0}" proof (clarsimp simp: g_def) fix y assume "y \ S" and f0: "f (y /\<^sub>R norm y) = 0" then have "y \ 0 \ y /\<^sub>R norm y \ sphere 0 1 \ S" by (auto simp: subspace_mul [OF \subspace S\]) then show "y = 0" by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one) qed ultimately show "g ` (S - {0}) \ T - {0}" by auto next have *: "sphere 0 1 \ T \ f ` (sphere 0 1 \ S)" using fim by (simp add: image_subset_iff) have "x \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" if "x \ T" "x \ 0" for x proof - have "x /\<^sub>R norm x \ T" using \subspace T\ subspace_mul that by blast then obtain u where u: "f u \ T" "x /\<^sub>R norm x = f u" "norm u = 1" "u \ S" using * [THEN subsetD, of "x /\<^sub>R norm x"] \x \ 0\ by auto with that have [simp]: "norm x *\<^sub>R f u = x" by (metis divideR_right norm_eq_zero) moreover have "norm x *\<^sub>R u \ S - {0}" using \subspace S\ subspace_scale that(2) u by auto with u show ?thesis by (simp add: image_eqI [where x="norm x *\<^sub>R u"]) qed then have "T - {0} \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" by force then show "T - {0} \ g ` (S - {0})" by (simp add: g_def) qed define T' where "T' \ {y. \x \ T. orthogonal x y}" have "subspace T'" by (simp add: subspace_orthogonal_to_vectors T'_def) have dim_eq: "dim T' + dim T = DIM('a)" using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ by (simp add: T'_def) have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x by (force intro: orthogonal_subspace_decomp_exists [of T x]) then obtain p1 p2 where p1span: "p1 x \ span T" and "\w. w \ span T \ orthogonal (p2 x) w" and eq: "p1 x + p2 x = x" for x by metis then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x using span_eq_iff \subspace T\ by blast+ then have p2: "\z. p2 z \ T'" by (simp add: T'_def orthogonal_commute) have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" using span_eq_iff p2 \subspace T'\ by blast show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: span_base) then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" proof - fix c :: real and x :: 'a have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x" by (metis eq pth_6) have f2: "c *\<^sub>R p2 x \ T'" by (simp add: \subspace T'\ p2 subspace_scale) have "c *\<^sub>R p1 x \ T" by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale) then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" using f2 f1 p12_eq by presburger qed moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" by (simp add: add.assoc add.left_commute eq) show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: p1span p2 span_base span_add) ultimately have "linear p1" "linear p2" by unfold_locales auto have "g differentiable_on p1 ` {x + y |x y. x \ S - {0} \ y \ T'}" using p12_eq \S \ T\ by (force intro: differentiable_on_subset [OF gdiff]) then have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (rule differentiable_on_compose [OF linear_imp_differentiable_on [OF \linear p1\]]) then have diff: "(\x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (intro derivative_intros linear_imp_differentiable_on [OF \linear p2\]) have "dim {x + y |x y. x \ S - {0} \ y \ T'} \ dim {x + y |x y. x \ S \ y \ T'}" by (blast intro: dim_subset) also have "... = dim S + dim T' - dim (S \ T')" using dim_sums_Int [OF \subspace S\ \subspace T'\] by (simp add: algebra_simps) also have "... < DIM('a)" using dimST dim_eq by auto finally have neg: "negligible {x + y |x y. x \ S - {0} \ y \ T'}" by (rule negligible_lowdim) have "negligible ((\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'})" by (rule negligible_differentiable_image_negligible [OF order_refl neg diff]) then have "negligible {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof (rule negligible_subset) have "\t' \ T'; s \ S; s \ 0\ \ g s + t' \ (\x. g (p1 x) + p2 x) ` {x + t' |x t'. x \ S \ x \ 0 \ t' \ T'}" for t' s using \S \ T\ p12_eq by (rule_tac x="s + t'" in image_eqI) auto then show "{x + y |x y. x \ g ` (S - {0}) \ y \ T'} \ (\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'}" by auto qed moreover have "- T' \ {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof clarsimp fix z assume "z \ T'" show "\x y. z = x + y \ x \ g ` (S - {0}) \ y \ T'" by (metis Diff_iff \z \ T'\ add.left_neutral eq geq p1 p2 singletonD) qed ultimately have "negligible (-T')" using negligible_subset by blast moreover have "negligible T'" using negligible_lowdim by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0) ultimately have "negligible (-T' \ T')" by (metis negligible_Un_eq) then show False using negligible_Un_eq non_negligible_UNIV by simp qed lemma spheremap_lemma2: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes ST: "subspace S" "subspace T" "dim S < dim T" and "S \ T" and contf: "continuous_on (sphere 0 1 \ S) f" and fim: "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" shows "\c. homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f (\x. c)" proof - have [simp]: "\x. \norm x = 1; x \ S\ \ norm (f x) = 1" using fim by (simp add: image_subset_iff) have "compact (sphere 0 1 \ S)" by (simp add: \subspace S\ closed_subspace compact_Int_closed) then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" using Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"] fim by auto have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x proof - have "norm (f x) = 1" using fim that by (simp add: image_subset_iff) then show ?thesis using g12 [OF that] by auto qed have diffg: "g differentiable_on sphere 0 1 \ S" by (metis pfg differentiable_on_polynomial_function) define h where "h \ \x. inverse(norm(g x)) *\<^sub>R g x" have h: "x \ sphere 0 1 \ S \ h x \ sphere 0 1 \ T" for x unfolding h_def using gnz [of x] by (auto simp: subspace_mul [OF \subspace T\] subsetD [OF gim]) have diffh: "h differentiable_on sphere 0 1 \ S" unfolding h_def using gnz by (fastforce intro: derivative_intros diffg differentiable_on_compose [OF diffg]) have homfg: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) f g" proof (rule homotopic_with_linear [OF contf]) show "continuous_on (sphere 0 1 \ S) g" using pfg by (simp add: differentiable_imp_continuous_on diffg) next have non0fg: "0 \ closed_segment (f x) (g x)" if "norm x = 1" "x \ S" for x proof - have "f x \ sphere 0 1" using fim that by (simp add: image_subset_iff) moreover have "norm(f x - g x) < 1/2" using g12 that by auto ultimately show ?thesis by (auto simp: norm_minus_commute dest: segment_bound) qed show "closed_segment (f x) (g x) \ T - {0}" if "x \ sphere 0 1 \ S" for x proof - have "convex T" by (simp add: \subspace T\ subspace_imp_convex) then have "convex hull {f x, g x} \ T" by (metis IntD2 closed_segment_subset fim gim image_subset_iff segment_convex_hull that) then show ?thesis using that non0fg segment_convex_hull by fastforce qed qed obtain d where d: "d \ (sphere 0 1 \ T) - h ` (sphere 0 1 \ S)" using h spheremap_lemma1 [OF ST \S \ T\ diffh] by force then have non0hd: "0 \ closed_segment (h x) (- d)" if "norm x = 1" "x \ S" for x using midpoint_between [of 0 "h x" "-d"] that h [of x] by (auto simp: between_mem_segment midpoint_def) have conth: "continuous_on (sphere 0 1 \ S) h" using differentiable_imp_continuous_on diffh by blast have hom_hd: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) h (\x. -d)" proof (rule homotopic_with_linear [OF conth continuous_on_const]) fix x assume x: "x \ sphere 0 1 \ S" have "convex hull {h x, - d} \ T" proof (rule hull_minimal) show "{h x, - d} \ T" using h d x by (force simp: subspace_neg [OF \subspace T\]) qed (simp add: subspace_imp_convex [OF \subspace T\]) with x segment_convex_hull show "closed_segment (h x) (- d) \ T - {0}" by (auto simp add: subset_Diff_insert non0hd) qed have conT0: "continuous_on (T - {0}) (\y. inverse(norm y) *\<^sub>R y)" by (intro continuous_intros) auto have sub0T: "(\y. y /\<^sub>R norm y) ` (T - {0}) \ sphere 0 1 \ T" by (fastforce simp: assms(2) subspace_mul) obtain c where homhc: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. c)" proof show "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. - d)" using d by (force simp: h_def intro: homotopic_with_eq homotopic_with_compose_continuous_left [OF hom_hd conT0 sub0T]) qed have "homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f h" by (force simp: h_def intro: homotopic_with_eq homotopic_with_compose_continuous_left [OF homfg conT0 sub0T]) then show ?thesis by (metis homotopic_with_trans [OF _ homhc]) qed lemma spheremap_lemma3: assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \ dim U" obtains T where "subspace T" "T \ U" "S \ {} \ aff_dim T = aff_dim S" "(rel_frontier S) homeomorphic (sphere 0 1 \ T)" proof (cases "S = {}") case True with \subspace U\ subspace_0 show ?thesis by (rule_tac T = "{0}" in that) auto next case False then obtain a where "a \ S" by auto then have affS: "aff_dim S = int (dim ((\x. -a+x) ` S))" by (metis hull_inc aff_dim_eq_dim) with affSU have "dim ((\x. -a+x) ` S) \ dim U" by linarith with choose_subspace_of_subspace obtain T where "subspace T" "T \ span U" and dimT: "dim T = dim ((\x. -a+x) ` S)" . show ?thesis proof (rule that [OF \subspace T\]) show "T \ U" using span_eq_iff \subspace U\ \T \ span U\ by blast show "aff_dim T = aff_dim S" using dimT \subspace T\ affS aff_dim_subspace by fastforce show "rel_frontier S homeomorphic sphere 0 1 \ T" proof - have "aff_dim (ball 0 1 \ T) = aff_dim (T)" by (metis IntI interior_ball \subspace T\ aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one) then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \ T)" using \aff_dim T = aff_dim S\ by simp have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \ T)" proof (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \convex S\ \bounded S\]) show "convex (ball 0 1 \ T)" by (simp add: \subspace T\ convex_Int subspace_imp_convex) show "bounded (ball 0 1 \ T)" by (simp add: bounded_Int) show "aff_dim S = aff_dim (ball 0 1 \ T)" by (rule affS_eq) qed also have "... = frontier (ball 0 1) \ T" proof (rule convex_affine_rel_frontier_Int [OF convex_ball]) show "affine T" by (simp add: \subspace T\ subspace_imp_affine) show "interior (ball 0 1) \ T \ {}" using \subspace T\ subspace_0 by force qed also have "... = sphere 0 1 \ T" by auto finally show ?thesis . qed qed qed proposition inessential_spheremap_lowdim_gen: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "convex S" "bounded S" "convex T" "bounded T" and affST: "aff_dim S < aff_dim T" and contf: "continuous_on (rel_frontier S) f" and fim: "f ` (rel_frontier S) \ rel_frontier T" obtains c where "homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" proof (cases "S = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis proof (cases "T = {}") case True then show ?thesis using fim that by auto next case False obtain T':: "'a set" where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" using \T \ {}\ spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a] by (force simp add: aff_dim_le_DIM) with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" using homotopy_equivalent_space_sym by blast have "aff_dim S \ int (dim T')" using affT' \subspace T'\ affST aff_dim_subspace by force with spheremap_lemma3 [OF \bounded S\ \convex S\ \subspace T'\] \S \ {}\ obtain S':: "'a set" where "subspace S'" "S' \ T'" and affS': "aff_dim S' = aff_dim S" and homT: "rel_frontier S homeomorphic sphere 0 1 \ S'" by metis with homeomorphic_imp_homotopy_eqv have relS: "sphere 0 1 \ S' homotopy_eqv rel_frontier S" using homotopy_equivalent_space_sym by blast have dimST': "dim S' < dim T'" by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) have "\c. homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" using spheremap_lemma2 homotopy_eqv_cohomotopic_triviality_null[OF relS] using homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim] by (metis \S' \ T'\ \subspace S'\ \subspace T'\ dimST') with that show ?thesis by blast qed qed lemma inessential_spheremap_lowdim: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis by (meson nullhomotopic_into_contractible f contractible_sphere that) next case False show ?thesis proof (cases "r \ 0") case True then show ?thesis by (meson f nullhomotopic_from_contractible contractible_sphere that) next case False with \\ s \ 0\ have "r > 0" "s > 0" by auto show thesis using inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f] using \0 < r\ \0 < s\ assms(1) that by (auto simp add: f aff_dim_cball) qed qed subsection\ Some technical lemmas about extending maps from cell complexes\ lemma extending_maps_Union_aux: assumes fin: "finite \" and "\S. S \ \ \ closed S" and "\S T. \S \ \; T \ \; S \ T\ \ S \ T \ K" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" using assms proof (induction \) case empty show ?case by simp next case (insert S \) then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \ T" and feq: "\x \ S \ K. f x = h x" by (meson insertI1) obtain g where contg: "continuous_on (\\) g" and gim: "g ` \\ \ T" and geq: "\x \ \\ \ K. g x = h x" using insert by auto have fg: "f x = g x" if "x \ T" "T \ \" "x \ S" for x T proof - have "T \ S \ K \ S = T" using that by (metis (no_types) insert.prems(2) insertCI) then show ?thesis using UnionI feq geq \S \ \\ subsetD that by fastforce qed moreover have "continuous_on (S \ \ \) (\x. if x \ S then f x else g x)" by (auto simp: insert closed_Union contf contg intro: fg continuous_on_cases) ultimately show ?case by (smt (verit, del_insts) Int_iff UnE complete_lattice_class.Sup_insert feq fim geq gim image_subset_iff) qed lemma extending_maps_Union: assumes fin: "finite \" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" apply (simp flip: Union_maximal_sets [OF fin]) apply (rule extending_maps_Union_aux) apply (simp_all add: Union_maximal_sets [OF fin] assms) by (metis K psubsetI) lemma extend_map_lemma: assumes "finite \" "\ \ \" "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ - \ \ aff_dim X < aff_dim T" and face: "\S T. \S \ \; T \ \\ \ (S \ T) face_of S" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof (cases "\ - \ = {}") case True show ?thesis proof show "continuous_on (\ \) f" using True \\ \ \\ contf by auto show "f ` \ \ \ rel_frontier T" using True fim by auto qed auto next case False then have "0 \ aff_dim T" by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less) then obtain i::nat where i: "int i = aff_dim T" by (metis nonneg_eq_int) have Union_empty_eq: "\{D. D = {} \ P D} = {}" for P :: "'a set \ bool" by auto have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have extendf: "\g. continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) g \ g ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) \ rel_frontier T \ (\x \ \\. g x = f x)" if "i \ aff_dim T" for i::nat using that proof (induction i) case 0 show ?case using 0 contf fim by (auto simp add: Union_empty_eq) next case (Suc p) with \bounded T\ have "rel_frontier T \ {}" by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T]) then obtain t where t: "t \ rel_frontier T" by auto have ple: "int p \ aff_dim T" using Suc.prems by force obtain h where conth: "continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) h" and him: "h ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) \ rel_frontier T" and heq: "\x. x \ \\ \ h x = f x" using Suc.IH [OF ple] by auto let ?Faces = "{D. \C \ \. D face_of C \ aff_dim D \ p}" have extendh: "\g. continuous_on D g \ g ` D \ rel_frontier T \ (\x \ D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p}). g x = h x)" if D: "D \ \ \ ?Faces" for D proof (cases "D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})") case True have "continuous_on D h" using True conth continuous_on_subset by blast moreover have "h ` D \ rel_frontier T" using True him by blast ultimately show ?thesis by blast next case False note notDsub = False show ?thesis proof (cases "\a. D = {a}") case True then obtain a where "D = {a}" by auto with notDsub t show ?thesis by (rule_tac x="\x. t" in exI) simp next case False have "D \ {}" using notDsub by auto have Dnotin: "D \ \ \ {D. \C \ \. D face_of C \ aff_dim D < p}" using notDsub by auto then have "D \ \" by simp have "D \ ?Faces - {D. \C \ \. D face_of C \ aff_dim D < p}" using Dnotin that by auto then obtain C where "C \ \" "D face_of C" and affD: "aff_dim D = int p" by auto then have "bounded D" using face_of_polytope_polytope poly polytope_imp_bounded by blast then have [simp]: "\ affine D" using affine_bounded_eq_trivial False \D \ {}\ \bounded D\ by blast have "{F. F facet_of D} \ {E. E face_of C \ aff_dim E < int p}" by clarify (metis \D face_of C\ affD eq_iff face_of_trans facet_of_def zle_diff1_eq) moreover have "polyhedron D" using \C \ \\ \D face_of C\ face_of_polytope_polytope poly polytope_imp_polyhedron by auto ultimately have relf_sub: "rel_frontier D \ \ {E. E face_of C \ aff_dim E < p}" by (simp add: rel_frontier_of_polyhedron Union_mono) then have him_relf: "h ` rel_frontier D \ rel_frontier T" using \C \ \\ him by blast have "convex D" by (simp add: \polyhedron D\ polyhedron_imp_convex) have affD_lessT: "aff_dim D < aff_dim T" using Suc.prems affD by linarith have contDh: "continuous_on (rel_frontier D) h" using \C \ \\ relf_sub by (blast intro: continuous_on_subset [OF conth]) then have *: "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c)) = (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x\rel_frontier D. g x = h x))" by (simp add: assms rel_frontier_eq_empty him_relf nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) have "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c))" by (metis inessential_spheremap_lowdim_gen [OF \convex D\ \bounded D\ \convex T\ \bounded T\ affD_lessT contDh him_relf]) then obtain g where contg: "continuous_on UNIV g" and gim: "range g \ rel_frontier T" and gh: "\x. x \ rel_frontier D \ g x = h x" by (metis *) have "D \ E \ rel_frontier D" if "E \ \ \ {D. Bex \ ((face_of) D) \ aff_dim D < int p}" for E proof (rule face_of_subset_rel_frontier) show "D \ E face_of D" using that proof safe assume "E \ \" then show "D \ E face_of D" by (meson \C \ \\ \D face_of C\ assms(2) face' face_of_Int_subface face_of_refl_eq poly polytope_imp_convex subsetD) next fix x assume "aff_dim E < int p" "x \ \" "E face_of x" then show "D \ E face_of D" by (meson \C \ \\ \D face_of C\ face' face_of_Int_subface that) qed show "D \ E \ D" using that notDsub by auto qed moreover have "continuous_on D g" using contg continuous_on_subset by blast ultimately show ?thesis by (rule_tac x=g in exI) (use gh gim in fastforce) qed qed have intle: "i < 1 + int j \ i \ int j" for i j by auto have "finite \" using \finite \\ \\ \ \\ rev_finite_subset by blast moreover have "finite (?Faces)" proof - have \
: "finite (\ {{D. D face_of C} |C. C \ \})" by (auto simp: \finite \\ finite_polytope_faces poly) show ?thesis by (auto intro: finite_subset [OF _ \
]) qed ultimately have fin: "finite (\ \ ?Faces)" by simp have clo: "closed S" if "S \ \ \ ?Faces" for S using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "\ Y \ X" for X Y proof - have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E by (rule face_of_Int_subface [OF _ _ XY]) (auto simp: face' DE) show ?thesis using that apply clarsimp by (smt (verit, ccfv_SIG) IntI face_of_aff_dim_lt face_of_imp_convex [of X] face_of_imp_convex [of Y] face_of_trans ff) qed obtain g where "continuous_on (\(\ \ ?Faces)) g" "g ` \(\ \ ?Faces) \ rel_frontier T" "(\x \ \(\ \ ?Faces) \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < p}). g x = h x)" by (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+) then show ?case by (simp add: intle local.heq [symmetric], blast) qed have eq: "\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i}) = \\" proof show "\(\ \ {D. \C\\. D face_of C \ aff_dim D < int i}) \ \\" using \\ \ \\ face_of_imp_subset by fastforce show "\\ \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < i})" proof (rule Union_mono) show "\ \ \ \ {D. \C\\. D face_of C \ aff_dim D < int i}" using face by (fastforce simp: aff i) qed qed have "int i \ aff_dim T" by (simp add: i) then show ?thesis using extendf [of i] unfolding eq by (metis that) qed lemma extend_map_lemma_cofinite0: assumes "finite \" and "pairwise (\S T. S \ T \ K) \" and "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" shows "\C g. finite C \ disjnt C U \ card C \ card \ \ continuous_on (\\ - C) g \ g ` (\\ - C) \ T \ (\x \ (\\ - C) \ K. g x = h x)" using assms proof induction case empty then show ?case by force next case (insert X \) then have "closed X" and clo: "\X. X \ \ \ closed X" and \: "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and pwX: "\Y. Y \ \ \ Y \ X \ X \ Y \ K \ Y \ X \ K" and pwF: "pairwise (\ S T. S \ T \ K) \" by (simp_all add: pairwise_insert) obtain C g where C: "finite C" "disjnt C U" "card C \ card \" and contg: "continuous_on (\\ - C) g" and gim: "g ` (\\ - C) \ T" and gh: "\x. x \ (\\ - C) \ K \ g x = h x" using insert.IH [OF pwF \ clo] by auto obtain a f where "a \ U" and contf: "continuous_on (X - {a}) f" and fim: "f ` (X - {a}) \ T" and fh: "(\x \ X \ K. f x = h x)" using insert.prems by (meson insertI1) show ?case proof (intro exI conjI) show "finite (insert a C)" by (simp add: C) show "disjnt (insert a C) U" using C \a \ U\ by simp show "card (insert a C) \ card (insert X \)" by (simp add: C card_insert_if insert.hyps le_SucI) have "closed (\\)" using clo insert.hyps by blast have "continuous_on (X - insert a C) f" using contf by (force simp: elim: continuous_on_subset) moreover have "continuous_on (\ \ - insert a C) g" using contg by (force simp: elim: continuous_on_subset) ultimately have "continuous_on (X - insert a C \ (\\ - insert a C)) (\x. if x \ X then f x else g x)" apply (intro continuous_on_cases_local; simp add: closedin_closed) using \closed X\ apply blast using \closed (\\)\ apply blast using fh gh insert.hyps pwX by fastforce then show "continuous_on (\(insert X \) - insert a C) (\a. if a \ X then f a else g a)" by (blast intro: continuous_on_subset) show "\x\(\(insert X \) - insert a C) \ K. (if x \ X then f x else g x) = h x" using gh by (auto simp: fh) show "(\a. if a \ X then f a else g a) ` (\(insert X \) - insert a C) \ T" using fim gim by auto force qed qed lemma extend_map_lemma_cofinite1: assumes "finite \" and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" and clo: "\X. X \ \ \ closed X" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ T" "\x. x \ (\\ - C) \ K \ g x = h x" proof - let ?\ = "{X \ \. \Y\\. \ X \ Y}" have [simp]: "\?\ = \\" by (simp add: Union_maximal_sets assms) have fin: "finite ?\" by (force intro: finite_subset [OF _ \finite \\]) have pw: "pairwise (\ S T. S \ T \ K) ?\" by (simp add: pairwise_def) (metis K psubsetI) have "card {X \ \. \Y\\. \ X \ Y} \ card \" by (simp add: \finite \\ card_mono) moreover obtain C g where "finite C \ disjnt C U \ card C \ card ?\ \ continuous_on (\?\ - C) g \ g ` (\?\ - C) \ T \ (\x \ (\?\ - C) \ K. g x = h x)" using extend_map_lemma_cofinite0 [OF fin pw, of U T h] by (fastforce intro!: clo \) ultimately show ?thesis by (rule_tac C=C and g=g in that) auto qed lemma extend_map_lemma_cofinite: assumes "finite \" "\ \ \" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and aff: "\X. X \ \ - \ \ aff_dim X \ aff_dim T" obtains C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\ \ - C) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof - define \ where "\ \ \ \ {D. \C \ \ - \. D face_of C \ aff_dim D < aff_dim T}" have "finite \" using assms finite_subset by blast have *: "finite (\{{D. D face_of C} |C. C \ \})" using finite_polytope_faces poly \finite \\ by force then have "finite \" by (auto simp: \_def \finite \\ intro: finite_subset [OF _ *]) have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" by (simp add: \_def) (smt (verit) \\ \ \\ DiffE face' face_of_Int_subface in_mono inf.idem) obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" and hf: "\x. x \ \\ \ h x = f x" proof (rule extend_map_lemma [OF \finite \\ [unfolded \_def] Un_upper1 T]) show "\X. \X \ \ \ {D. \C\\ - \. D face_of C \ aff_dim D < aff_dim T}\ \ polytope X" using \\ \ \\ face_of_polytope_polytope poly by fastforce qed (use * \_def contf fim in auto) have "bounded (\\)" using \finite \\ \\ \ \\ poly polytope_imp_bounded by blast then have "\\ \ UNIV" by auto then obtain a where a: "a \ \\" by blast have \: "\a g. a \ \\ \ continuous_on (D - {a}) g \ g ` (D - {a}) \ rel_frontier T \ (\x \ D \ \\. g x = h x)" if "D \ \" for D proof (cases "D \ \\") case True then have "h ` (D - {a}) \ rel_frontier T" "continuous_on (D - {a}) h" using him by (blast intro!: \a \ \\\ continuous_on_subset [OF conth])+ then show ?thesis using a by blast next case False note D_not_subset = False show ?thesis proof (cases "D \ \") case True with D_not_subset show ?thesis by (auto simp: \_def) next case False then have affD: "aff_dim D \ aff_dim T" by (simp add: \D \ \\ aff) show ?thesis proof (cases "rel_interior D = {}") case True with \D \ \\ poly a show ?thesis by (force simp: rel_interior_eq_empty polytope_imp_convex) next case False then obtain b where brelD: "b \ rel_interior D" by blast have "polyhedron D" by (simp add: poly polytope_imp_polyhedron that) have "rel_frontier D retract_of affine hull D - {b}" by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD) then obtain r where relfD: "rel_frontier D \ affine hull D - {b}" and contr: "continuous_on (affine hull D - {b}) r" and rim: "r ` (affine hull D - {b}) \ rel_frontier D" and rid: "\x. x \ rel_frontier D \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof (intro exI conjI ballI) show "b \ \\" proof clarify fix E assume "b \ E" "E \ \" then have "E \ D face_of E \ E \ D face_of D" using \\ \ \\ face' that by auto with face_of_subset_rel_frontier \E \ \\ \b \ E\ brelD rel_interior_subset [of D] D_not_subset rel_frontier_def \_def show False by blast qed have "r ` (D - {b}) \ r ` (affine hull D - {b})" by (simp add: Diff_mono hull_subset image_mono) also have "... \ rel_frontier D" by (rule rim) also have "... \ \{E. E face_of D \ aff_dim E < aff_dim T}" using affD by (force simp: rel_frontier_of_polyhedron [OF \polyhedron D\] facet_of_def) also have "... \ \(\)" using D_not_subset \_def that by fastforce finally have rsub: "r ` (D - {b}) \ \(\)" . show "continuous_on (D - {b}) (h \ r)" proof (rule continuous_on_compose) show "continuous_on (D - {b}) r" by (meson Diff_mono continuous_on_subset contr hull_subset order_refl) show "continuous_on (r ` (D - {b})) h" by (simp add: Diff_mono hull_subset continuous_on_subset [OF conth rsub]) qed show "(h \ r) ` (D - {b}) \ rel_frontier T" using brelD him rsub by fastforce show "(h \ r) x = h x" if x: "x \ D \ \\" for x proof - consider A where "x \ D" "A \ \" "x \ A" | A B where "x \ D" "A face_of B" "B \ \" "B \ \" "aff_dim A < aff_dim T" "x \ A" using x by (auto simp: \_def) then have xrel: "x \ rel_frontier D" proof cases case 1 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" using \A \ \\ \\ \ \\ face \D \ \\ by blast show "D \ A \ D" using \A \ \\ D_not_subset \_def by blast qed (auto simp: 1) next case 2 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) have "D face_of D" by (simp add: \polyhedron D\ polyhedron_imp_convex face_of_refl) then show "D \ A face_of D" by (meson "2"(2) "2"(3) \D \ \\ face' face_of_Int_Int face_of_face) show "D \ A \ D" using "2" D_not_subset \_def by blast qed (auto simp: 2) qed show ?thesis by (simp add: rid xrel) qed qed qed qed qed have clo: "\S. S \ \ \ closed S" by (simp add: poly polytope_imp_closed) obtain C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" and gh: "\x. x \ (\\ - C) \ \\ \ g x = h x" proof (rule extend_map_lemma_cofinite1 [OF \finite \\ \ clo]) show "X \ Y \ \\" if XY: "X \ \" "Y \ \" and "\ X \ Y" "\ Y \ X" for X Y proof (cases "X \ \") case True then show ?thesis by (auto simp: \_def) next case False have "X \ Y \ X" using \\ X \ Y\ by blast with XY show ?thesis by (clarsimp simp: \_def) (metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl not_le poly polytope_imp_convex) qed qed (blast)+ with \\ \ \\ show ?thesis by (rule_tac C=C and g=g in that) (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) qed text\The next two proofs are similar\ theorem extend_map_cell_complex_to_sphere: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X < aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T - 1" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" proof (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly _ face]) show "\X. X \ \ \ aff_dim X \ aff_dim T - 1" by (simp add: aff) qed auto obtain h where conth: "continuous_on (\\) h" and him: "h ` \\ \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) qed (use \finite \\ T polyG affG faceG gim in fastforce)+ show ?thesis proof show "continuous_on (\\) h" using \\\ = \\\ conth by auto show "h ` \\ \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "x \ \\" using \\\ = \\\ \S \ \\\ that by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce have "h x = g x" using \X \ \\ \X \ V\ \x \ X\ hg by auto also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed qed qed theorem extend_map_cell_complex_to_sphere_cofinite: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains C g where "finite C" "disjnt C S" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" by (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly aff face]) auto obtain C h where "finite C" and dis: "disjnt C (\(\ \ Pow V))" and card: "card C \ card \" and conth: "continuous_on (\\ - C) h" and him: "h ` (\\ - C) \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma_cofinite [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) show "g ` \(\ \ Pow V) \ rel_frontier T" using gim by force qed (auto intro: \finite \\ T polyG affG dest: faceG) have Ssub: "S \ \(\ \ Pow V)" proof fix x assume "x \ S" then have "x \ \\" using \\\ = \\\ \S \ \\\ by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce then show "x \ \(\ \ Pow V)" using \X \ \\ \x \ X\ by blast qed show ?thesis proof show "continuous_on (\\-C) h" using \\\ = \\\ conth by auto show "h ` (\\ - C) \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "h x = g x" using Ssub hg that by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed show "disjnt C S" using dis Ssub by (meson disjnt_iff subset_eq) qed (intro \finite C\) qed subsection\ Special cases and corollaries involving spheres\ proposition extend_map_affine_to_sphere_cofinite_simple: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" and aff: "aff_dim T \ aff_dim U" and "S \ T" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" obtains K g where "finite K" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - have "\K g. finite K \ disjnt K S \ continuous_on (T - K) g \ g ` (T - K) \ rel_frontier U \ (\x \ S. g x = f x)" if "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" for T proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with \bounded U\ have "aff_dim U \ 0" using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto with aff have "aff_dim T \ 0" by auto then obtain a where "T \ {a}" using \affine T\ affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto then show ?thesis using \S = {}\ fim by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset) next case False then obtain a where "a \ rel_frontier U" by auto then show ?thesis using continuous_on_const [of _ a] \S = {}\ by force qed next case False have "bounded S" by (simp add: \compact S\ compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define bbox where "bbox \ cbox (-(b+One)) (b+One)" have "cbox (-b) b \ bbox" by (auto simp: bbox_def algebra_simps intro!: subset_box_imp) with b \S \ T\ have "S \ bbox \ T" by auto then have Ssub: "S \ \{bbox \ T}" by auto then have "aff_dim (bbox \ T) \ aff_dim U" by (metis aff aff_dim_subset inf_commute inf_le1 order_trans) obtain K g where K: "finite K" "disjnt K S" and contg: "continuous_on (\{bbox \ T} - K) g" and gim: "g ` (\{bbox \ T} - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_cell_complex_to_sphere_cofinite [OF _ Ssub _ \convex U\ \bounded U\ _ _ _ contf fim]) show "closed S" using \compact S\ compact_eq_bounded_closed by auto show poly: "\X. X \ {bbox \ T} \ polytope X" by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \affine T\) show "\X Y. \X \ {bbox \ T}; Y \ {bbox \ T}\ \ X \ Y face_of X" by (simp add:poly face_of_refl polytope_imp_convex) show "\X. X \ {bbox \ T} \ aff_dim X \ aff_dim U" by (simp add: \aff_dim (bbox \ T) \ aff_dim U\) qed auto define fro where "fro \ \d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))" obtain d where d12: "1/2 \ d" "d \ 1" and dd: "disjnt K (fro d)" proof (rule disjoint_family_elem_disjnt [OF _ \finite K\]) show "infinite {1/2..1::real}" by (simp add: infinite_Icc) have dis1: "disjnt (fro x) (fro y)" if "x b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def) have clo_cbT: "closed (cbox (- c) c \ T)" by (simp add: affine_closed closed_Int closed_cbox \affine T\) have cpT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ b cbsub(2) \S \ T\ by fastforce have "closest_point (cbox (- c) c \ T) x \ K" if "x \ T" "x \ K" for x proof (cases "x \ cbox (-c) c") case True with that show ?thesis by (simp add: closest_point_self) next case False have int_ne: "interior (cbox (-c) c) \ T \ {}" using \S \ {}\ \S \ T\ b \cbox (- b) b \ box (- c) c\ by force have "convex T" by (meson \affine T\ affine_imp_convex) then have "x \ affine hull (cbox (- c) c \ T)" by (metis Int_commute Int_iff \S \ {}\ \S \ T\ cbsub(1) \x \ T\ affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox) then have "x \ affine hull (cbox (- c) c \ T) - rel_interior (cbox (- c) c \ T)" by (meson DiffI False Int_iff rel_interior_subset subsetCE) then have "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne]) moreover have "(rel_frontier (cbox (- c) c \ T)) \ fro d" by (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) (auto simp: fro_def c_def) ultimately show ?thesis using dd by (force simp: disjnt_def) qed then have cpt_subset: "closest_point (cbox (- c) c \ T) ` (T - K) \ \{bbox \ T} - K" using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force show ?thesis proof (intro conjI ballI exI) have "continuous_on (T - K) (closest_point (cbox (- c) c \ T))" proof (rule continuous_on_closest_point) show "convex (cbox (- c) c \ T)" by (simp add: affine_imp_convex convex_Int \affine T\) show "closed (cbox (- c) c \ T)" using clo_cbT by blast show "cbox (- c) c \ T \ {}" using \S \ {}\ cbsub(2) b that by auto qed then show "continuous_on (T - K) (g \ closest_point (cbox (- c) c \ T))" by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset]) have "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ g ` (\{bbox \ T} - K)" by (metis image_comp image_mono cpt_subset) also have "... \ rel_frontier U" by (rule gim) finally show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ rel_frontier U" . show "(g \ closest_point (cbox (- c) c \ T)) x = f x" if "x \ S" for x proof - have "(g \ closest_point (cbox (- c) c \ T)) x = g x" unfolding o_def by (metis IntI \S \ T\ b cbsub(2) closest_point_self subset_eq that) also have "... = f x" by (simp add: that gf) finally show ?thesis . qed qed (auto simp: K) qed then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (affine hull T - K) g" and gim: "g ` (affine hull T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (metis aff affine_affine_hull aff_dim_affine_hull order_trans [OF \S \ T\ hull_subset [of T affine]]) then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset) then show ?thesis by (rule_tac K="K \ T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg) qed subsection\Extending maps to spheres\ (*Up to extend_map_affine_to_sphere_cofinite_gen*) lemma extend_map_affine_to_sphere1: fixes f :: "'a::euclidean_space \ 'b::topological_space" assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" and fim: "f ` (U - K) \ T" and comps: "\C. \C \ components(U - S); C \ K \ {}\ \ C \ L \ {}" and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \ U" obtains g where "continuous_on (U - L) g" "g ` (U - L) \ T" "\x. x \ S \ g x = f x" proof (cases "K = {}") case True then show ?thesis by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI image_eqI subset_iff that) next case False have "S \ U" using clo closedin_limpt by blast then have "(U - S) \ K \ {}" by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute) then have "\(components (U - S)) \ K \ {}" using Union_components by simp then obtain C0 where C0: "C0 \ components (U - S)" "C0 \ K \ {}" by blast have "convex U" by (simp add: affine_imp_convex \affine U\) then have "locally connected U" by (rule convex_imp_locally_connected) have "\a g. a \ C \ a \ L \ continuous_on (S \ (C - {a})) g \ g ` (S \ (C - {a})) \ T \ (\x \ S. g x = f x)" if C: "C \ components (U - S)" and CK: "C \ K \ {}" for C proof - have "C \ U-S" "C \ L \ {}" by (simp_all add: in_components_subset comps that) then obtain a where a: "a \ C" "a \ L" by auto have opeUC: "openin (top_of_set U) C" by (metis C \locally connected U\ clo closedin_def locally_connected_open_component topspace_euclidean_subtopology) then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" using openin_contains_cball by (metis \a \ C\) then have "ball a d \ U \ C" by auto obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" and subC: "{x. (\ (h x = x \ k x = x))} \ C" and bou: "bounded {x. (\ (h x = x \ k x = x))}" and hin: "\x. x \ C \ K \ h x \ ball a d \ U" proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) show "openin (top_of_set C) (ball a d \ U)" by (metis open_ball \C \ U\ \ball a d \ U \ C\ inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology) show "openin (top_of_set (affine hull C)) C" by (metis \a \ C\ \openin (top_of_set U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) show "ball a d \ U \ {}" using \0 < d\ \C \ U\ \a \ C\ by force show "finite (C \ K)" by (simp add: \finite K\) show "S \ C \ affine hull C" by (metis \S \ U\ \a \ C\ affine_hull_eq affine_hull_openin assms(2) empty_iff hull_subset le_sup_iff opeUC) show "connected C" by (metis C in_components_connected) qed auto have a_BU: "a \ ball a d \ U" using \0 < d\ \C \ U\ \a \ C\ by auto have "rel_frontier (cball a d \ U) retract_of (affine hull (cball a d \ U) - {a})" proof (rule rel_frontier_retract_of_punctured_affine_hull) show "bounded (cball a d \ U)" "convex (cball a d \ U)" by (auto simp: \convex U\ convex_Int) show "a \ rel_interior (cball a d \ U)" by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) qed moreover have "rel_frontier (cball a d \ U) = frontier (cball a d) \ U" by (metis a_BU \affine U\ convex_affine_rel_frontier_Int convex_cball equals0D interior_cball) moreover have "affine hull (cball a d \ U) = U" by (metis \convex U\ a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \affine U\ equals0D inf.commute interior_cball) ultimately have "frontier (cball a d) \ U retract_of (U - {a})" by metis then obtain r where contr: "continuous_on (U - {a}) r" and rim: "r ` (U - {a}) \ sphere a d" "r ` (U - {a}) \ U" and req: "\x. x \ sphere a d \ U \ r x = x" using \affine U\ by (auto simp: retract_of_def retraction_def hull_same) define j where "j \ \x. if x \ ball a d then r x else x" have kj: "\x. x \ S \ k (j x) = x" using \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def subC by auto have Uaeq: "U - {a} = (cball a d - {a}) \ U \ (U - ball a d)" using \0 < d\ by auto have jim: "j ` (S \ (C - {a})) \ (S \ C) - ball a d" proof clarify fix y assume "y \ S \ (C - {a})" then have "y \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by auto then have "r y \ sphere a d" using rim by auto then show "j y \ S \ C - ball a d" unfolding j_def using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim by (metis Diff_iff Int_iff Un_iff subsetD cball_diff_eq_sphere image_subset_iff) qed have contj: "continuous_on (U - {a}) j" unfolding j_def Uaeq proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric]) show "\T. closed T \ (cball a d - {a}) \ U = (U - {a}) \ T" using affine_closed \affine U\ by (rule_tac x="(cball a d) \ U" in exI) blast show "\T. closed T \ U - ball a d = (U - {a}) \ T" using \0 < d\ \affine U\ by (rule_tac x="U - ball a d" in exI) (force simp: affine_closed) show "continuous_on ((cball a d - {a}) \ U) r" by (force intro: continuous_on_subset [OF contr]) qed have fT: "x \ U - K \ f x \ T" for x using fim by blast show ?thesis proof (intro conjI exI) show "continuous_on (S \ (C - {a})) (f \ k \ j)" proof (intro continuous_on_compose) have "S \ (C - {a}) \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by force then show "continuous_on (S \ (C - {a})) j" by (rule continuous_on_subset [OF contj]) have "j ` (S \ (C - {a})) \ S \ C" using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by blast then show "continuous_on (j ` (S \ (C - {a}))) k" by (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) show "continuous_on (k ` j ` (S \ (C - {a}))) f" proof (clarify intro!: continuous_on_subset [OF contf]) fix y assume "y \ S \ (C - {a})" have ky: "k y \ S \ C" using homeomorphism_image2 [OF homhk] \y \ S \ (C - {a})\ by blast have jy: "j y \ S \ C - ball a d" using Un_iff \y \ S \ (C - {a})\ jim by auto have "k (j y) \ U" using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy by blast moreover have "k (j y) \ K" using K unfolding disjnt_iff by (metis DiffE Int_iff Un_iff hin homeomorphism_def homhk image_eqI jy) ultimately show "k (j y) \ U - K" by blast qed qed have ST: "\x. x \ S \ (f \ k \ j) x \ T" proof (simp add: kj) show "\x. x \ S \ f x \ T" using K unfolding disjnt_iff by (metis DiffI \S \ U\ subsetD fim image_subset_iff) qed moreover have "(f \ k \ j) x \ T" if "x \ C" "x \ a" "x \ S" for x proof - have rx: "r x \ sphere a d" using \C \ U\ rim that by fastforce have jj: "j x \ S \ C - ball a d" using jim that by blast have "k (j x) = j x \ k (j x) \ C \ j x \ C" by (metis Diff_iff Int_iff Un_iff \S \ U\ subsetD d j_def jj rx sphere_cball that(1)) then have kj: "k (j x) \ C" using homeomorphism_apply2 [OF homhk, of "j x"] \C \ U\ \S \ U\ a rx by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC) then show ?thesis by (metis DiffE DiffI IntD1 IntI \C \ U\ comp_apply fT hin homeomorphism_apply2 homhk jj kj subset_eq) qed ultimately show "(f \ k \ j) ` (S \ (C - {a})) \ T" by force show "\x\S. (f \ k \ j) x = f x" using kj by simp qed (auto simp: a) qed then obtain a h where ah: "\C. \C \ components (U - S); C \ K \ {}\ \ a C \ C \ a C \ L \ continuous_on (S \ (C - {a C})) (h C) \ h C ` (S \ (C - {a C})) \ T \ (\x \ S. h C x = f x)" using that by metis define F where "F \ {C \ components (U - S). C \ K \ {}}" define G where "G \ {C \ components (U - S). C \ K = {}}" define UF where "UF \ (\C\F. C - {a C})" have "C0 \ F" by (auto simp: F_def C0) have "finite F" proof (subst finite_image_iff [of "\C. C \ K" F, symmetric]) show "inj_on (\C. C \ K) F" unfolding F_def inj_on_def using components_nonoverlap by blast show "finite ((\C. C \ K) ` F)" unfolding F_def by (rule finite_subset [of _ "Pow K"]) (auto simp: \finite K\) qed obtain g where contg: "continuous_on (S \ UF) g" and gh: "\x i. \i \ F; x \ (S \ UF) \ (S \ (i - {a i}))\ \ g x = h i x" proof (rule pasting_lemma_exists_closed [OF \finite F\]) let ?X = "top_of_set (S \ UF)" show "topspace ?X \ (\C\F. S \ (C - {a C}))" using \C0 \ F\ by (force simp: UF_def) show "closedin (top_of_set (S \ UF)) (S \ (C - {a C}))" if "C \ F" for C proof (rule closedin_closed_subset [of U "S \ C"]) have "C \ components (U - S)" using F_def that by blast then show "closedin (top_of_set U) (S \ C)" by (rule closedin_Un_complement_component [OF \locally connected U\ clo]) next have "x = a C'" if "C' \ F" "x \ C'" "x \ U" for x C' proof - have "\A. x \ \A \ C' \ A" using \x \ C'\ by blast with that show "x = a C'" by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq) qed then show "S \ UF \ U" using \S \ U\ by (force simp: UF_def) next show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" using F_def UF_def components_nonoverlap that by auto qed show "continuous_map (subtopology ?X (S \ (C' - {a C'}))) euclidean (h C')" if "C' \ F" for C' proof - have C': "C' \ components (U - S)" "C' \ K \ {}" using F_def that by blast+ show ?thesis using ah [OF C'] by (auto simp: F_def subtopology_subtopology intro: continuous_on_subset) qed show "\i j x. \i \ F; j \ F; x \ topspace ?X \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ \ h i x = h j x" using components_eq by (fastforce simp: components_eq F_def ah) qed auto have SU': "S \ \G \ (S \ UF) \ U" using \S \ U\ in_components_subset by (auto simp: F_def G_def UF_def) have clo1: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ \G)" proof (rule closedin_closed_subset [OF _ SU']) have *: "\C. C \ F \ openin (top_of_set U) C" unfolding F_def by (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_connected_open_component mem_Collect_eq topspace_euclidean_subtopology) show "closedin (top_of_set U) (U - UF)" unfolding UF_def by (force intro: openin_delete *) have "(\x\F. x - {a x}) \ S = {}" "\ G \ U" using in_components_subset by (auto simp: F_def G_def) moreover have "\ G \ UF = {}" using components_nonoverlap by (fastforce simp: F_def G_def UF_def) ultimately show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" using UF_def \S \ U\ by auto qed have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) show "closedin (top_of_set U) (\C\F. S \ C)" proof (rule closedin_Union) show "\T. T \ (\) S ` F \ closedin (top_of_set U) T" using F_def \locally connected U\ clo closedin_Un_complement_component by blast qed (simp add: \finite F\) show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" proof show "\ ((\) S ` F) \ (S \ \ G \ (S \ UF)) \ S \ UF" using components_eq [of _ "U-S"] by (auto simp add: F_def G_def UF_def disjoint_iff_not_equal) qed (use UF_def \C0 \ F\ in blast) qed have SUG: "S \ \G \ U - K" using \S \ U\ K in_components_subset[of _ "U-S"] by (force simp: G_def disjnt_iff) then have contf': "continuous_on (S \ \G) f" by (rule continuous_on_subset [OF contf]) have contg': "continuous_on (S \ UF) g" by (simp add: contg) have "\x. \S \ U; x \ S\ \ f x = g x" by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" using \S \ U\ components_eq [of _ "U-S"] by (fastforce simp add: F_def G_def UF_def) have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) show ?thesis proof have UF: "\F - L \ UF" unfolding F_def UF_def using ah by blast have "U - S - L = \(components (U - S)) - L" by simp also have "... = \F \ \G - L" unfolding F_def G_def by blast also have "... \ UF \ \G" using UF by blast finally have "U - L \ S \ \G \ (S \ UF)" by blast then show "continuous_on (U - L) (\x. if x \ S \ \G then f x else g x)" by (rule continuous_on_subset [OF cont]) have "((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ ((U - L) \ (-S \ UF))" using \U - L \ S \ \G \ (S \ UF)\ by auto moreover have "g ` ((U - L) \ (-S \ UF)) \ T" proof - have "g x \ T" if "x \ U" "x \ L" "x \ S" "C \ F" "x \ C" "x \ a C" for x C proof (subst gh) show "x \ (S \ UF) \ (S \ (C - {a C}))" using that by (auto simp: UF_def) show "h C x \ T" using ah that by (fastforce simp add: F_def) qed (rule that) then show ?thesis by (force simp: UF_def) qed ultimately have "g ` ((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ T" using image_mono order_trans by blast moreover have "f ` ((U - L) \ (S \ \G)) \ T" using fim SUG by blast ultimately show "(\x. if x \ S \ \G then f x else g x) ` (U - L) \ T" by force show "\x. x \ S \ (if x \ S \ \G then f x else g x) = f x" by (simp add: F_def G_def) qed qed lemma extend_map_affine_to_sphere2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" "affine T" "S \ T" and affTU: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and ovlap: "\C. C \ components(T - S) \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - obtain K g where K: "finite K" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" using assms extend_map_affine_to_sphere_cofinite_simple by metis have "(\y C. C \ components (T - S) \ x \ C \ y \ C \ y \ L)" if "x \ K" for x proof - have "x \ T-S" using \K \ T\ \disjnt K S\ disjnt_def that by fastforce then obtain C where "C \ components(T - S)" "x \ C" by (metis UnionE Union_components) with ovlap [of C] show ?thesis by blast qed then obtain \ where \: "\x. x \ K \ \C. C \ components (T - S) \ x \ C \ \ x \ C \ \ x \ L" by metis obtain h where conth: "continuous_on (T - \ ` K) h" and him: "h ` (T - \ ` K) \ rel_frontier U" and hg: "\x. x \ S \ h x = g x" proof (rule extend_map_affine_to_sphere1 [OF \finite K\ \affine T\ contg gim, of S "\ ` K"]) show cloTS: "closedin (top_of_set T) S" by (simp add: \compact S\ \S \ T\ closed_subset compact_imp_closed) show "\C. \C \ components (T - S); C \ K \ {}\ \ C \ \ ` K \ {}" using \ components_eq by blast qed (use K in auto) show ?thesis proof show *: "\ ` K \ L" using \ by blast show "finite (\ ` K)" by (simp add: K) show "\ ` K \ T" by clarify (meson \ Diff_iff contra_subsetD in_components_subset) show "continuous_on (T - \ ` K) h" by (rule conth) show "disjnt (\ ` K) S" using K \ in_components_subset by (fastforce simp: disjnt_def) qed (simp_all add: him hg gf) qed proposition extend_map_affine_to_sphere_cofinite_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with aff have "aff_dim T \ 0" using affine_bounded_eq_lowdim \bounded U\ order_trans by (auto simp add: rel_frontier_eq_empty) with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0" by linarith then show ?thesis proof cases assume "aff_dim T = -1" then have "T = {}" by (simp add: aff_dim_empty) then show ?thesis by (rule_tac K="{}" in that) auto next assume "aff_dim T = 0" then obtain a where "T = {a}" using aff_dim_eq_0 by blast then have "a \ L" using dis [of "{a}"] \S = {}\ by (auto simp: in_components_self) with \S = {}\ \T = {a}\ show ?thesis by (rule_tac K="{a}" and g=f in that) auto qed next case False then obtain y where "y \ rel_frontier U" by auto with \S = {}\ show ?thesis by (rule_tac K="{}" and g="\x. y" in that) (auto) qed next case False have "bounded S" by (simp add: assms compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define LU where "LU \ L \ (\ {C \ components (T - S). \bounded C} - cbox (-(b+One)) (b+One))" obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim]) show "C \ LU \ {}" if "C \ components (T - S)" for C proof (cases "bounded C") case True with dis that show ?thesis unfolding LU_def by fastforce next case False then have "\ bounded (\{C \ components (T - S). \ bounded C})" by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) then show ?thesis by (simp add: LU_def disjoint_iff) (meson False bounded_cbox bounded_subset subset_iff that) qed qed blast have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" "x \ box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)" "0 \ m" "m < n" "n \ 1" for m n x using that by (auto simp: mem_box algebra_simps) have "disjoint_family_on (\d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}" by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *) then obtain d where d12: "1/2 \ d" "d \ 1" and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))" using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"] by (auto simp: \finite K\) define c where "c \ b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ cbox (-(b+One)) (b+One)" using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib) have clo_cT: "closed (cbox (- c) c \ T)" using affine_closed \affine T\ by blast have cT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ \S \ T\ b cbsub by fastforce have S_sub_cc: "S \ cbox (- c) c" using \cbox (- b) b \ cbox (- c) c\ b by auto show ?thesis proof show "finite (K \ cbox (-(b+One)) (b+One))" using \finite K\ by blast show "K \ cbox (- (b + One)) (b + One) \ L" using \K \ LU\ by (auto simp: LU_def) show "K \ cbox (- (b + One)) (b + One) \ T" using \K \ T\ by auto show "disjnt (K \ cbox (- (b + One)) (b + One)) S" using \disjnt K S\ by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1) have cloTK: "closest_point (cbox (- c) c \ T) x \ T - K" if "x \ T" and Knot: "x \ K \ x \ cbox (- b - One) (b + One)" for x proof (cases "x \ cbox (- c) c") case True with \x \ T\ show ?thesis using cbsub(3) Knot by (force simp: closest_point_self) next case False have clo_in_rf: "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI) have "T \ interior (cbox (- c) c) \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then show "x \ affine hull (cbox (- c) c \ T)" by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \affine T\ hull_inc that(1)) next show "False" if "x \ rel_interior (cbox (- c) c \ T)" proof - have "interior (cbox (- c) c) \ T \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then have "affine hull (T \ cbox (- c) c) = T" using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"] by (simp add: affine_imp_convex \affine T\ inf_commute) then show ?thesis by (meson subsetD le_inf_iff rel_interior_subset that False) qed qed have "closest_point (cbox (- c) c \ T) x \ K" proof assume inK: "closest_point (cbox (- c) c \ T) x \ K" have "\x. x \ K \ x \ frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))" by (metis ddis disjnt_iff) then show False by (metis DiffI Int_iff \affine T\ cT_ne c_def clo_cT clo_in_rf closest_point_in_set convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox) qed then show ?thesis using cT_ne clo_cT closest_point_in_set by blast qed have "convex (cbox (- c) c \ T)" by (simp add: affine_imp_convex assms(4) convex_Int) then show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" using cloTK clo_cT cT_ne by (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]; force) have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x using gim [THEN subsetD] that cloTK by blast then show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K \ cbox (- (b + One)) (b + One)) \ rel_frontier U" by force show "\x. x \ S \ (g \ closest_point (cbox (- c) c \ T)) x = f x" by simp (metis (mono_tags, lifting) IntI \S \ T\ cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc) qed qed corollary extend_map_affine_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "affine T" "S \ T" and aff: "aff_dim T \ DIM('b)" and "0 \ r" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ sphere a r" "\x. x \ S \ g x = f x" proof (cases "r = 0") case True with fim show ?thesis by (rule_tac K="{}" and g = "\x. a" in that) (auto) next case False with assms have "0 < r" by auto then have "aff_dim T \ aff_dim (cball a r)" by (simp add: aff aff_dim_cball) then show ?thesis using extend_map_affine_to_sphere_cofinite_gen [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf] using fim by (fastforce simp: assms False that dest: dis) qed corollary extend_map_UNIV_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "DIM('a) \ DIM('b)" and "0 \ r" and "compact S" and "continuous_on S f" and "f ` S \ sphere a r" and "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "disjnt K S" "continuous_on (- K) g" "g ` (- K) \ sphere a r" "\x. x \ S \ g x = f x" using extend_map_affine_to_sphere_cofinite [OF \compact S\ affine_UNIV subset_UNIV] assms by (metis Compl_eq_Diff_UNIV aff_dim_UNIV of_nat_le_iff) corollary extend_map_UNIV_to_sphere_no_bounded_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. C \ components(- S) \ \ bounded C" obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" using extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"] by (metis Compl_empty_eq dis subset_empty) theorem Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "(\c \ components(- S). \bounded c) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof clarify fix f :: "'a \ 'a" assume contf: "continuous_on S f" and fim: "f ` S \ sphere 0 1" obtain g where contg: "continuous_on UNIV g" and gim: "range g \ sphere 0 1" and gf: "\x. x \ S \ g x = f x" by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \compact S\ contf fim L]) auto then obtain c where c: "homotopic_with_canon (\h. True) UNIV (sphere 0 1) g (\x. c)" using contractible_UNIV nullhomotopic_from_contractible by blast then show "\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)" by (metis assms compact_imp_closed contf contg contractible_empty fim gf gim nullhomotopic_from_contractible nullhomotopic_into_sphere_extension) qed next assume R [rule_format]: ?rhs show ?lhs unfolding components_def proof clarify fix a assume "a \ S" and a: "bounded (connected_component_set (- S) a)" have "\x\S. norm (x - a) \ 0" using \a \ S\ by auto then have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" by (intro continuous_intros) have im: "(\x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \ sphere 0 1" by clarsimp (metis \a \ S\ eq_iff_diff_eq_0 left_inverse norm_eq_zero) show False using R cont im Borsuk_map_essential_bounded_component [OF \compact S\ \a \ S\] a by blast qed qed corollary Borsuk_separation_theorem: fixes S :: "'a::euclidean_space set" assumes "compact S" and 2: "2 \ DIM('a)" shows "connected(- S) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "S = {}") case True then show ?thesis by auto next case False then have "(\c\components (- S). \ bounded c)" by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl) then show ?thesis by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed next assume R: ?rhs then have "\c\components (- S). \ bounded c \ connected (- S)" by (metis "2" assms(1) cobounded_has_bounded_component compact_imp_bounded double_complement) with R show ?lhs by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed lemma homotopy_eqv_separation: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "S homotopy_eqv T" and "compact S" and "compact T" shows "connected(- S) \ connected(- T)" proof - consider "DIM('a) = 1" | "2 \ DIM('a)" by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq) then show ?thesis proof cases case 1 then show ?thesis using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis next case 2 with assms show ?thesis by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null) qed qed proposition Jordan_Brouwer_separation: fixes S :: "'a::euclidean_space set" and a::'a assumes hom: "S homeomorphic sphere a r" and "0 < r" shows "\ connected(- S)" proof - have "- sphere a r \ ball a r \ {}" using \0 < r\ by (simp add: Int_absorb1 subset_eq) moreover have eq: "- sphere a r - ball a r = - cball a r" by auto have "- cball a r \ {}" proof - have "frontier (cball a r) \ {}" using \0 < r\ by auto then show ?thesis by (metis frontier_complement frontier_empty) qed with eq have "- sphere a r - ball a r \ {}" by auto moreover have "connected (- S) = connected (- sphere a r)" proof (rule homotopy_eqv_separation) show "S homotopy_eqv sphere a r" using hom homeomorphic_imp_homotopy_eqv by blast show "compact (sphere a r)" by simp then show " compact S" using hom homeomorphic_compactness by blast qed ultimately show ?thesis using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \0 < r\) qed proposition Jordan_Brouwer_frontier: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and T: "T \ components(- S)" and 2: "2 \ DIM('a)" shows "frontier T = S" proof (cases r rule: linorder_cases) assume "r < 0" with S T show ?thesis by auto next assume "r = 0" with S T card_eq_SucD obtain b where "S = {b}" by (auto simp: homeomorphic_finite [of "{a}" S]) have "components (- {b}) = { -{b}}" using T \S = {b}\ by (auto simp: components_eq_sing_iff connected_punctured_universe 2) with T show ?thesis by (metis \S = {b}\ cball_trivial frontier_cball frontier_complement singletonD sphere_trivial) next assume "r > 0" have "compact S" using homeomorphic_compactness compact_sphere S by blast show ?thesis proof (rule frontier_minimal_separating_closed) show "closed S" using \compact S\ compact_eq_bounded_closed by blast show "\ connected (- S)" using Jordan_Brouwer_separation S \0 < r\ by blast obtain f g where hom: "homeomorphism S (sphere a r) f g" using S by (auto simp: homeomorphic_def) show "connected (- T)" if "closed T" "T \ S" for T proof - have "f ` T \ sphere a r" using \T \ S\ hom homeomorphism_image1 by blast moreover have "f ` T \ sphere a r" using \T \ S\ hom by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE) ultimately have "f ` T \ sphere a r" by blast then have "connected (- f ` T)" by (rule psubset_sphere_Compl_connected [OF _ \0 < r\ 2]) moreover have "compact T" using \compact S\ bounded_subset compact_eq_bounded_closed that by blast moreover then have "compact (f ` T)" by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \T \ S\) moreover have "T homotopy_eqv f ` T" by (meson \f ` T \ sphere a r\ dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \T \ S\) ultimately show ?thesis using homotopy_eqv_separation [of T "f`T"] by blast qed qed (rule T) qed proposition Jordan_Brouwer_nonseparation: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and "T \ S" and 2: "2 \ DIM('a)" shows "connected(- T)" proof - have *: "connected(C \ (S - T))" if "C \ components(- S)" for C proof (rule connected_intermediate_closure) show "connected C" using in_components_connected that by auto have "S = frontier C" using "2" Jordan_Brouwer_frontier S that by blast with closure_subset show "C \ (S - T) \ closure C" by (auto simp: frontier_def) qed auto have "components(- S) \ {}" by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere components_eq_empty homeomorphic_compactness) then have "- T = (\C \ components(- S). C \ (S - T))" using Union_components [of "-S"] \T \ S\ by auto moreover have "connected ..." using \T \ S\ by (intro connected_Union) (auto simp: *) ultimately show ?thesis by simp qed subsection\ Invariance of domain and corollaries\ lemma invariance_of_domain_ball: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on (cball a r) f" and "0 < r" and inj: "inj_on f (cball a r)" shows "open(f ` ball a r)" proof (cases "DIM('a) = 1") case True obtain h::"'a\real" and k where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" "\x. norm(h x) = norm x" "\x. norm(k x) = norm x" and kh: "\x. k(h x) = x" and "\x. h(k x) = x" proof (rule isomorphisms_UNIV_UNIV) show "DIM('a) = DIM(real)" using True by force qed (metis UNIV_I UNIV_eq_I imageI) have cont: "continuous_on S h" "continuous_on T k" for S T by (simp_all add: \linear h\ \linear k\ linear_continuous_on linear_linear) have "continuous_on (h ` cball a r) (h \ f \ k)" by (intro continuous_on_compose cont continuous_on_subset [OF contf]) (auto simp: kh) moreover have "is_interval (h ` cball a r)" by (simp add: is_interval_connected_1 \linear h\ linear_continuous_on linear_linear connected_continuous_image) moreover have "inj_on (h \ f \ k) (h ` cball a r)" using inj by (simp add: inj_on_def) (metis \\x. k (h x) = x\) ultimately have *: "\T. \open T; T \ h ` cball a r\ \ open ((h \ f \ k) ` T)" using injective_eq_1d_open_map_UNIV by blast have "open ((h \ f \ k) ` (h ` ball a r))" by (rule *) (auto simp: \linear h\ \range h = UNIV\ open_surjective_linear_image) then have "open ((h \ f) ` ball a r)" by (simp add: image_comp \\x. k (h x) = x\ cong: image_cong) then show ?thesis unfolding image_comp [symmetric] by (metis open_bijective_linear_image_eq \linear h\ kh \range h = UNIV\ bijI inj_on_def) next case False then have 2: "DIM('a) \ 2" by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq) have fimsub: "f ` ball a r \ - f ` sphere a r" using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl) have hom: "f ` sphere a r homeomorphic sphere a r" by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball) then have nconn: "\ connected (- f ` sphere a r)" by (rule Jordan_Brouwer_separation) (auto simp: \0 < r\) have "bounded (f ` sphere a r)" by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball) then obtain C where C: "C \ components (- f ` sphere a r)" and "bounded C" using cobounded_has_bounded_component [OF _ nconn] "2" by auto moreover have "f ` (ball a r) = C" proof have "C \ {}" by (rule in_components_nonempty [OF C]) show "C \ f ` ball a r" proof (rule ccontr) assume nonsub: "\ C \ f ` ball a r" have "- f ` cball a r \ C" proof (rule components_maximal [OF C]) have "f ` cball a r homeomorphic cball a r" using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast then show "connected (- f ` cball a r)" by (auto intro: connected_complement_homeomorphic_convex_compact 2) show "- f ` cball a r \ - f ` sphere a r" by auto then show "C \ - f ` cball a r \ {}" using \C \ {}\ in_components_subset [OF C] nonsub using image_iff by fastforce qed then have "bounded (- f ` cball a r)" using bounded_subset \bounded C\ by auto then have "\ bounded (f ` cball a r)" using cobounded_imp_unbounded by blast then show "False" using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast qed with \C \ {}\ have "C \ f ` ball a r \ {}" by (simp add: inf.absorb_iff1) then show "f ` ball a r \ C" by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset) qed moreover have "open (- f ` sphere a r)" using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast ultimately show ?thesis using open_components by blast qed text\Proved by L. E. J. Brouwer (1912)\ theorem invariance_of_domain: fixes f :: "'a \ 'a::euclidean_space" assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)" unfolding open_subopen [of "f`S"] proof clarify fix a assume "a \ S" obtain \ where "\ > 0" and \: "cball a \ \ S" using \open S\ \a \ S\ open_contains_cball_eq by blast show "\T. open T \ f a \ T \ T \ f ` S" proof (intro exI conjI) show "open (f ` (ball a \))" by (meson \ \0 < \\ assms continuous_on_subset inj_on_subset invariance_of_domain_ball) show "f a \ f ` ball a \" by (simp add: \0 < \\) show "f ` ball a \ \ f ` S" using \ ball_subset_cball by blast qed qed lemma inv_of_domain_ss0: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - have "U \ S" using ope openin_imp_subset by blast have "(UNIV::'b set) homeomorphic S" by (simp add: \subspace S\ dimS homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" using homeomorphic_def by blast have homkh: "homeomorphism S (k ` S) k h" using homhk homeomorphism_image2 homeomorphism_sym by fastforce have "open ((k \ f \ h) ` k ` U)" proof (rule invariance_of_domain) show "continuous_on (k ` U) (k \ f \ h)" proof (intro continuous_intros) show "continuous_on (k ` U) h" by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest) have "h ` k ` U \ U" by (metis \U \ S\ dual_order.eq_iff homeomorphism_image2 homeomorphism_of_subsets homkh) then show "continuous_on (h ` k ` U) f" by (rule continuous_on_subset [OF contf]) have "f ` h ` k ` U \ S" using \h ` k ` U \ U\ fim by blast then show "continuous_on (f ` h ` k ` U) k" by (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) qed have ope_iff: "\T. open T \ openin (top_of_set (k ` S)) T" using homhk homeomorphism_image2 open_openin by fastforce show "open (k ` U)" by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) show "inj_on (k \ f \ h) (k ` U)" by (smt (verit) comp_apply inj_on_def \U \ S\ fim homeomorphism_apply2 homhk image_iff injf subsetD) qed moreover have eq: "f ` U = h ` (k \ f \ h \ k) ` U" unfolding image_comp [symmetric] using \U \ S\ fim by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff) ultimately show ?thesis by (metis (no_types, opaque_lifting) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV) qed lemma inv_of_domain_ss1: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - define S' where "S' \ {y. \x \ S. orthogonal x y}" have "subspace S'" by (simp add: S'_def subspace_orthogonal_to_vectors) define g where "g \ \z::'a*'a. ((f \ fst)z, snd z)" have "openin (top_of_set (S \ S')) (g ` (U \ S'))" proof (rule inv_of_domain_ss0) show "continuous_on (U \ S') g" unfolding g_def by (auto intro!: continuous_intros continuous_on_compose2 [OF contf continuous_on_fst]) show "g ` (U \ S') \ S \ S'" using fim by (auto simp: g_def) show "inj_on g (U \ S')" using injf by (auto simp: g_def inj_on_def) show "subspace (S \ S')" by (simp add: \subspace S'\ \subspace S\ subspace_Times) show "openin (top_of_set (S \ S')) (U \ S')" by (simp add: openin_Times [OF ope]) have "dim (S \ S') = dim S + dim S'" by (simp add: \subspace S'\ \subspace S\ dim_Times) also have "... = DIM('a)" using dim_subspace_orthogonal_to_vectors [OF \subspace S\ subspace_UNIV] by (simp add: add.commute S'_def) finally show "dim (S \ S') = DIM('a)" . qed moreover have "g ` (U \ S') = f ` U \ S'" by (auto simp: g_def image_iff) moreover have "0 \ S'" using \subspace S'\ subspace_affine by blast ultimately show ?thesis by (auto simp: openin_Times_eq) qed corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and VU: "dim V \ dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof - obtain V' where "subspace V'" "V' \ U" "dim V' = dim V" using choose_subspace_of_subspace [OF VU] by (metis span_eq_iff \subspace U\) then have "V homeomorphic V'" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V V' h k" using homeomorphic_def by blast have eq: "f ` S = k ` (h \ f) ` S" proof - have "k ` h ` f ` S = f ` S" by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl) then show ?thesis by (simp add: image_comp) qed show ?thesis unfolding eq proof (rule homeomorphism_imp_open_map) show homkh: "homeomorphism V' V k h" by (simp add: homeomorphism_symD homhk) have hfV': "(h \ f) ` S \ V'" using fim homeomorphism_image1 homhk by fastforce moreover have "openin (top_of_set U) ((h \ f) ` S)" proof (rule inv_of_domain_ss1) show "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) show "inj_on (h \ f) S" by (smt (verit, ccfv_SIG) comp_apply fim inj_on_def homeomorphism_apply2 [OF homkh] image_subset_iff injf) show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) ultimately show "openin (top_of_set V') ((h \ f) ` S)" using openin_subset_trans \V' \ U\ by force qed qed corollary invariance_of_dimension_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "dim U \ dim V" proof - have "False" if "dim V < dim U" proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" using homeomorphic_def by blast have "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" by (smt (verit, best) comp_apply inj_on_def fim homeomorphism_apply1 homhk image_subset_iff injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" using fim homeomorphism_image1 homhk by fastforce then have "dim ((h \ f) ` S) \ dim T" by (rule dim_subset) also have "dim ((h \ f) ` S) = dim U" using \S \ {}\ \subspace U\ by (blast intro: dim_openin ope_hf) finally show False using \dim V < dim U\ \dim T = dim V\ by simp qed then show ?thesis using not_less by blast qed corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" "aff_dim V \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case True then show ?thesis by auto next case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "dim ((+) (- b) ` V) \ dim ((+) (- a) ` U)" by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed then show ?thesis by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) qed corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "aff_dim U \ aff_dim V" proof - obtain a b where "a \ S" "a \ U" "b \ V" using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed (use \S \ {}\ in auto) then show ?thesis by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) qed corollary invariance_of_dimension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and "open S" and injf: "inj_on f S" and "S \ {}" shows "DIM('a) \ DIM('b)" using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms by auto corollary continuous_injective_image_subspace_dim_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace S" "subspace T" and contf: "continuous_on S f" and fim: "f ` S \ T" and injf: "inj_on f S" shows "dim S \ dim T" using invariance_of_dimension_subspaces [of S S _ f] assms by (auto simp: subspace_affine) lemma invariance_of_dimension_convex_domain: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "convex S" and contf: "continuous_on S f" and fim: "f ` S \ affine hull T" and injf: "inj_on f S" shows "aff_dim S \ aff_dim T" proof (cases "S = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" using fim rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast show "rel_interior S \ {}" by (simp add: False \convex S\ rel_interior_eq_empty) qed auto then show ?thesis by simp qed lemma homeomorphic_convex_sets_le: assumes "convex S" "S homeomorphic T" shows "aff_dim S \ aff_dim T" proof - obtain h k where homhk: "homeomorphism S T h k" using homeomorphic_def assms by blast show ?thesis proof (rule invariance_of_dimension_convex_domain [OF \convex S\]) show "continuous_on S h" using homeomorphism_def homhk by blast show "h ` S \ affine hull T" by (metis homeomorphism_def homhk hull_subset) show "inj_on h S" by (meson homeomorphism_apply1 homhk inj_on_inverseI) qed qed lemma homeomorphic_convex_sets: assumes "convex S" "convex T" "S homeomorphic T" shows "aff_dim S = aff_dim T" by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) lemma homeomorphic_convex_compact_sets_eq: assumes "convex S" "compact S" "convex T" "compact T" shows "S homeomorphic T \ aff_dim S = aff_dim T" by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) lemma invariance_of_domain_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" shows "open (f ` T)" proof - have "DIM(real) \ DIM('a)" by simp then show ?thesis using invariance_of_domain_gen assms continuous_on_subset subset_inj_on by metis qed lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" and gf: "\x. x \ S \ g(f x) = x" shows "continuous_on (f ` S) g" proof (clarsimp simp add: continuous_openin_preimage_eq) fix T :: "'a set" assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) have "open (f ` S)" by (rule invariance_of_domain_gen) (use assms inj_on_inverseI in auto) moreover have "open (f ` (S \ T))" using assms by (metis \open T\ continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq) ultimately show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" unfolding eq by (auto intro: open_openin_trans) qed lemma invariance_of_domain_homeomorphism: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" obtains g where "homeomorphism S (f ` S) f g" proof show "homeomorphism S (f ` S) f (inv_into S f)" by (simp add: assms continuous_on_inverse_open homeomorphism_def) qed corollary invariance_of_domain_homeomorphic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" shows "S homeomorphic (f ` S)" using invariance_of_domain_homeomorphism [OF assms] by (meson homeomorphic_def) lemma continuous_image_subset_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "f ` (interior S) \ interior(f ` S)" proof - have "open (f ` interior S)" using assms by (intro invariance_of_domain_gen) (auto simp: subset_inj_on interior_subset continuous_on_subset) then show ?thesis by (simp add: image_mono interior_maximal interior_subset) qed lemma homeomorphic_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)" shows "(interior S) homeomorphic (interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp have gim: "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp show "homeomorphism (interior S) (interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show "\x. x \ interior S \ g (f x) = x" by (meson \\x\S. f x \ T \ g (f x) = x\ subsetD interior_subset) have "interior T \ f ` interior S" proof fix x assume "x \ interior T" then have "g x \ interior S" using gim by blast then show "x \ f ` interior S" by (metis T \x \ interior T\ image_iff interior_subset subsetCE) qed then show "f ` interior S = interior T" using fim by blast show "continuous_on (interior S) f" by (metis interior_subset continuous_on_subset contf) show "\y. y \ interior T \ f (g y) = y" by (meson T subsetD interior_subset) have "interior S \ g ` interior T" proof fix x assume "x \ interior S" then have "f x \ interior T" using fim by blast then show "x \ g ` interior T" by (metis S \x \ interior S\ image_iff interior_subset subsetCE) qed then show "g ` interior T = interior S" using gim by blast show "continuous_on (interior T) g" by (metis interior_subset continuous_on_subset contg) qed qed lemma homeomorphic_open_imp_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done proposition homeomorphic_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") case False then have "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior) done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) qed (use assms in auto) lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)" shows "(frontier S) homeomorphic (frontier T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp then have fim: "f ` frontier S \ frontier T" unfolding frontier_def using continuous_image_subset_interior assms(2) assms(3) S by auto have "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp then have gim: "g ` frontier T \ frontier S" unfolding frontier_def using continuous_image_subset_interior T assms(2) assms(3) by auto show "homeomorphism (frontier S) (frontier T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ frontier S \ g (f x) = x" by (simp add: S assms(2) frontier_def) show fg: "\y. y \ frontier T \ f (g y) = y" by (simp add: T assms(3) frontier_def) have "frontier T \ f ` frontier S" proof fix x assume "x \ frontier T" then have "g x \ frontier S" using gim by blast then show "x \ f ` frontier S" by (metis fg \x \ frontier T\ imageI) qed then show "f ` frontier S = frontier T" using fim by blast show "continuous_on (frontier S) f" by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def) have "frontier S \ g ` frontier T" proof fix x assume "x \ frontier S" then have "f x \ frontier T" using fim by blast then show "x \ g ` frontier T" by (metis gf \x \ frontier S\ imageI) qed then show "g ` frontier T = frontier S" using gim by blast show "continuous_on (frontier T) g" by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def) qed qed lemma homeomorphic_frontiers: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" "interior S = {} \ interior T = {}" shows "(frontier S) homeomorphic (frontier T)" proof (cases "interior T = {}") case True then show ?thesis by (metis Diff_empty assms closure_eq frontier_def) next case False then have "DIM('a) = DIM('b)" using assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast then show ?thesis using assms homeomorphic_frontiers_same_dimension by blast qed lemma continuous_image_subset_rel_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and TS: "aff_dim T \ aff_dim S" shows "f ` (rel_interior S) \ rel_interior(f ` S)" proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "aff_dim (affine hull f ` S) \ aff_dim (affine hull S)" by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans) show "f ` rel_interior S \ affine hull f ` S" by (meson \f ` rel_interior S \ f ` S\ hull_subset order_trans) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast qed auto qed lemma homeomorphic_rel_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(rel_interior S) homeomorphic (rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (rel_interior S) (rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast have "rel_interior T \ f ` rel_interior S" proof fix x assume "x \ rel_interior T" then have "g x \ rel_interior S" using gim by blast then show "x \ f ` rel_interior S" by (metis fg \x \ rel_interior T\ imageI) qed moreover have "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) ultimately show "f ` rel_interior S = rel_interior T" by blast show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast have "rel_interior S \ g ` rel_interior T" proof fix x assume "x \ rel_interior S" then have "f x \ rel_interior T" using fim by blast then show "x \ g ` rel_interior T" by (metis gf \x \ rel_interior S\ imageI) qed then show "g ` rel_interior T = rel_interior S" using gim by blast show "continuous_on (rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_aff_dim_le: fixes S :: "'a::euclidean_space set" assumes "S homeomorphic T" "rel_interior S \ {}" shows "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof - obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto show ?thesis proof (rule invariance_of_dimension_affine_sets) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" by (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) show "inj_on f (rel_interior S)" by (metis S inj_on_inverseI inj_on_subset rel_interior_subset) qed (simp_all add: openin_rel_interior assms) qed lemma homeomorphic_rel_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(rel_interior S) homeomorphic (rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" using False assms homeomorphic_aff_dim_le by blast moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" using False assms(1) homeomorphic_aff_dim_le homeomorphic_sym by auto ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_rel_boundaries_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ S - rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ T - rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast show "f ` (S - rel_interior S) = T - rel_interior T" using S fST fim gim by auto show "continuous_on (S - rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "g ` (T - rel_interior T) = S - rel_interior S" using T gTS gim fim by auto show "continuous_on (T - rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_boundaries: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms by (auto simp: homeomorphic_minimal) have "aff_dim (affine hull S) \ aff_dim (affine hull T)" using False assms homeomorphic_aff_dim_le by blast moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" by (meson False assms(1) homeomorphic_aff_dim_le homeomorphic_sym) ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_boundaries_same_dimension [OF \S homeomorphic T\]) qed proposition uniformly_continuous_homeomorphism_UNIV_trivial: fixes f :: "'a::euclidean_space \ 'a" assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g" shows "S = UNIV" proof (cases "S = {}") case True then show ?thesis by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI) next case False have "inj g" by (metis UNIV_I hom homeomorphism_apply2 injI) then have "open (g ` UNIV)" by (blast intro: invariance_of_domain hom homeomorphism_cont2) then have "open S" using hom homeomorphism_image2 by blast moreover have "complete S" unfolding complete_def proof clarify fix \ assume \: "\n. \ n \ S" and "Cauchy \" have "Cauchy (f o \)" - using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf by blast + using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf + unfolding Cauchy_continuous_on_def by blast then obtain l where "(f \ \) \ l" by (auto simp: convergent_eq_Cauchy [symmetric]) show "\l\S. \ \ l" proof show "g l \ S" using hom homeomorphism_image2 by blast have "(g \ (f \ \)) \ g l" by (meson UNIV_I \(f \ \) \ l\ continuous_on_sequentially hom homeomorphism_cont2) then show "\ \ g l" proof - have "\n. \ n = (g \ (f \ \)) n" by (metis (no_types) \ comp_eq_dest_lhs hom homeomorphism_apply1) then show ?thesis by (metis (no_types) LIMSEQ_iff \(g \ (f \ \)) \ g l\) qed qed qed then have "closed S" by (simp add: complete_eq_closed) ultimately show ?thesis using clopen [of S] False by simp qed subsection\Formulation of loop homotopy in terms of maps out of type complex\ lemma homotopic_circlemaps_imp_homotopic_loops: assumes "homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "homotopic_loops S (f \ exp \ (\t. 2 * of_real pi * of_real t * \)) (g \ exp \ (\t. 2 * of_real pi * of_real t * \))" proof - have "homotopic_with_canon (\f. True) {z. cmod z = 1} S f g" using assms by (auto simp: sphere_def) moreover have "continuous_on {0..1} (exp \ (\t. 2 * of_real pi * of_real t * \))" by (intro continuous_intros) moreover have "(exp \ (\t. 2 * of_real pi * of_real t * \)) ` {0..1} \ {z. cmod z = 1}" by (auto simp: norm_mult) ultimately show ?thesis apply (simp add: homotopic_loops_def comp_assoc) apply (rule homotopic_with_compose_continuous_right) apply (auto simp: pathstart_def pathfinish_def) done qed lemma homotopic_loops_imp_homotopic_circlemaps: assumes "homotopic_loops S p q" shows "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. (Arg2pi z / (2 * pi)))) (q \ (\z. (Arg2pi z / (2 * pi))))" proof - obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "(\x. h (0, x) = p x)" and h1: "(\x. h (1, x) = q x)" and h01: "(\t\{0..1}. h (t, 1) = h (t, 0)) " using assms by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def) define j where "j \ \z. if 0 \ Im (snd z) then h (fst z, Arg2pi (snd z) / (2 * pi)) else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))" have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \ Arg2pi y = 0 \ Arg2pi (cnj y) = 0" if "cmod y = 1" for y using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps) show ?thesis proof (simp add: homotopic_with; intro conjI ballI exI) show "continuous_on ({0..1} \ sphere 0 1) (\w. h (fst w, Arg2pi (snd w) / (2 * pi)))" proof (rule continuous_on_eq) show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \ {0..1} \ sphere 0 1" for x using Arg2pi_eq that h01 by (force simp: j_def) have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" by auto have \
: "Arg2pi z \ 2 * pi" for z by (simp add: Arg2pi order_le_less) have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) by (auto simp: Arg2pi \
) have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) by (auto simp: Arg2pi \
) show "continuous_on ({0..1} \ sphere 0 1) j" apply (simp add: j_def) apply (subst eq) apply (rule continuous_on_cases_local) using Arg2pi_eq h01 by (force simp add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2)+ qed have "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ h ` ({0..1} \ {0..1})" by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le) also have "... \ S" using him by blast finally show "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ S" . qed (auto simp: h0 h1) qed lemma simply_connected_homotopic_loops: "simply_connected S \ (\p q. homotopic_loops S p p \ homotopic_loops S q q \ homotopic_loops S p q)" unfolding simply_connected_def using homotopic_loops_refl by metis lemma simply_connected_eq_homotopic_circlemaps1: fixes f :: "complex \ 'a::topological_space" and g :: "complex \ 'a" assumes S: "simply_connected S" and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \ S" and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" proof - let ?h = "(\t. complex_of_real (2 * pi * t) * \)" have "homotopic_loops S (f \ exp \ ?h) (f \ exp \ ?h) \ homotopic_loops S (g \ exp \ ?h) (g \ exp \ ?h)" by (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim) then have "homotopic_loops S (f \ exp \ ?h) (g \ exp \ ?h)" using S simply_connected_homotopic_loops by blast then show ?thesis apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) done qed lemma simply_connected_eq_homotopic_circlemaps2a: fixes h :: "complex \ 'a::topological_space" assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \ S" and hom: "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" by (metis conth continuous_on_const him hom image_subset_iff) lemma simply_connected_eq_homotopic_circlemaps2b: fixes S :: "'a::real_normed_vector set" assumes "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "path_connected S" proof (clarsimp simp add: path_connected_eq_homotopic_points) fix a b assume "a \ S" "b \ S" then show "homotopic_loops S (linepath a a) (linepath b b)" using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\x. a" "\x. b"]] by (auto simp: o_def linepath_def) qed lemma simply_connected_eq_homotopic_circlemaps3: fixes h :: "complex \ 'a::real_normed_vector" assumes "path_connected S" and hom: "\f::complex \ 'a. \continuous_on (sphere 0 1) f; f `(sphere 0 1) \ S\ \ \a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)" shows "simply_connected S" proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms) fix p assume p: "path p" "path_image p \ S" "pathfinish p = pathstart p" then have "homotopic_loops S p p" by (simp add: homotopic_loops_refl) then obtain a where homp: "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. Arg2pi z / (2 * pi))) (\x. a)" by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom) show "\a. a \ S \ homotopic_loops S p (linepath a a)" proof (intro exI conjI) show "a \ S" using homotopic_with_imp_subset2 [OF homp] by (metis dist_0_norm image_subset_iff mem_sphere norm_one) have teq: "\t. \0 \ t; t \ 1\ \ t = Arg2pi (exp (2 * of_real pi * of_real t * \)) / (2 * pi) \ t=1 \ Arg2pi (exp (2 * of_real pi * of_real t * \)) = 0" using Arg2pi_of_real [of 1] by (force simp: Arg2pi_exp) have "homotopic_loops S p (p \ (\z. Arg2pi z / (2 * pi)) \ exp \ (\t. 2 * complex_of_real pi * complex_of_real t * \))" using p teq by (fastforce simp: pathfinish_def pathstart_def intro: homotopic_loops_eq [OF p]) then show "homotopic_loops S p (linepath a a)" by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]]) qed qed proposition simply_connected_eq_homotopic_circlemaps: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ (\f g::complex \ 'a. continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" by (metis simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) proposition simply_connected_eq_contractible_circlemap: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\f::complex \ 'a. continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" by (metis simply_connected_eq_homotopic_circlemaps simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps3 simply_connected_imp_path_connected) corollary homotopy_eqv_simple_connectedness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) subsection\Homeomorphism of simple closed curves to circles\ proposition homeomorphic_simple_path_image_circle: fixes a :: complex and \ :: "real \ 'a::t2_space" assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" shows "(path_image \) homeomorphic sphere a r" proof - have "homotopic_loops (path_image \) \ \" by (simp add: assms homotopic_loops_refl simple_path_imp_path) then have hom: "homotopic_with_canon (\h. True) (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) (\ \ (\z. Arg2pi z / (2*pi)))" by (rule homotopic_loops_imp_homotopic_circlemaps) have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) g" proof (rule homeomorphism_compact) show "continuous_on (sphere 0 1) (\ \ (\z. Arg2pi z / (2*pi)))" using hom homotopic_with_imp_continuous by blast show "inj_on (\ \ (\z. Arg2pi z / (2*pi))) (sphere 0 1)" proof fix x y assume xy: "x \ sphere 0 1" "y \ sphere 0 1" and eq: "(\ \ (\z. Arg2pi z / (2*pi))) x = (\ \ (\z. Arg2pi z / (2*pi))) y" then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))" proof - have "(Arg2pi x / (2*pi)) \ {0..1}" "(Arg2pi y / (2*pi)) \ {0..1}" using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+ with eq show ?thesis using \simple_path \\ Arg2pi_lt_2pi unfolding simple_path_def loop_free_def o_def by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) qed with xy show "x = y" by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) qed have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg2pi z / (2*pi)) = \ x" by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) moreover have "\z\sphere 0 1. \ x = \ (Arg2pi z / (2*pi))" if "0 \ x" "x \ 1" for x proof (cases "x=1") case True with Arg2pi_of_real [of 1] loop show ?thesis by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \0 \ x\) next case False then have *: "(Arg2pi (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" using that by (auto simp: Arg2pi_exp field_split_simps) show ?thesis by (rule_tac x="exp(\ * of_real(2*pi*x))" in bexI) (auto simp: *) qed ultimately show "(\ \ (\z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \" by (auto simp: path_image_def image_iff) qed auto then have "path_image \ homeomorphic sphere (0::complex) 1" using homeomorphic_def homeomorphic_sym by blast also have "... homeomorphic sphere a r" by (simp add: assms homeomorphic_spheres) finally show ?thesis . qed lemma homeomorphic_simple_path_images: fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" shows "(path_image \1) homeomorphic (path_image \2)" by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) subsection\Dimension-based conditions for various homeomorphisms\ lemma homeomorphic_subspaces_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "subspace S" "subspace T" shows "S homeomorphic T \ dim S = dim T" proof assume "S homeomorphic T" then obtain f g where hom: "homeomorphism S T f g" using homeomorphic_def by blast show "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le) show "dim T \ dim S" by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le) qed next assume "dim S = dim T" then show "S homeomorphic T" by (simp add: assms homeomorphic_subspaces) qed lemma homeomorphic_affine_sets_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "affine S" "affine T" shows "S homeomorphic T \ aff_dim S = aff_dim T" proof (cases "S = {} \ T = {}") case True then show ?thesis using assms homeomorphic_affine_sets by force next case False then obtain a b where "a \ S" "b \ T" by blast then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ then show ?thesis by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets) qed lemma homeomorphic_hyperplanes_eq: fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('a) = DIM('b))" proof - have "DIM('a) - Suc 0 = DIM('b) - Suc 0 \ DIM('a) = DIM('b)" by (metis DIM_positive Suc_pred) then show ?thesis by (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) qed lemma homeomorphic_UNIV_UNIV: shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \ DIM('a::euclidean_space) = DIM('b::euclidean_space)" by (simp add: homeomorphic_subspaces_eq) lemma simply_connected_sphere_gen: assumes "convex S" "bounded S" and 3: "3 \ aff_dim S" shows "simply_connected(rel_frontier S)" proof - have pa: "path_connected (rel_frontier S)" using assms by (simp add: path_connected_sphere_gen) show ?thesis proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa) fix f assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \ rel_frontier S" have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)" by simp have "convex (cball (0::complex) 1)" by (rule convex_cball) then obtain c where "homotopic_with_canon (\z. True) (sphere (0::complex) 1) (rel_frontier S) f (\x. c)" apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \convex S\ \bounded S\, where f=f]) using f 3 by (auto simp: aff_dim_cball) then show "\a. homotopic_with_canon (\h. True) (sphere 0 1) (rel_frontier S) f (\x. a)" by blast qed qed subsection\more invariance of domain\(*FIX ME title? *) proposition invariance_of_domain_sphere_affine_set_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and U: "bounded U" "convex U" and "affine T" and affTU: "aff_dim T < aff_dim U" and ope: "openin (top_of_set (rel_frontier U)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "rel_frontier U = {}") case True then show ?thesis using ope openin_subset by force next case False obtain b c where b: "b \ rel_frontier U" and c: "c \ rel_frontier U" and "b \ c" using \bounded U\ rel_frontier_not_sing [of U] subset_singletonD False by fastforce obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1" proof (rule choose_affine_subset [OF affine_UNIV]) show "- 1 \ aff_dim U - 1" by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le) show "aff_dim U - 1 \ aff_dim (UNIV::'a set)" by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq) qed auto have SU: "S \ rel_frontier U" using ope openin_imp_subset by auto have homb: "rel_frontier U - {b} homeomorphic V" and homc: "rel_frontier U - {c} homeomorphic V" using homeomorphic_punctured_sphere_affine_gen [of U _ V] by (simp_all add: \affine V\ affV U b c) then obtain g h j k where gh: "homeomorphism (rel_frontier U - {b}) V g h" and jk: "homeomorphism (rel_frontier U - {c}) V j k" by (auto simp: homeomorphic_def) with SU have hgsub: "(h ` g ` (S - {b})) \ S" and kjsub: "(k ` j ` (S - {c})) \ S" by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) have "openin (top_of_set (rel_frontier U - {b})) (S - {b})" by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) then show "openin (top_of_set V) (g ` (S - {b}))" by (rule homeomorphism_imp_open_map [OF gh]) show "continuous_on (g ` (S - {b})) (f \ h)" proof (rule continuous_on_compose) show "continuous_on (g ` (S - {b})) h" by (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) qed (use contf continuous_on_subset hgsub in blast) show "inj_on (f \ h) (g ` (S - {b}))" by (smt (verit, del_insts) SU homeomorphism_def inj_on_def injf gh Diff_iff comp_apply imageE subset_iff) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (j ` (S - {c}))" by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl homeomorphism_imp_open_map [OF jk]) show "continuous_on (j ` (S - {c})) (f \ k)" proof (rule continuous_on_compose) show "continuous_on (j ` (S - {c})) k" by (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) qed (use contf continuous_on_subset kjsub in blast) show "inj_on (f \ k) (j ` (S - {c}))" by (smt (verit, del_insts) SU homeomorphism_def inj_on_def injf jk Diff_iff comp_apply imageE subset_iff) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) ultimately have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" by (rule openin_Un) moreover have "(f \ h) ` g ` (S - {b}) = f ` (S - {b})" proof - have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" by (metis Diff_mono SU gh homeomorphism_image2 homeomorphism_of_subsets set_eq_subset) qed then show ?thesis by (metis image_comp) qed moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" using homeomorphism_apply1 [OF jk] SU by (meson Diff_mono homeomorphism_def homeomorphism_of_subsets jk subset_refl) then show ?thesis by (metis image_comp) qed moreover have "f ` (S - {b}) \ f ` (S - {c}) = f ` (S)" using \b \ c\ by blast ultimately show ?thesis by simp qed lemma invariance_of_domain_sphere_affine_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and "r \ 0" "affine T" and affTU: "aff_dim T < DIM('a)" and ope: "openin (top_of_set (sphere a r)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "sphere a r = {}") case True then show ?thesis using ope openin_subset by force next case False show ?thesis proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \affine T\]) show "aff_dim T < aff_dim (cball a r)" by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty) show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed lemma no_embedding_sphere_lowdim: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0" shows "DIM('a) \ DIM('b)" proof - have "False" if "DIM('a) > DIM('b)" proof - have "compact (f ` sphere a r)" using compact_continuous_image by (simp add: compact_continuous_image contf) then have "\ open (f ` sphere a r)" using compact_open by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty) then show False using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \r > 0\ by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that) qed then show ?thesis using not_less by blast qed lemma simply_connected_sphere: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by simp next case equal then show ?thesis by (auto simp: convex_imp_simply_connected) next case greater then show ?thesis using simply_connected_sphere_gen [of "cball a r"] assms by (simp add: aff_dim_cball) qed lemma simply_connected_sphere_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(sphere a r) \ 3 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True have "simply_connected (sphere a r)" using True less_eq_real_def by (auto intro: convex_imp_simply_connected) with True show ?thesis by auto next case False show ?thesis proof assume L: ?lhs have "False" if "DIM('a) = 1 \ DIM('a) = 2" using that proof assume "DIM('a) = 1" with L show False using connected_sphere_eq simply_connected_imp_connected by (metis False Suc_1 not_less_eq_eq order_refl) next assume "DIM('a) = 2" then have "sphere a r homeomorphic sphere (0::complex) 1" by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one) then have "simply_connected(sphere (0::complex) 1)" using L homeomorphic_simply_connected_eq by blast then obtain a::complex where "homotopic_with_canon (\h. True) (sphere 0 1) (sphere 0 1) id (\x. a)" by (metis continuous_on_id' id_apply image_id subset_refl simply_connected_eq_contractible_circlemap) then show False using contractible_sphere contractible_def not_one_le_zero by blast qed with False show ?rhs by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq numeral_3_eq_3 order_antisym_conv) next assume ?rhs with False show ?lhs by (simp add: simply_connected_sphere) qed qed lemma simply_connected_punctured_universe_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(- {a}) \ 3 \ DIM('a)" proof - have [simp]: "a \ rel_interior (cball a 1)" by (simp add: rel_interior_nonempty_interior) have [simp]: "affine hull cball a 1 - {a} = -{a}" by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one) have "sphere a 1 homotopy_eqv - {a}" using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] by auto then have "simply_connected(- {a}) \ simply_connected(sphere a 1)" using homotopy_eqv_simple_connectedness by blast also have "... \ 3 \ DIM('a)" by (simp add: simply_connected_sphere_eq) finally show ?thesis . qed lemma not_simply_connected_circle: fixes a :: complex shows "0 < r \ \ simply_connected(sphere a r)" by (simp add: simply_connected_sphere_eq) proposition simply_connected_punctured_convex: fixes a :: "'a::euclidean_space" assumes "convex S" and 3: "3 \ aff_dim S" shows "simply_connected(S - {a})" proof (cases "a \ rel_interior S") case True then obtain e where "a \ S" "0 < e" and e: "cball a e \ affine hull S \ S" by (auto simp: rel_interior_cball) have con: "convex (cball a e \ affine hull S)" by (simp add: convex_Int) have bo: "bounded (cball a e \ affine hull S)" by (simp add: bounded_Int) have "affine hull S \ interior (cball a e) \ {}" using \0 < e\ \a \ S\ hull_subset by fastforce then have "3 \ aff_dim (affine hull S \ cball a e)" by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull]) also have "... = aff_dim (cball a e \ affine hull S)" by (simp add: Int_commute) finally have "3 \ aff_dim (cball a e \ affine hull S)" . moreover have "rel_frontier (cball a e \ affine hull S) homotopy_eqv S - {a}" proof (rule homotopy_eqv_rel_frontier_punctured_convex) show "a \ rel_interior (cball a e \ affine hull S)" by (meson IntI Int_mono \a \ S\ \0 < e\ e \cball a e \ affine hull S \ S\ ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball) have "closed (cball a e \ affine hull S)" by blast then show "rel_frontier (cball a e \ affine hull S) \ S" by (metis Diff_subset closure_closed dual_order.trans e rel_frontier_def) show "S \ affine hull (cball a e \ affine hull S)" by (metis (no_types, lifting) IntI \a \ S\ \0 < e\ affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI) qed (auto simp: assms con bo) ultimately show ?thesis using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo] by blast next case False then have "rel_interior S \ S - {a}" by (simp add: False rel_interior_subset subset_Diff_insert) moreover have "S - {a} \ closure S" by (meson Diff_subset closure_subset subset_trans) ultimately show ?thesis by (metis contractible_imp_simply_connected contractible_convex_tweak_boundary_points [OF \convex S\]) qed corollary simply_connected_punctured_universe: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(- {a})" proof - have [simp]: "affine hull cball a 1 = UNIV" by (simp add: aff_dim_cball affine_hull_UNIV) have "a \ rel_interior (cball a 1)" by (simp add: rel_interior_interior) then have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})" using homotopy_eqv_rel_frontier_punctured_affine_hull homotopy_eqv_simple_connectedness by blast then show ?thesis using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV) qed subsection\The power, squaring and exponential functions as covering maps\ proposition covering_space_power_punctured_plane: assumes "0 < n" shows "covering_space (- {0}) (\z::complex. z^n) (- {0})" proof - consider "n = 1" | "2 \ n" using assms by linarith then obtain e where "0 < e" and e: "\w z. cmod(w - z) < e * cmod z \ (w^n = z^n \ w = z)" proof cases assume "n = 1" then show ?thesis by (rule_tac e=1 in that) auto next assume "2 \ n" have eq_if_pow_eq: "w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z" and eq: "w^n = z^n" for w z proof (cases "z = 0") case True with eq assms show ?thesis by (auto simp: power_0_left) next case False then have "z \ 0" by auto have "(w/z)^n = 1" by (metis False divide_self_if eq power_divide power_one) then obtain j::nat where j: "w / z = exp (2 * of_real pi * \ * j / n)" and "j < n" using Suc_leI assms \2 \ n\ complex_roots_unity [THEN eqset_imp_iff, of n "w/z"] by force have "cmod (w/z - 1) < 2 * sin (pi / real n)" using lt assms \z \ 0\ by (simp add: field_split_simps norm_divide) then have "cmod (exp (\ * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)" by (simp add: j field_simps) then have "2 * \sin((2 * pi * j / n) / 2)\ < 2 * sin (pi / real n)" by (simp only: dist_exp_i_1) then have sin_less: "sin((pi * j / n)) < sin (pi / real n)" by (simp add: field_simps) then have "w / z = 1" proof (cases "j = 0") case True then show ?thesis by (auto simp: j) next case False then have "sin (pi / real n) \ sin((pi * j / n))" proof (cases "j / n \ 1/2") case True show ?thesis using \j \ 0 \ \j < n\ True by (intro sin_monotone_2pi_le) (auto simp: field_simps intro: order_trans [of _ 0]) next case False then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)" using \j < n\ by (simp add: algebra_simps diff_divide_distrib of_nat_diff) show ?thesis unfolding seq using \j < n\ False by (intro sin_monotone_2pi_le) (auto simp: field_simps intro: order_trans [of _ 0]) qed with sin_less show ?thesis by force qed then show ?thesis by simp qed show ?thesis proof show "0 < 2 * sin (pi / real n)" by (force simp: \2 \ n\ sin_pi_divide_n_gt_0) qed (meson eq_if_pow_eq) qed have zn1: "continuous_on (- {0}) (\z::complex. z^n)" by (rule continuous_intros)+ have zn2: "(\z::complex. z^n) ` (- {0}) = - {0}" using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n]) have zn3: "\T. z^n \ T \ open T \ 0 \ T \ (\v. \v = -{0} \ (\z. z ^ n) -` T \ (\u\v. open u \ 0 \ u) \ pairwise disjnt v \ (\u\v. Ex (homeomorphism u T (\z. z^n))))" if "z \ 0" for z::complex proof - define d where "d \ min (1/2) (e/4) * norm z" have "0 < d" by (simp add: d_def \0 < e\ \z \ 0\) have iff_x_eq_y: "x^n = y^n \ x = y" if eq: "w^n = z^n" and x: "x \ ball w d" and y: "y \ ball w d" for w x y proof - have [simp]: "norm z = norm w" using that by (simp add: assms power_eq_imp_eq_norm) show ?thesis proof (cases "w = 0") case True with \z \ 0\ assms eq show ?thesis by (auto simp: power_0_left) next case False have "cmod (x - y) < 2*d" using x y by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add) also have "... \ 2 * e / 4 * norm w" using \e > 0\ by (simp add: d_def min_mult_distrib_right) also have "... = e * (cmod w / 2)" by simp also have "... \ e * cmod y" proof (rule mult_left_mono) have "cmod (w - y) < cmod w / 2 \ cmod w / 2 \ cmod y" by (metis (no_types) dist_0_norm dist_norm norm_triangle_half_l not_le order_less_irrefl) then show "cmod w / 2 \ cmod y" using y by (simp add: dist_norm d_def min_mult_distrib_right) qed (use \e > 0\ in auto) finally have "cmod (x - y) < e * cmod y" . then show ?thesis by (rule e) qed qed then have inj: "inj_on (\w. w^n) (ball z d)" by (simp add: inj_on_def) have cont: "continuous_on (ball z d) (\w. w ^ n)" by (intro continuous_intros) have noncon: "\ (\w::complex. w^n) constant_on UNIV" by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power) have im_eq: "(\w. w^n) ` ball z' d = (\w. w^n) ` ball z d" if z': "z'^n = z^n" for z' proof - have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast have "(w \ (\w. w^n) ` ball z' d) = (w \ (\w. w^n) ` ball z d)" for w proof (cases "w=0") case True with assms show ?thesis by (simp add: image_def ball_def nz') next case False have "z' \ 0" using \z \ 0\ nz' by force have 1: "(z*x / z')^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have 2: "cmod (z - z * x / z') = cmod (z' - x)" if "x \ 0" for x proof - have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')" by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) divide_complex_def mult.right_neutral norm_mult right_diff_distrib') also have "... = cmod z' * cmod (1 - x / z')" by (simp add: nz') also have "... = cmod (z' - x)" by (simp add: \z' \ 0\ diff_divide_eq_iff norm_divide) finally show ?thesis . qed have 3: "(z'*x / z)^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have 4: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \ 0" for x proof - have "cmod (z * (1 - x * inverse z)) = cmod (z - x)" by (metis \z \ 0\ diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7)) then show ?thesis by (metis (no_types) mult.assoc divide_complex_def mult.right_neutral norm_mult nz' right_diff_distrib') qed show ?thesis by (simp add: set_eq_iff image_def ball_def) (metis 1 2 3 4 diff_zero dist_norm nz') qed then show ?thesis by blast qed have ex_ball: "\B. (\z'. B = ball z' d \ z'^n = z^n) \ x \ B" if "x \ 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w proof - have "w \ 0" by (metis assms power_eq_0_iff that(1) that(2)) have [simp]: "cmod x = cmod w" using assms power_eq_imp_eq_norm eq by blast have [simp]: "cmod (x * z / w - x) = cmod (z - w)" proof - have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)" by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right) also have "... = cmod w * cmod (z / w - 1)" by simp also have "... = cmod (z - w)" by (simp add: \w \ 0\ divide_diff_eq_iff nonzero_norm_divide) finally show ?thesis . qed show ?thesis proof (intro exI conjI) show "(z / w * x) ^ n = z ^ n" by (metis \w \ 0\ eq nonzero_eq_divide_eq power_mult_distrib) show "x \ ball (z / w * x) d" using \d > 0\ that by (simp add: ball_eq_ball_iff \z \ 0\ \w \ 0\ field_simps) (simp add: dist_norm) qed auto qed show ?thesis proof (rule exI, intro conjI) show "z ^ n \ (\w. w ^ n) ` ball z d" using \d > 0\ by simp show "open ((\w. w ^ n) ` ball z d)" by (rule invariance_of_domain [OF cont open_ball inj]) show "0 \ (\w. w ^ n) ` ball z d" using \z \ 0\ assms by (force simp: d_def) show "\v. \v = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d \ (\u\v. open u \ 0 \ u) \ disjoint v \ (\u\v. Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n)))" proof (rule exI, intro ballI conjI) show "\{ball z' d |z'. z'^n = z^n} = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d" (is "?l = ?r") proof have "\z'. cmod z' < d \ z' ^ n \ z ^ n" by (auto simp add: assms d_def power_eq_imp_eq_norm that) then show "?l \ ?r" by auto (metis im_eq image_eqI mem_ball) show "?r \ ?l" by auto (meson ex_ball) qed show "\u. u \ {ball z' d |z'. z' ^ n = z ^ n} \ 0 \ u" by (force simp add: assms d_def power_eq_imp_eq_norm that) show "disjoint {ball z' d |z'. z' ^ n = z ^ n}" proof (clarsimp simp add: pairwise_def disjnt_iff) fix \ \ x assume "\^n = z^n" "\^n = z^n" "ball \ d \ ball \ d" and "dist \ x < d" "dist \ x < d" then have "dist \ \ < d+d" using dist_triangle_less_add by blast then have "cmod (\ - \) < 2*d" by (simp add: dist_norm) also have "... \ e * cmod z" using mult_right_mono \0 < e\ that by (auto simp: d_def) finally have "cmod (\ - \) < e * cmod z" . with e have "\ = \" by (metis \\^n = z^n\ \\^n = z^n\ assms power_eq_imp_eq_norm) then show "False" using \ball \ d \ ball \ d\ by blast qed show "Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" if "u \ {ball z' d |z'. z' ^ n = z ^ n}" for u proof (rule invariance_of_domain_homeomorphism [of "u" "\z. z^n"]) show "open u" using that by auto show "continuous_on u (\z. z ^ n)" by (intro continuous_intros) show "inj_on (\z. z ^ n) u" using that by (auto simp: iff_x_eq_y inj_on_def) show "\g. homeomorphism u ((\z. z ^ n) ` u) (\z. z ^ n) g \ Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" using im_eq that by clarify metis qed auto qed auto qed qed show ?thesis using assms apply (simp add: covering_space_def zn1 zn2) apply (subst zn2 [symmetric]) apply (simp add: openin_open_eq open_Compl zn3) done qed corollary covering_space_square_punctured_plane: "covering_space (- {0}) (\z::complex. z^2) (- {0})" by (simp add: covering_space_power_punctured_plane) proposition covering_space_exp_punctured_plane: "covering_space UNIV (\z::complex. exp z) (- {0})" proof (simp add: covering_space_def, intro conjI ballI) show "continuous_on UNIV (\z::complex. exp z)" by (rule continuous_on_exp [OF continuous_on_id]) show "range exp = - {0::complex}" by auto (metis exp_Ln range_eqI) show "\T. z \ T \ openin (top_of_set (- {0})) T \ (\v. \v = exp -` T \ (\u\v. open u) \ disjoint v \ (\u\v. \q. homeomorphism u T exp q))" if "z \ - {0::complex}" for z proof - have "z \ 0" using that by auto have "ball (Ln z) 1 \ ball (Ln z) pi" using pi_ge_two by (simp add: ball_subset_ball_iff) then have inj_exp: "inj_on exp (ball (Ln z) 1)" using inj_on_exp_pi inj_on_subset by blast define twopi where "twopi \ \n. of_real (2 * of_int n * pi) * \" define \ where "\ \ range (\n. (\x. x + twopi n) ` (ball(Ln z) 1))" have exp_eq': "exp w = exp z \ (\n::int. w = z + twopi n)" for z w by (simp add: exp_eq twopi_def) show ?thesis proof (intro exI conjI) show "z \ exp ` (ball(Ln z) 1)" by (metis \z \ 0\ centre_in_ball exp_Ln rev_image_eqI zero_less_one) have "open (- {0::complex})" by blast with inj_exp show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id]) show UV: "\\ = exp -` exp ` ball (Ln z) 1" by (force simp: \_def twopi_def Complex_Transcendental.exp_eq image_iff) show "\V\\. open V" by (auto simp: \_def inj_on_def continuous_intros invariance_of_domain) have "2 \ cmod (twopi m -twopi n)" if "m \ n" for m n proof - have "1 \ abs (m - n)" using that by linarith then have "1 \ cmod (of_int m - of_int n) * 1" by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff) also have "... \ cmod (of_int m - of_int n) * of_real pi" using pi_ge_two by (intro mult_left_mono) auto also have "... \ cmod ((of_int m - of_int n) * of_real pi * \)" by (simp add: norm_mult) also have "... \ cmod (of_int m * of_real pi * \ - of_int n * of_real pi * \)" by (simp add: algebra_simps) finally have "1 \ cmod (of_int m * of_real pi * \ - of_int n * of_real pi * \)" . then have "2 * 1 \ cmod (2 * (of_int m * of_real pi * \ - of_int n * of_real pi * \))" by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral) then show ?thesis by (simp add: algebra_simps twopi_def) qed then show "disjoint \" unfolding \_def pairwise_def disjnt_iff by (smt (verit, best) add.commute add_diff_cancel_left' add_diff_eq dist_commute dist_complex_def dist_triangle imageE mem_ball) show "\u\\. \q. homeomorphism u (exp ` ball (Ln z) 1) exp q" proof fix u assume "u \ \" then obtain n where n: "u = (\x. x + twopi n) ` (ball(Ln z) 1)" by (auto simp: \_def) have "compact (cball (Ln z) 1)" by simp moreover have "continuous_on (cball (Ln z) 1) exp" by (rule continuous_on_exp [OF continuous_on_id]) moreover have "inj_on exp (cball (Ln z) 1)" using pi_ge_two inj_on_subset [OF inj_on_exp_pi [of "Ln z"]] by (simp add: subset_iff) ultimately obtain \ where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \" using homeomorphism_compact by blast have eq1: "exp ` u = exp ` ball (Ln z) 1" by (smt (verit) n exp_eq' image_cong image_image) have \exp: "\ (exp x) + twopi n = x" if "x \ u" for x proof - have "exp x = exp (x - twopi n)" using exp_eq' by auto then have "\ (exp x) = \ (exp (x - twopi n))" by simp also have "... = x - twopi n" using \x \ u\ by (auto simp: n intro: homeomorphism_apply1 [OF hom]) finally show ?thesis by simp qed have exp2n: "exp (\ (exp x) + twopi n) = exp x" if "dist (Ln z) x < 1" for x by (metis \exp exp_eq' imageI mem_ball n that) have "continuous_on (exp ` ball (Ln z) 1) \" by (meson ball_subset_cball continuous_on_subset hom homeomorphism_cont2 image_mono) then have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + twopi n)" by (intro continuous_intros) have "homeomorphism u (exp ` ball (Ln z) 1) exp ((\x. x + twopi n) \ \)" unfolding homeomorphism_def apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id]) apply (auto simp: \exp exp2n cont n) apply (force simp: image_iff homeomorphism_apply1 [OF hom])+ done then show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" by metis qed qed qed qed subsection\Hence the Borsukian results about mappings into circles\(*FIX ME title *) lemma inessential_eq_continuous_logarithm: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x)))" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane) next assume ?rhs then obtain g where contg: "continuous_on S g" and f: "\x. x \ S \ f x = exp(g x)" by metis obtain a where "homotopic_with_canon (\h. True) S (- {of_real 0}) (exp \ g) (\x. a)" proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV]) show "continuous_on (UNIV::complex set) exp" by (intro continuous_intros) show "range exp \ - {0}" by auto qed force then have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using f homotopic_with_eq by fastforce then show ?lhs .. qed corollary inessential_imp_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" assumes "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" obtains g where "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" proof - have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using assms homotopic_with_subset_right by fastforce then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma inessential_eq_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(\ * of_real(g x))))" (is "?lhs \ ?rhs") proof assume L: ?lhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(g x)" using inessential_imp_continuous_logarithm_circle by blast have "f ` S \ sphere 0 1" by (metis L homotopic_with_imp_subset1) then have "\x. x \ S \ Re (g x) = 0" using g by auto then show ?rhs by (rule_tac x="Im \ g" in exI) (auto simp: Euler g intro: contg continuous_intros) next assume ?rhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(\* of_real(g x))" by metis obtain a where "homotopic_with_canon (\h. True) S (sphere 0 1) ((exp \ (\z. \*z)) \ (of_real \ g)) (\x. a)" proof (rule nullhomotopic_through_contractible) show "continuous_on S (complex_of_real \ g)" by (intro conjI contg continuous_intros) show "(complex_of_real \ g) ` S \ \" by auto show "continuous_on \ (exp \ (*)\)" by (intro continuous_intros) show "(exp \ (*)\) ` \ \ sphere 0 1" by (auto simp: complex_is_Real_iff) qed (auto simp: convex_Reals convex_imp_contractible) moreover have "\x. x \ S \ (exp \ (*)\ \ (complex_of_real \ g)) x = f x" by (simp add: g) ultimately have "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" using homotopic_with_eq by force then show ?lhs .. qed proposition homotopic_with_sphere_times: fixes f :: "'a::real_normed_vector \ complex" assumes hom: "homotopic_with_canon (\x. True) S (sphere 0 1) f g" and conth: "continuous_on S h" and hin: "\x. x \ S \ h x \ sphere 0 1" shows "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x * h x) (\x. g x * h x)" proof - obtain k where contk: "continuous_on ({0..1::real} \ S) k" and kim: "k ` ({0..1} \ S) \ sphere 0 1" and k0: "\x. k(0, x) = f x" and k1: "\x. k(1, x) = g x" using hom by (auto simp: homotopic_with_def) show ?thesis apply (simp add: homotopic_with) apply (rule_tac x="\z. k z * (h \ snd)z" in exI) using kim hin by (fastforce simp: conth norm_mult k0 k1 intro!: contk continuous_intros)+ qed proposition homotopic_circlemaps_divide: fixes f :: "'a::real_normed_vector \ complex" shows "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c))" proof - have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" if "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)" for c proof - have "S = {} \ path_component (sphere 0 1) 1 c" using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1] by (auto simp: path_connected_component) then have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. 1) (\x. c)" by (simp add: homotopic_constant_maps) then show ?thesis using homotopic_with_symD homotopic_with_trans that by blast qed then have *: "(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)) \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" by auto have "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" (is "?lhs \ ?rhs") proof assume L: ?lhs have geq1 [simp]: "\x. x \ S \ cmod (g x) = 1" using homotopic_with_imp_subset2 [OF L] by (simp add: image_subset_iff) have cont: "continuous_on S (inverse \ g)" proof (rule continuous_intros) show "continuous_on S g" using homotopic_with_imp_continuous [OF L] by blast show "continuous_on (g ` S) inverse" by (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse]) auto qed have [simp]: "\x. x \ S \ g x \ 0" using geq1 by fastforce have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" using homotopic_with_eq [OF homotopic_with_sphere_times [OF L cont]] by (auto simp: divide_inverse norm_inverse) with L show ?rhs by (simp add: homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2) next assume ?rhs then show ?lhs by (elim conjE homotopic_with_eq [OF homotopic_with_sphere_times]; force) qed then show ?thesis by (simp add: *) qed subsection\Upper and lower hemicontinuous functions\ text\And relation in the case of preimage map to open and closed maps, and fact that upper and lower hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the function gives a bounded and nonempty set).\ text\Many similar proofs below.\ lemma upper_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}) \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ T - U}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ T - U} = {x \ S. f x \ U \ {}}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ (T - U) \ {}} = S - {x \ S. f x \ U}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma lower_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}) \ (\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ T-U}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ T-U} = S - {x \ S. f x \ U \ {}}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ (T - U) \ {}} = {x \ S. f x \ U}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma open_map_iff_lower_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)) \ (\U. closedin (top_of_set S) U \ closedin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "T - (f ` (S - U)) = {y \ T. {x \ S. f x = y} \ U}" using assms by blast ultimately show "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ U}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "{y \ T. {x \ S. f x = y} \ S - U} = T - (f ` U)" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) (f ` U)" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) qed lemma closed_map_iff_upper_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)) \ (\U. openin (top_of_set S) U \ openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "f ` (S - U) = T - {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and cloSU: "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "(f ` U) = T - {y \ T. {x \ S. f x = y} \ S - U}" using assms closedin_imp_subset [OF cloSU] by auto ultimately show "closedin (top_of_set T) (f ` U)" by (simp add: openin_closedin_eq) qed proposition upper_lower_hemicontinuous_explicit: fixes T :: "('b::{real_normed_vector,heine_borel}) set" assumes fST: "\x. x \ S \ f x \ T" and ope: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}" and clo: "\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}" and "x \ S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \ {}" obtains d where "0 < d" "\x'. \x' \ S; dist x x' < d\ \ (\y \ f x. \y'. y' \ f x' \ dist y y' < e) \ (\y' \ f x'. \y. y \ f x \ dist y' y < e)" proof - have "openin (top_of_set T) (T \ (\a\f x. \b\ball 0 e. {a + b}))" by (auto simp: open_sums openin_open_Int) with ope have "openin (top_of_set S) {u \ S. f u \ T \ (\a\f x. \b\ball 0 e. {a + b})}" by blast with \0 < e\ \x \ S\ obtain d1 where "d1 > 0" and d1: "\x'. \x' \ S; dist x' x < d1\ \ f x' \ T \ f x' \ (\a \ f x. \b \ ball 0 e. {a + b})" by (force simp: openin_euclidean_subtopology_iff dest: fST) have oo: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}" using lower_hemicontinuous fST clo by blast have "compact (closure(f x))" by (simp add: bofx) moreover have "closure(f x) \ (\a \ f x. ball a (e/2))" using \0 < e\ by (force simp: closure_approachable simp del: divide_const_simps) ultimately obtain C where "C \ f x" "finite C" "closure(f x) \ (\a \ C. ball a (e/2))" by (meson compactE finite_subset_image Elementary_Metric_Spaces.open_ball compactE_image) then have fx_cover: "f x \ (\a \ C. ball a (e/2))" by (meson closure_subset order_trans) with fx_ne have "C \ {}" by blast have xin: "x \ (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" using \x \ S\ \0 < e\ fST \C \ f x\ by force have "openin (top_of_set S) {x \ S. f x \ (T \ ball a (e/2)) \ {}}" for a by (simp add: openin_open_Int oo) then have "openin (top_of_set S) (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" by (simp add: Int_assoc openin_INT2 [OF \finite C\ \C \ {}\]) with xin obtain d2 where "d2>0" and d2: "\u v. \u \ S; dist u x < d2; v \ C\ \ f u \ T \ ball v (e/2) \ {}" unfolding openin_euclidean_subtopology_iff using xin by fastforce show ?thesis proof (intro that conjI ballI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by linarith next fix x' y assume "x' \ S" "dist x x' < min d1 d2" "y \ f x" then have dd2: "dist x' x < d2" by (auto simp: dist_commute) obtain a where "a \ C" "y \ ball a (e/2)" using fx_cover \y \ f x\ by auto then show "\y'. y' \ f x' \ dist y y' < e" using d2 [OF \x' \ S\ dd2] dist_triangle_half_r by fastforce next fix x' y' assume "x' \ S" "dist x x' < min d1 d2" "y' \ f x'" then have "dist x' x < d1" by (auto simp: dist_commute) then have "y' \ (\a\f x. \b\ball 0 e. {a + b})" using d1 [OF \x' \ S\] \y' \ f x'\ by force then show "\y. y \ f x \ dist y' y < e" by clarsimp (metis add_diff_cancel_left' dist_norm) qed qed subsection\Complex logs exist on various "well-behaved" sets\ lemma continuous_logarithm_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" proof - obtain c where hom: "homotopic_with_canon (\h. True) S (-{0}) f (\x. c)" using nullhomotopic_from_contractible assms by (metis imageE subset_Compl_singleton) then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma continuous_logarithm_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" using covering_space_lift [OF covering_space_exp_punctured_plane S contf] by (metis (full_types) f imageE subset_Compl_singleton) lemma continuous_logarithm_on_cball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (cball a r) f" and "\z. z \ cball a r \ f z \ 0" obtains h where "continuous_on (cball a r) h" "\z. z \ cball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_logarithm_on_ball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (ball a r) f" and "\z. z \ ball a r \ f z \ 0" obtains h where "continuous_on (ball a r) h" "\z. z \ ball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_sqrt_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" and "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_contractible [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed lemma continuous_sqrt_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_simply_connected [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed subsection\Another simple case where sphere maps are nullhomotopic\ lemma inessential_spheremap_2_aux: fixes f :: "'a::euclidean_space \ complex" assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere 0 1)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) f (\x. c)" proof - obtain g where contg: "continuous_on (sphere a r) g" and feq: "\x. x \ sphere a r \ f x = exp(g x)" proof (rule continuous_logarithm_on_simply_connected [OF contf]) show "simply_connected (sphere a r)" using 2 by (simp add: simply_connected_sphere_eq) show "locally path_connected (sphere a r)" by (simp add: locally_path_connected_sphere) show "\z. z \ sphere a r \ f z \ 0" using fim by force qed auto have "\g. continuous_on (sphere a r) g \ (\x\sphere a r. f x = exp (\ * complex_of_real (g x)))" proof (intro exI conjI) show "continuous_on (sphere a r) (Im \ g)" by (intro contg continuous_intros continuous_on_compose) show "\x\sphere a r. f x = exp (\ * complex_of_real ((Im \ g) x))" using exp_eq_polar feq fim norm_exp_eq_Re by auto qed with inessential_eq_continuous_logarithm_circle that show ?thesis by metis qed lemma inessential_spheremap_2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis using contf contractible_sphere fim nullhomotopic_into_contractible that by blast next case False then have "sphere b s homeomorphic sphere (0::complex) 1" using assms by (simp add: homeomorphic_spheres_gen) then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k" by (auto simp: homeomorphic_def) then have conth: "continuous_on (sphere b s) h" and contk: "continuous_on (sphere 0 1) k" and him: "h ` sphere b s \ sphere 0 1" and kim: "k ` sphere 0 1 \ sphere b s" by (simp_all add: homeomorphism_def) obtain c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) (h \ f) (\x. c)" proof (rule inessential_spheremap_2_aux [OF a2]) show "continuous_on (sphere a r) (h \ f)" by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim) show "(h \ f) ` sphere a r \ sphere 0 1" using fim him by force qed auto then have "homotopic_with_canon (\f. True) (sphere a r) (sphere b s) (k \ (h \ f)) (k \ (\x. c))" by (rule homotopic_with_compose_continuous_left [OF _ contk kim]) moreover have "\x. r = dist a x \ f x = k (h (f x))" by (metis fim hk homeomorphism_def image_subset_iff mem_sphere) ultimately have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" by (auto intro: homotopic_with_eq) then show ?thesis by (metis that) qed subsection\Holomorphic logarithms and square roots\ lemma g_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" by (simp add: LIM_offset_zero_iff DERIV_D cong: if_cong Lim_cong_within) qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" by (rule eventually_mono) (auto simp: exp_eq dist_norm norm_mult) then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed lemma contractible_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_contractible [OF contf S fnz]) then show thesis using fnz g_imp_holomorphic_log holf that by blast qed lemma simply_connected_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_simply_connected [OF contf S fnz]) then show thesis using fnz g_imp_holomorphic_log holf that by blast qed lemma contractible_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using contractible_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" by (simp add: feq flip: exp_double) qed qed lemma simply_connected_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using simply_connected_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" by (simp add: feq flip: exp_double) qed qed text\ Related theorems about holomorphic inverse cosines.\ lemma contractible_imp_holomorphic_arccos: assumes holf: "f holomorphic_on S" and S: "contractible S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = cos(g z)" proof - have hol1f: "(\z. 1 - f z ^ 2) holomorphic_on S" by (intro holomorphic_intros holf) obtain g where holg: "g holomorphic_on S" and eq: "\z. z \ S \ 1 - (f z)\<^sup>2 = (g z)\<^sup>2" using contractible_imp_holomorphic_sqrt [OF hol1f S] by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff) have holfg: "(\z. f z + \*g z) holomorphic_on S" by (intro holf holg holomorphic_intros) have "\z. z \ S \ f z + \*g z \ 0" by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff) then obtain h where holh: "h holomorphic_on S" and fgeq: "\z. z \ S \ f z + \*g z = exp (h z)" using contractible_imp_holomorphic_log [OF holfg S] by metis show ?thesis proof show "(\z. -\*h z) holomorphic_on S" by (intro holh holomorphic_intros) show "f z = cos (- \*h z)" if "z \ S" for z proof - have "(f z + \*g z)*(f z - \*g z) = 1" using that eq by (auto simp: algebra_simps power2_eq_square) then have "f z - \*g z = inverse (f z + \*g z)" using inverse_unique by force also have "... = exp (- h z)" by (simp add: exp_minus fgeq that) finally have "f z = exp (- h z) + \*g z" by (simp add: diff_eq_eq) with that show ?thesis by (simp add: cos_exp_eq flip: fgeq) qed qed qed lemma contractible_imp_holomorphic_arccos_bounded: assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \ S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "norm(g a) \ pi + norm(f a)" "\z. z \ S \ f z = cos(g z)" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = cos (g z)" using contractible_imp_holomorphic_arccos [OF holf S non1] by blast obtain b where "cos b = f a" "norm b \ pi + norm (f a)" using cos_Arccos norm_Arccos_bounded by blast then have "cos b = cos (g a)" by (simp add: \a \ S\ feq) then consider n where "n \ \" "b = g a + of_real(2*n*pi)" | n where "n \ \" "b = -g a + of_real(2*n*pi)" by (auto simp: complex_cos_eq) then show ?thesis proof cases case 1 show ?thesis proof show "(\z. g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "1" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed next case 2 show ?thesis proof show "(\z. -g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (-g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "2" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (-g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed qed qed subsection\The "Borsukian" property of sets\ text\This doesn't have a standard name. Kuratowski uses ``contractible with respect to \[S\<^sup>1]\'' while Whyburn uses ``property b''. It's closely related to unicoherence.\ definition\<^marker>\tag important\ Borsukian where "Borsukian S \ \f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\a. homotopic_with_canon (\h. True) S (- {0}) f (\x. a))" lemma Borsukian_retraction_gen: assumes "Borsukian S" "continuous_on S h" "h ` S = T" "continuous_on T k" "k ` T \ S" "\y. y \ T \ h(k y) = y" shows "Borsukian T" proof - interpret R: Retracts S h T k using assms by (simp add: Retracts.intro) show ?thesis using assms apply (clarsimp simp add: Borsukian_def) apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto) done qed lemma retract_of_Borsukian: "\Borsukian T; S retract_of T\ \ Borsukian S" by (smt (verit) Borsukian_retraction_gen retract_of_def retraction retraction_def retraction_subset) lemma homeomorphic_Borsukian: "\Borsukian S; S homeomorphic T\ \ Borsukian T" using Borsukian_retraction_gen order_refl by (fastforce simp add: homeomorphism_def homeomorphic_def) lemma homeomorphic_Borsukian_eq: "S homeomorphic T \ Borsukian S \ Borsukian T" by (meson homeomorphic_Borsukian homeomorphic_sym) lemma Borsukian_translation: fixes S :: "'a::real_normed_vector set" shows "Borsukian (image (\x. a + x) S) \ Borsukian S" using homeomorphic_Borsukian_eq homeomorphic_translation by blast lemma Borsukian_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "Borsukian(f ` S) \ Borsukian S" using assms homeomorphic_Borsukian_eq linear_homeomorphic_image by blast lemma homotopy_eqv_Borsukianness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "S homotopy_eqv T" shows "(Borsukian S \ Borsukian T)" by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null) lemma Borsukian_alt: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f g. continuous_on S f \ f ` S \ -{0} \ continuous_on S g \ g ` S \ -{0} \ homotopic_with_canon (\h. True) S (- {0::complex}) f g)" unfolding Borsukian_def homotopic_triviality by (simp add: path_connected_punctured_universe) lemma Borsukian_continuous_logarithm: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" by (simp add: Borsukian_def inessential_eq_continuous_logarithm) lemma Borsukian_continuous_logarithm_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp: Borsukian_continuous_logarithm) next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on S f" and 0: "0 \ f ` S" then have "continuous_on S (\x. f x / complex_of_real (cmod (f x)))" by (intro continuous_intros) auto moreover have "(\x. f x / complex_of_real (cmod (f x))) ` S \ sphere 0 1" using 0 by (auto simp: norm_divide) ultimately obtain g where contg: "continuous_on S g" and fg: "\x \ S. f x / complex_of_real (cmod (f x)) = exp(g x)" using RHS [of "\x. f x / of_real(norm(f x))"] by auto show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" proof (intro exI ballI conjI) show "continuous_on S (\x. (Ln \ of_real \ norm \ f)x + g x)" by (intro continuous_intros contf contg conjI) (use "0" in auto) show "f x = exp ((Ln \ complex_of_real \ cmod \ f) x + g x)" if "x \ S" for x using 0 that apply (simp add: exp_add) by (metis div_by_0 exp_Ln exp_not_eq_zero fg mult.commute nonzero_eq_divide_eq) qed qed qed lemma Borsukian_continuous_logarithm_circle_real: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S (complex_of_real \ g) \ (\x \ S. f x = exp(\ * of_real(g x)))))" (is "?lhs = ?rhs") proof assume LHS: ?lhs show ?rhs proof (clarify) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" using LHS by (auto simp: Borsukian_continuous_logarithm_circle) then have "\x\S. f x = exp (\ * complex_of_real ((Im \ g) x))" using f01 exp_eq_polar norm_exp_eq_Re by auto then show "\g. continuous_on S (complex_of_real \ g) \ (\x\S. f x = exp (\ * complex_of_real (g x)))" by (rule_tac x="Im \ g" in exI) (force intro: continuous_intros contg) qed next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm_circle) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S (complex_of_real \ g)" and "\x. x \ S \ f x = exp(\ * of_real(g x))" by (metis RHS) then show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" by (rule_tac x="\x. \* of_real(g x)" in exI) (auto simp: continuous_intros contg) qed qed lemma Borsukian_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\a. homotopic_with_canon (\h. True) S (sphere (0::complex) 1) f (\x. a)))" by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real) lemma contractible_imp_Borsukian: "contractible S \ Borsukian S" by (meson Borsukian_def nullhomotopic_from_contractible) lemma simply_connected_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "\simply_connected S; locally path_connected S\ \ Borsukian S" by (metis (no_types, lifting) Borsukian_continuous_logarithm continuous_logarithm_on_simply_connected image_eqI subset_Compl_singleton) lemma starlike_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "starlike S \ Borsukian S" by (simp add: contractible_imp_Borsukian starlike_imp_contractible) lemma Borsukian_empty: "Borsukian {}" by (auto simp: contractible_imp_Borsukian) lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)" by (auto simp: contractible_imp_Borsukian) lemma convex_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "convex S \ Borsukian S" by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible) proposition Borsukian_sphere: fixes a :: "'a::euclidean_space" shows "3 \ DIM('a) \ Borsukian (sphere a r)" using ENR_sphere by (blast intro: simply_connected_imp_Borsukian ENR_imp_locally_path_connected simply_connected_sphere) lemma Borsukian_Un_lemma: fixes S :: "'a::real_normed_vector set" assumes BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" and *: "\f g::'a \ complex. \continuous_on S f; continuous_on T g; \x. x \ S \ x \ T \ f x = g x\ \ continuous_on (S \ T) (\x. if x \ S then f x else g x)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" using True * [OF contg conth] by (meson disjoint_iff) show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" by (metis contg conth continuous_on_diff continuous_on_subset inf_le1 inf_le2) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by (auto simp: constant_on_def) with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro * contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) done qed qed proposition Borsukian_open_Un: fixes S :: "'a::real_normed_vector set" assumes opeS: "openin (top_of_set (S \ T)) S" and opeT: "openin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" by (force intro: Borsukian_Un_lemma [OF BS BT ST] continuous_on_cases_local_open [OF opeS opeT]) lemma Borsukian_closed_Un: fixes S :: "'a::real_normed_vector set" assumes cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" by (force intro: Borsukian_Un_lemma [OF BS BT ST] continuous_on_cases_local [OF cloS cloT]) lemma Borsukian_separation_compact: fixes S :: "complex set" assumes "compact S" shows "Borsukian S \ connected(- S)" by (simp add: Borsuk_separation_theorem Borsukian_circle assms) lemma Borsukian_monotone_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and conn: "\y. y \ T \ connected {x. x \ S \ f x = y}" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix g :: "'b \ complex" assume contg: "continuous_on T g" and 0: "0 \ g ` T" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ -{0}" using fim 0 by auto ultimately obtain h where conth: "continuous_on S h" and gfh: "\x. x \ S \ (g \ f) x = exp(h x)" using \Borsukian S\ by (auto simp: Borsukian_continuous_logarithm) have "\y. \x. y \ T \ x \ S \ f x = y" using fim by auto then obtain f' where f': "\y. y \ T \ f' y \ S \ f (f' y) = y" by metis have *: "(\x. h x - h(f' y)) constant_on {x. x \ S \ f x = y}" if "y \ T" for y proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\x. h x - h (f' y)"], simp_all add: algebra_simps) show "continuous_on {x \ S. f x = y} (\x. h x - h (f' y))" by (intro continuous_intros continuous_on_subset [OF conth]) auto show "\e>0. \u. u \ S \ f u = y \ h u \ h x \ e \ cmod (h u - h x)" if x: "x \ S \ f x = y" for x proof - have "h u = h x" if "u \ S" "f u = y" "cmod (h u - h x) < 2 * pi" for u proof (rule exp_complex_eqI) have "\Im (h u) - Im (h x)\ \ cmod (h u - h x)" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (h u) - Im (h x)\ < 2 * pi" using that by linarith show "exp (h u) = exp (h x)" by (simp add: gfh [symmetric] x that) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed show "\h. continuous_on T h \ (\x\T. g x = exp (h x))" proof (intro exI conjI) show "continuous_on T (h \ f')" proof (rule continuous_from_closed_graph [of "h ` S"]) show "compact (h ` S)" by (simp add: \compact S\ compact_continuous_image conth) show "(h \ f') ` T \ h ` S" by (auto simp: f') have "h x = h (f' (f x))" if "x \ S" for x using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq) moreover have "\x. x \ T \ \u. u \ S \ x = f u \ h (f' x) = h u" using f' by fastforce ultimately have eq: "((\x. (x, (h \ f') x)) ` T) = {p. \x. x \ S \ (x, p) \ (S \ UNIV) \ ((\z. snd z - ((f \ fst) z, (h \ fst) z)) -` {0})}" using fim by (auto simp: image_iff) moreover have "closed \" apply (intro closed_compact_projection [OF \compact S\] continuous_closed_preimage continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth]) by (auto simp: \compact S\ closed_Times compact_imp_closed) ultimately show "closed ((\x. (x, (h \ f') x)) ` T)" by simp qed qed (use f' gfh in fastforce) qed lemma Borsukian_open_map_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and ope: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real) fix g :: "'b \ complex" assume contg: "continuous_on T g" and gim: "g ` T \ sphere 0 1" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ sphere 0 1" using fim gim by auto ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \ h)" and gfh: "\x. x \ S \ (g \ f) x = exp(\ * of_real(h x))" using \Borsukian S\ Borsukian_continuous_logarithm_circle_real by metis then have conth: "continuous_on S h" by simp have "\x. x \ S \ f x = y \ (\x' \ S. f x' = y \ h x \ h x')" if "y \ T" for y proof - have 1: "compact (h ` {x \ S. f x = y})" proof (rule compact_continuous_image) show "continuous_on {x \ S. f x = y} h" by (rule continuous_on_subset [OF conth]) auto have "compact (S \ f -` {y})" by (rule proper_map_from_compact [OF contf _ \compact S\, of T]) (simp_all add: fim that) then show "compact {x \ S. f x = y}" by (auto simp: vimage_def Int_def) qed have 2: "h ` {x \ S. f x = y} \ {}" using fim that by auto have "\s \ h ` {x \ S. f x = y}. \t \ h ` {x \ S. f x = y}. s \ t" using compact_attains_inf [OF 1 2] by blast then show ?thesis by auto qed then obtain k where kTS: "\y. y \ T \ k y \ S" and fk: "\y. y \ T \ f (k y) = y " and hle: "\x' y. \y \ T; x' \ S; f x' = y\ \ h (k y) \ h x'" by metis have "continuous_on T (h \ k)" proof (clarsimp simp add: continuous_on_iff) fix y and e::real assume "y \ T" "0 < e" moreover have "uniformly_continuous_on S (complex_of_real \ h)" using \compact S\ cont_cxh compact_uniformly_continuous by blast ultimately obtain d where "0 < d" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (h x') (h x) < e" by (force simp: uniformly_continuous_on_def) obtain \ where "0 < \" and \: "\x'. \x' \ T; dist y x' < \\ \ (\v \ {z \ S. f z = y}. \v'. v' \ {z \ S. f z = x'} \ dist v v' < d) \ (\v' \ {z \ S. f z = x'}. \v. v \ {z \ S. f z = y} \ dist v' v < d)" proof (rule upper_lower_hemicontinuous_explicit [of T "\y. {z \ S. f z = y}" S]) show "\U. openin (top_of_set S) U \ openin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]] by (simp add: Abstract_Topology_2.continuous_imp_closed_map \compact S\ contf fim) show "\U. closedin (top_of_set S) U \ closedin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] by meson show "bounded {z \ S. f z = y}" by (metis (no_types, lifting) compact_imp_bounded [OF \compact S\] bounded_subset mem_Collect_eq subsetI) qed (use \y \ T\ \0 < d\ fk kTS in \force+\) have "dist (h (k y')) (h (k y)) < e" if "y' \ T" "dist y y' < \" for y' proof - have k1: "k y \ S" "f (k y) = y" and k2: "k y' \ S" "f (k y') = y'" by (auto simp: \y \ T\ \y' \ T\ kTS fk) have 1: "\v. \v \ S; f v = y\ \ \v'. v' \ {z \ S. f z = y'} \ dist v v' < d" and 2: "\v'. \v' \ S; f v' = y'\ \ \v. v \ {z \ S. f z = y} \ dist v' v < d" using \ [OF that] by auto then obtain w' w where "w' \ S" "f w' = y'" "dist (k y) w' < d" and "w \ S" "f w = y" "dist (k y') w < d" using 1 [OF k1] 2 [OF k2] by auto then show ?thesis using d [of w "k y'"] d [of w' "k y"] k1 k2 \y' \ T\ \y \ T\ hle by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps) qed then show "\d>0. \x'\T. dist x' y < d \ dist (h (k x')) (h (k y)) < e" using \0 < \\ by (auto simp: dist_commute) qed then show "\h. continuous_on T h \ (\x\T. g x = exp (\ * complex_of_real (h x)))" using fk gfh kTS by force qed text\If two points are separated by a closed set, there's a minimal one.\ proposition closed_irreducible_separator: fixes a :: "'a::real_normed_vector" assumes "closed S" and ab: "\ connected_component (- S) a b" obtains T where "T \ S" "closed T" "T \ {}" "\ connected_component (- T) a b" "\U. U \ T \ connected_component (- U) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis proof assume *: "a \ S" show ?thesis proof show "{a} \ S" using * by blast show "\ connected_component (- {a}) a b" using connected_component_in by auto show "\U. U \ {a} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto next assume *: "b \ S" show ?thesis proof show "{b} \ S" using * by blast show "\ connected_component (- {b}) a b" using connected_component_in by auto show "\U. U \ {b} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto qed next case False define A where "A \ connected_component_set (- S) a" define B where "B \ connected_component_set (- (closure A)) b" have "a \ A" using False A_def by auto have "b \ B" unfolding A_def B_def closure_Un_frontier using ab False \closed S\ frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force have "frontier B \ frontier (connected_component_set (- closure A) b)" using B_def by blast also have frsub: "... \ frontier A" proof - have "\A. closure (- closure (- A)) \ closure A" by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl) then show ?thesis by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans) qed finally have frBA: "frontier B \ frontier A" . show ?thesis proof show "frontier B \ S" proof - have "frontier S \ S" by (simp add: \closed S\ frontier_subset_closed) then show ?thesis using frsub frontier_complement frontier_of_connected_component_subset unfolding A_def B_def by blast qed show "closed (frontier B)" by simp show "\ connected_component (- frontier B) a b" unfolding connected_component_def proof clarify fix T assume "connected T" and TB: "T \ - frontier B" and "a \ T" and "b \ T" have "a \ B" by (metis A_def B_def ComplD \a \ A\ assms(1) closed_open connected_component_subset in_closure_connected_component subsetD) have "T \ B \ {}" using \b \ B\ \b \ T\ by blast moreover have "T - B \ {}" using \a \ B\ \a \ T\ by blast ultimately show "False" using connected_Int_frontier [of T B] TB \connected T\ by blast qed moreover have "connected_component (- frontier B) a b" if "frontier B = {}" using connected_component_eq_UNIV that by auto ultimately show "frontier B \ {}" by blast show "connected_component (- U) a b" if "U \ frontier B" for U proof - obtain p where Usub: "U \ frontier B" and p: "p \ frontier B" "p \ U" using \U \ frontier B\ by blast show ?thesis unfolding connected_component_def proof (intro exI conjI) have "connected ((insert p A) \ (insert p B))" proof (rule connected_Un) show "connected (insert p A)" by (metis A_def IntD1 frBA \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI) show "connected (insert p B)" by (metis B_def IntD1 \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI) qed blast then show "connected (insert p (B \ A))" by (simp add: sup.commute) have "A \ - U" using A_def Usub \frontier B \ S\ connected_component_subset by fastforce moreover have "B \ - U" using B_def Usub connected_component_subset frBA frontier_closures by fastforce ultimately show "insert p (B \ A) \ - U" using p by auto qed (auto simp: \a \ A\ \b \ B\) qed qed qed lemma frontier_minimal_separating_closed_pointwise: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" "a \ S" and nconn: "\ connected_component (- S) a b" and conn: "\T. \closed T; T \ S\ \ connected_component (- T) a b" shows "frontier(connected_component_set (- S) a) = S" (is "?F = S") proof - have "?F \ S" by (simp add: S componentsI frontier_of_components_closed_complement) moreover have False if "?F \ S" proof - have "connected_component (- ?F) a b" by (simp add: conn that) then obtain T where "connected T" "T \ -?F" "a \ T" "b \ T" by (auto simp: connected_component_def) moreover have "T \ ?F \ {}" proof (rule connected_Int_frontier [OF \connected T\]) show "T \ connected_component_set (- S) a \ {}" using \a \ S\ \a \ T\ by fastforce show "T - connected_component_set (- S) a \ {}" using \b \ T\ nconn by blast qed ultimately show ?thesis by blast qed ultimately show ?thesis by blast qed subsection\Unicoherence (closed)\ definition\<^marker>\tag important\ unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ closedin (top_of_set U) S \ closedin (top_of_set U) T \ connected (S \ T)" lemma unicoherentI [intro?]: assumes "\S T. \connected S; connected T; U = S \ T; closedin (top_of_set U) S; closedin (top_of_set U) T\ \ connected (S \ T)" shows "unicoherent U" using assms unfolding unicoherent_def by blast lemma unicoherentD: assumes "unicoherent U" "connected S" "connected T" "U = S \ T" "closedin (top_of_set U) S" "closedin (top_of_set U) T" shows "connected (S \ T)" using assms unfolding unicoherent_def by blast proposition homeomorphic_unicoherent: assumes ST: "S homeomorphic T" and S: "unicoherent S" shows "unicoherent T" proof - obtain f g where gf: "\x. x \ S \ g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S" and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g" using ST by (auto simp: homeomorphic_def homeomorphism_def) show ?thesis proof fix U V assume "connected U" "connected V" and T: "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" have "f ` (g ` U \ g ` V) \ U" "f ` (g ` U \ g ` V) \ V" using gf fim T by auto (metis UnCI image_iff)+ moreover have "U \ V \ f ` (g ` U \ g ` V)" using gf fim by (force simp: image_iff T) ultimately have "U \ V = f ` (g ` U \ g ` V)" by blast moreover have "connected (f ` (g ` U \ g ` V))" proof (rule connected_continuous_image) show "continuous_on (g ` U \ g ` V) f" using T fim gfim by (metis Un_upper1 contf continuous_on_subset image_mono inf_le1) show "connected (g ` U \ g ` V)" proof (intro conjI unicoherentD [OF S]) show "connected (g ` U)" "connected (g ` V)" using \connected U\ cloU \connected V\ cloV by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+ show "S = g ` U \ g ` V" using T fim gfim by auto have hom: "homeomorphism T S g f" by (simp add: contf contg fim gf gfim homeomorphism_def) have "closedin (top_of_set T) U" "closedin (top_of_set T) V" by (simp_all add: cloU cloV) then show "closedin (top_of_set S) (g ` U)" "closedin (top_of_set S) (g ` V)" by (blast intro: homeomorphism_imp_closed_map [OF hom])+ qed qed ultimately show "connected (U \ V)" by metis qed qed lemma homeomorphic_unicoherent_eq: "S homeomorphic T \ (unicoherent S \ unicoherent T)" by (meson homeomorphic_sym homeomorphic_unicoherent) lemma unicoherent_translation: fixes S :: "'a::real_normed_vector set" shows "unicoherent (image (\x. a + x) S) \ unicoherent S" using homeomorphic_translation homeomorphic_unicoherent_eq by blast lemma unicoherent_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(unicoherent(f ` S) \ unicoherent S)" using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast lemma Borsukian_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "Borsukian U" shows "unicoherent U" unfolding unicoherent_def proof clarify fix S T assume "connected S" "connected T" "U = S \ T" and cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" show "connected (S \ T)" unfolding connected_closedin_eq proof clarify fix V W assume "closedin (top_of_set (S \ T)) V" and "closedin (top_of_set (S \ T)) W" and VW: "V \ W = S \ T" "V \ W = {}" and "V \ {}" "W \ {}" then have cloV: "closedin (top_of_set U) V" and cloW: "closedin (top_of_set U) W" using \U = S \ T\ cloS cloT closedin_trans by blast+ obtain q where contq: "continuous_on U q" and q01: "\x. x \ U \ q x \ {0..1::real}" and qV: "\x. x \ V \ q x = 0" and qW: "\x. x \ W \ q x = 1" by (rule Urysohn_local [OF cloV cloW \V \ W = {}\, of 0 1]) (fastforce simp: closed_segment_eq_real_ivl) let ?h = "\x. if x \ S then exp(pi * \ * q x) else 1 / exp(pi * \ * q x)" have eqST: "exp(pi * \ * q x) = 1 / exp(pi * \ * q x)" if "x \ S \ T" for x proof - have "x \ V \ W" using that \V \ W = S \ T\ by blast with qV qW show ?thesis by force qed obtain g where contg: "continuous_on U g" and circle: "g ` U \ sphere 0 1" and S: "\x. x \ S \ g x = exp(pi * \ * q x)" and T: "\x. x \ T \ g x = 1 / exp(pi * \ * q x)" proof show "continuous_on U ?h" unfolding \U = S \ T\ proof (rule continuous_on_cases_local [OF cloS cloT]) show "continuous_on S (\x. exp (pi * \ * q x))" proof (intro continuous_intros) show "continuous_on S q" using \U = S \ T\ continuous_on_subset contq by blast qed show "continuous_on T (\x. 1 / exp (pi * \ * q x))" proof (intro continuous_intros) show "continuous_on T q" using \U = S \ T\ continuous_on_subset contq by auto qed auto qed (use eqST in auto) qed (use eqST in \auto simp: norm_divide\) then obtain h where conth: "continuous_on U h" and heq: "\x. x \ U \ g x = exp (h x)" by (metis Borsukian_continuous_logarithm_circle assms) obtain v w where "v \ V" "w \ W" using \V \ {}\ \W \ {}\ by blast then have vw: "v \ S \ T" "w \ S \ T" using VW by auto have iff: "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 1 \ abs (m - n)" for m n proof - have "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 2 * pi \ cmod ((2 * pi * \) * (of_int m - of_int n))" by (simp add: algebra_simps) also have "... \ 2 * pi \ 2 * pi * cmod (of_int m - of_int n)" by (simp add: norm_mult) also have "... \ 1 \ abs (m - n)" by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff) finally show ?thesis . qed have *: "\n::int. h x - (pi * \ * q x) = (of_int(2*n) * pi) * \" if "x \ S" for x using that S \U = S \ T\ heq exp_eq [symmetric] by (simp add: algebra_simps) moreover have "(\x. h x - (pi * \ * q x)) constant_on S" proof (rule continuous_discrete_range_constant [OF \connected S\]) have "continuous_on S h" "continuous_on S q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on S (\x. h x - (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" if "x \ S" "y \ S" and ne: "h y - (pi * \ * q y) \ h x - (pi * \ * q x)" for x y using * [OF \x \ S\] * [OF \y \ S\] ne by (auto simp: iff) then show "\x. x \ S \ \e>0. \y. y \ S \ h y - (pi * \ * q y) \ h x - (pi * \ * q x) \ e \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain m where m: "\x. x \ S \ h x - (pi * \ * q x) = (of_int(2*m) * pi) * \" using vw by (force simp: constant_on_def) have *: "\n::int. h x = - (pi * \ * q x) + (of_int(2*n) * pi) * \" if "x \ T" for x unfolding exp_eq [symmetric] using that T \U = S \ T\ by (simp add: exp_minus field_simps heq [symmetric]) moreover have "(\x. h x + (pi * \ * q x)) constant_on T" proof (rule continuous_discrete_range_constant [OF \connected T\]) have "continuous_on T h" "continuous_on T q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on T (\x. h x + (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" if "x \ T" "y \ T" and ne: "h y + (pi * \ * q y) \ h x + (pi * \ * q x)" for x y using * [OF \x \ T\] * [OF \y \ T\] ne by (auto simp: iff) then show "\x. x \ T \ \e>0. \y. y \ T \ h y + (pi * \ * q y) \ h x + (pi * \ * q x) \ e \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain n where n: "\x. x \ T \ h x + (pi * \ * q x) = (of_int(2*n) * pi) * \" using vw by (force simp: constant_on_def) show "False" using m [of v] m [of w] n [of v] n [of w] vw by (auto simp: algebra_simps \v \ V\ \w \ W\ qV qW) qed qed corollary contractible_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "contractible U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian) corollary convex_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "convex U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian) text\If the type class constraint can be relaxed, I don't know how!\ corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)" by (simp add: convex_imp_unicoherent) lemma unicoherent_monotone_image_compact: fixes T :: "'b :: t2_space set" assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T" and conn: "\y. y \ T \ connected (S \ f -` {y})" shows "unicoherent T" proof fix U V assume UV: "connected U" "connected V" "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" moreover have "compact T" using \compact S\ compact_continuous_image contf fim by blast ultimately have "closed U" "closed V" by (auto simp: closedin_closed_eq compact_imp_closed) let ?SUV = "(S \ f -` U) \ (S \ f -` V)" have UV_eq: "f ` ?SUV = U \ V" using \T = U \ V\ fim by force+ have "connected (f ` ?SUV)" proof (rule connected_continuous_image) show "continuous_on ?SUV f" by (meson contf continuous_on_subset inf_le1) show "connected ?SUV" proof (rule unicoherentD [OF \unicoherent S\, of "S \ f -` U" "S \ f -` V"]) have "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" by (metis \compact S\ closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono) then show "connected (S \ f -` U)" "connected (S \ f -` V)" using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim]) show "S = (S \ f -` U) \ (S \ f -` V)" using UV fim by blast show "closedin (top_of_set S) (S \ f -` U)" "closedin (top_of_set S) (S \ f -` V)" by (auto simp: continuous_on_imp_closedin cloU cloV contf fim) qed qed with UV_eq show "connected (U \ V)" by simp qed subsection\Several common variants of unicoherence\ lemma connected_frontier_simple: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected(- S)" shows "connected(frontier S)" unfolding frontier_closures by (rule unicoherentD [OF unicoherent_UNIV]; simp add: assms connected_imp_connected_closure flip: closure_Un) lemma connected_frontier_component_complement: fixes S :: "'a :: euclidean_space set" assumes "connected S" "C \ components(- S)" shows "connected(frontier C)" by (meson assms component_complement_connected connected_frontier_simple in_components_connected) lemma connected_frontier_disjoint: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \ frontier T" shows "connected(frontier S)" proof (cases "S = UNIV") case True then show ?thesis by simp next case False then have "-S \ {}" by blast then obtain C where C: "C \ components(- S)" and "T \ C" by (metis ComplI disjnt_iff subsetI exists_component_superset \disjnt S T\ \connected T\) moreover have "frontier S = frontier C" proof - have "frontier C \ frontier S" using C frontier_complement frontier_of_components_subset by blast moreover have "x \ frontier C" if "x \ frontier S" for x proof - have "x \ closure C" using that unfolding frontier_def by (metis (no_types) Diff_eq ST \T \ C\ closure_mono contra_subsetD frontier_def le_inf_iff that) moreover have "x \ interior C" using that unfolding frontier_def by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono) ultimately show ?thesis by (auto simp: frontier_def) qed ultimately show ?thesis by blast qed ultimately show ?thesis using \connected S\ connected_frontier_component_complement by auto qed subsection\Some separation results\ lemma separation_by_component_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected_component (- S) a b" obtains C where "C \ components S" "\ connected_component(- C) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis using connected_component_in componentsI that by fastforce next case False obtain T where "T \ S" "closed T" "T \ {}" and nab: "\ connected_component (- T) a b" and conn: "\U. U \ T \ connected_component (- U) a b" using closed_irreducible_separator [OF assms] by metis moreover have "connected T" proof - have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T" using frontier_minimal_separating_closed_pointwise by (metis False \T \ S\ \closed T\ connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+ have "connected (frontier (connected_component_set (- T) a))" proof (rule connected_frontier_disjoint) show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)" unfolding disjnt_iff by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab) show "frontier (connected_component_set (- T) a) \ frontier (connected_component_set (- T) b)" by (simp add: ab) qed auto with ab \closed T\ show ?thesis by simp qed ultimately obtain C where "C \ components S" "T \ C" using exists_component_superset [of T S] by blast then show ?thesis by (meson Compl_anti_mono connected_component_of_subset nab that) qed lemma separation_by_component_closed: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain x y where "closed S" "x \ S" "y \ S" and "\ connected_component (- S) x y" using assms by (auto simp: connected_iff_connected_component) then obtain C where "C \ components S" "\ connected_component(- C) x y" using separation_by_component_closed_pointwise by metis then show "thesis" by (metis Compl_iff \x \ S\ \y \ S\ connected_component_eq_self in_components_subset mem_Collect_eq subsetD that) qed lemma separation_by_Un_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b" shows "connected_component (- (S \ T)) a b" proof (rule ccontr) have "a \ S" "b \ S" "a \ T" "b \ T" using conS conT connected_component_in by auto assume "\ connected_component (- (S \ T)) a b" then obtain C where "C \ components (S \ T)" and C: "\ connected_component(- C) a b" using separation_by_component_closed_pointwise assms by blast then have "C \ S \ C \ T" proof - have "connected C" "C \ S \ T" using \C \ components (S \ T)\ in_components_subset by (blast elim: componentsE)+ moreover then have "C \ T = {} \ C \ S = {}" by (metis Int_empty_right ST inf.commute connected_closed) ultimately show ?thesis by blast qed then show False by (meson Compl_anti_mono C conS conT connected_component_of_subset) qed lemma separation_by_Un_closed: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected(- S)" and conT: "connected(- T)" shows "connected(- (S \ T))" using assms separation_by_Un_closed_pointwise by (fastforce simp add: connected_iff_connected_component) lemma open_unicoherent_UNIV: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "connected S" "connected T" "S \ T = UNIV" shows "connected(S \ T)" proof - have "connected(- (-S \ -T))" by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms) then show ?thesis by simp qed lemma separation_by_component_open_aux: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and "S \ {}" "T \ {}" obtains C where "C \ components(-(S \ T))" "C \ {}" "frontier C \ S \ {}" "frontier C \ T \ {}" proof (rule ccontr) let ?S = "S \ \{C \ components(- (S \ T)). frontier C \ S}" let ?T = "T \ \{C \ components(- (S \ T)). frontier C \ T}" assume "\ thesis" with that have *: "frontier C \ S = {} \ frontier C \ T = {}" if C: "C \ components (- (S \ T))" "C \ {}" for C using C by blast have "\A B::'a set. closed A \ closed B \ UNIV \ A \ B \ A \ B = {} \ A \ {} \ B \ {}" proof (intro exI conjI) have "frontier (\{C \ components (- S \ - T). frontier C \ S}) \ S" using subset_trans [OF frontier_Union_subset_closure] by (metis (no_types, lifting) SUP_least \closed S\ closure_minimal mem_Collect_eq) then have "frontier ?S \ S" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?S" using frontier_subset_eq by fastforce have "frontier (\{C \ components (- S \ - T). frontier C \ T}) \ T" using subset_trans [OF frontier_Union_subset_closure] by (metis (no_types, lifting) SUP_least \closed T\ closure_minimal mem_Collect_eq) then have "frontier ?T \ T" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?T" using frontier_subset_eq by fastforce have "UNIV \ (S \ T) \ \(components(- (S \ T)))" using Union_components by blast also have "... \ ?S \ ?T" proof - have "C \ components (-(S \ T)) \ frontier C \ S \ C \ components (-(S \ T)) \ frontier C \ T" if "C \ components (- (S \ T))" "C \ {}" for C using * [OF that] that by clarify (metis (no_types, lifting) UnE \closed S\ \closed T\ closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE) then show ?thesis by blast qed finally show "UNIV \ ?S \ ?T" . have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} \ - (S \ T)" using in_components_subset by fastforce moreover have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} = {}" proof - have "C \ C' = {}" if "C \ components (- (S \ T))" "frontier C \ S" "C' \ components (- (S \ T))" "frontier C' \ T" for C C' proof - have NUN: "- S \ - T \ UNIV" using \T \ {}\ by blast have "C \ C'" proof assume "C = C'" with that have "frontier C' \ S \ T" by simp also have "... = {}" using \S \ T = {}\ by blast finally have "C' = {} \ C' = UNIV" using frontier_eq_empty by auto then show False using \C = C'\ NUN that by (force simp: dest: in_components_nonempty in_components_subset) qed with that show ?thesis by (simp add: components_nonoverlap [of _ "-(S \ T)"]) qed then show ?thesis by blast qed ultimately show "?S \ ?T = {}" using ST by blast show "?S \ {}" "?T \ {}" using \S \ {}\ \T \ {}\ by blast+ qed then show False by (metis Compl_disjoint connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI) qed proposition separation_by_component_open: fixes S :: "'a :: euclidean_space set" assumes "open S" and non: "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain T U where "closed T" "closed U" and TU: "T \ U = - S" "T \ U = {}" "T \ {}" "U \ {}" using assms by (auto simp: connected_closed_set closed_def) then obtain C where C: "C \ components(-(T \ U))" "C \ {}" and "frontier C \ T \ {}" "frontier C \ U \ {}" using separation_by_component_open_aux [OF \closed T\ \closed U\ \T \ U = {}\] by force show "thesis" proof show "C \ components S" using C(1) TU(1) by auto show "\ connected (- C)" proof assume "connected (- C)" then have "connected (frontier C)" using connected_frontier_simple [of C] \C \ components S\ in_components_connected by blast then show False unfolding connected_closed by (metis C(1) TU(2) \closed T\ \closed U\ \frontier C \ T \ {}\ \frontier C \ U \ {}\ closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute) qed qed qed lemma separation_by_Un_open: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "S \ T = {}" and cS: "connected(-S)" and cT: "connected(-T)" shows "connected(- (S \ T))" using assms unicoherent_UNIV unfolding unicoherent_def by force lemma nonseparation_by_component_eq: fixes S :: "'a :: euclidean_space set" assumes "open S \ closed S" shows "((\C \ components S. connected(-C)) \ connected(- S))" (is "?lhs = ?rhs") proof assume ?lhs with assms show ?rhs by (meson separation_by_component_closed separation_by_component_open) next assume ?rhs with assms show ?lhs using component_complement_connected by force qed text\Another interesting equivalent of an inessential mapping into C-{0}\ proposition inessential_eq_extensible: fixes f :: "'a::euclidean_space \ complex" assumes "closed S" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on UNIV g \ (\x \ S. g x = f x) \ (\x. g x \ 0))" (is "?lhs = ?rhs") proof assume ?lhs then obtain a where a: "homotopic_with_canon (\h. True) S (-{0}) f (\t. a)" .. show ?rhs proof (cases "S = {}") case True with a show ?thesis by force next case False have anr: "ANR (-{0::complex})" by (simp add: ANR_delete open_Compl open_imp_ANR) obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \ -{0}" and gf: "\x. x \ S \ g x = f x" proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]]) show "closedin (top_of_set UNIV) S" using assms by auto show "range (\t. a) \ - {0}" using a homotopic_with_imp_subset2 False by blast qed (use anr that in \force+\) then show ?thesis by force qed next assume ?rhs then obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ S \ g x = f x" and non0: "\x. g x \ 0" by metis obtain h k::"'a\'a" where hk: "homeomorphism (ball 0 1) UNIV h k" using homeomorphic_ball01_UNIV homeomorphic_def by blast then have "continuous_on (ball 0 1) (g \ h)" by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest) then obtain j where contj: "continuous_on (ball 0 1) j" and j: "\z. z \ ball 0 1 \ exp(j z) = (g \ h) z" by (metis (mono_tags, opaque_lifting) continuous_logarithm_on_ball comp_apply non0) have [simp]: "\x. x \ S \ h (k x) = x" using hk homeomorphism_apply2 by blast have "\\. continuous_on S \\ (\x\S. f x = exp (\ x))" proof (intro exI conjI ballI) show "continuous_on S (j \ k)" proof (rule continuous_on_compose) show "continuous_on S k" by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest) show "continuous_on (k ` S) j" by (auto intro: continuous_on_subset [OF contj] simp flip: homeomorphism_image2 [OF hk]) qed show "f x = exp ((j \ k) x)" if "x \ S" for x proof - have "f x = (g \ h) (k x)" by (simp add: gf that) also have "... = exp (j (k x))" by (metis rangeI homeomorphism_image2 [OF hk] j) finally show ?thesis by simp qed qed then show ?lhs by (simp add: inessential_eq_continuous_logarithm) qed lemma inessential_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes T: "path_connected T" and "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and hom: "\S. S \ \ \ \a. homotopic_with_canon (\x. True) S T f (\x. a)" obtains a where "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (cases "\\ = {}") case True with that show ?thesis by force next case False then obtain C where "C \ \" "C \ {}" by blast then obtain a where clo: "closedin (top_of_set (\\)) C" and ope: "openin (top_of_set (\\)) C" and "homotopic_with_canon (\x. True) C T f (\x. a)" using assms by blast with \C \ {}\ have "f ` C \ T" "a \ T" using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+ have "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (rule homotopic_on_clopen_Union) show "\S. S \ \ \ closedin (top_of_set (\\)) S" "\S. S \ \ \ openin (top_of_set (\\)) S" by (simp_all add: assms) show "homotopic_with_canon (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") case False then obtain b where "b \ S" by blast obtain c where c: "homotopic_with_canon (\x. True) S T f (\x. c)" using \S \ \\ hom by blast then have "c \ T" using \b \ S\ homotopic_with_imp_subset2 by blast then have "homotopic_with_canon (\x. True) S T (\x. a) (\x. c)" using T \a \ T\ by (simp add: homotopic_constant_maps path_connected_component) then show ?thesis using c homotopic_with_symD homotopic_with_trans by blast qed auto qed then show ?thesis .. qed proposition Janiszewski_dual: fixes S :: "complex set" assumes "compact S" "compact T" "connected S" "connected T" "connected(- (S \ T))" shows "connected(S \ T)" proof - have ST: "compact (S \ T)" by (simp add: assms compact_Un) with Borsukian_imp_unicoherent [of "S \ T"] ST assms show ?thesis by (simp add: Borsukian_separation_compact closed_subset compact_imp_closed unicoherentD) qed end diff --git a/src/HOL/Analysis/Lipschitz.thy b/src/HOL/Analysis/Lipschitz.thy --- a/src/HOL/Analysis/Lipschitz.thy +++ b/src/HOL/Analysis/Lipschitz.thy @@ -1,839 +1,850 @@ (* Title: HOL/Analysis/Lipschitz.thy Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr Author: Fabian Immler, TU München *) section \Lipschitz Continuity\ theory Lipschitz imports - Derivative + Derivative Abstract_Metric_Spaces begin definition\<^marker>\tag important\ lipschitz_on where "lipschitz_on C U f \ (0 \ C \ (\x \ U. \y\U. dist (f x) (f y) \ C * dist x y))" bundle lipschitz_syntax begin notation\<^marker>\tag important\ lipschitz_on ("_-lipschitz'_on" [1000]) end bundle no_lipschitz_syntax begin no_notation lipschitz_on ("_-lipschitz'_on" [1000]) end unbundle lipschitz_syntax lemma lipschitz_onI: "L-lipschitz_on X f" if "\x y. x \ X \ y \ X \ dist (f x) (f y) \ L * dist x y" "0 \ L" using that by (auto simp: lipschitz_on_def) lemma lipschitz_onD: "dist (f x) (f y) \ L * dist x y" if "L-lipschitz_on X f" "x \ X" "y \ X" using that by (auto simp: lipschitz_on_def) lemma lipschitz_on_nonneg: "0 \ L" if "L-lipschitz_on X f" using that by (auto simp: lipschitz_on_def) lemma lipschitz_on_normD: "norm (f x - f y) \ L * norm (x - y)" if "lipschitz_on L X f" "x \ X" "y \ X" using lipschitz_onD[OF that] by (simp add: dist_norm) lemma lipschitz_on_mono: "L-lipschitz_on D f" if "M-lipschitz_on E f" "D \ E" "M \ L" using that by (force simp: lipschitz_on_def intro: order_trans[OF _ mult_right_mono]) lemmas lipschitz_on_subset = lipschitz_on_mono[OF _ _ order_refl] and lipschitz_on_le = lipschitz_on_mono[OF _ order_refl] lemma lipschitz_on_leI: "L-lipschitz_on X f" if "\x y. x \ X \ y \ X \ x \ y \ dist (f x) (f y) \ L * dist x y" "0 \ L" for f::"'a::{linorder_topology, ordered_real_vector, metric_space} \ 'b::metric_space" proof (rule lipschitz_onI) fix x y assume xy: "x \ X" "y \ X" consider "y \ x" | "x \ y" by (rule le_cases) then show "dist (f x) (f y) \ L * dist x y" proof cases case 1 then have "dist (f y) (f x) \ L * dist y x" by (auto intro!: that xy) then show ?thesis by (simp add: dist_commute) qed (auto intro!: that xy) qed fact lemma lipschitz_on_concat: fixes a b c::real assumes f: "L-lipschitz_on {a .. b} f" assumes g: "L-lipschitz_on {b .. c} g" assumes fg: "f b = g b" shows "lipschitz_on L {a .. c} (\x. if x \ b then f x else g x)" (is "lipschitz_on _ _ ?f") proof (rule lipschitz_on_leI) fix x y assume x: "x \ {a..c}" and y: "y \ {a..c}" and xy: "x \ y" consider "x \ b \ b < y" | "x \ b \ y \ b" by arith then show "dist (?f x) (?f y) \ L * dist x y" proof cases case 1 have "dist (f x) (g y) \ dist (f x) (f b) + dist (g b) (g y)" unfolding fg by (rule dist_triangle) also have "dist (f x) (f b) \ L * dist x b" using 1 x by (auto intro!: lipschitz_onD[OF f]) also have "dist (g b) (g y) \ L * dist b y" using 1 x y by (auto intro!: lipschitz_onD[OF g] lipschitz_onD[OF f]) finally have "dist (f x) (g y) \ L * dist x b + L * dist b y" by simp also have "\ = L * (dist x b + dist b y)" by (simp add: algebra_simps) also have "dist x b + dist b y = dist x y" using 1 x y by (auto simp: dist_real_def abs_real_def) finally show ?thesis using 1 by simp next case 2 with lipschitz_onD[OF f, of x y] lipschitz_onD[OF g, of x y] x y xy show ?thesis by (auto simp: fg) qed qed (rule lipschitz_on_nonneg[OF f]) lemma lipschitz_on_concat_max: fixes a b c::real assumes f: "L-lipschitz_on {a .. b} f" assumes g: "M-lipschitz_on {b .. c} g" assumes fg: "f b = g b" shows "(max L M)-lipschitz_on {a .. c} (\x. if x \ b then f x else g x)" proof - have "lipschitz_on (max L M) {a .. b} f" "lipschitz_on (max L M) {b .. c} g" by (auto intro!: lipschitz_on_mono[OF f order_refl] lipschitz_on_mono[OF g order_refl]) from lipschitz_on_concat[OF this fg] show ?thesis . qed +text \Equivalence between "abstract" and "type class" Lipschitz notions, +for all types in the metric space class\ +lemma Lipschitz_map_euclidean [simp]: + "Lipschitz_continuous_map euclidean_metric euclidean_metric f + \ (\B. lipschitz_on B UNIV f)" (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (force simp add: Lipschitz_continuous_map_pos lipschitz_on_def less_le_not_le) + show "?rhs \ ?lhs" + by (metis Lipschitz_continuous_map_def lipschitz_onD mdist_euclidean_metric mspace_euclidean_metric top_greatest) +qed subsubsection \Continuity\ proposition lipschitz_on_uniformly_continuous: assumes "L-lipschitz_on X f" shows "uniformly_continuous_on X f" unfolding uniformly_continuous_on_def proof safe fix e::real assume "0 < e" from assms have l: "(L+1)-lipschitz_on X f" by (rule lipschitz_on_mono) auto show "\d>0. \x\X. \x'\X. dist x' x < d \ dist (f x') (f x) < e" using lipschitz_onD[OF l] lipschitz_on_nonneg[OF assms] \0 < e\ by (force intro!: exI[where x="e/(L + 1)"] simp: field_simps) qed proposition lipschitz_on_continuous_on: "continuous_on X f" if "L-lipschitz_on X f" by (rule uniformly_continuous_imp_continuous[OF lipschitz_on_uniformly_continuous[OF that]]) lemma lipschitz_on_continuous_within: "continuous (at x within X) f" if "L-lipschitz_on X f" "x \ X" using lipschitz_on_continuous_on[OF that(1)] that(2) by (auto simp: continuous_on_eq_continuous_within) subsubsection \Differentiable functions\ proposition bounded_derivative_imp_lipschitz: assumes "\x. x \ X \ (f has_derivative f' x) (at x within X)" assumes convex: "convex X" assumes "\x. x \ X \ onorm (f' x) \ C" "0 \ C" shows "C-lipschitz_on X f" proof (rule lipschitz_onI) show "\x y. x \ X \ y \ X \ dist (f x) (f y) \ C * dist x y" by (auto intro!: assms differentiable_bound[unfolded dist_norm[symmetric], OF convex]) qed fact subsubsection\<^marker>\tag unimportant\ \Structural introduction rules\ named_theorems lipschitz_intros "structural introduction rules for Lipschitz controls" lemma lipschitz_on_compose [lipschitz_intros]: "(D * C)-lipschitz_on U (g o f)" if f: "C-lipschitz_on U f" and g: "D-lipschitz_on (f`U) g" proof (rule lipschitz_onI) show "D* C \ 0" using lipschitz_on_nonneg[OF f] lipschitz_on_nonneg[OF g] by auto fix x y assume H: "x \ U" "y \ U" have "dist (g (f x)) (g (f y)) \ D * dist (f x) (f y)" apply (rule lipschitz_onD[OF g]) using H by auto also have "... \ D * C * dist x y" using mult_left_mono[OF lipschitz_onD(1)[OF f H] lipschitz_on_nonneg[OF g]] by auto finally show "dist ((g \ f) x) ((g \ f) y) \ D * C* dist x y" unfolding comp_def by (auto simp add: mult.commute) qed lemma lipschitz_on_compose2: "(D * C)-lipschitz_on U (\x. g (f x))" if "C-lipschitz_on U f" "D-lipschitz_on (f`U) g" using lipschitz_on_compose[OF that] by (simp add: o_def) lemma lipschitz_on_cong[cong]: "C-lipschitz_on U g \ D-lipschitz_on V f" if "C = D" "U = V" "\x. x \ V \ g x = f x" using that by (auto simp: lipschitz_on_def) lemma lipschitz_on_transform: "C-lipschitz_on U g" if "C-lipschitz_on U f" "\x. x \ U \ g x = f x" using that by simp lemma lipschitz_on_empty_iff[simp]: "C-lipschitz_on {} f \ C \ 0" by (auto simp: lipschitz_on_def) lemma lipschitz_on_insert_iff[simp]: "C-lipschitz_on (insert y X) f \ C-lipschitz_on X f \ (\x \ X. dist (f x) (f y) \ C * dist x y)" by (auto simp: lipschitz_on_def dist_commute) lemma lipschitz_on_singleton [lipschitz_intros]: "C \ 0 \ C-lipschitz_on {x} f" and lipschitz_on_empty [lipschitz_intros]: "C \ 0 \ C-lipschitz_on {} f" by simp_all lemma lipschitz_on_id [lipschitz_intros]: "1-lipschitz_on U (\x. x)" by (auto simp: lipschitz_on_def) lemma lipschitz_on_constant [lipschitz_intros]: "0-lipschitz_on U (\x. c)" by (auto simp: lipschitz_on_def) lemma lipschitz_on_add [lipschitz_intros]: fixes f::"'a::metric_space \'b::real_normed_vector" assumes "C-lipschitz_on U f" "D-lipschitz_on U g" shows "(C+D)-lipschitz_on U (\x. f x + g x)" proof (rule lipschitz_onI) show "C + D \ 0" using lipschitz_on_nonneg[OF assms(1)] lipschitz_on_nonneg[OF assms(2)] by auto fix x y assume H: "x \ U" "y \ U" have "dist (f x + g x) (f y + g y) \ dist (f x) (f y) + dist (g x) (g y)" by (simp add: dist_triangle_add) also have "... \ C * dist x y + D * dist x y" using lipschitz_onD(1)[OF assms(1) H] lipschitz_onD(1)[OF assms(2) H] by auto finally show "dist (f x + g x) (f y + g y) \ (C+D) * dist x y" by (auto simp add: algebra_simps) qed lemma lipschitz_on_cmult [lipschitz_intros]: fixes f::"'a::metric_space \ 'b::real_normed_vector" assumes "C-lipschitz_on U f" shows "(abs(a) * C)-lipschitz_on U (\x. a *\<^sub>R f x)" proof (rule lipschitz_onI) show "abs(a) * C \ 0" using lipschitz_on_nonneg[OF assms(1)] by auto fix x y assume H: "x \ U" "y \ U" have "dist (a *\<^sub>R f x) (a *\<^sub>R f y) = abs(a) * dist (f x) (f y)" by (metis dist_norm norm_scaleR real_vector.scale_right_diff_distrib) also have "... \ abs(a) * C * dist x y" using lipschitz_onD(1)[OF assms(1) H] by (simp add: Groups.mult_ac(1) mult_left_mono) finally show "dist (a *\<^sub>R f x) (a *\<^sub>R f y) \ \a\ * C * dist x y" by auto qed lemma lipschitz_on_cmult_real [lipschitz_intros]: fixes f::"'a::metric_space \ real" assumes "C-lipschitz_on U f" shows "(abs(a) * C)-lipschitz_on U (\x. a * f x)" using lipschitz_on_cmult[OF assms] by auto lemma lipschitz_on_cmult_nonneg [lipschitz_intros]: fixes f::"'a::metric_space \ 'b::real_normed_vector" assumes "C-lipschitz_on U f" "a \ 0" shows "(a * C)-lipschitz_on U (\x. a *\<^sub>R f x)" using lipschitz_on_cmult[OF assms(1), of a] assms(2) by auto lemma lipschitz_on_cmult_real_nonneg [lipschitz_intros]: fixes f::"'a::metric_space \ real" assumes "C-lipschitz_on U f" "a \ 0" shows "(a * C)-lipschitz_on U (\x. a * f x)" using lipschitz_on_cmult_nonneg[OF assms] by auto lemma lipschitz_on_cmult_upper [lipschitz_intros]: fixes f::"'a::metric_space \ 'b::real_normed_vector" assumes "C-lipschitz_on U f" "abs(a) \ D" shows "(D * C)-lipschitz_on U (\x. a *\<^sub>R f x)" apply (rule lipschitz_on_mono[OF lipschitz_on_cmult[OF assms(1), of a], of _ "D * C"]) using assms(2) lipschitz_on_nonneg[OF assms(1)] mult_right_mono by auto lemma lipschitz_on_cmult_real_upper [lipschitz_intros]: fixes f::"'a::metric_space \ real" assumes "C-lipschitz_on U f" "abs(a) \ D" shows "(D * C)-lipschitz_on U (\x. a * f x)" using lipschitz_on_cmult_upper[OF assms] by auto lemma lipschitz_on_minus[lipschitz_intros]: fixes f::"'a::metric_space \'b::real_normed_vector" assumes "C-lipschitz_on U f" shows "C-lipschitz_on U (\x. - f x)" by (metis (mono_tags, lifting) assms dist_minus lipschitz_on_def) lemma lipschitz_on_minus_iff[simp]: "L-lipschitz_on X (\x. - f x) \ L-lipschitz_on X f" "L-lipschitz_on X (- f) \ L-lipschitz_on X f" for f::"'a::metric_space \'b::real_normed_vector" using lipschitz_on_minus[of L X f] lipschitz_on_minus[of L X "-f"] by auto lemma lipschitz_on_diff[lipschitz_intros]: fixes f::"'a::metric_space \'b::real_normed_vector" assumes "C-lipschitz_on U f" "D-lipschitz_on U g" shows "(C + D)-lipschitz_on U (\x. f x - g x)" using lipschitz_on_add[OF assms(1) lipschitz_on_minus[OF assms(2)]] by auto lemma lipschitz_on_closure [lipschitz_intros]: assumes "C-lipschitz_on U f" "continuous_on (closure U) f" shows "C-lipschitz_on (closure U) f" proof (rule lipschitz_onI) show "C \ 0" using lipschitz_on_nonneg[OF assms(1)] by simp fix x y assume "x \ closure U" "y \ closure U" obtain u v::"nat \ 'a" where *: "\n. u n \ U" "u \ x" "\n. v n \ U" "v \ y" using \x \ closure U\ \y \ closure U\ unfolding closure_sequential by blast have a: "(\n. f (u n)) \ f x" using *(1) *(2) \x \ closure U\ \continuous_on (closure U) f\ unfolding comp_def continuous_on_closure_sequentially[of U f] by auto have b: "(\n. f (v n)) \ f y" using *(3) *(4) \y \ closure U\ \continuous_on (closure U) f\ unfolding comp_def continuous_on_closure_sequentially[of U f] by auto have l: "(\n. C * dist (u n) (v n) - dist (f (u n)) (f (v n))) \ C * dist x y - dist (f x) (f y)" by (intro tendsto_intros * a b) have "C * dist (u n) (v n) - dist (f (u n)) (f (v n)) \ 0" for n using lipschitz_onD(1)[OF assms(1) \u n \ U\ \v n \ U\] by simp then have "C * dist x y - dist (f x) (f y) \ 0" using LIMSEQ_le_const[OF l, of 0] by auto then show "dist (f x) (f y) \ C * dist x y" by auto qed lemma lipschitz_on_Pair[lipschitz_intros]: assumes f: "L-lipschitz_on A f" assumes g: "M-lipschitz_on A g" shows "(sqrt (L\<^sup>2 + M\<^sup>2))-lipschitz_on A (\a. (f a, g a))" proof (rule lipschitz_onI, goal_cases) case (1 x y) have "dist (f x, g x) (f y, g y) = sqrt ((dist (f x) (f y))\<^sup>2 + (dist (g x) (g y))\<^sup>2)" by (auto simp add: dist_Pair_Pair real_le_lsqrt) also have "\ \ sqrt ((L * dist x y)\<^sup>2 + (M * dist x y)\<^sup>2)" by (auto intro!: real_sqrt_le_mono add_mono power_mono 1 lipschitz_onD f g) also have "\ \ sqrt (L\<^sup>2 + M\<^sup>2) * dist x y" by (auto simp: power_mult_distrib ring_distribs[symmetric] real_sqrt_mult) finally show ?case . qed simp lemma lipschitz_extend_closure: fixes f::"('a::metric_space) \ ('b::complete_space)" assumes "C-lipschitz_on U f" shows "\g. C-lipschitz_on (closure U) g \ (\x\U. g x = f x)" proof - obtain g where g: "\x. x \ U \ g x = f x" "uniformly_continuous_on (closure U) g" using uniformly_continuous_on_extension_on_closure[OF lipschitz_on_uniformly_continuous[OF assms]] by metis have "C-lipschitz_on (closure U) g" apply (rule lipschitz_on_closure, rule lipschitz_on_transform[OF assms]) using g uniformly_continuous_imp_continuous[OF g(2)] by auto then show ?thesis using g(1) by auto qed lemma (in bounded_linear) lipschitz_boundE: obtains B where "B-lipschitz_on A f" proof - from nonneg_bounded obtain B where B: "B \ 0" "\x. norm (f x) \ B * norm x" by (auto simp: ac_simps) have "B-lipschitz_on A f" by (auto intro!: lipschitz_onI B simp: dist_norm diff[symmetric]) thus ?thesis .. qed subsection \Local Lipschitz continuity\ text \Given a function defined on a real interval, it is Lipschitz-continuous if and only if it is locally so, as proved in the following lemmas. It is useful especially for piecewise-defined functions: if each piece is Lipschitz, then so is the whole function. The same goes for functions defined on geodesic spaces, or more generally on geodesic subsets in a metric space (for instance convex subsets in a real vector space), and this follows readily from the real case, but we will not prove it explicitly. We give several variations around this statement. This is essentially a connectedness argument.\ lemma locally_lipschitz_imp_lipschitz_aux: fixes f::"real \ ('a::metric_space)" assumes "a \ b" "continuous_on {a..b} f" "\x. x \ {a.. \y \ {x<..b}. dist (f y) (f x) \ M * (y-x)" shows "dist (f b) (f a) \ M * (b-a)" proof - define A where "A = {x \ {a..b}. dist (f x) (f a) \ M * (x-a)}" have *: "A = (\x. M * (x-a) - dist (f x) (f a))-`{0..} \ {a..b}" unfolding A_def by auto have "a \ A" unfolding A_def using \a \ b\ by auto then have "A \ {}" by auto moreover have "bdd_above A" unfolding A_def by auto moreover have "closed A" unfolding * by (rule closed_vimage_Int, auto intro!: continuous_intros assms) ultimately have "Sup A \ A" by (rule closed_contains_Sup) have "Sup A = b" proof (rule ccontr) assume "Sup A \ b" define x where "x = Sup A" have I: "dist (f x) (f a) \ M * (x-a)" using \Sup A \ A\ x_def A_def by auto have "x \ {a..Sup A \ A\ \Sup A \ b\ A_def by auto then obtain y where J: "y \ {x<..b}" "dist (f y) (f x) \ M * (y-x)" using assms(3) by blast have "dist (f y) (f a) \ dist (f y) (f x) + dist (f x) (f a)" by (rule dist_triangle) also have "... \ M * (y-x) + M * (x-a)" using I J(2) by auto finally have "dist (f y) (f a) \ M * (y-a)" by (auto simp add: algebra_simps) then have "y \ A" unfolding A_def using \y \ {x<..b}\ \x \ {a.. by auto then have "y \ Sup A" by (rule cSup_upper, auto simp: A_def) then show False using \y \ {x<..b}\ x_def by auto qed then show ?thesis using \Sup A \ A\ A_def by auto qed lemma locally_lipschitz_imp_lipschitz: fixes f::"real \ ('a::metric_space)" assumes "continuous_on {a..b} f" "\x y. x \ {a.. y > x \ \z \ {x<..y}. dist (f z) (f x) \ M * (z-x)" "M \ 0" shows "lipschitz_on M {a..b} f" proof (rule lipschitz_onI[OF _ \M \ 0\]) have *: "dist (f t) (f s) \ M * (t-s)" if "s \ t" "s \ {a..b}" "t \ {a..b}" for s t proof (rule locally_lipschitz_imp_lipschitz_aux, simp add: \s \ t\) show "continuous_on {s..t} f" using continuous_on_subset[OF assms(1)] that by auto fix x assume "x \ {s.. {a..z\{x<..t}. dist (f z) (f x) \ M * (z - x)" using assms(2)[OF \x \ {a.., of t] \x \ {s.. by auto qed fix x y assume "x \ {a..b}" "y \ {a..b}" consider "x \ y" | "y \ x" by linarith then show "dist (f x) (f y) \ M * dist x y" apply (cases) using *[OF _ \x \ {a..b}\ \y \ {a..b}\] *[OF _ \y \ {a..b}\ \x \ {a..b}\] by (auto simp add: dist_commute dist_real_def) qed text \We deduce that if a function is Lipschitz on finitely many closed sets on the real line, then it is Lipschitz on any interval contained in their union. The difficulty in the proof is to show that any point \z\ in this interval (except the maximum) has a point arbitrarily close to it on its right which is contained in a common initial closed set. Otherwise, we show that there is a small interval \(z, T)\ which does not intersect any of the initial closed sets, a contradiction.\ proposition lipschitz_on_closed_Union: assumes "\i. i \ I \ lipschitz_on M (U i) f" "\i. i \ I \ closed (U i)" "finite I" "M \ 0" "{u..(v::real)} \ (\i\I. U i)" shows "lipschitz_on M {u..v} f" proof (rule locally_lipschitz_imp_lipschitz[OF _ _ \M \ 0\]) have *: "continuous_on (U i) f" if "i \ I" for i by (rule lipschitz_on_continuous_on[OF assms(1)[OF \i\ I\]]) have "continuous_on (\i\I. U i) f" apply (rule continuous_on_closed_Union) using \finite I\ * assms(2) by auto then show "continuous_on {u..v} f" using \{u..(v::real)} \ (\i\I. U i)\ continuous_on_subset by auto fix z Z assume z: "z \ {u.. v" by auto define T where "T = min Z v" then have T: "T > z" "T \ v" "T \ u" "T \ Z" using z by auto define A where "A = (\i\ I \ {i. U i \ {z<..T} \ {}}. U i \ {z..T})" have a: "closed A" unfolding A_def apply (rule closed_UN) using \finite I\ \\i. i \ I \ closed (U i)\ by auto have b: "bdd_below A" unfolding A_def using \finite I\ by auto have "\i \ I. T \ U i" using \{u..v} \ (\i\I. U i)\ T by auto then have c: "T \ A" unfolding A_def using T by (auto, fastforce) have "Inf A \ z" apply (rule cInf_greatest, auto) using c unfolding A_def by auto moreover have "Inf A \ z" proof (rule ccontr) assume "\(Inf A \ z)" then obtain w where w: "w > z" "w < Inf A" by (meson dense not_le_imp_less) have "Inf A \ T" using a b c by (simp add: cInf_lower) then have "w \ T" using w by auto then have "w \ {u..v}" using w \z \ {u.. T by auto then obtain j where j: "j \ I" "w \ U j" using \{u..v} \ (\i\I. U i)\ by fastforce then have "w \ U j \ {z..T}" "U j \ {z<..T} \ {}" using j T w \w \ T\ by auto then have "w \ A" unfolding A_def using \j \ I\ by auto then have "Inf A \ w" using a b by (simp add: cInf_lower) then show False using w by auto qed ultimately have "Inf A = z" by simp moreover have "Inf A \ A" apply (rule closed_contains_Inf) using a b c by auto ultimately have "z \ A" by simp then obtain i where i: "i \ I" "U i \ {z<..T} \ {}" "z \ U i" unfolding A_def by auto then obtain t where "t \ U i \ {z<..T}" by blast then have "dist (f t) (f z) \ M * (t - z)" using lipschitz_onD(1)[OF assms(1)[of i], of t z] i dist_real_def by auto then show "\t\{z<..Z}. dist (f t) (f z) \ M * (t - z)" using \T \ Z\ \t \ U i \ {z<..T}\ by auto qed subsection \Local Lipschitz continuity (uniform for a family of functions)\ definition\<^marker>\tag important\ local_lipschitz:: "'a::metric_space set \ 'b::metric_space set \ ('a \ 'b \ 'c::metric_space) \ bool" where "local_lipschitz T X f \ \x \ X. \t \ T. \u>0. \L. \t \ cball t u \ T. L-lipschitz_on (cball x u \ X) (f t)" lemma local_lipschitzI: assumes "\t x. t \ T \ x \ X \ \u>0. \L. \t \ cball t u \ T. L-lipschitz_on (cball x u \ X) (f t)" shows "local_lipschitz T X f" using assms unfolding local_lipschitz_def by auto lemma local_lipschitzE: assumes local_lipschitz: "local_lipschitz T X f" assumes "t \ T" "x \ X" obtains u L where "u > 0" "\s. s \ cball t u \ T \ L-lipschitz_on (cball x u \ X) (f s)" using assms local_lipschitz_def by metis lemma local_lipschitz_continuous_on: assumes local_lipschitz: "local_lipschitz T X f" assumes "t \ T" shows "continuous_on X (f t)" unfolding continuous_on_def proof safe fix x assume "x \ X" from local_lipschitzE[OF local_lipschitz \t \ T\ \x \ X\] obtain u L where "0 < u" and L: "\s. s \ cball t u \ T \ L-lipschitz_on (cball x u \ X) (f s)" by metis have "x \ ball x u" using \0 < u\ by simp from lipschitz_on_continuous_on[OF L] have tendsto: "(f t \ f t x) (at x within cball x u \ X)" using \0 < u\ \x \ X\ \t \ T\ by (auto simp: continuous_on_def) moreover have "\\<^sub>F xa in at x. (xa \ cball x u \ X) = (xa \ X)" using eventually_at_ball[OF \0 < u\, of x UNIV] by eventually_elim auto ultimately show "(f t \ f t x) (at x within X)" by (rule Lim_transform_within_set) qed lemma local_lipschitz_compose1: assumes ll: "local_lipschitz (g ` T) X (\t. f t)" assumes g: "continuous_on T g" shows "local_lipschitz T X (\t. f (g t))" proof (rule local_lipschitzI) fix t x assume "t \ T" "x \ X" then have "g t \ g ` T" by simp from local_lipschitzE[OF assms(1) this \x \ X\] obtain u L where "0 < u" and l: "(\s. s \ cball (g t) u \ g ` T \ L-lipschitz_on (cball x u \ X) (f s))" by auto from g[unfolded continuous_on_eq_continuous_within, rule_format, OF \t \ T\, unfolded continuous_within_eps_delta, rule_format, OF \0 < u\] obtain d where d: "d>0" "\x'. x'\T \ dist x' t < d \ dist (g x') (g t) < u" by (auto) show "\u>0. \L. \t\cball t u \ T. L-lipschitz_on (cball x u \ X) (f (g t))" using d \0 < u\ by (fastforce intro: exI[where x="(min d u)/2"] exI[where x=L] intro!: less_imp_le[OF d(2)] lipschitz_on_subset[OF l] simp: dist_commute) qed context fixes T::"'a::metric_space set" and X f assumes local_lipschitz: "local_lipschitz T X f" begin lemma continuous_on_TimesI: assumes y: "\x. x \ X \ continuous_on T (\t. f t x)" shows "continuous_on (T \ X) (\(t, x). f t x)" unfolding continuous_on_iff proof (safe, simp) fix a b and e::real assume H: "a \ T" "b \ X" "0 < e" hence "0 < e/2" by simp from y[unfolded continuous_on_iff, OF \b \ X\, rule_format, OF \a \ T\ \0 < e/2\] obtain d where d: "d > 0" "\t. t \ T \ dist t a < d \ dist (f t b) (f a b) < e/2" by auto from \a : T\ \b \ X\ obtain u L where u: "0 < u" and L: "\t. t \ cball a u \ T \ L-lipschitz_on (cball b u \ X) (f t)" by (erule local_lipschitzE[OF local_lipschitz]) have "a \ cball a u \ T" by (auto simp: \0 < u\ \a \ T\ less_imp_le) from lipschitz_on_nonneg[OF L[OF \a \ cball _ _ \ _\]] have "0 \ L" . let ?d = "Min {d, u, (e/2/(L + 1))}" show "\d>0. \x\T. \y\X. dist (x, y) (a, b) < d \ dist (f x y) (f a b) < e" proof (rule exI[where x = ?d], safe) show "0 < ?d" using \0 \ L\ \0 < u\ \0 < e\ \0 < d\ by (auto intro!: divide_pos_pos ) fix x y assume "x \ T" "y \ X" assume dist_less: "dist (x, y) (a, b) < ?d" have "dist y b \ dist (x, y) (a, b)" using dist_snd_le[of "(x, y)" "(a, b)"] by auto also note dist_less also { note calculation also have "?d \ u" by simp finally have "dist y b < u" . } have "?d \ e/2/(L + 1)" by simp also have "(L + 1) * \ \ e / 2" using \0 < e\ \L \ 0\ by (auto simp: field_split_simps) finally have le1: "(L + 1) * dist y b < e / 2" using \L \ 0\ by simp have "dist x a \ dist (x, y) (a, b)" using dist_fst_le[of "(x, y)" "(a, b)"] by auto also note dist_less finally have "dist x a < ?d" . also have "?d \ d" by simp finally have "dist x a < d" . note \dist x a < ?d\ also have "?d \ u" by simp finally have "dist x a < u" . then have "x \ cball a u \ T" using \x \ T\ by (auto simp: dist_commute) have "dist (f x y) (f a b) \ dist (f x y) (f x b) + dist (f x b) (f a b)" by (rule dist_triangle) also have "(L + 1)-lipschitz_on (cball b u \ X) (f x)" using L[OF \x \ cball a u \ T\] by (rule lipschitz_on_le) simp then have "dist (f x y) (f x b) \ (L + 1) * dist y b" apply (rule lipschitz_onD) subgoal using \y \ X\ \dist y b < u\ by (simp add: dist_commute) subgoal using \0 < u\ \b \ X\ by simp done also have "(L + 1) * dist y b \ e / 2" using le1 \0 \ L\ by simp also have "dist (f x b) (f a b) < e / 2" by (rule d; fact) also have "e / 2 + e / 2 = e" by simp finally show "dist (f x y) (f a b) < e" by simp qed qed lemma local_lipschitz_compact_implies_lipschitz: assumes "compact X" "compact T" assumes cont: "\x. x \ X \ continuous_on T (\t. f t x)" obtains L where "\t. t \ T \ L-lipschitz_on X (f t)" proof - { assume *: "\n::nat. \(\t\T. n-lipschitz_on X (f t))" { fix n::nat from *[of n] have "\x y t. t \ T \ x \ X \ y \ X \ dist (f t y) (f t x) > n * dist y x" by (force simp: lipschitz_on_def) } then obtain t and x y::"nat \ 'b" where xy: "\n. x n \ X" "\n. y n \ X" and t: "\n. t n \ T" and d: "\n. dist (f (t n) (y n)) (f (t n) (x n)) > n * dist (y n) (x n)" by metis from xy assms obtain lx rx where lx': "lx \ X" "strict_mono (rx :: nat \ nat)" "(x o rx) \ lx" by (metis compact_def) with xy have "\n. (y o rx) n \ X" by auto with assms obtain ly ry where ly': "ly \ X" "strict_mono (ry :: nat \ nat)" "((y o rx) o ry) \ ly" by (metis compact_def) with t have "\n. ((t o rx) o ry) n \ T" by simp with assms obtain lt rt where lt': "lt \ T" "strict_mono (rt :: nat \ nat)" "(((t o rx) o ry) o rt) \ lt" by (metis compact_def) from lx' ly' have lx: "(x o (rx o ry o rt)) \ lx" (is "?x \ _") and ly: "(y o (rx o ry o rt)) \ ly" (is "?y \ _") and lt: "(t o (rx o ry o rt)) \ lt" (is "?t \ _") subgoal by (simp add: LIMSEQ_subseq_LIMSEQ o_assoc lt'(2)) subgoal by (simp add: LIMSEQ_subseq_LIMSEQ ly'(3) o_assoc lt'(2)) subgoal by (simp add: o_assoc lt'(3)) done hence "(\n. dist (?y n) (?x n)) \ dist ly lx" by (metis tendsto_dist) moreover let ?S = "(\(t, x). f t x) ` (T \ X)" have "eventually (\n::nat. n > 0) sequentially" by (metis eventually_at_top_dense) hence "eventually (\n. norm (dist (?y n) (?x n)) \ norm (\diameter ?S\ / n) * 1) sequentially" proof eventually_elim case (elim n) have "0 < rx (ry (rt n))" using \0 < n\ by (metis dual_order.strict_trans1 lt'(2) lx'(2) ly'(2) seq_suble) have compact: "compact ?S" by (auto intro!: compact_continuous_image continuous_on_subset[OF continuous_on_TimesI] compact_Times \compact X\ \compact T\ cont) have "norm (dist (?y n) (?x n)) = dist (?y n) (?x n)" by simp also from this elim d[of "rx (ry (rt n))"] have "\ < dist (f (?t n) (?y n)) (f (?t n) (?x n)) / rx (ry (rt (n)))" using lx'(2) ly'(2) lt'(2) \0 < rx _\ by (auto simp add: field_split_simps strict_mono_def) also have "\ \ diameter ?S / n" proof (rule frac_le) show "diameter ?S \ 0" using compact compact_imp_bounded diameter_ge_0 by blast show "dist (f (?t n) (?y n)) (f (?t n) (?x n)) \ diameter ((\(t,x). f t x) ` (T \ X))" by (metis (no_types) compact compact_imp_bounded diameter_bounded_bound image_eqI mem_Sigma_iff o_apply split_conv t xy(1) xy(2)) show "real n \ real (rx (ry (rt n)))" by (meson le_trans lt'(2) lx'(2) ly'(2) of_nat_mono strict_mono_imp_increasing) qed (use \n > 0\ in auto) also have "\ \ abs (diameter ?S) / n" by (auto intro!: divide_right_mono) finally show ?case by simp qed with _ have "(\n. dist (?y n) (?x n)) \ 0" by (rule tendsto_0_le) (metis tendsto_divide_0[OF tendsto_const] filterlim_at_top_imp_at_infinity filterlim_real_sequentially) ultimately have "lx = ly" using LIMSEQ_unique by fastforce with assms lx' have "lx \ X" by auto from \lt \ T\ this obtain u L where L: "u > 0" "\t. t \ cball lt u \ T \ L-lipschitz_on (cball lx u \ X) (f t)" by (erule local_lipschitzE[OF local_lipschitz]) hence "L \ 0" by (force intro!: lipschitz_on_nonneg \lt \ T\) from L lt ly lx \lx = ly\ have "eventually (\n. ?t n \ ball lt u) sequentially" "eventually (\n. ?y n \ ball lx u) sequentially" "eventually (\n. ?x n \ ball lx u) sequentially" by (auto simp: dist_commute Lim) moreover have "eventually (\n. n > L) sequentially" by (metis filterlim_at_top_dense filterlim_real_sequentially) ultimately have "eventually (\_. False) sequentially" proof eventually_elim case (elim n) hence "dist (f (?t n) (?y n)) (f (?t n) (?x n)) \ L * dist (?y n) (?x n)" using assms xy t unfolding dist_norm[symmetric] by (intro lipschitz_onD[OF L(2)]) (auto) also have "\ \ n * dist (?y n) (?x n)" using elim by (intro mult_right_mono) auto also have "\ \ rx (ry (rt n)) * dist (?y n) (?x n)" by (intro mult_right_mono[OF _ zero_le_dist]) (meson lt'(2) lx'(2) ly'(2) of_nat_le_iff order_trans seq_suble) also have "\ < dist (f (?t n) (?y n)) (f (?t n) (?x n))" by (auto intro!: d) finally show ?case by simp qed hence False by simp } then obtain L where "\t. t \ T \ L-lipschitz_on X (f t)" by metis thus ?thesis .. qed lemma local_lipschitz_subset: assumes "S \ T" "Y \ X" shows "local_lipschitz S Y f" proof (rule local_lipschitzI) fix t x assume "t \ S" "x \ Y" then have "t \ T" "x \ X" using assms by auto from local_lipschitzE[OF local_lipschitz, OF this] obtain u L where u: "0 < u" and L: "\s. s \ cball t u \ T \ L-lipschitz_on (cball x u \ X) (f s)" by blast show "\u>0. \L. \t\cball t u \ S. L-lipschitz_on (cball x u \ Y) (f t)" using assms by (auto intro: exI[where x=u] exI[where x=L] intro!: u lipschitz_on_subset[OF _ Int_mono[OF order_refl \Y \ X\]] L) qed end lemma local_lipschitz_minus: fixes f::"'a::metric_space \ 'b::metric_space \ 'c::real_normed_vector" shows "local_lipschitz T X (\t x. - f t x) = local_lipschitz T X f" by (auto simp: local_lipschitz_def lipschitz_on_minus) lemma local_lipschitz_PairI: assumes f: "local_lipschitz A B (\a b. f a b)" assumes g: "local_lipschitz A B (\a b. g a b)" shows "local_lipschitz A B (\a b. (f a b, g a b))" proof (rule local_lipschitzI) fix t x assume "t \ A" "x \ B" from local_lipschitzE[OF f this] local_lipschitzE[OF g this] obtain u L v M where "0 < u" "(\s. s \ cball t u \ A \ L-lipschitz_on (cball x u \ B) (f s))" "0 < v" "(\s. s \ cball t v \ A \ M-lipschitz_on (cball x v \ B) (g s))" by metis then show "\u>0. \L. \t\cball t u \ A. L-lipschitz_on (cball x u \ B) (\b. (f t b, g t b))" by (intro exI[where x="min u v"]) (force intro: lipschitz_on_subset intro!: lipschitz_on_Pair) qed lemma local_lipschitz_constI: "local_lipschitz S T (\t x. f t)" by (auto simp: intro!: local_lipschitzI lipschitz_on_constant intro: exI[where x=1]) lemma (in bounded_linear) local_lipschitzI: shows "local_lipschitz A B (\_. f)" proof (rule local_lipschitzI, goal_cases) case (1 t x) from lipschitz_boundE[of "(cball x 1 \ B)"] obtain C where "C-lipschitz_on (cball x 1 \ B) f" by auto then show ?case by (auto intro: exI[where x=1]) qed proposition c1_implies_local_lipschitz: fixes T::"real set" and X::"'a::{banach,heine_borel} set" and f::"real \ 'a \ 'a" assumes f': "\t x. t \ T \ x \ X \ (f t has_derivative blinfun_apply (f' (t, x))) (at x)" assumes cont_f': "continuous_on (T \ X) f'" assumes "open T" assumes "open X" shows "local_lipschitz T X f" proof (rule local_lipschitzI) fix t x assume "t \ T" "x \ X" from open_contains_cball[THEN iffD1, OF \open X\, rule_format, OF \x \ X\] obtain u where u: "u > 0" "cball x u \ X" by auto moreover from open_contains_cball[THEN iffD1, OF \open T\, rule_format, OF \t \ T\] obtain v where v: "v > 0" "cball t v \ T" by auto ultimately have "compact (cball t v \ cball x u)" "cball t v \ cball x u \ T \ X" by (auto intro!: compact_Times) then have "compact (f' ` (cball t v \ cball x u))" by (auto intro!: compact_continuous_image continuous_on_subset[OF cont_f']) then obtain B where B: "B > 0" "\s y. s \ cball t v \ y \ cball x u \ norm (f' (s, y)) \ B" by (auto dest!: compact_imp_bounded simp: bounded_pos) have lipschitz: "B-lipschitz_on (cball x (min u v) \ X) (f s)" if s: "s \ cball t v" for s proof - note s also note \cball t v \ T\ finally have deriv: "\y. y \ cball x u \ (f s has_derivative blinfun_apply (f' (s, y))) (at y within cball x u)" using \_ \ X\ by (auto intro!: has_derivative_at_withinI[OF f']) have "norm (f s y - f s z) \ B * norm (y - z)" if "y \ cball x u" "z \ cball x u" for y z using s that by (intro differentiable_bound[OF convex_cball deriv]) (auto intro!: B simp: norm_blinfun.rep_eq[symmetric]) then show ?thesis using \0 < B\ by (auto intro!: lipschitz_onI simp: dist_norm) qed show "\u>0. \L. \t\cball t u \ T. L-lipschitz_on (cball x u \ X) (f t)" by (force intro: exI[where x="min u v"] exI[where x=B] intro!: lipschitz simp: u v) qed end diff --git a/src/HOL/Homology/Invariance_of_Domain.thy b/src/HOL/Homology/Invariance_of_Domain.thy --- a/src/HOL/Homology/Invariance_of_Domain.thy +++ b/src/HOL/Homology/Invariance_of_Domain.thy @@ -1,3032 +1,3033 @@ section\Invariance of Domain\ theory Invariance_of_Domain imports Brouwer_Degree "HOL-Analysis.Continuous_Extension" "HOL-Analysis.Homeomorphism" begin subsection\Degree invariance mod 2 for map between pairs\ theorem Borsuk_odd_mapping_degree_step: assumes cmf: "continuous_map (nsphere n) (nsphere n) f" and f: "\x. x \ topspace(nsphere n) \ (f \ (\x i. -x i)) x = ((\x i. -x i) \ f) x" and fim: "f ` (topspace(nsphere(n - Suc 0))) \ topspace(nsphere(n - Suc 0))" shows "even (Brouwer_degree2 n f - Brouwer_degree2 (n - Suc 0) f)" proof (cases "n = 0") case False define neg where "neg \ \x::nat\real. \i. -x i" define upper where "upper \ \n. {x::nat\real. x n \ 0}" define lower where "lower \ \n. {x::nat\real. x n \ 0}" define equator where "equator \ \n. {x::nat\real. x n = 0}" define usphere where "usphere \ \n. subtopology (nsphere n) (upper n)" define lsphere where "lsphere \ \n. subtopology (nsphere n) (lower n)" have [simp]: "neg x i = -x i" for x i by (force simp: neg_def) have equator_upper: "equator n \ upper n" by (force simp: equator_def upper_def) have upper_usphere: "subtopology (nsphere n) (upper n) = usphere n" by (simp add: usphere_def) let ?rhgn = "relative_homology_group n (nsphere n)" let ?hi_ee = "hom_induced n (nsphere n) (equator n) (nsphere n) (equator n)" interpret GE: comm_group "?rhgn (equator n)" by simp interpret HB: group_hom "?rhgn (equator n)" "homology_group (int n - 1) (subtopology (nsphere n) (equator n))" "hom_boundary n (nsphere n) (equator n)" by (simp add: group_hom_axioms_def group_hom_def hom_boundary_hom) interpret HIU: group_hom "?rhgn (equator n)" "?rhgn (upper n)" "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id" by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom) have subt_eq: "subtopology (nsphere n) {x. x n = 0} = nsphere (n - Suc 0)" by (metis False Suc_pred le_zero_eq not_le subtopology_nsphere_equator) then have equ: "subtopology (nsphere n) (equator n) = nsphere(n - Suc 0)" "subtopology (lsphere n) (equator n) = nsphere(n - Suc 0)" "subtopology (usphere n) (equator n) = nsphere(n - Suc 0)" using False by (auto simp: lsphere_def usphere_def equator_def lower_def upper_def subtopology_subtopology simp flip: Collect_conj_eq cong: rev_conj_cong) have cmr: "continuous_map (nsphere(n - Suc 0)) (nsphere(n - Suc 0)) f" by (metis (no_types, lifting) IntE cmf fim continuous_map_from_subtopology continuous_map_in_subtopology equ(1) image_subset_iff topspace_subtopology) have "f x n = 0" if "x \ topspace (nsphere n)" "x n = 0" for x proof - have "x \ topspace (nsphere (n - Suc 0))" by (simp add: that topspace_nsphere_minus1) moreover have "topspace (nsphere n) \ {f. f n = 0} = topspace (nsphere (n - Suc 0))" by (metis subt_eq topspace_subtopology) ultimately show ?thesis using cmr continuous_map_def by fastforce qed then have fimeq: "f ` (topspace (nsphere n) \ equator n) \ topspace (nsphere n) \ equator n" using fim cmf by (auto simp: equator_def continuous_map_def image_subset_iff) have "\k. continuous_map (powertop_real UNIV) euclideanreal (\x. - x k)" by (metis UNIV_I continuous_map_product_projection continuous_map_minus) then have cm_neg: "continuous_map (nsphere m) (nsphere m) neg" for m by (force simp: nsphere continuous_map_in_subtopology neg_def continuous_map_componentwise_UNIV intro: continuous_map_from_subtopology) then have cm_neg_lu: "continuous_map (lsphere n) (usphere n) neg" by (auto simp: lsphere_def usphere_def lower_def upper_def continuous_map_from_subtopology continuous_map_in_subtopology) have neg_in_top_iff: "neg x \ topspace(nsphere m) \ x \ topspace(nsphere m)" for m x by (simp add: nsphere_def neg_def topspace_Euclidean_space) obtain z where zcarr: "z \ carrier (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))" and zeq: "subgroup_generated (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) {z} = reduced_homology_group (int n - 1) (nsphere (n - Suc 0))" using cyclic_reduced_homology_group_nsphere [of "int n - 1" "n - Suc 0"] by (auto simp: cyclic_group_def) have "hom_boundary n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0} \ Group.iso (relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}) (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))" using iso_lower_hemisphere_reduced_homology_group [of "int n - 1" "n - Suc 0"] False by simp then obtain gp where g: "group_isomorphisms (relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}) (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) (hom_boundary n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}) gp" by (auto simp: group.iso_iff_group_isomorphisms) then interpret gp: group_hom "reduced_homology_group (int n - 1) (nsphere (n - Suc 0))" "relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}" gp by (simp add: group_hom_axioms_def group_hom_def group_isomorphisms_def) obtain zp where zpcarr: "zp \ carrier(relative_homology_group n (lsphere n) (equator n))" and zp_z: "hom_boundary n (lsphere n) (equator n) zp = z" and zp_sg: "subgroup_generated (relative_homology_group n (lsphere n) (equator n)) {zp} = relative_homology_group n (lsphere n) (equator n)" proof show "gp z \ carrier (relative_homology_group n (lsphere n) (equator n))" "hom_boundary n (lsphere n) (equator n) (gp z) = z" using g zcarr by (auto simp: lsphere_def equator_def lower_def group_isomorphisms_def) have giso: "gp \ Group.iso (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) (relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0})" by (metis (mono_tags, lifting) g group_isomorphisms_imp_iso group_isomorphisms_sym) show "subgroup_generated (relative_homology_group n (lsphere n) (equator n)) {gp z} = relative_homology_group n (lsphere n) (equator n)" apply (rule monoid.equality) using giso gp.subgroup_generated_by_image [of "{z}"] zcarr by (auto simp: lsphere_def equator_def lower_def zeq gp.iso_iff) qed have hb_iso: "hom_boundary n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0} \ iso (relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}) (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))" using iso_upper_hemisphere_reduced_homology_group [of "int n - 1" "n - Suc 0"] False by simp then obtain gn where g: "group_isomorphisms (relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}) (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) (hom_boundary n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}) gn" by (auto simp: group.iso_iff_group_isomorphisms) then interpret gn: group_hom "reduced_homology_group (int n - 1) (nsphere (n - Suc 0))" "relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0}" gn by (simp add: group_hom_axioms_def group_hom_def group_isomorphisms_def) obtain zn where zncarr: "zn \ carrier(relative_homology_group n (usphere n) (equator n))" and zn_z: "hom_boundary n (usphere n) (equator n) zn = z" and zn_sg: "subgroup_generated (relative_homology_group n (usphere n) (equator n)) {zn} = relative_homology_group n (usphere n) (equator n)" proof show "gn z \ carrier (relative_homology_group n (usphere n) (equator n))" "hom_boundary n (usphere n) (equator n) (gn z) = z" using g zcarr by (auto simp: usphere_def equator_def upper_def group_isomorphisms_def) have giso: "gn \ Group.iso (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) (relative_homology_group n (subtopology (nsphere n) {x. x n \ 0}) {x. x n = 0})" by (metis (mono_tags, lifting) g group_isomorphisms_imp_iso group_isomorphisms_sym) show "subgroup_generated (relative_homology_group n (usphere n) (equator n)) {gn z} = relative_homology_group n (usphere n) (equator n)" apply (rule monoid.equality) using giso gn.subgroup_generated_by_image [of "{z}"] zcarr by (auto simp: usphere_def equator_def upper_def zeq gn.iso_iff) qed let ?hi_lu = "hom_induced n (lsphere n) (equator n) (nsphere n) (upper n) id" interpret gh_lu: group_hom "relative_homology_group n (lsphere n) (equator n)" "?rhgn (upper n)" ?hi_lu by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom) interpret gh_eef: group_hom "?rhgn (equator n)" "?rhgn (equator n)" "?hi_ee f" by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom) define wp where "wp \ ?hi_lu zp" then have wpcarr: "wp \ carrier(?rhgn (upper n))" by (simp add: hom_induced_carrier) have "hom_induced n (nsphere n) {} (nsphere n) {x. x n \ 0} id \ iso (reduced_homology_group n (nsphere n)) (?rhgn {x. x n \ 0})" using iso_reduced_homology_group_upper_hemisphere [of n n n] by auto then have "carrier(?rhgn {x. x n \ 0}) \ (hom_induced n (nsphere n) {} (nsphere n) {x. x n \ 0} id) ` carrier(reduced_homology_group n (nsphere n))" by (simp add: iso_iff) then obtain vp where vpcarr: "vp \ carrier(reduced_homology_group n (nsphere n))" and eqwp: "hom_induced n (nsphere n) {} (nsphere n) (upper n) id vp = wp" using wpcarr by (auto simp: upper_def) define wn where "wn \ hom_induced n (usphere n) (equator n) (nsphere n) (lower n) id zn" then have wncarr: "wn \ carrier(?rhgn (lower n))" by (simp add: hom_induced_carrier) have "hom_induced n (nsphere n) {} (nsphere n) {x. x n \ 0} id \ iso (reduced_homology_group n (nsphere n)) (?rhgn {x. x n \ 0})" using iso_reduced_homology_group_lower_hemisphere [of n n n] by auto then have "carrier(?rhgn {x. x n \ 0}) \ (hom_induced n (nsphere n) {} (nsphere n) {x. x n \ 0} id) ` carrier(reduced_homology_group n (nsphere n))" by (simp add: iso_iff) then obtain vn where vpcarr: "vn \ carrier(reduced_homology_group n (nsphere n))" and eqwp: "hom_induced n (nsphere n) {} (nsphere n) (lower n) id vn = wn" using wncarr by (auto simp: lower_def) define up where "up \ hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id zp" then have upcarr: "up \ carrier(?rhgn (equator n))" by (simp add: hom_induced_carrier) define un where "un \ hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id zn" then have uncarr: "un \ carrier(?rhgn (equator n))" by (simp add: hom_induced_carrier) have *: "(\(x, y). hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id x \\<^bsub>?rhgn (equator n)\<^esub> hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id y) \ Group.iso (relative_homology_group n (lsphere n) (equator n) \\ relative_homology_group n (usphere n) (equator n)) (?rhgn (equator n))" proof (rule conjunct1 [OF exact_sequence_sum_lemma [OF abelian_relative_homology_group]]) show "hom_induced n (lsphere n) (equator n) (nsphere n) (upper n) id \ Group.iso (relative_homology_group n (lsphere n) (equator n)) (?rhgn (upper n))" apply (simp add: lsphere_def usphere_def equator_def lower_def upper_def) using iso_relative_homology_group_lower_hemisphere by blast show "hom_induced n (usphere n) (equator n) (nsphere n) (lower n) id \ Group.iso (relative_homology_group n (usphere n) (equator n)) (?rhgn (lower n))" apply (simp_all add: lsphere_def usphere_def equator_def lower_def upper_def) using iso_relative_homology_group_upper_hemisphere by blast+ show "exact_seq ([?rhgn (lower n), ?rhgn (equator n), relative_homology_group n (lsphere n) (equator n)], [hom_induced n (nsphere n) (equator n) (nsphere n) (lower n) id, hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id])" unfolding lsphere_def usphere_def equator_def lower_def upper_def by (rule homology_exactness_triple_3) force show "exact_seq ([?rhgn (upper n), ?rhgn (equator n), relative_homology_group n (usphere n) (equator n)], [hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id, hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id])" unfolding lsphere_def usphere_def equator_def lower_def upper_def by (rule homology_exactness_triple_3) force next fix x assume "x \ carrier (relative_homology_group n (lsphere n) (equator n))" show "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id (hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id x) = hom_induced n (lsphere n) (equator n) (nsphere n) (upper n) id x" by (simp add: hom_induced_compose' subset_iff lsphere_def usphere_def equator_def lower_def upper_def) next fix x assume "x \ carrier (relative_homology_group n (usphere n) (equator n))" show "hom_induced n (nsphere n) (equator n) (nsphere n) (lower n) id (hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id x) = hom_induced n (usphere n) (equator n) (nsphere n) (lower n) id x" by (simp add: hom_induced_compose' subset_iff lsphere_def usphere_def equator_def lower_def upper_def) qed then have sb: "carrier (?rhgn (equator n)) \ (\(x, y). hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id x \\<^bsub>?rhgn (equator n)\<^esub> hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id y) ` carrier (relative_homology_group n (lsphere n) (equator n) \\ relative_homology_group n (usphere n) (equator n))" by (simp add: iso_iff) obtain a b::int where up_ab: "?hi_ee f up = up [^]\<^bsub>?rhgn (equator n)\<^esub> a\\<^bsub>?rhgn (equator n)\<^esub> un [^]\<^bsub>?rhgn (equator n)\<^esub> b" proof - have hiupcarr: "?hi_ee f up \ carrier(?rhgn (equator n))" by (simp add: hom_induced_carrier) obtain u v where u: "u \ carrier (relative_homology_group n (lsphere n) (equator n))" and v: "v \ carrier (relative_homology_group n (usphere n) (equator n))" and eq: "?hi_ee f up = hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id u \\<^bsub>?rhgn (equator n)\<^esub> hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id v" using subsetD [OF sb hiupcarr] by auto have "u \ carrier (subgroup_generated (relative_homology_group n (lsphere n) (equator n)) {zp})" by (simp_all add: u zp_sg) then obtain a::int where a: "u = zp [^]\<^bsub>relative_homology_group n (lsphere n) (equator n)\<^esub> a" by (metis group.carrier_subgroup_generated_by_singleton group_relative_homology_group rangeE zpcarr) have ae: "hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id (pow (relative_homology_group n (lsphere n) (equator n)) zp a) = pow (?rhgn (equator n)) (hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id zp) a" by (meson group_hom.hom_int_pow group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced zpcarr) have "v \ carrier (subgroup_generated (relative_homology_group n (usphere n) (equator n)) {zn})" by (simp_all add: v zn_sg) then obtain b::int where b: "v = zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> b" by (metis group.carrier_subgroup_generated_by_singleton group_relative_homology_group rangeE zncarr) have be: "hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id (zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> b) = hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id zn [^]\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub> b" by (meson group_hom.hom_int_pow group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced zncarr) show thesis proof show "?hi_ee f up = up [^]\<^bsub>?rhgn (equator n)\<^esub> a \\<^bsub>?rhgn (equator n)\<^esub> un [^]\<^bsub>?rhgn (equator n)\<^esub> b" using a ae b be eq local.up_def un_def by auto qed qed have "(hom_boundary n (nsphere n) (equator n) \ hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id) zp = z" using zp_z equ apply (simp add: lsphere_def naturality_hom_induced) by (metis hom_boundary_carrier hom_induced_id) then have up_z: "hom_boundary n (nsphere n) (equator n) up = z" by (simp add: up_def) have "(hom_boundary n (nsphere n) (equator n) \ hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id) zn = z" using zn_z equ apply (simp add: usphere_def naturality_hom_induced) by (metis hom_boundary_carrier hom_induced_id) then have un_z: "hom_boundary n (nsphere n) (equator n) un = z" by (simp add: un_def) have Bd_ab: "Brouwer_degree2 (n - Suc 0) f = a + b" proof (rule Brouwer_degree2_unique_generator; use False int_ops in simp_all) show "continuous_map (nsphere (n - Suc 0)) (nsphere (n - Suc 0)) f" using cmr by auto show "subgroup_generated (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) {z} = reduced_homology_group (int n - 1) (nsphere (n - Suc 0))" using zeq by blast have "(hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} f \ hom_boundary n (nsphere n) (equator n)) up = (hom_boundary n (nsphere n) (equator n) \ ?hi_ee f) up" using naturality_hom_induced [OF cmf fimeq, of n, symmetric] by (simp add: subtopology_restrict equ fun_eq_iff) also have "\ = hom_boundary n (nsphere n) (equator n) (up [^]\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub> a \\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub> un [^]\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub> b)" by (simp add: o_def up_ab) also have "\ = z [^]\<^bsub>reduced_homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> (a + b)" using zcarr apply (simp add: HB.hom_int_pow reduced_homology_group_def group.int_pow_subgroup_generated upcarr uncarr) by (metis equ(1) group.int_pow_mult group_relative_homology_group hom_boundary_carrier un_z up_z) finally show "hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} f z = z [^]\<^bsub>reduced_homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> (a + b)" by (simp add: up_z) qed define u where "u \ up \\<^bsub>?rhgn (equator n)\<^esub> inv\<^bsub>?rhgn (equator n)\<^esub> un" have ucarr: "u \ carrier (?rhgn (equator n))" by (simp add: u_def uncarr upcarr) then have "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = u [^]\<^bsub>?rhgn (equator n)\<^esub> (a - b) \ (GE.ord u) dvd a - b - Brouwer_degree2 n f" by (simp add: GE.int_pow_eq) moreover have "GE.ord u = 0" proof (clarsimp simp add: GE.ord_eq_0 ucarr) fix d :: nat assume "0 < d" and "u [^]\<^bsub>?rhgn (equator n)\<^esub> d = singular_relboundary_set n (nsphere n) (equator n)" then have "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id u [^]\<^bsub>?rhgn (upper n)\<^esub> d = \\<^bsub>?rhgn (upper n)\<^esub>" by (metis HIU.hom_one HIU.hom_nat_pow one_relative_homology_group ucarr) moreover have "?hi_lu = hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id \ hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id" by (simp add: lsphere_def image_subset_iff equator_upper flip: hom_induced_compose) then have p: "wp = hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id up" by (simp add: local.up_def wp_def) have n: "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id un = \\<^bsub>?rhgn (upper n)\<^esub>" using homology_exactness_triple_3 [OF equator_upper, of n "nsphere n"] using un_def zncarr by (auto simp: upper_usphere kernel_def) have "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id u = wp" unfolding u_def using p n HIU.inv_one HIU.r_one uncarr upcarr by auto ultimately have "(wp [^]\<^bsub>?rhgn (upper n)\<^esub> d) = \\<^bsub>?rhgn (upper n)\<^esub>" by simp moreover have "infinite (carrier (subgroup_generated (?rhgn (upper n)) {wp}))" proof - have "?rhgn (upper n) \ reduced_homology_group n (nsphere n)" unfolding upper_def using iso_reduced_homology_group_upper_hemisphere [of n n n] by (blast intro: group.iso_sym group_reduced_homology_group is_isoI) also have "\ \ integer_group" by (simp add: reduced_homology_group_nsphere) finally have iso: "?rhgn (upper n) \ integer_group" . have "carrier (subgroup_generated (?rhgn (upper n)) {wp}) = carrier (?rhgn (upper n))" using gh_lu.subgroup_generated_by_image [of "{zp}"] zpcarr HIU.carrier_subgroup_generated_subset gh_lu.iso_iff iso_relative_homology_group_lower_hemisphere zp_sg by (auto simp: lower_def lsphere_def upper_def equator_def wp_def) then show ?thesis using infinite_UNIV_int iso_finite [OF iso] by simp qed ultimately show False using HIU.finite_cyclic_subgroup \0 < d\ wpcarr by blast qed ultimately have iff: "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = u [^]\<^bsub>?rhgn (equator n)\<^esub> (a - b) \ Brouwer_degree2 n f = a - b" by auto have "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = ?hi_ee f u" proof - have ne: "topspace (nsphere n) \ equator n \ {}" by (metis equ(1) nonempty_nsphere topspace_subtopology) have eq1: "hom_boundary n (nsphere n) (equator n) u = \\<^bsub>reduced_homology_group (int n - 1) (subtopology (nsphere n) (equator n))\<^esub>" using one_reduced_homology_group u_def un_z uncarr up_z upcarr by force then have uhom: "u \ hom_induced n (nsphere n) {} (nsphere n) (equator n) id ` carrier (reduced_homology_group (int n) (nsphere n))" using homology_exactness_reduced_1 [OF ne, of n] eq1 ucarr by (auto simp: kernel_def) then obtain v where vcarr: "v \ carrier (reduced_homology_group (int n) (nsphere n))" and ueq: "u = hom_induced n (nsphere n) {} (nsphere n) (equator n) id v" by blast interpret GH_hi: group_hom "homology_group n (nsphere n)" "?rhgn (equator n)" "hom_induced n (nsphere n) {} (nsphere n) (equator n) id" by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom) have poweq: "pow (homology_group n (nsphere n)) x i = pow (reduced_homology_group n (nsphere n)) x i" for x and i::int by (simp add: False un_reduced_homology_group) have vcarr': "v \ carrier (homology_group n (nsphere n))" using carrier_reduced_homology_group_subset vcarr by blast have "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = hom_induced n (nsphere n) {} (nsphere n) (equator n) f v" using vcarr vcarr' by (simp add: ueq poweq hom_induced_compose' cmf flip: GH_hi.hom_int_pow Brouwer_degree2) also have "\ = hom_induced n (nsphere n) (topspace(nsphere n) \ equator n) (nsphere n) (equator n) f (hom_induced n (nsphere n) {} (nsphere n) (topspace(nsphere n) \ equator n) id v)" using fimeq by (simp add: hom_induced_compose' cmf) also have "\ = ?hi_ee f u" by (metis hom_induced inf.left_idem ueq) finally show ?thesis . qed moreover interpret gh_een: group_hom "?rhgn (equator n)" "?rhgn (equator n)" "?hi_ee neg" by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom) have hi_up_eq_un: "?hi_ee neg up = un [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" proof - have "?hi_ee neg (hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id zp) = hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) (neg \ id) zp" by (intro hom_induced_compose') (auto simp: lsphere_def equator_def cm_neg) also have "\ = hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id (hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp)" by (subst hom_induced_compose' [OF cm_neg_lu]) (auto simp: usphere_def equator_def) also have "hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp = zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" proof - let ?hb = "hom_boundary n (usphere n) (equator n)" have eq: "subtopology (nsphere n) {x. x n \ 0} = usphere n \ {x. x n = 0} = equator n" by (auto simp: usphere_def upper_def equator_def) with hb_iso have inj: "inj_on (?hb) (carrier (relative_homology_group n (usphere n) (equator n)))" by (simp add: iso_iff) interpret hb_hom: group_hom "relative_homology_group n (usphere n) (equator n)" "reduced_homology_group (int n - 1) (nsphere (n - Suc 0))" "?hb" using hb_iso iso_iff eq group_hom_axioms_def group_hom_def by fastforce show ?thesis proof (rule inj_onD [OF inj]) have *: "hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} neg z = z [^]\<^bsub>homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> Brouwer_degree2 (n - Suc 0) neg" using Brouwer_degree2 [of z "n - Suc 0" neg] False zcarr by (simp add: int_ops group.int_pow_subgroup_generated reduced_homology_group_def) have "?hb \ hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg = hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} neg \ hom_boundary n (lsphere n) (equator n)" apply (subst naturality_hom_induced [OF cm_neg_lu]) apply (force simp: equator_def neg_def) by (simp add: equ) then have "?hb (hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp) = (z [^]\<^bsub>homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> Brouwer_degree2 (n - Suc 0) neg)" by (metis "*" comp_apply zp_z) also have "\ = ?hb (zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg)" by (metis group.int_pow_subgroup_generated group_relative_homology_group hb_hom.hom_int_pow reduced_homology_group_def zcarr zn_z zncarr) finally show "?hb (hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp) = ?hb (zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg)" by simp qed (auto simp: hom_induced_carrier group.int_pow_closed zncarr) qed finally show ?thesis by (metis (no_types, lifting) group_hom.hom_int_pow group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced local.up_def un_def zncarr) qed have "continuous_map (nsphere (n - Suc 0)) (nsphere (n - Suc 0)) neg" using cm_neg by blast then have "homeomorphic_map (nsphere (n - Suc 0)) (nsphere (n - Suc 0)) neg" apply (auto simp: homeomorphic_map_maps homeomorphic_maps_def) apply (rule_tac x=neg in exI, auto) done then have Brouwer_degree2_21: "Brouwer_degree2 (n - Suc 0) neg ^ 2 = 1" using Brouwer_degree2_homeomorphic_map power2_eq_1_iff by force have hi_un_eq_up: "?hi_ee neg un = up [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" (is "?f un = ?y") proof - have [simp]: "neg \ neg = id" by force have "?f (?f ?y) = ?y" apply (subst hom_induced_compose' [OF cm_neg _ cm_neg]) apply(force simp: equator_def) apply (simp add: upcarr hom_induced_id_gen) done moreover have "?f ?y = un" using upcarr apply (simp only: gh_een.hom_int_pow hi_up_eq_un) by (metis (no_types, lifting) Brouwer_degree2_21 GE.group_l_invI GE.l_inv_ex group.int_pow_1 group.int_pow_pow power2_eq_1_iff uncarr zmult_eq_1_iff) ultimately show "?f un = ?y" by simp qed have "?hi_ee f un = un [^]\<^bsub>?rhgn (equator n)\<^esub> a \\<^bsub>?rhgn (equator n)\<^esub> up [^]\<^bsub>?rhgn (equator n)\<^esub> b" proof - let ?TE = "topspace (nsphere n) \ equator n" have fneg: "(f \ neg) x = (neg \ f) x" if "x \ topspace (nsphere n)" for x using f [OF that] by (force simp: neg_def) have neg_im: "neg ` (topspace (nsphere n) \ equator n) \ topspace (nsphere n) \ equator n" by (metis cm_neg continuous_map_image_subset_topspace equ(1) topspace_subtopology) have 1: "hom_induced n (nsphere n) ?TE (nsphere n) ?TE f \ hom_induced n (nsphere n) ?TE (nsphere n) ?TE neg = hom_induced n (nsphere n) ?TE (nsphere n) ?TE neg \ hom_induced n (nsphere n) ?TE (nsphere n) ?TE f" using neg_im fimeq cm_neg cmf apply (simp add: flip: hom_induced_compose del: hom_induced_restrict) using fneg by (auto intro: hom_induced_eq) have "(un [^]\<^bsub>?rhgn (equator n)\<^esub> a) \\<^bsub>?rhgn (equator n)\<^esub> (up [^]\<^bsub>?rhgn (equator n)\<^esub> b) = un [^]\<^bsub>?rhgn (equator n)\<^esub> (Brouwer_degree2 (n - 1) neg * a * Brouwer_degree2 (n - 1) neg) \\<^bsub>?rhgn (equator n)\<^esub> up [^]\<^bsub>?rhgn (equator n)\<^esub> (Brouwer_degree2 (n - 1) neg * b * Brouwer_degree2 (n - 1) neg)" proof - have "Brouwer_degree2 (n - Suc 0) neg = 1 \ Brouwer_degree2 (n - Suc 0) neg = - 1" using Brouwer_degree2_21 power2_eq_1_iff by blast then show ?thesis by fastforce qed also have "\ = ((un [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - 1) neg) [^]\<^bsub>?rhgn (equator n)\<^esub> a \\<^bsub>?rhgn (equator n)\<^esub> (up [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - 1) neg) [^]\<^bsub>?rhgn (equator n)\<^esub> b) [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - 1) neg" by (simp add: GE.int_pow_distrib GE.int_pow_pow uncarr upcarr) also have "\ = ?hi_ee neg (?hi_ee f up) [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" by (simp add: gh_een.hom_int_pow hi_un_eq_up hi_up_eq_un uncarr up_ab upcarr) finally have 2: "(un [^]\<^bsub>?rhgn (equator n)\<^esub> a) \\<^bsub>?rhgn (equator n)\<^esub> (up [^]\<^bsub>?rhgn (equator n)\<^esub> b) = ?hi_ee neg (?hi_ee f up) [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" . have "un = ?hi_ee neg up [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" by (metis (no_types, opaque_lifting) Brouwer_degree2_21 GE.int_pow_1 GE.int_pow_pow hi_up_eq_un power2_eq_1_iff uncarr zmult_eq_1_iff) moreover have "?hi_ee f ((?hi_ee neg up) [^]\<^bsub>?rhgn (equator n)\<^esub> (Brouwer_degree2 (n - Suc 0) neg)) = un [^]\<^bsub>?rhgn (equator n)\<^esub> a \\<^bsub>?rhgn (equator n)\<^esub> up [^]\<^bsub>?rhgn (equator n)\<^esub> b" using 1 2 by (simp add: hom_induced_carrier gh_eef.hom_int_pow fun_eq_iff) ultimately show ?thesis by blast qed then have "?hi_ee f u = u [^]\<^bsub>?rhgn (equator n)\<^esub> (a - b)" by (simp add: u_def upcarr uncarr up_ab GE.int_pow_diff GE.m_ac GE.int_pow_distrib GE.int_pow_inv GE.inv_mult_group) ultimately have "Brouwer_degree2 n f = a - b" using iff by blast with Bd_ab show ?thesis by simp qed simp subsection \General Jordan-Brouwer separation theorem and invariance of dimension\ proposition relative_homology_group_Euclidean_complement_step: assumes "closedin (Euclidean_space n) S" shows "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S) \ relative_homology_group (p + k) (Euclidean_space (n+k)) (topspace(Euclidean_space (n+k)) - S)" proof - have *: "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S) \ relative_homology_group (p + 1) (Euclidean_space (Suc n)) (topspace(Euclidean_space (Suc n)) - {x \ S. x n = 0})" (is "?lhs \ ?rhs") if clo: "closedin (Euclidean_space (Suc n)) S" and cong: "\x y. \x \ S; \i. i \ n \ x i = y i\ \ y \ S" for p n S proof - have Ssub: "S \ topspace (Euclidean_space (Suc n))" by (meson clo closedin_def) define lo where "lo \ {x \ topspace(Euclidean_space (Suc n)). x n < (if x \ S then 0 else 1)}" define hi where "hi = {x \ topspace(Euclidean_space (Suc n)). x n > (if x \ S then 0 else -1)}" have lo_hi_Int: "lo \ hi = {x \ topspace(Euclidean_space (Suc n)) - S. x n \ {-1<..<1}}" by (auto simp: hi_def lo_def) have lo_hi_Un: "lo \ hi = topspace(Euclidean_space (Suc n)) - {x \ S. x n = 0}" by (auto simp: hi_def lo_def) define ret where "ret \ \c::real. \x i. if i = n then c else x i" have cm_ret: "continuous_map (powertop_real UNIV) (powertop_real UNIV) (ret t)" for t by (auto simp: ret_def continuous_map_componentwise_UNIV intro: continuous_map_product_projection) let ?ST = "\t. subtopology (Euclidean_space (Suc n)) {x. x n = t}" define squashable where "squashable \ \t S. \x t'. x \ S \ (x n \ t' \ t' \ t \ t \ t' \ t' \ x n) \ ret t' x \ S" have squashable: "squashable t (topspace(Euclidean_space(Suc n)))" for t by (simp add: squashable_def topspace_Euclidean_space ret_def) have squashableD: "\squashable t S; x \ S; x n \ t' \ t' \ t \ t \ t' \ t' \ x n\ \ ret t' x \ S" for x t' t S by (auto simp: squashable_def) have "squashable 1 hi" by (force simp: squashable_def hi_def ret_def topspace_Euclidean_space intro: cong) have "squashable t UNIV" for t by (force simp: squashable_def hi_def ret_def topspace_Euclidean_space intro: cong) have squashable_0_lohi: "squashable 0 (lo \ hi)" using Ssub by (auto simp: squashable_def hi_def lo_def ret_def topspace_Euclidean_space intro: cong) have rm_ret: "retraction_maps (subtopology (Euclidean_space (Suc n)) U) (subtopology (Euclidean_space (Suc n)) {x. x \ U \ x n = t}) (ret t) id" if "squashable t U" for t U unfolding retraction_maps_def proof (intro conjI ballI) show "continuous_map (subtopology (Euclidean_space (Suc n)) U) (subtopology (Euclidean_space (Suc n)) {x \ U. x n = t}) (ret t)" apply (simp add: cm_ret continuous_map_in_subtopology continuous_map_from_subtopology Euclidean_space_def) using that by (fastforce simp: squashable_def ret_def) next show "continuous_map (subtopology (Euclidean_space (Suc n)) {x \ U. x n = t}) (subtopology (Euclidean_space (Suc n)) U) id" using continuous_map_in_subtopology by fastforce show "ret t (id x) = x" if "x \ topspace (subtopology (Euclidean_space (Suc n)) {x \ U. x n = t})" for x using that by (simp add: topspace_Euclidean_space ret_def fun_eq_iff) qed have cm_snd: "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (powertop_real UNIV) S)) euclideanreal (\x. snd x k)" for k::nat and S using continuous_map_componentwise_UNIV continuous_map_into_fulltopology continuous_map_snd by fastforce have cm_fstsnd: "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (powertop_real UNIV) S)) euclideanreal (\x. fst x * snd x k)" for k::nat and S by (intro continuous_intros continuous_map_into_fulltopology [OF continuous_map_fst] cm_snd) have hw_sub: "homotopic_with (\k. k ` V \ V) (subtopology (Euclidean_space (Suc n)) U) (subtopology (Euclidean_space (Suc n)) U) (ret t) id" if "squashable t U" "squashable t V" for U V t unfolding homotopic_with_def proof (intro exI conjI allI ballI) let ?h = "\(z,x). ret ((1 - z) * t + z * x n) x" show "(\x. ?h (u, x)) ` V \ V" if "u \ {0..1}" for u using that by clarsimp (metis squashableD [OF \squashable t V\] convex_bound_le diff_ge_0_iff_ge eq_diff_eq' le_cases less_eq_real_def segment_bound_lemma) have 1: "?h ` ({0..1} \ ({x. \i\Suc n. x i = 0} \ U)) \ U" by clarsimp (metis squashableD [OF \squashable t U\] convex_bound_le diff_ge_0_iff_ge eq_diff_eq' le_cases less_eq_real_def segment_bound_lemma) show "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (Euclidean_space (Suc n)) U)) (subtopology (Euclidean_space (Suc n)) U) ?h" apply (simp add: continuous_map_in_subtopology Euclidean_space_def subtopology_subtopology 1) apply (auto simp: case_prod_unfold ret_def continuous_map_componentwise_UNIV) apply (intro continuous_map_into_fulltopology [OF continuous_map_fst] cm_snd continuous_intros) by (auto simp: cm_snd) qed (auto simp: ret_def) have cs_hi: "contractible_space(subtopology (Euclidean_space(Suc n)) hi)" proof - have "homotopic_with (\x. True) (?ST 1) (?ST 1) id (\x. (\i. if i = n then 1 else 0))" apply (subst homotopic_with_sym) apply (simp add: homotopic_with) apply (rule_tac x="(\(z,x) i. if i=n then 1 else z * x i)" in exI) apply (auto simp: Euclidean_space_def subtopology_subtopology continuous_map_in_subtopology case_prod_unfold continuous_map_componentwise_UNIV cm_fstsnd) done then have "contractible_space (?ST 1)" unfolding contractible_space_def by metis moreover have "?thesis = contractible_space (?ST 1)" proof (intro deformation_retract_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) have "{x. \i\Suc n. x i = 0} \ {x \ hi. x n = 1} = {x. \i\Suc n. x i = 0} \ {x. x n = 1}" by (auto simp: hi_def topspace_Euclidean_space) then have eq: "subtopology (Euclidean_space (Suc n)) {x. x \ hi \ x n = 1} = ?ST 1" by (simp add: Euclidean_space_def subtopology_subtopology) show "homotopic_with (\x. True) (subtopology (Euclidean_space (Suc n)) hi) (subtopology (Euclidean_space (Suc n)) hi) (ret 1) id" using hw_sub [OF \squashable 1 hi\ \squashable 1 UNIV\] eq by simp show "retraction_maps (subtopology (Euclidean_space (Suc n)) hi) (?ST 1) (ret 1) id" using rm_ret [OF \squashable 1 hi\] eq by simp qed ultimately show ?thesis by metis qed have "?lhs \ relative_homology_group p (Euclidean_space (Suc n)) (lo \ hi)" proof (rule group.iso_sym [OF _ deformation_retract_imp_isomorphic_relative_homology_groups]) have "{x. \i\Suc n. x i = 0} \ {x. x n = 0} = {x. \i\n. x i = (0::real)}" by auto (metis le_less_Suc_eq not_le) then have "?ST 0 = Euclidean_space n" by (simp add: Euclidean_space_def subtopology_subtopology) then show "retraction_maps (Euclidean_space (Suc n)) (Euclidean_space n) (ret 0) id" using rm_ret [OF \squashable 0 UNIV\] by auto then have "ret 0 x \ topspace (Euclidean_space n)" if "x \ topspace (Euclidean_space (Suc n))" "-1 < x n" "x n < 1" for x using that by (simp add: continuous_map_def retraction_maps_def) then show "(ret 0) ` (lo \ hi) \ topspace (Euclidean_space n) - S" by (auto simp: local.cong ret_def hi_def lo_def) show "homotopic_with (\h. h ` (lo \ hi) \ lo \ hi) (Euclidean_space (Suc n)) (Euclidean_space (Suc n)) (ret 0) id" using hw_sub [OF squashable squashable_0_lohi] by simp qed (auto simp: lo_def hi_def Euclidean_space_def) also have "\ \ relative_homology_group p (subtopology (Euclidean_space (Suc n)) hi) (lo \ hi)" proof (rule group.iso_sym [OF _ isomorphic_relative_homology_groups_inclusion_contractible]) show "contractible_space (subtopology (Euclidean_space (Suc n)) hi)" by (simp add: cs_hi) show "topspace (Euclidean_space (Suc n)) \ hi \ {}" apply (simp add: hi_def topspace_Euclidean_space set_eq_iff) apply (rule_tac x="\i. if i = n then 1 else 0" in exI, auto) done qed auto also have "\ \ relative_homology_group p (subtopology (Euclidean_space (Suc n)) (lo \ hi)) lo" proof - have oo: "openin (Euclidean_space (Suc n)) {x \ topspace (Euclidean_space (Suc n)). x n \ A}" if "open A" for A proof (rule openin_continuous_map_preimage) show "continuous_map (Euclidean_space (Suc n)) euclideanreal (\x. x n)" proof - have "\n f. continuous_map (product_topology f UNIV) (f (n::nat)) (\f. f n::real)" by (simp add: continuous_map_product_projection) then show ?thesis using Euclidean_space_def continuous_map_from_subtopology by (metis (mono_tags)) qed qed (auto intro: that) have "openin (Euclidean_space(Suc n)) lo" apply (simp add: openin_subopen [of _ lo]) apply (simp add: lo_def, safe) apply (force intro: oo [of "lessThan 0", simplified] open_Collect_less) apply (rule_tac x="{x \ topspace(Euclidean_space(Suc n)). x n < 1} \ (topspace(Euclidean_space(Suc n)) - S)" in exI) using clo apply (force intro: oo [of "lessThan 1", simplified] open_Collect_less) done moreover have "openin (Euclidean_space(Suc n)) hi" apply (simp add: openin_subopen [of _ hi]) apply (simp add: hi_def, safe) apply (force intro: oo [of "greaterThan 0", simplified] open_Collect_less) apply (rule_tac x="{x \ topspace(Euclidean_space(Suc n)). x n > -1} \ (topspace(Euclidean_space(Suc n)) - S)" in exI) using clo apply (force intro: oo [of "greaterThan (-1)", simplified] open_Collect_less) done ultimately have *: "subtopology (Euclidean_space (Suc n)) (lo \ hi) closure_of (topspace (subtopology (Euclidean_space (Suc n)) (lo \ hi)) - hi) \ subtopology (Euclidean_space (Suc n)) (lo \ hi) interior_of lo" by (metis (no_types, lifting) Diff_idemp Diff_subset_conv Un_commute Un_upper2 closure_of_interior_of interior_of_closure_of interior_of_complement interior_of_eq lo_hi_Un openin_Un openin_open_subtopology topspace_subtopology_subset) have eq: "((lo \ hi) \ (lo \ hi - (topspace (Euclidean_space (Suc n)) \ (lo \ hi) - hi))) = hi" "(lo - (topspace (Euclidean_space (Suc n)) \ (lo \ hi) - hi)) = lo \ hi" by (auto simp: lo_def hi_def Euclidean_space_def) show ?thesis using homology_excision_axiom [OF *, of "lo \ hi" p] by (force simp: subtopology_subtopology eq is_iso_def) qed also have "\ \ relative_homology_group (p + 1 - 1) (subtopology (Euclidean_space (Suc n)) (lo \ hi)) lo" by simp also have "\ \ relative_homology_group (p + 1) (Euclidean_space (Suc n)) (lo \ hi)" proof (rule group.iso_sym [OF _ isomorphic_relative_homology_groups_relboundary_contractible]) have proj: "continuous_map (powertop_real UNIV) euclideanreal (\f. f n)" by (metis UNIV_I continuous_map_product_projection) have hilo: "\x. x \ hi \ (\i. if i = n then - x i else x i) \ lo" "\x. x \ lo \ (\i. if i = n then - x i else x i) \ hi" using local.cong by (auto simp: hi_def lo_def topspace_Euclidean_space split: if_split_asm) have "subtopology (Euclidean_space (Suc n)) hi homeomorphic_space subtopology (Euclidean_space (Suc n)) lo" unfolding homeomorphic_space_def apply (rule_tac x="\x i. if i = n then -(x i) else x i" in exI)+ using proj apply (auto simp: homeomorphic_maps_def Euclidean_space_def continuous_map_in_subtopology hilo continuous_map_componentwise_UNIV continuous_map_from_subtopology continuous_map_minus intro: continuous_map_from_subtopology continuous_map_product_projection) done then have "contractible_space(subtopology (Euclidean_space(Suc n)) hi) \ contractible_space (subtopology (Euclidean_space (Suc n)) lo)" by (rule homeomorphic_space_contractibility) then show "contractible_space (subtopology (Euclidean_space (Suc n)) lo)" using cs_hi by auto show "topspace (Euclidean_space (Suc n)) \ lo \ {}" apply (simp add: lo_def Euclidean_space_def set_eq_iff) apply (rule_tac x="\i. if i = n then -1 else 0" in exI, auto) done qed auto also have "\ \ ?rhs" by (simp flip: lo_hi_Un) finally show ?thesis . qed show ?thesis proof (induction k) case (Suc m) with assms obtain T where cloT: "closedin (powertop_real UNIV) T" and SeqT: "S = T \ {x. \i\n. x i = 0}" by (auto simp: Euclidean_space_def closedin_subtopology) then have "closedin (Euclidean_space (m + n)) S" apply (simp add: Euclidean_space_def closedin_subtopology) apply (rule_tac x="T \ topspace(Euclidean_space n)" in exI) using closedin_Euclidean_space topspace_Euclidean_space by force moreover have "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - S) \ relative_homology_group (p + 1) (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S)" if "closedin (Euclidean_space n) S" for p n proof - define S' where "S' \ {x \ topspace(Euclidean_space(Suc n)). (\i. if i < n then x i else 0) \ S}" have Ssub_n: "S \ topspace (Euclidean_space n)" by (meson that closedin_def) have "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S') \ relative_homology_group (p + 1) (Euclidean_space (Suc n)) (topspace(Euclidean_space (Suc n)) - {x \ S'. x n = 0})" proof (rule *) have cm: "continuous_map (powertop_real UNIV) euclideanreal (\f. f u)" for u by (metis UNIV_I continuous_map_product_projection) have "continuous_map (subtopology (powertop_real UNIV) {x. \i>n. x i = 0}) euclideanreal (\x. if k \ n then x k else 0)" for k by (simp add: continuous_map_from_subtopology [OF cm]) moreover have "\i\n. (if i < n then x i else 0) = 0" if "x \ topspace (subtopology (powertop_real UNIV) {x. \i>n. x i = 0})" for x using that by simp ultimately have "continuous_map (Euclidean_space (Suc n)) (Euclidean_space n) (\x i. if i < n then x i else 0)" by (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV continuous_map_from_subtopology [OF cm] image_subset_iff) then show "closedin (Euclidean_space (Suc n)) S'" unfolding S'_def using that by (rule closedin_continuous_map_preimage) next fix x y assume xy: "\i. i \ n \ x i = y i" "x \ S'" then have "(\i. if i < n then x i else 0) = (\i. if i < n then y i else 0)" by (simp add: S'_def Euclidean_space_def fun_eq_iff) with xy show "y \ S'" by (simp add: S'_def Euclidean_space_def) qed moreover have abs_eq: "(\i. if i < n then x i else 0) = x" if "\i. i \ n \ x i = 0" for x :: "nat \ real" and n using that by auto then have "topspace (Euclidean_space n) - S' = topspace (Euclidean_space n) - S" by (simp add: S'_def Euclidean_space_def set_eq_iff cong: conj_cong) moreover have "topspace (Euclidean_space (Suc n)) - {x \ S'. x n = 0} = topspace (Euclidean_space (Suc n)) - S" using Ssub_n apply (auto simp: S'_def subset_iff Euclidean_space_def set_eq_iff abs_eq cong: conj_cong) by (metis abs_eq le_antisym not_less_eq_eq) ultimately show ?thesis by simp qed ultimately have "relative_homology_group (p + m)(Euclidean_space (m + n))(topspace (Euclidean_space (m + n)) - S) \ relative_homology_group (p + m + 1) (Euclidean_space (Suc (m + n))) (topspace (Euclidean_space (Suc (m + n))) - S)" by (metis \closedin (Euclidean_space (m + n)) S\) then show ?case using Suc.IH iso_trans by (force simp: algebra_simps) qed (simp add: iso_refl) qed lemma iso_Euclidean_complements_lemma1: assumes S: "closedin (Euclidean_space m) S" and cmf: "continuous_map(subtopology (Euclidean_space m) S) (Euclidean_space n) f" obtains g where "continuous_map (Euclidean_space m) (Euclidean_space n) g" "\x. x \ S \ g x = f x" proof - have cont: "continuous_on (topspace (Euclidean_space m) \ S) (\x. f x i)" for i by (metis (no_types) continuous_on_product_then_coordinatewise cm_Euclidean_space_iff_continuous_on cmf topspace_subtopology) have "f ` (topspace (Euclidean_space m) \ S) \ topspace (Euclidean_space n)" using cmf continuous_map_image_subset_topspace by fastforce then have "\g. continuous_on (topspace (Euclidean_space m)) g \ (\x \ S. g x = f x i)" for i using S Tietze_unbounded [OF cont [of i]] by (metis closedin_Euclidean_space_iff closedin_closed_Int topspace_subtopology topspace_subtopology_subset) then obtain g where cmg: "\i. continuous_map (Euclidean_space m) euclideanreal (g i)" and gf: "\i x. x \ S \ g i x = f x i" unfolding continuous_map_Euclidean_space_iff by metis let ?GG = "\x i. if i < n then g i x else 0" show thesis proof show "continuous_map (Euclidean_space m) (Euclidean_space n) ?GG" unfolding Euclidean_space_def [of n] by (auto simp: continuous_map_in_subtopology continuous_map_componentwise cmg) show "?GG x = f x" if "x \ S" for x proof - have "S \ topspace (Euclidean_space m)" by (meson S closedin_def) then have "f x \ topspace (Euclidean_space n)" using cmf that unfolding continuous_map_def topspace_subtopology by blast then show ?thesis by (force simp: topspace_Euclidean_space gf that) qed qed qed lemma iso_Euclidean_complements_lemma2: assumes S: "closedin (Euclidean_space m) S" and T: "closedin (Euclidean_space n) T" and hom: "homeomorphic_map (subtopology (Euclidean_space m) S) (subtopology (Euclidean_space n) T) f" obtains g where "homeomorphic_map (prod_topology (Euclidean_space m) (Euclidean_space n)) (prod_topology (Euclidean_space n) (Euclidean_space m)) g" "\x. x \ S \ g(x,(\i. 0)) = (f x,(\i. 0))" proof - obtain g where cmf: "continuous_map (subtopology (Euclidean_space m) S) (subtopology (Euclidean_space n) T) f" and cmg: "continuous_map (subtopology (Euclidean_space n) T) (subtopology (Euclidean_space m) S) g" and gf: "\x. x \ S \ g (f x) = x" and fg: "\y. y \ T \ f (g y) = y" using hom S T closedin_subset unfolding homeomorphic_map_maps homeomorphic_maps_def by fastforce obtain f' where cmf': "continuous_map (Euclidean_space m) (Euclidean_space n) f'" and f'f: "\x. x \ S \ f' x = f x" using iso_Euclidean_complements_lemma1 S cmf continuous_map_into_fulltopology by metis obtain g' where cmg': "continuous_map (Euclidean_space n) (Euclidean_space m) g'" and g'g: "\x. x \ T \ g' x = g x" using iso_Euclidean_complements_lemma1 T cmg continuous_map_into_fulltopology by metis define p where "p \ \(x,y). (x,(\i. y i + f' x i))" define p' where "p' \ \(x,y). (x,(\i. y i - f' x i))" define q where "q \ \(x,y). (x,(\i. y i + g' x i))" define q' where "q' \ \(x,y). (x,(\i. y i - g' x i))" have "homeomorphic_maps (prod_topology (Euclidean_space m) (Euclidean_space n)) (prod_topology (Euclidean_space m) (Euclidean_space n)) p p'" "homeomorphic_maps (prod_topology (Euclidean_space n) (Euclidean_space m)) (prod_topology (Euclidean_space n) (Euclidean_space m)) q q'" "homeomorphic_maps (prod_topology (Euclidean_space m) (Euclidean_space n)) (prod_topology (Euclidean_space n) (Euclidean_space m)) (\(x,y). (y,x)) (\(x,y). (y,x))" apply (simp_all add: p_def p'_def q_def q'_def homeomorphic_maps_def continuous_map_pairwise) apply (force simp: case_prod_unfold continuous_map_of_fst [unfolded o_def] cmf' cmg' intro: continuous_intros)+ done then have "homeomorphic_maps (prod_topology (Euclidean_space m) (Euclidean_space n)) (prod_topology (Euclidean_space n) (Euclidean_space m)) (q' \ (\(x,y). (y,x)) \ p) (p' \ ((\(x,y). (y,x)) \ q))" using homeomorphic_maps_compose homeomorphic_maps_sym by (metis (no_types, lifting)) moreover have "\x. x \ S \ (q' \ (\(x,y). (y,x)) \ p) (x, \i. 0) = (f x, \i. 0)" apply (simp add: q'_def p_def f'f) apply (simp add: fun_eq_iff) by (metis S T closedin_subset g'g gf hom homeomorphic_imp_surjective_map image_eqI topspace_subtopology_subset) ultimately show thesis using homeomorphic_map_maps that by blast qed proposition isomorphic_relative_homology_groups_Euclidean_complements: assumes S: "closedin (Euclidean_space n) S" and T: "closedin (Euclidean_space n) T" and hom: "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)" shows "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S) \ relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - T)" proof - have subST: "S \ topspace(Euclidean_space n)" "T \ topspace(Euclidean_space n)" by (meson S T closedin_def)+ have "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - S) \ relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - S)" using relative_homology_group_Euclidean_complement_step [OF S] by blast moreover have "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - T) \ relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - T)" using relative_homology_group_Euclidean_complement_step [OF T] by blast moreover have "relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - S) \ relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - T)" proof - obtain f where f: "homeomorphic_map (subtopology (Euclidean_space n) S) (subtopology (Euclidean_space n) T) f" using hom unfolding homeomorphic_space by blast obtain g where g: "homeomorphic_map (prod_topology (Euclidean_space n) (Euclidean_space n)) (prod_topology (Euclidean_space n) (Euclidean_space n)) g" and gf: "\x. x \ S \ g(x,(\i. 0)) = (f x,(\i. 0))" using S T f iso_Euclidean_complements_lemma2 by blast define h where "h \ \x::nat \real. ((\i. if i < n then x i else 0), (\j. if j < n then x(n + j) else 0))" define k where "k \ \(x,y) i. if i < 2 * n then if i < n then x i else y(i - n) else (0::real)" have hk: "homeomorphic_maps (Euclidean_space(2 * n)) (prod_topology (Euclidean_space n) (Euclidean_space n)) h k" unfolding homeomorphic_maps_def proof safe show "continuous_map (Euclidean_space (2 * n)) (prod_topology (Euclidean_space n) (Euclidean_space n)) h" apply (simp add: h_def continuous_map_pairwise o_def continuous_map_componentwise_Euclidean_space) unfolding Euclidean_space_def by (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection) have "continuous_map (prod_topology (Euclidean_space n) (Euclidean_space n)) euclideanreal (\p. fst p i)" for i using Euclidean_space_def continuous_map_into_fulltopology continuous_map_fst by fastforce moreover have "continuous_map (prod_topology (Euclidean_space n) (Euclidean_space n)) euclideanreal (\p. snd p (i - n))" for i using Euclidean_space_def continuous_map_into_fulltopology continuous_map_snd by fastforce ultimately show "continuous_map (prod_topology (Euclidean_space n) (Euclidean_space n)) (Euclidean_space (2 * n)) k" by (simp add: k_def continuous_map_pairwise o_def continuous_map_componentwise_Euclidean_space case_prod_unfold) qed (auto simp: k_def h_def fun_eq_iff topspace_Euclidean_space) define kgh where "kgh \ k \ g \ h" let ?i = "hom_induced (p + n) (Euclidean_space(2 * n)) (topspace(Euclidean_space(2 * n)) - S) (Euclidean_space(2 * n)) (topspace(Euclidean_space(2 * n)) - T) kgh" have "?i \ iso (relative_homology_group (p + int n) (Euclidean_space (2 * n)) (topspace (Euclidean_space (2 * n)) - S)) (relative_homology_group (p + int n) (Euclidean_space (2 * n)) (topspace (Euclidean_space (2 * n)) - T))" proof (rule homeomorphic_map_relative_homology_iso) show hm: "homeomorphic_map (Euclidean_space (2 * n)) (Euclidean_space (2 * n)) kgh" unfolding kgh_def by (meson hk g homeomorphic_map_maps homeomorphic_maps_compose homeomorphic_maps_sym) have Teq: "T = f ` S" using f homeomorphic_imp_surjective_map subST(1) subST(2) topspace_subtopology_subset by blast have khf: "\x. x \ S \ k(h(f x)) = f x" by (metis (no_types, lifting) Teq hk homeomorphic_maps_def image_subset_iff le_add1 mult_2 subST(2) subsetD subset_Euclidean_space) have gh: "g(h x) = h(f x)" if "x \ S" for x proof - have [simp]: "(\i. if i < n then x i else 0) = x" using subST(1) that topspace_Euclidean_space by (auto simp: fun_eq_iff) have "f x \ topspace(Euclidean_space n)" using Teq subST(2) that by blast moreover have "(\j. if j < n then x (n + j) else 0) = (\j. 0::real)" using Euclidean_space_def subST(1) that by force ultimately show ?thesis by (simp add: topspace_Euclidean_space h_def gf \x \ S\ fun_eq_iff) qed have *: "\S \ U; T \ U; kgh ` U = U; inj_on kgh U; kgh ` S = T\ \ kgh ` (U - S) = U - T" for U unfolding inj_on_def set_eq_iff by blast show "kgh ` (topspace (Euclidean_space (2 * n)) - S) = topspace (Euclidean_space (2 * n)) - T" proof (rule *) show "kgh ` topspace (Euclidean_space (2 * n)) = topspace (Euclidean_space (2 * n))" by (simp add: hm homeomorphic_imp_surjective_map) show "inj_on kgh (topspace (Euclidean_space (2 * n)))" using hm homeomorphic_map_def by auto show "kgh ` S = T" by (simp add: Teq kgh_def gh khf) qed (use subST topspace_Euclidean_space in \fastforce+\) qed auto then show ?thesis by (simp add: is_isoI mult_2) qed ultimately show ?thesis by (meson group.iso_sym iso_trans group_relative_homology_group) qed lemma lemma_iod: assumes "S \ T" "S \ {}" and Tsub: "T \ topspace(Euclidean_space n)" and S: "\a b u. \a \ S; b \ T; 0 < u; u < 1\ \ (\i. (1 - u) * a i + u * b i) \ S" shows "path_connectedin (Euclidean_space n) T" proof - obtain a where "a \ S" using assms by blast have "path_component_of (subtopology (Euclidean_space n) T) a b" if "b \ T" for b unfolding path_component_of_def proof (intro exI conjI) have [simp]: "\i\n. a i = 0" using Tsub \a \ S\ assms(1) topspace_Euclidean_space by auto have [simp]: "\i\n. b i = 0" using Tsub that topspace_Euclidean_space by auto have inT: "(\i. (1 - x) * a i + x * b i) \ T" if "0 \ x" "x \ 1" for x proof (cases "x = 0 \ x = 1") case True with \a \ S\ \b \ T\ \S \ T\ show ?thesis by force next case False then show ?thesis using subsetD [OF \S \ T\ S] \a \ S\ \b \ T\ that by auto qed have "continuous_on {0..1} (\x. (1 - x) * a k + x * b k)" for k by (intro continuous_intros) then show "pathin (subtopology (Euclidean_space n) T) (\t i. (1 - t) * a i + t * b i)" apply (simp add: Euclidean_space_def subtopology_subtopology pathin_subtopology) apply (simp add: pathin_def continuous_map_componentwise_UNIV inT) done qed auto then have "path_connected_space (subtopology (Euclidean_space n) T)" by (metis Tsub path_component_of_equiv path_connected_space_iff_path_component topspace_subtopology_subset) then show ?thesis by (simp add: Tsub path_connectedin_def) qed lemma invariance_of_dimension_closedin_Euclidean_space: assumes "closedin (Euclidean_space n) S" shows "subtopology (Euclidean_space n) S homeomorphic_space Euclidean_space n \ S = topspace(Euclidean_space n)" (is "?lhs = ?rhs") proof assume L: ?lhs have Ssub: "S \ topspace (Euclidean_space n)" by (meson assms closedin_def) moreover have False if "a \ S" and "a \ topspace (Euclidean_space n)" for a proof - have cl_n: "closedin (Euclidean_space (Suc n)) (topspace(Euclidean_space n))" using Euclidean_space_def closedin_Euclidean_space closedin_subtopology by fastforce then have sub: "subtopology (Euclidean_space(Suc n)) (topspace(Euclidean_space n)) = Euclidean_space n" by (metis (no_types, lifting) Euclidean_space_def closedin_subset subtopology_subtopology topspace_Euclidean_space topspace_subtopology topspace_subtopology_subset) then have cl_S: "closedin (Euclidean_space(Suc n)) S" using cl_n assms closedin_closed_subtopology by fastforce have sub_SucS: "subtopology (Euclidean_space (Suc n)) S = subtopology (Euclidean_space n) S" by (metis Ssub sub subtopology_subtopology topspace_subtopology topspace_subtopology_subset) have non0: "{y. \x::nat\real. (\i\Suc n. x i = 0) \ (\i\n. x i \ 0) \ y = x n} = -{0}" proof safe show "False" if "\i\Suc n. f i = 0" "0 = f n" "n \ i" "f i \ 0" for f::"nat\real" and i by (metis that le_antisym not_less_eq_eq) show "\f::nat\real. (\i\Suc n. f i = 0) \ (\i\n. f i \ 0) \ a = f n" if "a \ 0" for a by (rule_tac x="(\i. 0)(n:= a)" in exI) (force simp: that) qed have "homology_group 0 (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S)) \ homology_group 0 (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n)))" proof (rule isomorphic_relative_contractible_space_imp_homology_groups) show "(topspace (Euclidean_space (Suc n)) - S = {}) = (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n) = {})" using cl_n closedin_subset that by auto next fix p show "relative_homology_group p (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S) \ relative_homology_group p (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n))" by (simp add: L sub_SucS cl_S cl_n isomorphic_relative_homology_groups_Euclidean_complements sub) qed (auto simp: L) moreover have "continuous_map (powertop_real UNIV) euclideanreal (\x. x n)" by (metis (no_types) UNIV_I continuous_map_product_projection) then have cm: "continuous_map (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n))) euclideanreal (\x. x n)" by (simp add: Euclidean_space_def continuous_map_from_subtopology) have False if "path_connected_space (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n)))" using path_connectedin_continuous_map_image [OF cm that [unfolded path_connectedin_topspace [symmetric]]] bounded_path_connected_Compl_real [of "{0}"] by (simp add: topspace_Euclidean_space image_def Bex_def non0 flip: path_connectedin_topspace) moreover have eq: "T = T \ {x. x n \ 0} \ T \ {x. x n \ 0}" for T :: "(nat \ real) set" by auto have "path_connectedin (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S)" proof (subst eq, rule path_connectedin_Un) have "topspace(Euclidean_space(Suc n)) \ {x. x n = 0} = topspace(Euclidean_space n)" apply (auto simp: topspace_Euclidean_space) by (metis Suc_leI inf.absorb_iff2 inf.orderE leI) let ?S = "topspace(Euclidean_space(Suc n)) \ {x. x n < 0}" show "path_connectedin (Euclidean_space (Suc n)) ((topspace (Euclidean_space (Suc n)) - S) \ {x. x n \ 0})" proof (rule lemma_iod) show "?S \ (topspace (Euclidean_space (Suc n)) - S) \ {x. x n \ 0}" using Ssub topspace_Euclidean_space by auto show "?S \ {}" apply (simp add: topspace_Euclidean_space set_eq_iff) apply (rule_tac x="(\i. 0)(n:= -1)" in exI) apply auto done fix a b and u::real assume "a \ ?S" "0 < u" "u < 1" "b \ (topspace (Euclidean_space (Suc n)) - S) \ {x. x n \ 0}" then show "(\i. (1 - u) * a i + u * b i) \ ?S" by (simp add: topspace_Euclidean_space add_neg_nonpos less_eq_real_def mult_less_0_iff) qed (simp add: topspace_Euclidean_space subset_iff) let ?T = "topspace(Euclidean_space(Suc n)) \ {x. x n > 0}" show "path_connectedin (Euclidean_space (Suc n)) ((topspace (Euclidean_space (Suc n)) - S) \ {x. 0 \ x n})" proof (rule lemma_iod) show "?T \ (topspace (Euclidean_space (Suc n)) - S) \ {x. 0 \ x n}" using Ssub topspace_Euclidean_space by auto show "?T \ {}" apply (simp add: topspace_Euclidean_space set_eq_iff) apply (rule_tac x="(\i. 0)(n:= 1)" in exI) apply auto done fix a b and u::real assume "a \ ?T" "0 < u" "u < 1" "b \ (topspace (Euclidean_space (Suc n)) - S) \ {x. 0 \ x n}" then show "(\i. (1 - u) * a i + u * b i) \ ?T" by (simp add: topspace_Euclidean_space add_pos_nonneg) qed (simp add: topspace_Euclidean_space subset_iff) show "(topspace (Euclidean_space (Suc n)) - S) \ {x. x n \ 0} \ ((topspace (Euclidean_space (Suc n)) - S) \ {x. 0 \ x n}) \ {}" using that apply (auto simp: Set.set_eq_iff topspace_Euclidean_space) by (metis Suc_leD order_refl) qed then have "path_connected_space (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S))" apply (simp add: path_connectedin_subtopology flip: path_connectedin_topspace) by (metis Int_Diff inf_idem) ultimately show ?thesis using isomorphic_homology_imp_path_connectedness by blast qed ultimately show ?rhs by blast qed (simp add: homeomorphic_space_refl) lemma isomorphic_homology_groups_Euclidean_complements: assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T" "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)" shows "homology_group p (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - S)) \ homology_group p (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - T))" proof (rule isomorphic_relative_contractible_space_imp_homology_groups) show "topspace (Euclidean_space n) - S \ topspace (Euclidean_space n)" using assms homeomorphic_space_sym invariance_of_dimension_closedin_Euclidean_space subtopology_superset by fastforce show "topspace (Euclidean_space n) - T \ topspace (Euclidean_space n)" using assms invariance_of_dimension_closedin_Euclidean_space subtopology_superset by force show "(topspace (Euclidean_space n) - S = {}) = (topspace (Euclidean_space n) - T = {})" by (metis Diff_eq_empty_iff assms closedin_subset homeomorphic_space_sym invariance_of_dimension_closedin_Euclidean_space subset_antisym subtopology_topspace) show "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - S) \ relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - T)" for p using assms isomorphic_relative_homology_groups_Euclidean_complements by blast qed auto lemma eqpoll_path_components_Euclidean_complements: assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T" "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)" shows "path_components_of (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - S)) \ path_components_of (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - T))" by (simp add: assms isomorphic_homology_groups_Euclidean_complements isomorphic_homology_imp_path_components) lemma path_connectedin_Euclidean_complements: assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T" "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)" shows "path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - S) \ path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - T)" by (meson Diff_subset assms isomorphic_homology_groups_Euclidean_complements isomorphic_homology_imp_path_connectedness path_connectedin_def) lemma eqpoll_connected_components_Euclidean_complements: assumes S: "closedin (Euclidean_space n) S" and T: "closedin (Euclidean_space n) T" and ST: "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)" shows "connected_components_of (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - S)) \ connected_components_of (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - T))" using eqpoll_path_components_Euclidean_complements [OF assms] by (metis S T closedin_def locally_path_connected_Euclidean_space locally_path_connected_space_open_subset path_components_eq_connected_components_of) lemma connected_in_Euclidean_complements: assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T" "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)" shows "connectedin (Euclidean_space n) (topspace(Euclidean_space n) - S) \ connectedin (Euclidean_space n) (topspace(Euclidean_space n) - T)" apply (simp add: connectedin_def connected_space_iff_components_subset_singleton subset_singleton_iff_lepoll) using eqpoll_connected_components_Euclidean_complements [OF assms] by (meson eqpoll_sym lepoll_trans1) theorem invariance_of_dimension_Euclidean_space: "Euclidean_space m homeomorphic_space Euclidean_space n \ m = n" proof (cases m n rule: linorder_cases) case less then have *: "topspace (Euclidean_space m) \ topspace (Euclidean_space n)" by (meson le_cases not_le subset_Euclidean_space) then have "Euclidean_space m = subtopology (Euclidean_space n) (topspace(Euclidean_space m))" by (simp add: Euclidean_space_def inf.absorb_iff2 subtopology_subtopology) then show ?thesis by (metis (no_types, lifting) * Euclidean_space_def closedin_Euclidean_space closedin_closed_subtopology eq_iff invariance_of_dimension_closedin_Euclidean_space subset_Euclidean_space topspace_Euclidean_space) next case equal then show ?thesis by (simp add: homeomorphic_space_refl) next case greater then have *: "topspace (Euclidean_space n) \ topspace (Euclidean_space m)" by (meson le_cases not_le subset_Euclidean_space) then have "Euclidean_space n = subtopology (Euclidean_space m) (topspace(Euclidean_space n))" by (simp add: Euclidean_space_def inf.absorb_iff2 subtopology_subtopology) then show ?thesis by (metis (no_types, lifting) "*" Euclidean_space_def closedin_Euclidean_space closedin_closed_subtopology eq_iff homeomorphic_space_sym invariance_of_dimension_closedin_Euclidean_space subset_Euclidean_space topspace_Euclidean_space) qed lemma biglemma: assumes "n \ 0" and S: "compactin (Euclidean_space n) S" and cmh: "continuous_map (subtopology (Euclidean_space n) S) (Euclidean_space n) h" and "inj_on h S" shows "path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - h ` S) \ path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - S)" proof (rule path_connectedin_Euclidean_complements) have hS_sub: "h ` S \ topspace(Euclidean_space n)" by (metis (no_types) S cmh compactin_subspace continuous_map_image_subset_topspace topspace_subtopology_subset) show clo_S: "closedin (Euclidean_space n) S" using assms by (simp add: continuous_map_in_subtopology Hausdorff_Euclidean_space compactin_imp_closedin) show clo_hS: "closedin (Euclidean_space n) (h ` S)" using Hausdorff_Euclidean_space S cmh compactin_absolute compactin_imp_closedin image_compactin by blast have "homeomorphic_map (subtopology (Euclidean_space n) S) (subtopology (Euclidean_space n) (h ` S)) h" proof (rule continuous_imp_homeomorphic_map) show "compact_space (subtopology (Euclidean_space n) S)" by (simp add: S compact_space_subtopology) show "Hausdorff_space (subtopology (Euclidean_space n) (h ` S))" using hS_sub by (simp add: Hausdorff_Euclidean_space Hausdorff_space_subtopology) show "continuous_map (subtopology (Euclidean_space n) S) (subtopology (Euclidean_space n) (h ` S)) h" using cmh continuous_map_in_subtopology by fastforce show "h ` topspace (subtopology (Euclidean_space n) S) = topspace (subtopology (Euclidean_space n) (h ` S))" using clo_hS clo_S closedin_subset by auto show "inj_on h (topspace (subtopology (Euclidean_space n) S))" by (metis \inj_on h S\ clo_S closedin_def topspace_subtopology_subset) qed then show "subtopology (Euclidean_space n) (h ` S) homeomorphic_space subtopology (Euclidean_space n) S" using homeomorphic_space homeomorphic_space_sym by blast qed lemma lemmaIOD: assumes "\T. T \ U \ c \ T" "\T. T \ U \ d \ T" "\U = c \ d" "\T. T \ U \ T \ {}" "pairwise disjnt U" "~(\T. U \ {T})" shows "c \ U" using assms apply safe subgoal for C' D' proof (cases "C'=D'") show "c \ U" if UU: "\ U = c \ d" and U: "\T. T \ U \ T \ {}" "disjoint U" and "\T. U \ {T}" "c \ C'" "D' \ U" "d \ D'" "C' = D'" proof - have "c \ d = D'" using Union_upper sup_mono UU that(5) that(6) that(7) that(8) by auto then have "\U = D'" by (simp add: UU) with U have "U = {D'}" by (metis (no_types, lifting) disjnt_Union1 disjnt_self_iff_empty insertCI pairwiseD subset_iff that(4) that(6)) then show ?thesis using that(4) by auto qed show "c \ U" if "\ U = c \ d""disjoint U" "C' \ U" "c \ C'""D' \ U" "d \ D'" "C' \ D'" proof - have "C' \ D' = {}" using \disjoint U\ \C' \ U\ \D' \ U\ \C' \ D'\unfolding disjnt_iff pairwise_def by blast then show ?thesis using subset_antisym that(1) \C' \ U\ \c \ C'\ \d \ D'\ by fastforce qed qed done theorem invariance_of_domain_Euclidean_space: assumes U: "openin (Euclidean_space n) U" and cmf: "continuous_map (subtopology (Euclidean_space n) U) (Euclidean_space n) f" and "inj_on f U" shows "openin (Euclidean_space n) (f ` U)" (is "openin ?E (f ` U)") proof (cases "n = 0") case True have [simp]: "Euclidean_space 0 = discrete_topology {\i. 0}" by (auto simp: subtopology_eq_discrete_topology_sing topspace_Euclidean_space) show ?thesis using cmf True U by auto next case False define enorm where "enorm \ \x. sqrt(\ii. if i = k then d else 0) = (if k < n then \d\ else 0)" for k d using \n \ 0\ by (auto simp: enorm_def power2_eq_square if_distrib [of "\x. x * _"] cong: if_cong) define zero::"nat\real" where "zero \ \i. 0" have zero_in [simp]: "zero \ topspace ?E" using False by (simp add: zero_def topspace_Euclidean_space) have enorm_eq_0 [simp]: "enorm x = 0 \ x = zero" if "x \ topspace(Euclidean_space n)" for x using that unfolding zero_def enorm_def apply (simp add: sum_nonneg_eq_0_iff fun_eq_iff topspace_Euclidean_space) using le_less_linear by blast have [simp]: "enorm zero = 0" by (simp add: zero_def enorm_def) have cm_enorm: "continuous_map ?E euclideanreal enorm" unfolding enorm_def proof (intro continuous_intros) show "continuous_map ?E euclideanreal (\x. x i)" if "i \ {.. enorm x" for x by (auto simp: enorm_def sum_nonneg) have le_enorm: "\x i\ \ enorm x" if "i < n" for i x proof - have "\x i\ \ sqrt (\k\{i}. (x k)\<^sup>2)" by auto also have "\ \ sqrt (\k2)" by (rule real_sqrt_le_mono [OF sum_mono2]) (use that in auto) finally show ?thesis by (simp add: enorm_def) qed define B where "B \ \r. {x \ topspace ?E. enorm x < r}" define C where "C \ \r. {x \ topspace ?E. enorm x \ r}" define S where "S \ \r. {x \ topspace ?E. enorm x = r}" have BC: "B r \ C r" and SC: "S r \ C r" and disjSB: "disjnt (S r) (B r)" and eqC: "B r \ S r = C r" for r by (auto simp: B_def C_def S_def disjnt_def) consider "n = 1" | "n \ 2" using False by linarith then have **: "openin ?E (h ` (B r))" if "r > 0" and cmh: "continuous_map(subtopology ?E (C r)) ?E h" and injh: "inj_on h (C r)" for r h proof cases case 1 define e :: "[real,nat]\real" where "e \ \x i. if i = 0 then x else 0" define e' :: "(nat\real)\real" where "e' \ \x. x 0" have "continuous_map euclidean euclideanreal (\f. f (0::nat))" by auto then have "continuous_map (subtopology (powertop_real UNIV) {f. \n\Suc 0. f n = 0}) euclideanreal (\f. f 0)" by (metis (mono_tags) continuous_map_from_subtopology euclidean_product_topology) then have hom_ee': "homeomorphic_maps euclideanreal (Euclidean_space 1) e e'" by (auto simp: homeomorphic_maps_def e_def e'_def continuous_map_in_subtopology Euclidean_space_def) have eBr: "e ` {-r<..x. x * _"] cong: if_cong) have in_Cr: "\x. \-r < x; x < r\ \ (\i. if i = 0 then x else 0) \ C r" using \n \ 0\ by (auto simp: C_def topspace_Euclidean_space) have inj: "inj_on (e' \ h \ e) {- r<..i. if i = 0 then x else 0) 0 = h (\i. if i = 0 then y else 0) 0" and "-r < x" "x < r" "-r < y" "y < r" for x y :: real proof - have x: "(\i. if i = 0 then x else 0) \ C r" and y: "(\i. if i = 0 then y else 0) \ C r" by (blast intro: inj_onD [OF \inj_on h (C r)\] that in_Cr)+ have "continuous_map (subtopology (Euclidean_space (Suc 0)) (C r)) (Euclidean_space (Suc 0)) h" using cmh by (simp add: 1) then have "h ` ({x. \i\Suc 0. x i = 0} \ C r) \ {x. \i\Suc 0. x i = 0}" by (force simp: Euclidean_space_def subtopology_subtopology continuous_map_def) have "h (\i. if i = 0 then x else 0) j = h (\i. if i = 0 then y else 0) j" for j proof (cases j) case (Suc j') have "h ` ({x. \i\Suc 0. x i = 0} \ C r) \ {x. \i\Suc 0. x i = 0}" using continuous_map_image_subset_topspace [OF cmh] by (simp add: 1 Euclidean_space_def subtopology_subtopology) with Suc f x y show ?thesis by (simp add: "1" image_subset_iff) qed (use f in blast) then have "(\i. if i = 0 then x else 0) = (\i::nat. if i = 0 then y else 0)" by (blast intro: inj_onD [OF \inj_on h (C r)\] that in_Cr) then show ?thesis by (simp add: fun_eq_iff) presburger qed qed have hom_e': "homeomorphic_map (Euclidean_space 1) euclideanreal e'" using hom_ee' homeomorphic_maps_map by blast have "openin (Euclidean_space n) (h ` e ` {- r<.. topspace (Euclidean_space 1)" using "1" C_def \\r. B r \ C r\ cmh continuous_map_image_subset_topspace eBr by fastforce have cont: "continuous_on {- r<.. h \ e)" proof (intro continuous_on_compose) have "\i. continuous_on {- r<..x. if i = 0 then x else 0)" by (auto simp: continuous_on_topological) then show "continuous_on {- r<.. topspace (subtopology ?E (C r))" by (auto simp: eBr \\r. B r \ C r\) (auto simp: B_def) with cmh show "continuous_on (e ` {- r<.. topspace ?E" using subCr cmh by (simp add: continuous_map_def image_subset_iff) moreover have "continuous_on (topspace ?E) e'" by (metis "1" continuous_map_Euclidean_space_iff hom_ee' homeomorphic_maps_def) ultimately show "continuous_on (h ` e ` {- r<..r. closedin (Euclidean_space n) (C r)" unfolding C_def by (rule closedin_continuous_map_preimage [OF cm_enorm, of concl: "{.._}", simplified]) have cloS: "\r. closedin (Euclidean_space n) (S r)" unfolding S_def by (rule closedin_continuous_map_preimage [OF cm_enorm, of concl: "{_}", simplified]) have C_subset: "C r \ UNIV \\<^sub>E {- \r\..\r\}" using le_enorm \r > 0\ apply (auto simp: C_def topspace_Euclidean_space abs_le_iff) apply (metis add.inverse_neutral le_cases less_minus_iff not_le order_trans) by (metis enorm_ge0 not_le order.trans) have compactinC: "compactin (Euclidean_space n) (C r)" unfolding Euclidean_space_def compactin_subtopology proof show "compactin (powertop_real UNIV) (C r)" proof (rule closed_compactin [OF _ C_subset]) show "closedin (powertop_real UNIV) (C r)" by (metis Euclidean_space_def cloC closedin_Euclidean_space closedin_closed_subtopology topspace_Euclidean_space) qed (simp add: compactin_PiE) qed (auto simp: C_def topspace_Euclidean_space) have compactinS: "compactin (Euclidean_space n) (S r)" unfolding Euclidean_space_def compactin_subtopology proof show "compactin (powertop_real UNIV) (S r)" proof (rule closed_compactin) show "S r \ UNIV \\<^sub>E {- \r\..\r\}" using C_subset \\r. S r \ C r\ by blast show "closedin (powertop_real UNIV) (S r)" by (metis Euclidean_space_def cloS closedin_Euclidean_space closedin_closed_subtopology topspace_Euclidean_space) qed (simp add: compactin_PiE) qed (auto simp: S_def topspace_Euclidean_space) have h_if_B: "\y. y \ B r \ h y \ topspace ?E" using B_def \\r. B r \ S r = C r\ cmh continuous_map_image_subset_topspace by fastforce have com_hSr: "compactin (Euclidean_space n) (h ` S r)" by (meson \\r. S r \ C r\ cmh compactinS compactin_subtopology image_compactin) have ope_comp_hSr: "openin (Euclidean_space n) (topspace (Euclidean_space n) - h ` S r)" proof (rule openin_diff) show "closedin (Euclidean_space n) (h ` S r)" using Hausdorff_Euclidean_space com_hSr compactin_imp_closedin by blast qed auto have h_pcs: "h ` (B r) \ path_components_of (subtopology ?E (topspace ?E - h ` (S r)))" proof (rule lemmaIOD) have pc_interval: "path_connectedin (Euclidean_space n) {x \ topspace(Euclidean_space n). enorm x \ T}" if T: "is_interval T" for T proof - define mul :: "[real, nat \ real, nat] \ real" where "mul \ \a x i. a * x i" let ?neg = "mul (-1)" have neg_neg [simp]: "?neg (?neg x) = x" for x by (simp add: mul_def) have enorm_mul [simp]: "enorm(mul a x) = abs a * enorm x" for a x by (simp add: enorm_def mul_def power_mult_distrib) (metis real_sqrt_abs real_sqrt_mult sum_distrib_left) have mul_in_top: "mul a x \ topspace ?E" if "x \ topspace ?E" for a x using mul_def that topspace_Euclidean_space by auto have neg_in_S: "?neg x \ S r" if "x \ S r" for x r using that topspace_Euclidean_space S_def by simp (simp add: mul_def) have *: "path_connectedin ?E (S d)" if "d \ 0" for d proof (cases "d = 0") let ?ES = "subtopology ?E (S d)" case False then have "d > 0" using that by linarith moreover have "path_connected_space ?ES" unfolding path_connected_space_iff_path_component proof clarify have **: "path_component_of ?ES x y" if x: "x \ topspace ?ES" and y: "y \ topspace ?ES" "x \ ?neg y" for x y proof - show ?thesis unfolding path_component_of_def pathin_def S_def proof (intro exI conjI) let ?g = "(\x. mul (d / enorm x) x) \ (\t i. (1 - t) * x i + t * y i)" show "continuous_map (top_of_set {0::real..1}) (subtopology ?E {x \ topspace ?E. enorm x = d}) ?g" proof (rule continuous_map_compose) let ?Y = "subtopology ?E (- {zero})" have **: False if eq0: "\j. (1 - r) * x j + r * y j = 0" and ne: "x i \ - y i" and d: "enorm x = d" "enorm y = d" and r: "0 \ r" "r \ 1" for i r proof - have "mul (1-r) x = ?neg (mul r y)" using eq0 by (simp add: mul_def fun_eq_iff algebra_simps) then have "enorm (mul (1-r) x) = enorm (?neg (mul r y))" by metis with r have "(1-r) * enorm x = r * enorm y" by simp then have r12: "r = 1/2" using \d \ 0\ d by auto show ?thesis using ne eq0 [of i] unfolding r12 by (simp add: algebra_simps) qed show "continuous_map (top_of_set {0..1}) ?Y (\t i. (1 - t) * x i + t * y i)" using x y unfolding continuous_map_componentwise_UNIV Euclidean_space_def continuous_map_in_subtopology apply (intro conjI allI continuous_intros) apply (auto simp: zero_def mul_def S_def Euclidean_space_def fun_eq_iff) using ** by blast have cm_enorm': "continuous_map (subtopology (powertop_real UNIV) A) euclideanreal enorm" for A unfolding enorm_def by (intro continuous_intros) auto have "continuous_map ?Y (subtopology ?E {x. enorm x = d}) (\x. mul (d / enorm x) x)" unfolding continuous_map_in_subtopology proof (intro conjI) show "continuous_map ?Y (Euclidean_space n) (\x. mul (d / enorm x) x)" unfolding continuous_map_in_subtopology Euclidean_space_def mul_def zero_def subtopology_subtopology continuous_map_componentwise_UNIV proof (intro conjI allI cm_enorm' continuous_intros) show "enorm x \ 0" if "x \ topspace (subtopology (powertop_real UNIV) ({x. \i\n. x i = 0} \ - {\i. 0}))" for x using that by simp (metis abs_le_zero_iff le_enorm not_less) qed auto qed (use \d > 0\ enorm_ge0 in auto) moreover have "subtopology ?E {x \ topspace ?E. enorm x = d} = subtopology ?E {x. enorm x = d}" by (simp add: subtopology_restrict Collect_conj_eq) ultimately show "continuous_map ?Y (subtopology (Euclidean_space n) {x \ topspace (Euclidean_space n). enorm x = d}) (\x. mul (d / enorm x) x)" by metis qed show "?g (0::real) = x" "?g (1::real) = y" using that by (auto simp: S_def zero_def mul_def fun_eq_iff) qed qed obtain a b where a: "a \ topspace ?ES" and b: "b \ topspace ?ES" and "a \ b" and negab: "?neg a \ b" proof let ?v = "\j i::nat. if i = j then d else 0" show "?v 0 \ topspace (subtopology ?E (S d))" "?v 1 \ topspace (subtopology ?E (S d))" using \n \ 2\ \d \ 0\ by (auto simp: S_def topspace_Euclidean_space) show "?v 0 \ ?v 1" "?neg (?v 0) \ (?v 1)" using \d > 0\ by (auto simp: mul_def fun_eq_iff) qed show "path_component_of ?ES x y" if x: "x \ topspace ?ES" and y: "y \ topspace ?ES" for x y proof - have "path_component_of ?ES x (?neg x)" proof - have "path_component_of ?ES x a" by (metis (no_types, opaque_lifting) ** a b \a \ b\ negab path_component_of_trans path_component_of_sym x) moreover have pa_ab: "path_component_of ?ES a b" using "**" a b negab neg_neg by blast then have "path_component_of ?ES a (?neg x)" by (metis "**" \a \ b\ cloS closedin_def neg_in_S path_component_of_equiv topspace_subtopology_subset x) ultimately show ?thesis by (meson path_component_of_trans) qed then show ?thesis using "**" x y by force qed qed ultimately show ?thesis by (simp add: cloS closedin_subset path_connectedin_def) qed (simp add: S_def cong: conj_cong) have "path_component_of (subtopology ?E {x \ topspace ?E. enorm x \ T}) x y" if "enorm x = a" "x \ topspace ?E" "enorm x \ T" "enorm y = b" "y \ topspace ?E" "enorm y \ T" for x y a b using that proof (induction a b arbitrary: x y rule: linorder_less_wlog) case (less a b) then have "a \ 0" using enorm_ge0 by blast with less.hyps have "b > 0" by linarith show ?case proof (rule path_component_of_trans) have y'_ts: "mul (a / b) y \ topspace ?E" using \y \ topspace ?E\ mul_in_top by blast moreover have "enorm (mul (a / b) y) = a" unfolding enorm_mul using \0 < b\ \0 \ a\ less.prems by simp ultimately have y'_S: "mul (a / b) y \ S a" using S_def by blast have "x \ S a" using S_def less.prems by blast with \x \ topspace ?E\ y'_ts y'_S have "path_component_of (subtopology ?E (S a)) x (mul (a / b) y)" by (metis * [OF \a \ 0\] path_connected_space_iff_path_component path_connectedin_def topspace_subtopology_subset) moreover have "{f \ topspace ?E. enorm f = a} \ {f \ topspace ?E. enorm f \ T}" using \enorm x = a\ \enorm x \ T\ by force ultimately show "path_component_of (subtopology ?E {x. x \ topspace ?E \ enorm x \ T}) x (mul (a / b) y)" by (simp add: S_def path_component_of_mono) have "pathin ?E (\t. mul (((1 - t) * b + t * a) / b) y)" using \b > 0\ \y \ topspace ?E\ unfolding pathin_def Euclidean_space_def mul_def continuous_map_in_subtopology continuous_map_componentwise_UNIV by (intro allI conjI continuous_intros) auto moreover have "mul (((1 - t) * b + t * a) / b) y \ topspace ?E" if "t \ {0..1}" for t using \y \ topspace ?E\ mul_in_top by blast moreover have "enorm (mul (((1 - t) * b + t * a) / b) y) \ T" if "t \ {0..1}" for t proof - have "a \ T" "b \ T" using less.prems by auto then have "\(1 - t) * b + t * a\ \ T" proof (rule mem_is_interval_1_I [OF T]) show "a \ \(1 - t) * b + t * a\" using that \a \ 0\ less.hyps segment_bound_lemma by auto show "\(1 - t) * b + t * a\ \ b" using that \a \ 0\ less.hyps by (auto intro: convex_bound_le) qed then show ?thesis unfolding enorm_mul \enorm y = b\ using that \b > 0\ by simp qed ultimately have pa: "pathin (subtopology ?E {x \ topspace ?E. enorm x \ T}) (\t. mul (((1 - t) * b + t * a) / b) y)" by (auto simp: pathin_subtopology) have ex_pathin: "\g. pathin (subtopology ?E {x \ topspace ?E. enorm x \ T}) g \ g 0 = y \ g 1 = mul (a / b) y" apply (rule_tac x="\t. mul (((1 - t) * b + t * a) / b) y" in exI) using \b > 0\ pa by (auto simp: mul_def) show "path_component_of (subtopology ?E {x. x \ topspace ?E \ enorm x \ T}) (mul (a / b) y) y" by (rule path_component_of_sym) (simp add: path_component_of_def ex_pathin) qed next case (refl a) then have pc: "path_component_of (subtopology ?E (S (enorm u))) u v" if "u \ topspace ?E \ S (enorm x)" "v \ topspace ?E \ S (enorm u)" for u v using * [of a] enorm_ge0 that by (auto simp: path_connectedin_def path_connected_space_iff_path_component S_def) have sub: "{u \ topspace ?E. enorm u = enorm x} \ {u \ topspace ?E. enorm u \ T}" using \enorm x \ T\ by auto show ?case using pc [of x y] refl by (auto simp: S_def path_component_of_mono [OF _ sub]) next case (sym a b) then show ?case by (blast intro: path_component_of_sym) qed then show ?thesis by (simp add: path_connectedin_def path_connected_space_iff_path_component) qed have "h ` S r \ topspace ?E" by (meson SC cmh compact_imp_compactin_subtopology compactinS compactin_subset_topspace image_compactin) moreover have "\ compact_space ?E " by (metis compact_Euclidean_space \n \ 0\) then have "\ compactin ?E (topspace ?E)" by (simp add: compact_space_def topspace_Euclidean_space) then have "h ` S r \ topspace ?E" using com_hSr by auto ultimately have top_hSr_ne: "topspace (subtopology ?E (topspace ?E - h ` S r)) \ {}" by auto show pc1: "\T. T \ path_components_of (subtopology ?E (topspace ?E - h ` S r)) \ h ` B r \ T" proof (rule exists_path_component_of_superset [OF _ top_hSr_ne]) have "path_connectedin ?E (h ` B r)" proof (rule path_connectedin_continuous_map_image) show "continuous_map (subtopology ?E (C r)) ?E h" by (simp add: cmh) have "path_connectedin ?E (B r)" using pc_interval[of "{.. topspace ?E - h ` S r" apply (auto simp: h_if_B) by (metis BC SC disjSB disjnt_iff inj_onD [OF injh] subsetD) ultimately show "path_connectedin (subtopology ?E (topspace ?E - h ` S r)) (h ` B r)" by (simp add: path_connectedin_subtopology) qed metis show "\T. T \ path_components_of (subtopology ?E (topspace ?E - h ` S r)) \ topspace ?E - h ` (C r) \ T" proof (rule exists_path_component_of_superset [OF _ top_hSr_ne]) have eq: "topspace ?E - {x \ topspace ?E. enorm x \ r} = {x \ topspace ?E. r < enorm x}" by auto have "path_connectedin ?E (topspace ?E - C r)" using pc_interval[of "{r<..}"] is_interval_convex_1 unfolding C_def eq by auto then have "path_connectedin ?E (topspace ?E - h ` C r)" by (metis biglemma [OF \n \ 0\ compactinC cmh injh]) then show "path_connectedin (subtopology ?E (topspace ?E - h ` S r)) (topspace ?E - h ` C r)" by (simp add: Diff_mono SC image_mono path_connectedin_subtopology) qed metis have "topspace ?E \ (topspace ?E - h ` S r) = h ` B r \ (topspace ?E - h ` C r)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using \\r. B r \ S r = C r\ by auto have "h ` B r \ h ` S r = {}" by (metis Diff_triv \\r. B r \ S r = C r\ \\r. disjnt (S r) (B r)\ disjnt_def inf_commute inj_on_Un injh) then show "?rhs \ ?lhs" using path_components_of_subset pc1 \\r. B r \ S r = C r\ by (fastforce simp add: h_if_B) qed then show "\ (path_components_of (subtopology ?E (topspace ?E - h ` S r))) = h ` B r \ (topspace ?E - h ` (C r))" by (simp add: Union_path_components_of) show "T \ {}" if "T \ path_components_of (subtopology ?E (topspace ?E - h ` S r))" for T using that by (simp add: nonempty_path_components_of) show "disjoint (path_components_of (subtopology ?E (topspace ?E - h ` S r)))" by (simp add: pairwise_disjoint_path_components_of) have "\ path_connectedin ?E (topspace ?E - h ` S r)" proof (subst biglemma [OF \n \ 0\ compactinS]) show "continuous_map (subtopology ?E (S r)) ?E h" by (metis Un_commute Un_upper1 cmh continuous_map_from_subtopology_mono eqC) show "inj_on h (S r)" using SC inj_on_subset injh by blast show "\ path_connectedin ?E (topspace ?E - S r)" proof have "topspace ?E - S r = {x \ topspace ?E. enorm x \ r}" by (auto simp: S_def) moreover have "enorm ` {x \ topspace ?E. enorm x \ r} = {0..} - {r}" proof have "\x. x \ topspace ?E \ enorm x \ r \ d = enorm x" if "d \ r" "d \ 0" for d proof (intro exI conjI) show "(\i. if i = 0 then d else 0) \ topspace ?E" using \n \ 0\ by (auto simp: Euclidean_space_def) show "enorm (\i. if i = 0 then d else 0) \ r" "d = enorm (\i. if i = 0 then d else 0)" using \n \ 0\ that by simp_all qed then show "{0..} - {r} \ enorm ` {x \ topspace ?E. enorm x \ r}" by (auto simp: image_def) qed (auto simp: enorm_ge0) ultimately have non_r: "enorm ` (topspace ?E - S r) = {0..} - {r}" by simp have "\x\0. x \ r \ r \ x" by (metis gt_ex le_cases not_le order_trans) then have "\ is_interval ({0..} - {r})" unfolding is_interval_1 using \r > 0\ by (auto simp: Bex_def) then show False if "path_connectedin ?E (topspace ?E - S r)" using path_connectedin_continuous_map_image [OF cm_enorm that] by (simp add: is_interval_path_connected_1 non_r) qed qed then have "\ path_connected_space (subtopology ?E (topspace ?E - h ` S r))" by (simp add: path_connectedin_def) then show "\T. path_components_of (subtopology ?E (topspace ?E - h ` S r)) \ {T}" by (simp add: path_components_of_subset_singleton) qed moreover have "openin ?E A" if "A \ path_components_of (subtopology ?E (topspace ?E - h ` (S r)))" for A using locally_path_connected_Euclidean_space [of n] that ope_comp_hSr by (simp add: locally_path_connected_space_open_path_components) ultimately show ?thesis by metis qed have "\T. openin ?E T \ f x \ T \ T \ f ` U" if "x \ U" for x proof - have x: "x \ topspace ?E" by (meson U in_mono openin_subset that) obtain V where V: "openin (powertop_real UNIV) V" and Ueq: "U = V \ {x. \i\n. x i = 0}" using U by (auto simp: openin_subtopology Euclidean_space_def) with \x \ U\ have "x \ V" by blast then obtain T where Tfin: "finite {i. T i \ UNIV}" and Topen: "\i. open (T i)" and Tx: "x \ Pi\<^sub>E UNIV T" and TV: "Pi\<^sub>E UNIV T \ V" using V by (force simp: openin_product_topology_alt) have "\e>0. \x'. \x' - x i\ < e \ x' \ T i" for i using Topen [of i] Tx by (auto simp: open_real) then obtain \ where B0: "\i. \ i > 0" and BT: "\i x'. \x' - x i\ < \ i \ x' \ T i" by metis define r where "r \ Min (insert 1 (\ ` {i. T i \ UNIV}))" have "r > 0" by (simp add: B0 Tfin r_def) have inU: "y \ U" if y: "y \ topspace ?E" and yxr: "\i. i \y i - x i\ < r" for y proof - have "y i \ T i" for i proof (cases "T i = UNIV") show "y i \ T i" if "T i \ UNIV" proof (cases "i < n") case True then show ?thesis using yxr [OF True] that by (simp add: r_def BT Tfin) next case False then show ?thesis using B0 Ueq \x \ U\ topspace_Euclidean_space y by (force intro: BT) qed qed auto with TV have "y \ V" by auto then show ?thesis using that by (auto simp: Ueq topspace_Euclidean_space) qed have xinU: "(\i. x i + y i) \ U" if "y \ C(r/2)" for y proof (rule inU) have y: "y \ topspace ?E" using C_def that by blast show "(\i. x i + y i) \ topspace ?E" using x y by (simp add: topspace_Euclidean_space) have "enorm y \ r/2" using that by (simp add: C_def) then show "\x i + y i - x i\ < r" if "i < n" for i using le_enorm enorm_ge0 that \0 < r\ leI order_trans by fastforce qed show ?thesis proof (intro exI conjI) show "openin ?E ((f \ (\y i. x i + y i)) ` B (r/2))" proof (rule **) have "continuous_map (subtopology ?E (C(r/2))) (subtopology ?E U) (\y i. x i + y i)" by (auto simp: xinU continuous_map_in_subtopology intro!: continuous_intros continuous_map_Euclidean_space_add x) then show "continuous_map (subtopology ?E (C(r/2))) ?E (f \ (\y i. x i + y i))" by (rule continuous_map_compose) (simp add: cmf) show "inj_on (f \ (\y i. x i + y i)) (C(r/2))" proof (clarsimp simp add: inj_on_def C_def topspace_Euclidean_space simp del: divide_const_simps) show "y' = y" if ey: "enorm y \ r / 2" and ey': "enorm y' \ r / 2" and y0: "\i\n. y i = 0" and y'0: "\i\n. y' i = 0" and feq: "f (\i. x i + y' i) = f (\i. x i + y i)" for y' y :: "nat \ real" proof - have "(\i. x i + y i) \ U" proof (rule inU) show "(\i. x i + y i) \ topspace ?E" using topspace_Euclidean_space x y0 by auto show "\x i + y i - x i\ < r" if "i < n" for i using ey le_enorm [of _ y] \r > 0\ that by fastforce qed moreover have "(\i. x i + y' i) \ U" proof (rule inU) show "(\i. x i + y' i) \ topspace ?E" using topspace_Euclidean_space x y'0 by auto show "\x i + y' i - x i\ < r" if "i < n" for i using ey' le_enorm [of _ y'] \r > 0\ that by fastforce qed ultimately have "(\i. x i + y' i) = (\i. x i + y i)" using feq by (meson \inj_on f U\ inj_on_def) then show ?thesis by (auto simp: fun_eq_iff) qed qed qed (simp add: \0 < r\) have "x \ (\y i. x i + y i) ` B (r / 2)" proof show "x = (\i. x i + zero i)" by (simp add: zero_def) qed (auto simp: B_def \r > 0\) then show "f x \ (f \ (\y i. x i + y i)) ` B (r/2)" by (metis image_comp image_eqI) show "(f \ (\y i. x i + y i)) ` B (r/2) \ f ` U" using \\r. B r \ C r\ xinU by fastforce qed qed then show ?thesis using openin_subopen by force qed corollary invariance_of_domain_Euclidean_space_embedding_map: assumes "openin (Euclidean_space n) U" and cmf: "continuous_map(subtopology (Euclidean_space n) U) (Euclidean_space n) f" and "inj_on f U" shows "embedding_map(subtopology (Euclidean_space n) U) (Euclidean_space n) f" proof (rule injective_open_imp_embedding_map [OF cmf]) show "open_map (subtopology (Euclidean_space n) U) (Euclidean_space n) f" unfolding open_map_def by (meson assms continuous_map_from_subtopology_mono inj_on_subset invariance_of_domain_Euclidean_space openin_imp_subset openin_trans_full) show "inj_on f (topspace (subtopology (Euclidean_space n) U))" using assms openin_subset topspace_subtopology_subset by fastforce qed corollary invariance_of_domain_Euclidean_space_gen: assumes "n \ m" and U: "openin (Euclidean_space m) U" and cmf: "continuous_map(subtopology (Euclidean_space m) U) (Euclidean_space n) f" and "inj_on f U" shows "openin (Euclidean_space n) (f ` U)" proof - have *: "Euclidean_space n = subtopology (Euclidean_space m) (topspace(Euclidean_space n))" by (metis Euclidean_space_def \n \ m\ inf.absorb_iff2 subset_Euclidean_space subtopology_subtopology topspace_Euclidean_space) moreover have "U \ topspace (subtopology (Euclidean_space m) U)" by (metis U inf.absorb_iff2 openin_subset openin_subtopology openin_topspace) ultimately show ?thesis by (metis (no_types) U \inj_on f U\ cmf continuous_map_in_subtopology inf.absorb_iff2 inf.orderE invariance_of_domain_Euclidean_space openin_imp_subset openin_subtopology openin_topspace) qed corollary invariance_of_domain_Euclidean_space_embedding_map_gen: assumes "n \ m" and U: "openin (Euclidean_space m) U" and cmf: "continuous_map(subtopology (Euclidean_space m) U) (Euclidean_space n) f" and "inj_on f U" shows "embedding_map(subtopology (Euclidean_space m) U) (Euclidean_space n) f" proof (rule injective_open_imp_embedding_map [OF cmf]) show "open_map (subtopology (Euclidean_space m) U) (Euclidean_space n) f" by (meson U \n \ m\ \inj_on f U\ cmf continuous_map_from_subtopology_mono invariance_of_domain_Euclidean_space_gen open_map_def openin_open_subtopology subset_inj_on) show "inj_on f (topspace (subtopology (Euclidean_space m) U))" using assms openin_subset topspace_subtopology_subset by fastforce qed subsection\Relating two variants of Euclidean space, one within product topology. \ proposition homeomorphic_maps_Euclidean_space_euclidean_gen_OLD: fixes B :: "'n::euclidean_space set" assumes "finite B" "independent B" and orth: "pairwise orthogonal B" and n: "card B = n" obtains f g where "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g" proof - note representation_basis [OF \independent B\, simp] obtain b where injb: "inj_on b {..finite B\] by (metis n card_Collect_less_nat card_image lessThan_def) then have biB: "\i. i < n \ b i \ B" by force have repr: "\v. v \ span B \ (\iR b i) = v" using real_vector.sum_representation_eq [OF \independent B\ _ \finite B\] by (metis (no_types, lifting) injb beq order_refl sum.reindex_cong) let ?f = "\x. \iR b i" let ?g = "\v i. if i < n then representation B v (b i) else 0" show thesis proof show "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) ?f ?g" unfolding homeomorphic_maps_def proof (intro conjI) have *: "continuous_map euclidean (top_of_set (span B)) ?f" by (metis (mono_tags) biB continuous_map_span_sum lessThan_iff) show "continuous_map (Euclidean_space n) (top_of_set (span B)) ?f" unfolding Euclidean_space_def by (rule continuous_map_from_subtopology) (simp add: euclidean_product_topology *) show "continuous_map (top_of_set (span B)) (Euclidean_space n) ?g" unfolding Euclidean_space_def by (auto simp: continuous_map_in_subtopology continuous_map_componentwise_UNIV continuous_on_representation \independent B\ biB orth pairwise_orthogonal_imp_finite) have [simp]: "\x i. i x i *\<^sub>R b i \ span B" by (simp add: biB span_base span_scale) have "representation B (?f x) (b j) = x j" if 0: "\i\n. x i = (0::real)" and "j < n" for x j proof - have "representation B (?f x) (b j) = (\iR b i) (b j))" by (subst real_vector.representation_sum) (auto simp add: \independent B\) also have "... = (\iix\topspace (Euclidean_space n). ?g (?f x) = x" by (auto simp: Euclidean_space_def) show "\y\topspace (top_of_set (span B)). ?f (?g y) = y" using repr by (auto simp: Euclidean_space_def) qed qed qed proposition homeomorphic_maps_Euclidean_space_euclidean_gen: fixes B :: "'n::euclidean_space set" assumes "independent B" and orth: "pairwise orthogonal B" and n: "card B = n" and 1: "\u. u \ B \ norm u = 1" obtains f g where "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g" and "\x. x \ topspace (Euclidean_space n) \ (norm (f x))\<^sup>2 = (\i2)" proof - note representation_basis [OF \independent B\, simp] have "finite B" using \independent B\ finiteI_independent by metis obtain b where injb: "inj_on b {..finite B\] by (metis n card_Collect_less_nat card_image lessThan_def) then have biB: "\i. i < n \ b i \ B" by force have "0 \ B" using \independent B\ dependent_zero by blast have [simp]: "b i \ b j = (if j = i then 1 else 0)" if "i < n" "j < n" for i j proof (cases "i = j") case True with 1 that show ?thesis by (auto simp: norm_eq_sqrt_inner biB) next case False then have "b i \ b j" by (meson inj_onD injb lessThan_iff that) then show ?thesis using orth by (auto simp: orthogonal_def pairwise_def norm_eq_sqrt_inner that biB) qed have [simp]: "\x i. i x i *\<^sub>R b i \ span B" by (simp add: biB span_base span_scale) have repr: "\v. v \ span B \ (\iR b i) = v" using real_vector.sum_representation_eq [OF \independent B\ _ \finite B\] by (metis (no_types, lifting) injb beq order_refl sum.reindex_cong) define f where "f \ \x. \iR b i" define g where "g \ \v i. if i < n then representation B v (b i) else 0" show thesis proof show "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g" unfolding homeomorphic_maps_def proof (intro conjI) have *: "continuous_map euclidean (top_of_set (span B)) f" unfolding f_def by (rule continuous_map_span_sum) (use biB \0 \ B\ in auto) show "continuous_map (Euclidean_space n) (top_of_set (span B)) f" unfolding Euclidean_space_def by (rule continuous_map_from_subtopology) (simp add: euclidean_product_topology *) show "continuous_map (top_of_set (span B)) (Euclidean_space n) g" unfolding Euclidean_space_def g_def by (auto simp: continuous_map_in_subtopology continuous_map_componentwise_UNIV continuous_on_representation \independent B\ biB orth pairwise_orthogonal_imp_finite) have "representation B (f x) (b j) = x j" if 0: "\i\n. x i = (0::real)" and "j < n" for x j proof - have "representation B (f x) (b j) = (\iR b i) (b j))" unfolding f_def by (subst real_vector.representation_sum) (auto simp add: \independent B\) also have "... = (\iindependent B\ biB representation_scale span_base) also have "... = (\ix\topspace (Euclidean_space n). g (f x) = x" by (auto simp: Euclidean_space_def f_def g_def) show "\y\topspace (top_of_set (span B)). f (g y) = y" using repr by (auto simp: Euclidean_space_def f_def g_def) qed show normeq: "(norm (f x))\<^sup>2 = (\i2)" if "x \ topspace (Euclidean_space n)" for x unfolding f_def dot_square_norm [symmetric] by (simp add: power2_eq_square inner_sum_left inner_sum_right if_distrib biB cong: if_cong) qed qed corollary homeomorphic_maps_Euclidean_space_euclidean: obtains f :: "(nat \ real) \ 'n::euclidean_space" and g where "homeomorphic_maps (Euclidean_space (DIM('n))) euclidean f g" by (force intro: homeomorphic_maps_Euclidean_space_euclidean_gen [OF independent_Basis orthogonal_Basis refl norm_Basis]) lemma homeomorphic_maps_nsphere_euclidean_sphere: fixes B :: "'n::euclidean_space set" assumes B: "independent B" and orth: "pairwise orthogonal B" and n: "card B = n" and "n \ 0" and 1: "\u. u \ B \ norm u = 1" obtains f :: "(nat \ real) \ 'n::euclidean_space" and g where "homeomorphic_maps (nsphere(n - 1)) (top_of_set (sphere 0 1 \ span B)) f g" proof - have "finite B" using \independent B\ finiteI_independent by metis obtain f g where fg: "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g" and normf: "\x. x \ topspace (Euclidean_space n) \ (norm (f x))\<^sup>2 = (\i2)" using homeomorphic_maps_Euclidean_space_euclidean_gen [OF B orth n 1] by blast obtain b where injb: "inj_on b {..finite B\] by (metis n card_Collect_less_nat card_image lessThan_def) then have biB: "\i. i < n \ b i \ B" by force have [simp]: "\i. i < n \ b i \ 0" using \independent B\ biB dependent_zero by fastforce have [simp]: "b i \ b j = (if j = i then (norm (b i))\<^sup>2 else 0)" if "i < n" "j < n" for i j proof (cases "i = j") case False then have "b i \ b j" by (meson inj_onD injb lessThan_iff that) then show ?thesis using orth by (auto simp: orthogonal_def pairwise_def norm_eq_sqrt_inner that biB) qed (auto simp: norm_eq_sqrt_inner) have [simp]: "Suc (n - Suc 0) = n" using Suc_pred \n \ 0\ by blast then have [simp]: "{..card B - Suc 0} = {..i2) = (1::real)" "x \ topspace (Euclidean_space n)" for x proof - have "norm (f x)^2 = 1" using normf that by (simp add: n) with that show ?thesis by (simp add: power2_eq_imp_eq) qed have "homeomorphic_maps (nsphere (n - 1)) (top_of_set (span B \ sphere 0 1)) f g" unfolding nsphere_def subtopology_subtopology [symmetric] proof (rule homeomorphic_maps_subtopologies_alt) show "homeomorphic_maps (Euclidean_space (Suc (n - 1))) (top_of_set (span B)) f g" using fg by (force simp add: ) show "f ` (topspace (Euclidean_space (Suc (n - 1))) \ {x. (\i\n - 1. (x i)\<^sup>2) = 1}) \ sphere 0 1" using n by (auto simp: image_subset_iff Euclidean_space_def 1) have "(\i\n - Suc 0. (g u i)\<^sup>2) = 1" if "u \ span B" and "norm (u::'n) = 1" for u proof - obtain v where [simp]: "u = f v" "v \ topspace (Euclidean_space n)" using fg unfolding homeomorphic_maps_map subset_iff by (metis \u \ span B\ homeomorphic_imp_surjective_map image_eqI topspace_euclidean_subtopology) then have [simp]: "g (f v) = v" by (meson fg homeomorphic_maps_map) have fv21: "norm (f v) ^ 2 = 1" using that by simp show ?thesis using that normf fv21 \v \ topspace (Euclidean_space n)\ n by force qed then show "g ` (topspace (top_of_set (span B)) \ sphere 0 1) \ {x. (\i\n - 1. (x i)\<^sup>2) = 1}" by auto qed then show "homeomorphic_maps (nsphere(n - 1)) (top_of_set (sphere 0 1 \ span B)) f g" by (simp add: inf_commute) qed qed subsection\ Invariance of dimension and domain\ lemma homeomorphic_maps_iff_homeomorphism [simp]: "homeomorphic_maps (top_of_set S) (top_of_set T) f g \ homeomorphism S T f g" unfolding homeomorphic_maps_def homeomorphism_def by force lemma homeomorphic_space_iff_homeomorphic [simp]: "(top_of_set S) homeomorphic_space (top_of_set T) \ S homeomorphic T" by (simp add: homeomorphic_def homeomorphic_space_def) lemma homeomorphic_subspace_Euclidean_space: fixes S :: "'a::euclidean_space set" assumes "subspace S" shows "top_of_set S homeomorphic_space Euclidean_space n \ dim S = n" proof - obtain B where B: "B \ S" "independent B" "span B = S" "card B = dim S" and orth: "pairwise orthogonal B" and 1: "\x. x \ B \ norm x = 1" by (metis assms orthonormal_basis_subspace) then have "finite B" by (simp add: pairwise_orthogonal_imp_finite) have "top_of_set S homeomorphic_space top_of_set (span B)" unfolding homeomorphic_space_iff_homeomorphic by (auto simp: assms B intro: homeomorphic_subspaces) also have "\ homeomorphic_space Euclidean_space (dim S)" unfolding homeomorphic_space_def using homeomorphic_maps_Euclidean_space_euclidean_gen [OF \independent B\ orth] homeomorphic_maps_sym 1 B by metis finally have "top_of_set S homeomorphic_space Euclidean_space (dim S)" . then show ?thesis using homeomorphic_space_sym homeomorphic_space_trans invariance_of_dimension_Euclidean_space by blast qed lemma homeomorphic_subspace_Euclidean_space_dim: fixes S :: "'a::euclidean_space set" assumes "subspace S" shows "top_of_set S homeomorphic_space Euclidean_space (dim S)" by (simp add: homeomorphic_subspace_Euclidean_space assms) lemma homeomorphic_subspaces_eq: fixes S T:: "'a::euclidean_space set" assumes "subspace S" "subspace T" shows "S homeomorphic T \ dim S = dim T" proof show "dim S = dim T" if "S homeomorphic T" proof - have "Euclidean_space (dim S) homeomorphic_space top_of_set S" using \subspace S\ homeomorphic_space_sym homeomorphic_subspace_Euclidean_space_dim by blast also have "\ homeomorphic_space top_of_set T" by (simp add: that) also have "\ homeomorphic_space Euclidean_space (dim T)" by (simp add: homeomorphic_subspace_Euclidean_space assms) finally have "Euclidean_space (dim S) homeomorphic_space Euclidean_space (dim T)" . then show ?thesis by (simp add: invariance_of_dimension_Euclidean_space) qed next show "S homeomorphic T" if "dim S = dim T" by (metis that assms homeomorphic_subspaces) qed lemma homeomorphic_affine_Euclidean_space: assumes "affine S" shows "top_of_set S homeomorphic_space Euclidean_space n \ aff_dim S = n" (is "?X homeomorphic_space ?E \ aff_dim S = n") proof (cases "S = {}") case True with assms show ?thesis using homeomorphic_empty_space nonempty_Euclidean_space by fastforce next case False then obtain a where "a \ S" by force have "(?X homeomorphic_space ?E) = (top_of_set (image (\x. -a + x) S) homeomorphic_space ?E)" proof show "top_of_set ((+) (- a) ` S) homeomorphic_space ?E" if "?X homeomorphic_space ?E" using that by (meson homeomorphic_space_iff_homeomorphic homeomorphic_space_sym homeomorphic_space_trans homeomorphic_translation) show "?X homeomorphic_space ?E" if "top_of_set ((+) (- a) ` S) homeomorphic_space ?E" using that by (meson homeomorphic_space_iff_homeomorphic homeomorphic_space_trans homeomorphic_translation) qed also have "\ \ aff_dim S = n" by (metis \a \ S\ aff_dim_eq_dim affine_diffs_subspace affine_hull_eq assms homeomorphic_subspace_Euclidean_space of_nat_eq_iff) finally show ?thesis . qed corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and VU: "dim V \ dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof - have "S \ U" using openin_imp_subset [OF ope] . have Uhom: "top_of_set U homeomorphic_space Euclidean_space (dim U)" and Vhom: "top_of_set V homeomorphic_space Euclidean_space (dim V)" by (simp_all add: assms homeomorphic_subspace_Euclidean_space_dim) then obtain \ \' where hom: "homeomorphic_maps (top_of_set U) (Euclidean_space (dim U)) \ \'" by (auto simp: homeomorphic_space_def) obtain \ \' where \: "homeomorphic_map (top_of_set V) (Euclidean_space (dim V)) \" and \'\: "\x\V. \' (\ x) = x" using Vhom by (auto simp: homeomorphic_space_def homeomorphic_maps_map) have "((\ \ f \ \') o \) ` S = (\ o f) ` S" proof (rule image_cong [OF refl]) show "(\ \ f \ \' \ \) x = (\ \ f) x" if "x \ S" for x using that unfolding o_def by (metis \S \ U\ hom homeomorphic_maps_map in_mono topspace_euclidean_subtopology) qed moreover have "openin (Euclidean_space (dim V)) ((\ \ f \ \') ` \ ` S)" proof (rule invariance_of_domain_Euclidean_space_gen [OF VU]) show "openin (Euclidean_space (dim U)) (\ ` S)" using homeomorphic_map_openness_eq hom homeomorphic_maps_map ope by blast show "continuous_map (subtopology (Euclidean_space (dim U)) (\ ` S)) (Euclidean_space (dim V)) (\ \ f \ \')" proof (intro continuous_map_compose) have "continuous_on ({x. \i\dim U. x i = 0} \ \ ` S) \'" if "continuous_on {x. \i\dim U. x i = 0} \'" using that by (force elim: continuous_on_subset) moreover have "\' ` ({x. \i\dim U. x i = 0} \ \ ` S) \ S" if "\x\U. \' (\ x) = x" using that \S \ U\ by fastforce ultimately show "continuous_map (subtopology (Euclidean_space (dim U)) (\ ` S)) (top_of_set S) \'" using hom unfolding homeomorphic_maps_def by (simp add: Euclidean_space_def subtopology_subtopology euclidean_product_topology) show "continuous_map (top_of_set S) (top_of_set V) f" by (simp add: contf fim) show "continuous_map (top_of_set V) (Euclidean_space (dim V)) \" by (simp add: \ homeomorphic_imp_continuous_map) qed show "inj_on (\ \ f \ \') (\ ` S)" using injf hom unfolding inj_on_def homeomorphic_maps_map by simp (metis \S \ U\ \'\ fim imageI subsetD) qed ultimately have "openin (Euclidean_space (dim V)) (\ ` f ` S)" by (simp add: image_comp) then show ?thesis by (simp add: fim homeomorphic_map_openness_eq [OF \]) qed lemma invariance_of_domain: fixes f :: "'a \ 'a::euclidean_space" assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV] assms by (force simp add: ) corollary invariance_of_dimension_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "dim U \ dim V" proof - have "False" if "dim V < dim U" proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" using homeomorphic_def by blast have "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" using fim homeomorphism_image1 homhk by fastforce then have "dim ((h \ f) ` S) \ dim T" by (rule dim_subset) also have "dim ((h \ f) ` S) = dim U" using \S \ {}\ \subspace U\ by (blast intro: dim_openin ope_hf) finally show False using \dim V < dim U\ \dim T = dim V\ by simp qed then show ?thesis using not_less by blast qed corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" "aff_dim V \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "dim ((+) (- b) ` V) \ dim ((+) (- a) ` U)" by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed then show ?thesis by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) qed auto corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "aff_dim U \ aff_dim V" proof - obtain a b where "a \ S" "a \ U" "b \ V" using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed (use \S \ {}\ in auto) then show ?thesis by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) qed corollary invariance_of_dimension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and "open S" and injf: "inj_on f S" and "S \ {}" shows "DIM('a) \ DIM('b)" using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms by auto corollary continuous_injective_image_subspace_dim_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace S" "subspace T" and contf: "continuous_on S f" and fim: "f ` S \ T" and injf: "inj_on f S" shows "dim S \ dim T" apply (rule invariance_of_dimension_subspaces [of S S _ f]) using assms by (auto simp: subspace_affine) lemma invariance_of_dimension_convex_domain: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "convex S" and contf: "continuous_on S f" and fim: "f ` S \ affine hull T" and injf: "inj_on f S" shows "aff_dim S \ aff_dim T" proof (cases "S = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" using fim rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast show "rel_interior S \ {}" by (simp add: False \convex S\ rel_interior_eq_empty) qed auto then show ?thesis by simp qed lemma homeomorphic_convex_sets_le: assumes "convex S" "S homeomorphic T" shows "aff_dim S \ aff_dim T" proof - obtain h k where homhk: "homeomorphism S T h k" using homeomorphic_def assms by blast show ?thesis proof (rule invariance_of_dimension_convex_domain [OF \convex S\]) show "continuous_on S h" using homeomorphism_def homhk by blast show "h ` S \ affine hull T" by (metis homeomorphism_def homhk hull_subset) show "inj_on h S" by (meson homeomorphism_apply1 homhk inj_on_inverseI) qed qed lemma homeomorphic_convex_sets: assumes "convex S" "convex T" "S homeomorphic T" shows "aff_dim S = aff_dim T" by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) lemma homeomorphic_convex_compact_sets_eq: assumes "convex S" "compact S" "convex T" "compact T" shows "S homeomorphic T \ aff_dim S = aff_dim T" by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) lemma invariance_of_domain_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" shows "open (f ` T)" apply (rule invariance_of_domain_gen [OF \open T\]) using assms apply (auto simp: elim: continuous_on_subset subset_inj_on) done lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" and gf: "\x. x \ S \ g(f x) = x" shows "continuous_on (f ` S) g" proof (clarsimp simp add: continuous_openin_preimage_eq) fix T :: "'a set" assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) have "openin (top_of_set (f ` S)) (f ` (S \ T))" proof (rule open_openin_trans [OF invariance_of_domain_gen]) show "inj_on f S" using inj_on_inverseI gf by auto show "open (f ` (S \ T))" by (meson \inj_on f S\ \open T\ assms(1-3) continuous_on_subset inf_le1 inj_on_subset invariance_of_domain_gen open_Int) qed (use assms in auto) then show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" by (simp add: eq) qed lemma invariance_of_domain_homeomorphism: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" obtains g where "homeomorphism S (f ` S) f g" proof show "homeomorphism S (f ` S) f (inv_into S f)" by (simp add: assms continuous_on_inverse_open homeomorphism_def) qed corollary invariance_of_domain_homeomorphic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" shows "S homeomorphic (f ` S)" using invariance_of_domain_homeomorphism [OF assms] by (meson homeomorphic_def) lemma continuous_image_subset_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "f ` (interior S) \ interior(f ` S)" proof (rule interior_maximal) show "f ` interior S \ f ` S" by (simp add: image_mono interior_subset) show "open (f ` interior S)" using assms by (auto simp: subset_inj_on interior_subset continuous_on_subset invariance_of_domain_gen) qed lemma homeomorphic_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)" shows "(interior S) homeomorphic (interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp have gim: "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp show "homeomorphism (interior S) (interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show "\x. x \ interior S \ g (f x) = x" by (meson \\x\S. f x \ T \ g (f x) = x\ subsetD interior_subset) have "interior T \ f ` interior S" proof fix x assume "x \ interior T" then have "g x \ interior S" using gim by blast then show "x \ f ` interior S" by (metis T \x \ interior T\ image_iff interior_subset subsetCE) qed then show "f ` interior S = interior T" using fim by blast show "continuous_on (interior S) f" by (metis interior_subset continuous_on_subset contf) show "\y. y \ interior T \ f (g y) = y" by (meson T subsetD interior_subset) have "interior S \ g ` interior T" proof fix x assume "x \ interior S" then have "f x \ interior T" using fim by blast then show "x \ g ` interior T" by (metis S \x \ interior S\ image_iff interior_subset subsetCE) qed then show "g ` interior T = interior S" using gim by blast show "continuous_on (interior T) g" by (metis interior_subset continuous_on_subset contg) qed qed lemma homeomorphic_open_imp_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done proposition homeomorphic_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") case True with assms show ?thesis by auto next case False then have "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior) done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)" shows "(frontier S) homeomorphic (frontier T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp then have fim: "f ` frontier S \ frontier T" apply (simp add: frontier_def) using continuous_image_subset_interior assms(2) assms(3) S by auto have "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp then have gim: "g ` frontier T \ frontier S" apply (simp add: frontier_def) using continuous_image_subset_interior T assms(2) assms(3) by auto show "homeomorphism (frontier S) (frontier T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ frontier S \ g (f x) = x" by (simp add: S assms(2) frontier_def) show fg: "\y. y \ frontier T \ f (g y) = y" by (simp add: T assms(3) frontier_def) have "frontier T \ f ` frontier S" proof fix x assume "x \ frontier T" then have "g x \ frontier S" using gim by blast then show "x \ f ` frontier S" by (metis fg \x \ frontier T\ imageI) qed then show "f ` frontier S = frontier T" using fim by blast show "continuous_on (frontier S) f" by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def) have "frontier S \ g ` frontier T" proof fix x assume "x \ frontier S" then have "f x \ frontier T" using fim by blast then show "x \ g ` frontier T" by (metis gf \x \ frontier S\ imageI) qed then show "g ` frontier T = frontier S" using gim by blast show "continuous_on (frontier T) g" by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def) qed qed lemma homeomorphic_frontiers: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" "interior S = {} \ interior T = {}" shows "(frontier S) homeomorphic (frontier T)" proof (cases "interior T = {}") case True then show ?thesis by (metis Diff_empty assms closure_eq frontier_def) next case False show ?thesis apply (rule homeomorphic_frontiers_same_dimension) apply (simp_all add: assms) using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast qed lemma continuous_image_subset_rel_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and TS: "aff_dim T \ aff_dim S" shows "f ` (rel_interior S) \ rel_interior(f ` S)" proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "aff_dim (affine hull f ` S) \ aff_dim (affine hull S)" by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans) show "f ` rel_interior S \ affine hull f ` S" by (meson \f ` rel_interior S \ f ` S\ hull_subset order_trans) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast qed auto qed lemma homeomorphic_rel_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(rel_interior S) homeomorphic (rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (rel_interior S) (rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast have "rel_interior T \ f ` rel_interior S" proof fix x assume "x \ rel_interior T" then have "g x \ rel_interior S" using gim by blast then show "x \ f ` rel_interior S" by (metis fg \x \ rel_interior T\ imageI) qed moreover have "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) ultimately show "f ` rel_interior S = rel_interior T" by blast show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast have "rel_interior S \ g ` rel_interior T" proof fix x assume "x \ rel_interior S" then have "f x \ rel_interior T" using fim by blast then show "x \ g ` rel_interior T" by (metis gf \x \ rel_interior S\ imageI) qed then show "g ` rel_interior T = rel_interior S" using gim by blast show "continuous_on (rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(rel_interior S) homeomorphic (rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_rel_boundaries_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ S - rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ T - rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast show "f ` (S - rel_interior S) = T - rel_interior T" using S fST fim gim by auto show "continuous_on (S - rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "g ` (T - rel_interior T) = S - rel_interior S" using T gTS gim fim by auto show "continuous_on (T - rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_boundaries: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_boundaries_same_dimension [OF \S homeomorphic T\]) qed proposition uniformly_continuous_homeomorphism_UNIV_trivial: fixes f :: "'a::euclidean_space \ 'a" assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g" shows "S = UNIV" proof (cases "S = {}") case True then show ?thesis by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI) next case False have "inj g" by (metis UNIV_I hom homeomorphism_apply2 injI) then have "open (g ` UNIV)" by (blast intro: invariance_of_domain hom homeomorphism_cont2) then have "open S" using hom homeomorphism_image2 by blast moreover have "complete S" unfolding complete_def proof clarify fix \ assume \: "\n. \ n \ S" and "Cauchy \" have "Cauchy (f o \)" - using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf by blast + using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf + unfolding Cauchy_continuous_on_def by blast then obtain l where "(f \ \) \ l" by (auto simp: convergent_eq_Cauchy [symmetric]) show "\l\S. \ \ l" proof show "g l \ S" using hom homeomorphism_image2 by blast have "(g \ (f \ \)) \ g l" by (meson UNIV_I \(f \ \) \ l\ continuous_on_sequentially hom homeomorphism_cont2) then show "\ \ g l" proof - have "\n. \ n = (g \ (f \ \)) n" by (metis (no_types) \ comp_eq_dest_lhs hom homeomorphism_apply1) then show ?thesis by (metis (no_types) LIMSEQ_iff \(g \ (f \ \)) \ g l\) qed qed qed then have "closed S" by (simp add: complete_eq_closed) ultimately show ?thesis using clopen [of S] False by simp qed proposition invariance_of_domain_sphere_affine_set_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and U: "bounded U" "convex U" and "affine T" and affTU: "aff_dim T < aff_dim U" and ope: "openin (top_of_set (rel_frontier U)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "rel_frontier U = {}") case True then show ?thesis using ope openin_subset by force next case False obtain b c where b: "b \ rel_frontier U" and c: "c \ rel_frontier U" and "b \ c" using \bounded U\ rel_frontier_not_sing [of U] subset_singletonD False by fastforce obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1" proof (rule choose_affine_subset [OF affine_UNIV]) show "- 1 \ aff_dim U - 1" by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le) show "aff_dim U - 1 \ aff_dim (UNIV::'a set)" by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq) qed auto have SU: "S \ rel_frontier U" using ope openin_imp_subset by auto have homb: "rel_frontier U - {b} homeomorphic V" and homc: "rel_frontier U - {c} homeomorphic V" using homeomorphic_punctured_sphere_affine_gen [of U _ V] by (simp_all add: \affine V\ affV U b c) then obtain g h j k where gh: "homeomorphism (rel_frontier U - {b}) V g h" and jk: "homeomorphism (rel_frontier U - {c}) V j k" by (auto simp: homeomorphic_def) with SU have hgsub: "(h ` g ` (S - {b})) \ S" and kjsub: "(k ` j ` (S - {c})) \ S" by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (g ` (S - {b}))" apply (rule homeomorphism_imp_open_map [OF gh]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (g ` (S - {b})) (f \ h)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) using contf continuous_on_subset hgsub by blast show "inj_on (f \ h) (g ` (S - {b}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (j ` (S - {c}))" apply (rule homeomorphism_imp_open_map [OF jk]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (j ` (S - {c})) (f \ k)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) using contf continuous_on_subset kjsub by blast show "inj_on (f \ k) (j ` (S - {c}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) ultimately have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" by (rule openin_Un) moreover have "(f \ h) ` g ` (S - {b}) = f ` (S - {b})" proof - have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" proof show "k ` j ` (S - {c}) \ S - {c}" using homeomorphism_apply1 [OF jk] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {c} \ k ` j ` (S - {c})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "f ` (S - {b}) \ f ` (S - {c}) = f ` (S)" using \b \ c\ by blast ultimately show ?thesis by simp qed lemma invariance_of_domain_sphere_affine_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and "r \ 0" "affine T" and affTU: "aff_dim T < DIM('a)" and ope: "openin (top_of_set (sphere a r)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "sphere a r = {}") case True then show ?thesis using ope openin_subset by force next case False show ?thesis proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \affine T\]) show "aff_dim T < aff_dim (cball a r)" by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty) show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed lemma no_embedding_sphere_lowdim: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0" shows "DIM('a) \ DIM('b)" proof - have "False" if "DIM('a) > DIM('b)" proof - have "compact (f ` sphere a r)" using compact_continuous_image by (simp add: compact_continuous_image contf) then have "\ open (f ` sphere a r)" using compact_open by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty) then show False using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \r > 0\ by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that) qed then show ?thesis using not_less by blast qed lemma empty_interior_lowdim_gen: fixes S :: "'N::euclidean_space set" and T :: "'M::euclidean_space set" assumes dim: "DIM('M) < DIM('N)" and ST: "S homeomorphic T" shows "interior S = {}" proof - obtain h :: "'M \ 'N" where "linear h" "\x. norm(h x) = norm x" by (rule isometry_subset_subspace [OF subspace_UNIV subspace_UNIV, where ?'a = 'M and ?'b = 'N]) (use dim in auto) then have "inj h" by (metis linear_inj_iff_eq_0 norm_eq_zero) then have "h ` T homeomorphic T" using \linear h\ homeomorphic_sym linear_homeomorphic_image by blast then have "interior (h ` T) homeomorphic interior S" using homeomorphic_interiors_same_dimension by (metis ST homeomorphic_sym homeomorphic_trans) moreover have "interior (range h) = {}" by (simp add: \inj h\ \linear h\ dim dim_image_eq empty_interior_lowdim) then have "interior (h ` T) = {}" by (metis image_mono interior_mono subset_empty top_greatest) ultimately show ?thesis by simp qed lemma empty_interior_lowdim_gen_le: fixes S :: "'N::euclidean_space set" and T :: "'M::euclidean_space set" assumes "DIM('M) \ DIM('N)" "interior T = {}" "S homeomorphic T" shows "interior S = {}" by (metis assms empty_interior_lowdim_gen homeomorphic_empty(1) homeomorphic_interiors_same_dimension less_le) lemma homeomorphic_affine_sets_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "affine S" "affine T" shows "S homeomorphic T \ aff_dim S = aff_dim T" proof (cases "S = {} \ T = {}") case True then show ?thesis using assms homeomorphic_affine_sets by force next case False then obtain a b where "a \ S" "b \ T" by blast then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ then show ?thesis by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets) qed lemma homeomorphic_hyperplanes_eq: fixes a :: "'M::euclidean_space" and c :: "'N::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('M) = DIM('N))" (is "?lhs = ?rhs") proof - have "(DIM('M) - Suc 0 = DIM('N) - Suc 0) \ (DIM('M) = DIM('N))" by auto (metis DIM_positive Suc_pred) then show ?thesis using assms by (simp add: homeomorphic_affine_sets_eq affine_hyperplane) qed end diff --git a/src/HOL/Limits.thy b/src/HOL/Limits.thy --- a/src/HOL/Limits.thy +++ b/src/HOL/Limits.thy @@ -1,3271 +1,3266 @@ (* Title: HOL/Limits.thy Author: Brian Huffman Author: Jacques D. Fleuriot, University of Cambridge Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Limits on Real Vector Spaces\ theory Limits imports Real_Vector_Spaces begin lemma range_mult [simp]: fixes a::"real" shows "range ((*) a) = (if a=0 then {0} else UNIV)" by (simp add: surj_def) (meson dvdE dvd_field_iff) subsection \Filter going to infinity norm\ definition at_infinity :: "'a::real_normed_vector filter" where "at_infinity = (INF r. principal {x. r \ norm x})" lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def by (subst eventually_INF_base) (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) corollary eventually_at_infinity_pos: "eventually p at_infinity \ (\b. 0 < b \ (\x. norm x \ b \ p x))" unfolding eventually_at_infinity by (meson le_less_trans norm_ge_zero not_le zero_less_one) lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" proof - have 1: "\\n\u. A n; \n\v. A n\ \ \b. \x. b \ \x\ \ A x" for A and u v::real by (rule_tac x="max (- v) u" in exI) (auto simp: abs_real_def) have 2: "\x. u \ \x\ \ A x \ \N. \n\N. A n" for A and u::real by (meson abs_less_iff le_cases less_le_not_le) have 3: "\x. u \ \x\ \ A x \ \N. \n\N. A n" for A and u::real by (metis (full_types) abs_ge_self abs_minus_cancel le_minus_iff order_trans) show ?thesis by (auto simp: filter_eq_iff eventually_sup eventually_at_infinity eventually_at_top_linorder eventually_at_bot_linorder intro: 1 2 3) qed lemma at_top_le_at_infinity: "at_top \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma at_bot_le_at_infinity: "at_bot \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \ filterlim f at_infinity F" for f :: "_ \ real" by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) lemma filterlim_real_at_infinity_sequentially: "filterlim real at_infinity sequentially" by (simp add: filterlim_at_top_imp_at_infinity filterlim_real_sequentially) lemma lim_infinity_imp_sequentially: "(f \ l) at_infinity \ ((\n. f(n)) \ l) sequentially" by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) subsubsection \Boundedness\ definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where "Bseq X \ Bfun X sequentially" lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" unfolding Bfun_metric_def by (subst eventually_sequentially_seg) lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) lemma Bfun_def: "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" unfolding Bfun_metric_def norm_conv_dist proof safe fix y K assume K: "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" moreover have "eventually (\x. dist (f x) 0 \ dist (f x) y + dist 0 y) F" by (intro always_eventually) (metis dist_commute dist_triangle) with * have "eventually (\x. dist (f x) 0 \ K + dist 0 y) F" by eventually_elim auto with \0 < K\ show "\K>0. eventually (\x. dist (f x) 0 \ K) F" by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto qed (force simp del: norm_conv_dist [symmetric]) lemma BfunI: assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp show "eventually (\x. norm (f x) \ max K 1) F" using K by (rule eventually_mono) simp qed lemma BfunE: assumes "Bfun f F" obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by blast lemma Cauchy_Bseq: assumes "Cauchy X" shows "Bseq X" proof - have "\y K. 0 < K \ (\N. \n\N. dist (X n) y \ K)" if "\m n. \m \ M; n \ M\ \ dist (X m) (X n) < 1" for M by (meson order.order_iff_strict that zero_less_one) with assms show ?thesis by (force simp: Cauchy_def Bfun_metric_def eventually_sequentially) qed subsubsection \Bounded Sequences\ lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" by (intro BfunI) (auto simp: eventually_sequentially) lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" unfolding Bfun_def eventually_sequentially proof safe fix N K assume "0 < K" "\n\N. norm (X n) \ K" then show "\K>0. \n. norm (X n) \ K" by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2) (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) qed auto lemma BseqE: "Bseq X \ (\K. 0 < K \ \n. norm (X n) \ K \ Q) \ Q" unfolding Bseq_def by auto lemma BseqD: "Bseq X \ \K. 0 < K \ (\n. norm (X n) \ K)" by (simp add: Bseq_def) lemma BseqI: "0 < K \ \n. norm (X n) \ K \ Bseq X" by (auto simp: Bseq_def) lemma Bseq_bdd_above: "Bseq X \ bdd_above (range X)" for X :: "nat \ real" proof (elim BseqE, intro bdd_aboveI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "X n \ K" by (auto elim!: allE[of _ n]) qed lemma Bseq_bdd_above': "Bseq X \ bdd_above (range (\n. norm (X n)))" for X :: "nat \ 'a :: real_normed_vector" proof (elim BseqE, intro bdd_aboveI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "norm (X n) \ K" by (auto elim!: allE[of _ n]) qed lemma Bseq_bdd_below: "Bseq X \ bdd_below (range X)" for X :: "nat \ real" proof (elim BseqE, intro bdd_belowI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "- K \ X n" by (auto elim!: allE[of _ n]) qed lemma Bseq_eventually_mono: assumes "eventually (\n. norm (f n) \ norm (g n)) sequentially" "Bseq g" shows "Bseq f" proof - from assms(2) obtain K where "0 < K" and "eventually (\n. norm (g n) \ K) sequentially" unfolding Bfun_def by fast with assms(1) have "eventually (\n. norm (f n) \ K) sequentially" by (fast elim: eventually_elim2 order_trans) with \0 < K\ show "Bseq f" unfolding Bfun_def by fast qed lemma lemma_NBseq_def: "(\K > 0. \n. norm (X n) \ K) \ (\N. \n. norm (X n) \ real(Suc N))" proof safe fix K :: real from reals_Archimedean2 obtain n :: nat where "K < real n" .. then have "K \ real (Suc n)" by auto moreover assume "\m. norm (X m) \ K" ultimately have "\m. norm (X m) \ real (Suc n)" by (blast intro: order_trans) then show "\N. \n. norm (X n) \ real (Suc N)" .. next show "\N. \n. norm (X n) \ real (Suc N) \ \K>0. \n. norm (X n) \ K" using of_nat_0_less_iff by blast qed text \Alternative definition for \Bseq\.\ lemma Bseq_iff: "Bseq X \ (\N. \n. norm (X n) \ real(Suc N))" by (simp add: Bseq_def) (simp add: lemma_NBseq_def) lemma lemma_NBseq_def2: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" proof - have *: "\N. \n. norm (X n) \ 1 + real N \ \N. \n. norm (X n) < 1 + real N" by (metis add.commute le_less_trans less_add_one of_nat_Suc) then show ?thesis unfolding lemma_NBseq_def by (metis less_le_not_le not_less_iff_gr_or_eq of_nat_Suc) qed text \Yet another definition for Bseq.\ lemma Bseq_iff1a: "Bseq X \ (\N. \n. norm (X n) < real (Suc N))" by (simp add: Bseq_def lemma_NBseq_def2) subsubsection \A Few More Equivalence Theorems for Boundedness\ text \Alternative formulation for boundedness.\ lemma Bseq_iff2: "Bseq X \ (\k > 0. \x. \n. norm (X n + - x) \ k)" by (metis BseqE BseqI' add.commute add_cancel_right_left add_uminus_conv_diff norm_add_leD norm_minus_cancel norm_minus_commute) text \Alternative formulation for boundedness.\ lemma Bseq_iff3: "Bseq X \ (\k>0. \N. \n. norm (X n + - X N) \ k)" (is "?P \ ?Q") proof assume ?P then obtain K where *: "0 < K" and **: "\n. norm (X n) \ K" by (auto simp: Bseq_def) from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp from ** have "\n. norm (X n - X 0) \ K + norm (X 0)" by (auto intro: order_trans norm_triangle_ineq4) then have "\n. norm (X n + - X 0) \ K + norm (X 0)" by simp with \0 < K + norm (X 0)\ show ?Q by blast next assume ?Q then show ?P by (auto simp: Bseq_iff2) qed subsubsection \Upper Bounds and Lubs of Bounded Sequences\ lemma Bseq_minus_iff: "Bseq (\n. - (X n) :: 'a::real_normed_vector) \ Bseq X" by (simp add: Bseq_def) lemma Bseq_add: fixes f :: "nat \ 'a::real_normed_vector" assumes "Bseq f" shows "Bseq (\x. f x + c)" proof - from assms obtain K where K: "\x. norm (f x) \ K" unfolding Bseq_def by blast { fix x :: nat have "norm (f x + c) \ norm (f x) + norm c" by (rule norm_triangle_ineq) also have "norm (f x) \ K" by (rule K) finally have "norm (f x + c) \ K + norm c" by simp } then show ?thesis by (rule BseqI') qed lemma Bseq_add_iff: "Bseq (\x. f x + c) \ Bseq f" for f :: "nat \ 'a::real_normed_vector" using Bseq_add[of f c] Bseq_add[of "\x. f x + c" "-c"] by auto lemma Bseq_mult: fixes f g :: "nat \ 'a::real_normed_field" assumes "Bseq f" and "Bseq g" shows "Bseq (\x. f x * g x)" proof - from assms obtain K1 K2 where K: "norm (f x) \ K1" "K1 > 0" "norm (g x) \ K2" "K2 > 0" for x unfolding Bseq_def by blast then have "norm (f x * g x) \ K1 * K2" for x by (auto simp: norm_mult intro!: mult_mono) then show ?thesis by (rule BseqI') qed lemma Bfun_const [simp]: "Bfun (\_. c) F" unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) lemma Bseq_cmult_iff: fixes c :: "'a::real_normed_field" assumes "c \ 0" shows "Bseq (\x. c * f x) \ Bseq f" proof assume "Bseq (\x. c * f x)" with Bfun_const have "Bseq (\x. inverse c * (c * f x))" by (rule Bseq_mult) with \c \ 0\ show "Bseq f" by (simp add: field_split_simps) qed (intro Bseq_mult Bfun_const) lemma Bseq_subseq: "Bseq f \ Bseq (\x. f (g x))" for f :: "nat \ 'a::real_normed_vector" unfolding Bseq_def by auto lemma Bseq_Suc_iff: "Bseq (\n. f (Suc n)) \ Bseq f" for f :: "nat \ 'a::real_normed_vector" using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) lemma increasing_Bseq_subseq_iff: assumes "\x y. x \ y \ norm (f x :: 'a::real_normed_vector) \ norm (f y)" "strict_mono g" shows "Bseq (\x. f (g x)) \ Bseq f" proof assume "Bseq (\x. f (g x))" then obtain K where K: "\x. norm (f (g x)) \ K" unfolding Bseq_def by auto { fix x :: nat from filterlim_subseq[OF assms(2)] obtain y where "g y \ x" by (auto simp: filterlim_at_top eventually_at_top_linorder) then have "norm (f x) \ norm (f (g y))" using assms(1) by blast also have "norm (f (g y)) \ K" by (rule K) finally have "norm (f x) \ K" . } then show "Bseq f" by (rule BseqI') qed (use Bseq_subseq[of f g] in simp_all) lemma nonneg_incseq_Bseq_subseq_iff: fixes f :: "nat \ real" and g :: "nat \ nat" assumes "\x. f x \ 0" "incseq f" "strict_mono g" shows "Bseq (\x. f (g x)) \ Bseq f" using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) lemma Bseq_eq_bounded: "range f \ {a..b} \ Bseq f" for a b :: real proof (rule BseqI'[where K="max (norm a) (norm b)"]) fix n assume "range f \ {a..b}" then have "f n \ {a..b}" by blast then show "norm (f n) \ max (norm a) (norm b)" by auto qed lemma incseq_bounded: "incseq X \ \i. X i \ B \ Bseq X" for B :: real by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) lemma decseq_bounded: "decseq X \ \i. B \ X i \ Bseq X" for B :: real by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) subsubsection\<^marker>\tag unimportant\ \Polynomal function extremal theorem, from HOL Light\ lemma polyfun_extremal_lemma: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "0 < e" shows "\M. \z. M \ norm(z) \ norm (\i\n. c(i) * z^i) \ e * norm(z) ^ (Suc n)" proof (induct n) case 0 with assms show ?case apply (rule_tac x="norm (c 0) / e" in exI) apply (auto simp: field_simps) done next case (Suc n) obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using Suc assms by blast show ?case proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" and "1 + norm (c (Suc n)) / e \ norm z" then have z2: "e + norm (c (Suc n)) \ e * norm z" using assms by (simp add: field_simps) have "norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using M [OF z1] by simp then have "norm (\i\n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by simp then have "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by (blast intro: norm_triangle_le elim: ) also have "... \ (e + norm (c (Suc n))) * norm z ^ Suc n" by (simp add: norm_power norm_mult algebra_simps) also have "... \ (e * norm z) * norm z ^ Suc n" by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) finally show "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc (Suc n)" by simp qed qed lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) fixes c :: "nat \ 'a::real_normed_div_algebra" assumes k: "c k \ 0" "1\k" and kn: "k\n" shows "eventually (\z. norm (\i\n. c(i) * z^i) \ B) at_infinity" using kn proof (induction n) case 0 then show ?case using k by simp next case (Suc m) show ?case proof (cases "c (Suc m) = 0") case True then show ?thesis using Suc k by auto (metis antisym_conv less_eq_Suc_le not_le) next case False then obtain M where M: "\z. M \ norm z \ norm (\i\m. c i * z^i) \ norm (c (Suc m)) / 2 * norm z ^ Suc m" using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc by auto have "\b. \z. b \ norm z \ B \ norm (\i\Suc m. c i * z^i)" proof (rule exI [where x="max M (max 1 (\B\ / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" "1 \ norm z" and "\B\ * 2 / norm (c (Suc m)) \ norm z" then have z2: "\B\ \ norm (c (Suc m)) * norm z / 2" using False by (simp add: field_simps) have nz: "norm z \ norm z ^ Suc m" by (metis \1 \ norm z\ One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) have *: "\y x. norm (c (Suc m)) * norm z / 2 \ norm y - norm x \ B \ norm (x + y)" by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) have "norm z * norm (c (Suc m)) + 2 * norm (\i\m. c i * z^i) \ norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" using M [of z] Suc z1 by auto also have "... \ 2 * (norm (c (Suc m)) * norm z ^ Suc m)" using nz by (simp add: mult_mono del: power_Suc) finally show "B \ norm ((\i\m. c i * z^i) + c (Suc m) * z ^ Suc m)" using Suc.IH apply (auto simp: eventually_at_infinity) apply (rule *) apply (simp add: field_simps norm_mult norm_power) done qed then show ?thesis by (simp add: eventually_at_infinity) qed qed subsection \Convergence to Zero\ definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" where "Zfun f F = (\r>0. eventually (\x. norm (f x) < r) F)" lemma ZfunI: "(\r. 0 < r \ eventually (\x. norm (f x) < r) F) \ Zfun f F" by (simp add: Zfun_def) lemma ZfunD: "Zfun f F \ 0 < r \ eventually (\x. norm (f x) < r) F" by (simp add: Zfun_def) lemma Zfun_ssubst: "eventually (\x. f x = g x) F \ Zfun g F \ Zfun f F" unfolding Zfun_def by (auto elim!: eventually_rev_mp) lemma Zfun_zero: "Zfun (\x. 0) F" unfolding Zfun_def by simp lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) F = Zfun (\x. f x) F" unfolding Zfun_def by simp lemma Zfun_imp_Zfun: assumes f: "Zfun f F" and g: "eventually (\x. norm (g x) \ norm (f x) * K) F" shows "Zfun (\x. g x) F" proof (cases "0 < K") case K: True show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" then have "0 < r / K" using K by simp then have "eventually (\x. norm (f x) < r / K) F" using ZfunD [OF f] by blast with g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) then have "norm (f x) * K < r" by (simp add: pos_less_divide_eq K) then show ?case by (simp add: order_le_less_trans [OF elim(1)]) qed qed next case False then have K: "K \ 0" by (simp only: not_less) show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" from g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) also have "norm (f x) * K \ norm (f x) * 0" using K norm_ge_zero by (rule mult_left_mono) finally show ?case using \0 < r\ by simp qed qed qed lemma Zfun_le: "Zfun g F \ \x. norm (f x) \ norm (g x) \ Zfun f F" by (erule Zfun_imp_Zfun [where K = 1]) simp lemma Zfun_add: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x + g x) F" proof (rule ZfunI) fix r :: real assume "0 < r" then have r: "0 < r / 2" by simp have "eventually (\x. norm (f x) < r/2) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < r/2) F" using g r by (rule ZfunD) ultimately show "eventually (\x. norm (f x + g x) < r) F" proof eventually_elim case (elim x) have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) also have "\ < r/2 + r/2" using elim by (rule add_strict_mono) finally show ?case by simp qed qed lemma Zfun_minus: "Zfun f F \ Zfun (\x. - f x) F" unfolding Zfun_def by simp lemma Zfun_diff: "Zfun f F \ Zfun g F \ Zfun (\x. f x - g x) F" using Zfun_add [of f F "\x. - g x"] by (simp add: Zfun_minus) lemma (in bounded_linear) Zfun: assumes g: "Zfun g F" shows "Zfun (\x. f (g x)) F" proof - obtain K where "norm (f x) \ norm x * K" for x using bounded by blast then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) F" by simp with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" proof (rule ZfunI) fix r :: real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "norm (x ** y) \ norm x * norm y * K" for x y using pos_bounded by blast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) have "eventually (\x. norm (f x) < r) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < inverse K) F" using g K' by (rule ZfunD) ultimately show "eventually (\x. norm (f x ** g x) < r) F" proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "norm (f x) * norm (g x) * K < r * inverse K * K" by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) also from K have "r * inverse K * K = r" by simp finally show ?case . qed qed lemma (in bounded_bilinear) Zfun_left: "Zfun f F \ Zfun (\x. f x ** a) F" by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: "Zfun f F \ Zfun (\x. a ** f x) F" by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] lemma tendsto_Zfun_iff: "(f \ a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) lemma tendsto_0_le: "(f \ 0) F \ eventually (\x. norm (g x) \ norm (f x) * K) F \ (g \ 0) F" by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) subsubsection \Distance and norms\ lemma tendsto_dist [tendsto_intros]: fixes l m :: "'a::metric_space" assumes f: "(f \ l) F" and g: "(g \ m) F" shows "((\x. dist (f x) (g x)) \ dist l m) F" proof (rule tendstoI) fix e :: real assume "0 < e" then have e2: "0 < e/2" by simp from tendstoD [OF f e2] tendstoD [OF g e2] show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" proof (eventually_elim) case (elim x) then show "dist (dist (f x) (g x)) (dist l m) < e" unfolding dist_real_def using dist_triangle2 [of "f x" "g x" "l"] and dist_triangle2 [of "g x" "l" "m"] and dist_triangle3 [of "l" "m" "f x"] and dist_triangle [of "f x" "m" "g x"] by arith qed qed lemma continuous_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" unfolding continuous_def by (rule tendsto_dist) lemma continuous_on_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" unfolding continuous_on_def by (auto intro: tendsto_dist) lemma continuous_at_dist: "isCont (dist a) b" using continuous_on_dist [OF continuous_on_const continuous_on_id] continuous_on_eq_continuous_within by blast lemma tendsto_norm [tendsto_intros]: "(f \ a) F \ ((\x. norm (f x)) \ norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) lemma continuous_norm [continuous_intros]: "continuous F f \ continuous F (\x. norm (f x))" unfolding continuous_def by (rule tendsto_norm) lemma continuous_on_norm [continuous_intros]: "continuous_on s f \ continuous_on s (\x. norm (f x))" unfolding continuous_on_def by (auto intro: tendsto_norm) lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" by (intro continuous_on_id continuous_on_norm) lemma tendsto_norm_zero: "(f \ 0) F \ ((\x. norm (f x)) \ 0) F" by (drule tendsto_norm) simp lemma tendsto_norm_zero_cancel: "((\x. norm (f x)) \ 0) F \ (f \ 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_norm_zero_iff: "((\x. norm (f x)) \ 0) F \ (f \ 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_rabs [tendsto_intros]: "(f \ l) F \ ((\x. \f x\) \ \l\) F" for l :: real by (fold real_norm_def) (rule tendsto_norm) lemma continuous_rabs [continuous_intros]: "continuous F f \ continuous F (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_norm) lemma continuous_on_rabs [continuous_intros]: "continuous_on s f \ continuous_on s (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_on_norm) lemma tendsto_rabs_zero: "(f \ (0::real)) F \ ((\x. \f x\) \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero) lemma tendsto_rabs_zero_cancel: "((\x. \f x\) \ (0::real)) F \ (f \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero_cancel) lemma tendsto_rabs_zero_iff: "((\x. \f x\) \ (0::real)) F \ (f \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero_iff) subsection \Topological Monoid\ class topological_monoid_add = topological_space + monoid_add + assumes tendsto_add_Pair: "LIM x (nhds a \\<^sub>F nhds b). fst x + snd x :> nhds (a + b)" class topological_comm_monoid_add = topological_monoid_add + comm_monoid_add lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::topological_monoid_add" shows "(f \ a) F \ (g \ b) F \ ((\x. f x + g x) \ a + b) F" using filterlim_compose[OF tendsto_add_Pair, of "\x. (f x, g x)" a b F] by (simp add: nhds_prod[symmetric] tendsto_Pair) lemma continuous_add [continuous_intros]: fixes f g :: "_ \ 'b::topological_monoid_add" shows "continuous F f \ continuous F g \ continuous F (\x. f x + g x)" unfolding continuous_def by (rule tendsto_add) lemma continuous_on_add [continuous_intros]: fixes f g :: "_ \ 'b::topological_monoid_add" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x + g x)" unfolding continuous_on_def by (auto intro: tendsto_add) lemma tendsto_add_zero: fixes f g :: "_ \ 'b::topological_monoid_add" shows "(f \ 0) F \ (g \ 0) F \ ((\x. f x + g x) \ 0) F" by (drule (1) tendsto_add) simp lemma tendsto_sum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ (f i \ a i) F) \ ((\x. \i\I. f i x) \ (\i\I. a i)) F" by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_add) lemma tendsto_null_sum: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" assumes "\i. i \ I \ ((\x. f x i) \ 0) F" shows "((\i. sum (f i) I) \ 0) F" using tendsto_sum [of I "\x y. f y x" "\x. 0"] assms by simp lemma continuous_sum [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ continuous F (f i)) \ continuous F (\x. \i\I. f i x)" unfolding continuous_def by (rule tendsto_sum) lemma continuous_on_sum [continuous_intros]: fixes f :: "'a \ 'b::topological_space \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ continuous_on S (f i)) \ continuous_on S (\x. \i\I. f i x)" unfolding continuous_on_def by (auto intro: tendsto_sum) instance nat :: topological_comm_monoid_add by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_add by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) subsubsection \Topological group\ class topological_group_add = topological_monoid_add + group_add + assumes tendsto_uminus_nhds: "(uminus \ - a) (nhds a)" begin lemma tendsto_minus [tendsto_intros]: "(f \ a) F \ ((\x. - f x) \ - a) F" by (rule filterlim_compose[OF tendsto_uminus_nhds]) end class topological_ab_group_add = topological_group_add + ab_group_add instance topological_ab_group_add < topological_comm_monoid_add .. lemma continuous_minus [continuous_intros]: "continuous F f \ continuous F (\x. - f x)" for f :: "'a::t2_space \ 'b::topological_group_add" unfolding continuous_def by (rule tendsto_minus) lemma continuous_on_minus [continuous_intros]: "continuous_on s f \ continuous_on s (\x. - f x)" for f :: "_ \ 'b::topological_group_add" unfolding continuous_on_def by (auto intro: tendsto_minus) lemma tendsto_minus_cancel: "((\x. - f x) \ - a) F \ (f \ a) F" for a :: "'a::topological_group_add" by (drule tendsto_minus) simp lemma tendsto_minus_cancel_left: "(f \ - (y::_::topological_group_add)) F \ ((\x. - f x) \ y) F" using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] by auto lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::topological_group_add" shows "(f \ a) F \ (g \ b) F \ ((\x. f x - g x) \ a - b) F" using tendsto_add [of f a F "\x. - g x" "- b"] by (simp add: tendsto_minus) lemma continuous_diff [continuous_intros]: fixes f g :: "'a::t2_space \ 'b::topological_group_add" shows "continuous F f \ continuous F g \ continuous F (\x. f x - g x)" unfolding continuous_def by (rule tendsto_diff) lemma continuous_on_diff [continuous_intros]: fixes f g :: "_ \ 'b::topological_group_add" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x - g x)" unfolding continuous_on_def by (auto intro: tendsto_diff) lemma continuous_on_op_minus: "continuous_on (s::'a::topological_group_add set) ((-) x)" by (rule continuous_intros | simp)+ instance real_normed_vector < topological_ab_group_add proof fix a b :: 'a show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" unfolding tendsto_Zfun_iff add_diff_add using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (intro Zfun_add) (auto simp: tendsto_Zfun_iff[symmetric] nhds_prod[symmetric] intro!: tendsto_fst) show "(uminus \ - a) (nhds a)" unfolding tendsto_Zfun_iff minus_diff_minus using filterlim_ident[of "nhds a"] by (intro Zfun_minus) (simp add: tendsto_Zfun_iff) qed lemmas real_tendsto_sandwich = tendsto_sandwich[where 'a=real] subsubsection \Linear operators and multiplication\ lemma linear_times [simp]: "linear (\x. c * x)" for c :: "'a::real_algebra" by (auto simp: linearI distrib_left) lemma (in bounded_linear) tendsto: "(g \ a) F \ ((\x. f (g x)) \ f a) F" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_linear) continuous: "continuous F g \ continuous F (\x. f (g x))" using tendsto[of g _ F] by (auto simp: continuous_def) lemma (in bounded_linear) continuous_on: "continuous_on s g \ continuous_on s (\x. f (g x))" using tendsto[of g] by (auto simp: continuous_on_def) lemma (in bounded_linear) tendsto_zero: "(g \ 0) F \ ((\x. f (g x)) \ 0) F" by (drule tendsto) (simp only: zero) lemma (in bounded_bilinear) tendsto: "(f \ a) F \ (g \ b) F \ ((\x. f x ** g x) \ a ** b) F" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) lemma (in bounded_bilinear) continuous: "continuous F f \ continuous F g \ continuous F (\x. f x ** g x)" using tendsto[of f _ F g] by (auto simp: continuous_def) lemma (in bounded_bilinear) continuous_on: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x ** g x)" using tendsto[of f _ _ g] by (auto simp: continuous_on_def) lemma (in bounded_bilinear) tendsto_zero: assumes f: "(f \ 0) F" and g: "(g \ 0) F" shows "((\x. f x ** g x) \ 0) F" using tendsto [OF f g] by (simp add: zero_left) lemma (in bounded_bilinear) tendsto_left_zero: "(f \ 0) F \ ((\x. f x ** c) \ 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) lemma (in bounded_bilinear) tendsto_right_zero: "(f \ 0) F \ ((\x. c ** f x) \ 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) lemmas tendsto_of_real [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_of_real] lemmas tendsto_scaleR [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] text\Analogous type class for multiplication\ class topological_semigroup_mult = topological_space + semigroup_mult + assumes tendsto_mult_Pair: "LIM x (nhds a \\<^sub>F nhds b). fst x * snd x :> nhds (a * b)" instance real_normed_algebra < topological_semigroup_mult proof fix a b :: 'a show "((\x. fst x * snd x) \ a * b) (nhds a \\<^sub>F nhds b)" unfolding nhds_prod[symmetric] using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (simp add: bounded_bilinear.tendsto [OF bounded_bilinear_mult]) qed lemma tendsto_mult [tendsto_intros]: fixes a b :: "'a::topological_semigroup_mult" shows "(f \ a) F \ (g \ b) F \ ((\x. f x * g x) \ a * b) F" using filterlim_compose[OF tendsto_mult_Pair, of "\x. (f x, g x)" a b F] by (simp add: nhds_prod[symmetric] tendsto_Pair) lemma tendsto_mult_left: "(f \ l) F \ ((\x. c * (f x)) \ c * l) F" for c :: "'a::topological_semigroup_mult" by (rule tendsto_mult [OF tendsto_const]) lemma tendsto_mult_right: "(f \ l) F \ ((\x. (f x) * c) \ l * c) F" for c :: "'a::topological_semigroup_mult" by (rule tendsto_mult [OF _ tendsto_const]) lemma tendsto_mult_left_iff [simp]: "c \ 0 \ tendsto(\x. c * f x) (c * l) F \ tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" by (auto simp: tendsto_mult_left dest: tendsto_mult_left [where c = "1/c"]) lemma tendsto_mult_right_iff [simp]: "c \ 0 \ tendsto(\x. f x * c) (l * c) F \ tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" by (auto simp: tendsto_mult_right dest: tendsto_mult_left [where c = "1/c"]) lemma tendsto_zero_mult_left_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. c * a n)\ 0 \ a \ 0" using assms tendsto_mult_left tendsto_mult_left_iff by fastforce lemma tendsto_zero_mult_right_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. a n * c)\ 0 \ a \ 0" using assms tendsto_mult_right tendsto_mult_right_iff by fastforce lemma tendsto_zero_divide_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. a n / c)\ 0 \ a \ 0" using tendsto_zero_mult_right_iff [of "1/c" a] assms by (simp add: field_simps) lemma lim_const_over_n [tendsto_intros]: fixes a :: "'a::real_normed_field" shows "(\n. a / of_nat n) \ 0" using tendsto_mult [OF tendsto_const [of a] lim_1_over_n] by simp lemmas continuous_of_real [continuous_intros] = bounded_linear.continuous [OF bounded_linear_of_real] lemmas continuous_scaleR [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_scaleR] lemmas continuous_mult [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_mult] lemmas continuous_on_of_real [continuous_intros] = bounded_linear.continuous_on [OF bounded_linear_of_real] lemmas continuous_on_scaleR [continuous_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] lemmas continuous_on_mult [continuous_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_mult] lemmas tendsto_mult_zero = bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_left_zero = bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_right_zero = bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] lemma continuous_mult_left: fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. c * f x)" by (rule continuous_mult [OF continuous_const]) lemma continuous_mult_right: fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. f x * c)" by (rule continuous_mult [OF _ continuous_const]) lemma continuous_on_mult_left: fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. c * f x)" by (rule continuous_on_mult [OF continuous_on_const]) lemma continuous_on_mult_right: fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. f x * c)" by (rule continuous_on_mult [OF _ continuous_on_const]) lemma continuous_on_mult_const [simp]: fixes c::"'a::real_normed_algebra" shows "continuous_on s ((*) c)" by (intro continuous_on_mult_left continuous_on_id) lemma tendsto_divide_zero: fixes c :: "'a::real_normed_field" shows "(f \ 0) F \ ((\x. f x / c) \ 0) F" by (cases "c=0") (simp_all add: divide_inverse tendsto_mult_left_zero) lemma tendsto_power [tendsto_intros]: "(f \ a) F \ ((\x. f x ^ n) \ a ^ n) F" for f :: "'a \ 'b::{power,real_normed_algebra}" by (induct n) (simp_all add: tendsto_mult) lemma tendsto_null_power: "\(f \ 0) F; 0 < n\ \ ((\x. f x ^ n) \ 0) F" for f :: "'a \ 'b::{power,real_normed_algebra_1}" using tendsto_power [of f 0 F n] by (simp add: power_0_left) lemma continuous_power [continuous_intros]: "continuous F f \ continuous F (\x. (f x)^n)" for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" unfolding continuous_def by (rule tendsto_power) lemma continuous_on_power [continuous_intros]: fixes f :: "_ \ 'b::{power,real_normed_algebra}" shows "continuous_on s f \ continuous_on s (\x. (f x)^n)" unfolding continuous_on_def by (auto intro: tendsto_power) lemma tendsto_prod [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ (f i \ L i) F) \ ((\x. \i\S. f i x) \ (\i\S. L i)) F" by (induct S rule: infinite_finite_induct) (simp_all add: tendsto_mult) lemma continuous_prod [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" unfolding continuous_def by (rule tendsto_prod) lemma continuous_on_prod [continuous_intros]: fixes f :: "'a \ _ \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" unfolding continuous_on_def by (auto intro: tendsto_prod) lemma tendsto_of_real_iff: "((\x. of_real (f x) :: 'a::real_normed_div_algebra) \ of_real c) F \ (f \ c) F" unfolding tendsto_iff by simp lemma tendsto_add_const_iff: "((\x. c + f x :: 'a::topological_group_add) \ c + d) F \ (f \ d) F" using tendsto_add[OF tendsto_const[of c], of f d] and tendsto_add[OF tendsto_const[of "-c"], of "\x. c + f x" "c + d"] by auto class topological_monoid_mult = topological_semigroup_mult + monoid_mult class topological_comm_monoid_mult = topological_monoid_mult + comm_monoid_mult lemma tendsto_power_strong [tendsto_intros]: fixes f :: "_ \ 'b :: topological_monoid_mult" assumes "(f \ a) F" "(g \ b) F" shows "((\x. f x ^ g x) \ a ^ b) F" proof - have "((\x. f x ^ b) \ a ^ b) F" by (induction b) (auto intro: tendsto_intros assms) also from assms(2) have "eventually (\x. g x = b) F" by (simp add: nhds_discrete filterlim_principal) hence "eventually (\x. f x ^ b = f x ^ g x) F" by eventually_elim simp hence "((\x. f x ^ b) \ a ^ b) F \ ((\x. f x ^ g x) \ a ^ b) F" by (intro filterlim_cong refl) finally show ?thesis . qed lemma continuous_mult' [continuous_intros]: fixes f g :: "_ \ 'b::topological_semigroup_mult" shows "continuous F f \ continuous F g \ continuous F (\x. f x * g x)" unfolding continuous_def by (rule tendsto_mult) lemma continuous_power' [continuous_intros]: fixes f :: "_ \ 'b::topological_monoid_mult" shows "continuous F f \ continuous F g \ continuous F (\x. f x ^ g x)" unfolding continuous_def by (rule tendsto_power_strong) auto lemma continuous_on_mult' [continuous_intros]: fixes f g :: "_ \ 'b::topological_semigroup_mult" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x * g x)" unfolding continuous_on_def by (auto intro: tendsto_mult) lemma continuous_on_power' [continuous_intros]: fixes f :: "_ \ 'b::topological_monoid_mult" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x ^ g x)" unfolding continuous_on_def by (auto intro: tendsto_power_strong) lemma tendsto_mult_one: fixes f g :: "_ \ 'b::topological_monoid_mult" shows "(f \ 1) F \ (g \ 1) F \ ((\x. f x * g x) \ 1) F" by (drule (1) tendsto_mult) simp lemma tendsto_prod' [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ (f i \ a i) F) \ ((\x. \i\I. f i x) \ (\i\I. a i)) F" by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_mult) lemma tendsto_one_prod': fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_mult" assumes "\i. i \ I \ ((\x. f x i) \ 1) F" shows "((\i. prod (f i) I) \ 1) F" using tendsto_prod' [of I "\x y. f y x" "\x. 1"] assms by simp lemma LIMSEQ_prod_0: fixes f :: "nat \ 'a::{semidom,topological_space}" assumes "f i = 0" shows "(\n. prod f {..n}) \ 0" proof (subst tendsto_cong) show "\\<^sub>F n in sequentially. prod f {..n} = 0" using assms eventually_at_top_linorder by auto qed auto lemma LIMSEQ_prod_nonneg: fixes f :: "nat \ 'a::{linordered_semidom,linorder_topology}" assumes 0: "\n. 0 \ f n" and a: "(\n. prod f {..n}) \ a" shows "a \ 0" by (simp add: "0" prod_nonneg LIMSEQ_le_const [OF a]) lemma continuous_prod' [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ continuous F (f i)) \ continuous F (\x. \i\I. f i x)" unfolding continuous_def by (rule tendsto_prod') lemma continuous_on_prod' [continuous_intros]: fixes f :: "'a \ 'b::topological_space \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ continuous_on S (f i)) \ continuous_on S (\x. \i\I. f i x)" unfolding continuous_on_def by (auto intro: tendsto_prod') instance nat :: topological_comm_monoid_mult by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_mult by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) class comm_real_normed_algebra_1 = real_normed_algebra_1 + comm_monoid_mult context real_normed_field begin subclass comm_real_normed_algebra_1 proof from norm_mult[of "1 :: 'a" 1] show "norm 1 = 1" by simp qed (simp_all add: norm_mult) end subsubsection \Inverse and division\ lemma (in bounded_bilinear) Zfun_prod_Bfun: assumes f: "Zfun f F" and g: "Bfun g F" shows "Zfun (\x. f x ** g x) F" proof - obtain K where K: "0 \ K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using nonneg_bounded by blast obtain B where B: "0 < B" and norm_g: "eventually (\x. norm (g x) \ B) F" using g by (rule BfunE) have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) F" using norm_g proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "\ \ norm (f x) * B * K" by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) also have "\ = norm (f x) * (B * K)" by (rule mult.assoc) finally show "norm (f x ** g x) \ norm (f x) * (B * K)" . qed with f show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Bfun_prod_Zfun: assumes f: "Bfun f F" and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" assumes a: "a \ 0" shows "Bfun (\x. inverse (f x)) F" proof - from a have "0 < norm a" by simp then have "\r>0. r < norm a" by (rule dense) then obtain r where r1: "0 < r" and r2: "r < norm a" by blast have "eventually (\x. dist (f x) a < r) F" using tendstoD [OF f r1] by blast then have "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" proof eventually_elim case (elim x) then have 1: "norm (f x - a) < r" by (simp add: dist_norm) then have 2: "f x \ 0" using r2 by auto then have "norm (inverse (f x)) = inverse (norm (f x))" by (rule nonzero_norm_inverse) also have "\ \ inverse (norm a - r)" proof (rule le_imp_inverse_le) show "0 < norm a - r" using r2 by simp have "norm a - norm (f x) \ norm (a - f x)" by (rule norm_triangle_ineq2) also have "\ = norm (f x - a)" by (rule norm_minus_commute) also have "\ < r" using 1 . finally show "norm a - r \ norm (f x)" by simp qed finally show "norm (inverse (f x)) \ inverse (norm a - r)" . qed then show ?thesis by (rule BfunI) qed lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0" shows "((\x. inverse (f x)) \ inverse a) F" proof - from a have "0 < norm a" by simp with f have "eventually (\x. dist (f x) a < norm a) F" by (rule tendstoD) then have "eventually (\x. f x \ 0) F" unfolding dist_norm by (auto elim!: eventually_mono) with a have "eventually (\x. inverse (f x) - inverse a = - (inverse (f x) * (f x - a) * inverse a)) F" by (auto elim!: eventually_mono simp: inverse_diff_inverse) moreover have "Zfun (\x. - (inverse (f x) * (f x - a) * inverse a)) F" by (intro Zfun_minus Zfun_mult_left bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) ultimately show ?thesis unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) qed lemma continuous_inverse: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. inverse (f x))" using assms unfolding continuous_def by (rule tendsto_inverse) lemma continuous_at_within_inverse[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. inverse (f x))" using assms unfolding continuous_within by (rule tendsto_inverse) lemma continuous_on_inverse[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. inverse (f x))" using assms unfolding continuous_on_def by (blast intro: tendsto_inverse) lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" shows "(f \ a) F \ (g \ b) F \ b \ 0 \ ((\x. f x / g x) \ a / b) F" by (simp add: tendsto_mult tendsto_inverse divide_inverse) lemma continuous_divide: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous F f" and "continuous F g" and "g (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) / (g x))" using assms unfolding continuous_def by (rule tendsto_divide) lemma continuous_at_within_divide[continuous_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \ 0" shows "continuous (at a within s) (\x. (f x) / (g x))" using assms unfolding continuous_within by (rule tendsto_divide) lemma isCont_divide[continuous_intros, simp]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "isCont f a" "isCont g a" "g a \ 0" shows "isCont (\x. (f x) / g x) a" using assms unfolding continuous_at by (rule tendsto_divide) lemma continuous_on_divide[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_field" assumes "continuous_on s f" "continuous_on s g" and "\x\s. g x \ 0" shows "continuous_on s (\x. (f x) / (g x))" using assms unfolding continuous_on_def by (blast intro: tendsto_divide) lemma tendsto_power_int [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0" shows "((\x. power_int (f x) n) \ power_int a n) F" using assms by (cases n rule: int_cases4) (auto intro!: tendsto_intros simp: power_int_minus) lemma continuous_power_int: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. power_int (f x) n)" using assms unfolding continuous_def by (rule tendsto_power_int) lemma continuous_at_within_power_int[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. power_int (f x) n)" using assms unfolding continuous_within by (rule tendsto_power_int) lemma continuous_on_power_int [continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. power_int (f x) n)" using assms unfolding continuous_on_def by (blast intro: tendsto_power_int) lemma tendsto_power_int' [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0 \ n \ 0" shows "((\x. power_int (f x) n) \ power_int a n) F" using assms by (cases n rule: int_cases4) (auto intro!: tendsto_intros simp: power_int_minus) lemma tendsto_sgn [tendsto_intros]: "(f \ l) F \ l \ 0 \ ((\x. sgn (f x)) \ sgn l) F" for l :: "'a::real_normed_vector" unfolding sgn_div_norm by (simp add: tendsto_intros) lemma continuous_sgn: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. sgn (f x))" using assms unfolding continuous_def by (rule tendsto_sgn) lemma continuous_at_within_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. sgn (f x))" using assms unfolding continuous_within by (rule tendsto_sgn) lemma isCont_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "isCont f a" and "f a \ 0" shows "isCont (\x. sgn (f x)) a" using assms unfolding continuous_at by (rule tendsto_sgn) lemma continuous_on_sgn[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. sgn (f x))" using assms unfolding continuous_on_def by (blast intro: tendsto_sgn) lemma filterlim_at_infinity: fixes f :: "_ \ 'a::real_normed_vector" assumes "0 \ c" shows "(LIM x F. f x :> at_infinity) \ (\r>c. eventually (\x. r \ norm (f x)) F)" unfolding filterlim_iff eventually_at_infinity proof safe fix P :: "'a \ bool" fix b assume *: "\r>c. eventually (\x. r \ norm (f x)) F" assume P: "\x. b \ norm x \ P x" have "max b (c + 1) > c" by auto with * have "eventually (\x. max b (c + 1) \ norm (f x)) F" by auto then show "eventually (\x. P (f x)) F" proof eventually_elim case (elim x) with P show "P (f x)" by auto qed qed force lemma filterlim_at_infinity_imp_norm_at_top: fixes F assumes "filterlim f at_infinity F" shows "filterlim (\x. norm (f x)) at_top F" proof - { fix r :: real have "\\<^sub>F x in F. r \ norm (f x)" using filterlim_at_infinity[of 0 f F] assms by (cases "r > 0") (auto simp: not_less intro: always_eventually order.trans[OF _ norm_ge_zero]) } thus ?thesis by (auto simp: filterlim_at_top) qed lemma filterlim_norm_at_top_imp_at_infinity: fixes F assumes "filterlim (\x. norm (f x)) at_top F" shows "filterlim f at_infinity F" using filterlim_at_infinity[of 0 f F] assms by (auto simp: filterlim_at_top) lemma filterlim_norm_at_top: "filterlim norm at_top at_infinity" by (rule filterlim_at_infinity_imp_norm_at_top) (rule filterlim_ident) lemma filterlim_at_infinity_conv_norm_at_top: "filterlim f at_infinity G \ filterlim (\x. norm (f x)) at_top G" by (auto simp: filterlim_at_infinity[OF order.refl] filterlim_at_top_gt[of _ _ 0]) lemma eventually_not_equal_at_infinity: "eventually (\x. x \ (a :: 'a :: {real_normed_vector})) at_infinity" proof - from filterlim_norm_at_top[where 'a = 'a] have "\\<^sub>F x in at_infinity. norm a < norm (x::'a)" by (auto simp: filterlim_at_top_dense) thus ?thesis by eventually_elim auto qed lemma filterlim_int_of_nat_at_topD: fixes F assumes "filterlim (\x. f (int x)) F at_top" shows "filterlim f F at_top" proof - have "filterlim (\x. f (int (nat x))) F at_top" by (rule filterlim_compose[OF assms filterlim_nat_sequentially]) also have "?this \ filterlim f F at_top" by (intro filterlim_cong refl eventually_mono [OF eventually_ge_at_top[of "0::int"]]) auto finally show ?thesis . qed lemma filterlim_int_sequentially [tendsto_intros]: "filterlim int at_top sequentially" unfolding filterlim_at_top proof fix C :: int show "eventually (\n. int n \ C) at_top" using eventually_ge_at_top[of "nat \C\"] by eventually_elim linarith qed lemma filterlim_real_of_int_at_top [tendsto_intros]: "filterlim real_of_int at_top at_top" unfolding filterlim_at_top proof fix C :: real show "eventually (\n. real_of_int n \ C) at_top" using eventually_ge_at_top[of "\C\"] by eventually_elim linarith qed lemma filterlim_abs_real: "filterlim (abs::real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) from eventually_ge_at_top[of "0::real"] show "eventually (\x::real. \x\ = x) at_top" by eventually_elim simp qed (simp_all add: filterlim_ident) lemma filterlim_of_real_at_infinity [tendsto_intros]: "filterlim (of_real :: real \ 'a :: real_normed_algebra_1) at_infinity at_top" by (intro filterlim_norm_at_top_imp_at_infinity) (auto simp: filterlim_abs_real) lemma not_tendsto_and_filterlim_at_infinity: fixes c :: "'a::real_normed_vector" assumes "F \ bot" and "(f \ c) F" and "filterlim f at_infinity F" shows False proof - from tendstoD[OF assms(2), of "1/2"] have "eventually (\x. dist (f x) c < 1/2) F" by simp moreover from filterlim_at_infinity[of "norm c" f F] assms(3) have "eventually (\x. norm (f x) \ norm c + 1) F" by simp ultimately have "eventually (\x. False) F" proof eventually_elim fix x assume A: "dist (f x) c < 1/2" assume "norm (f x) \ norm c + 1" also have "norm (f x) = dist (f x) 0" by simp also have "\ \ dist (f x) c + dist c 0" by (rule dist_triangle) finally show False using A by simp qed with assms show False by simp qed lemma filterlim_at_infinity_imp_not_convergent: assumes "filterlim f at_infinity sequentially" shows "\ convergent f" by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms]) (simp_all add: convergent_LIMSEQ_iff) lemma filterlim_at_infinity_imp_eventually_ne: assumes "filterlim f at_infinity F" shows "eventually (\z. f z \ c) F" proof - have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all with filterlim_at_infinity[OF order.refl, of f F] assms have "eventually (\z. norm (f z) \ norm c + 1) F" by blast then show ?thesis by eventually_elim auto qed lemma tendsto_of_nat [tendsto_intros]: "filterlim (of_nat :: nat \ 'a::real_normed_algebra_1) at_infinity sequentially" proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) fix r :: real assume r: "r > 0" define n where "n = nat \r\" from r have n: "\m\n. of_nat m \ r" unfolding n_def by linarith from eventually_ge_at_top[of n] show "eventually (\m. norm (of_nat m :: 'a) \ r) sequentially" by eventually_elim (use n in simp_all) qed subsection \Relate \<^const>\at\, \<^const>\at_left\ and \<^const>\at_right\\ text \ This lemmas are useful for conversion between \<^term>\at x\ to \<^term>\at_left x\ and \<^term>\at_right x\ and also \<^term>\at_right 0\. \ lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d)" for a d :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g="\x. x + d"]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_nhds_minus: "filtermap (\x. - x) (nhds a) = nhds (- a)" for a :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g=uminus]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_at_shift: "filtermap (\x. x - d) (at a) = at (a - d)" for a d :: "'a::real_normed_vector" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma filtermap_at_right_shift: "filtermap (\x. x - d) (at_right a) = at_right (a - d)" for a d :: "real" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma filterlim_shift: fixes d :: "'a::real_normed_vector" assumes "filterlim f F (at a)" shows "filterlim (f \ (+) d) F (at (a - d))" unfolding filterlim_iff proof (intro strip) fix P assume "eventually P F" then have "\\<^sub>F x in filtermap (\y. y - d) (at a). P (f (d + x))" using assms by (force simp add: filterlim_iff eventually_filtermap) then show "(\\<^sub>F x in at (a - d). P ((f \ (+) d) x))" by (force simp add: filtermap_at_shift) qed lemma filterlim_shift_iff: fixes d :: "'a::real_normed_vector" shows "filterlim (f \ (+) d) F (at (a - d)) = filterlim f F (at a)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs using filterlim_shift [OF L, of "-d"] by (simp add: filterlim_iff) qed (metis filterlim_shift) lemma at_right_to_0: "at_right a = filtermap (\x. x + a) (at_right 0)" for a :: real using filtermap_at_right_shift[of "-a" 0] by simp lemma filterlim_at_right_to_0: "filterlim f F (at_right a) \ filterlim (\x. f (x + a)) F (at_right 0)" for a :: real unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. lemma eventually_at_right_to_0: "eventually P (at_right a) \ eventually (\x. P (x + a)) (at_right 0)" for a :: real unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) lemma at_to_0: "at a = filtermap (\x. x + a) (at 0)" for a :: "'a::real_normed_vector" using filtermap_at_shift[of "-a" 0] by simp lemma filterlim_at_to_0: "filterlim f F (at a) \ filterlim (\x. f (x + a)) F (at 0)" for a :: "'a::real_normed_vector" unfolding filterlim_def filtermap_filtermap at_to_0[of a] .. lemma eventually_at_to_0: "eventually P (at a) \ eventually (\x. P (x + a)) (at 0)" for a :: "'a::real_normed_vector" unfolding at_to_0[of a] by (simp add: eventually_filtermap) lemma filtermap_at_minus: "filtermap (\x. - x) (at a) = at (- a)" for a :: "'a::real_normed_vector" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_left_minus: "at_left a = filtermap (\x. - x) (at_right (- a))" for a :: real by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_right_minus: "at_right a = filtermap (\x. - x) (at_left (- a))" for a :: real by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma filterlim_at_left_to_right: "filterlim f F (at_left a) \ filterlim (\x. f (- x)) F (at_right (-a))" for a :: real unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. lemma eventually_at_left_to_right: "eventually P (at_left a) \ eventually (\x. P (- x)) (at_right (-a))" for a :: real unfolding at_left_minus[of a] by (simp add: eventually_filtermap) lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" unfolding filterlim_at_top eventually_at_bot_dense by (metis leI minus_less_iff order_less_asym) lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" unfolding filterlim_at_bot eventually_at_top_dense by (metis leI less_minus_iff order_less_asym) lemma at_bot_mirror : shows "(at_bot::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_top" proof (rule filtermap_fun_inverse[symmetric]) show "filterlim uminus at_top (at_bot::'a filter)" using eventually_at_bot_linorder filterlim_at_top le_minus_iff by force show "filterlim uminus (at_bot::'a filter) at_top" by (simp add: filterlim_at_bot minus_le_iff) qed auto lemma at_top_mirror : shows "(at_top::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_bot" apply (subst at_bot_mirror) by (auto simp: filtermap_filtermap) lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \ (LIM x at_bot. f (-x::real) :> F)" unfolding filterlim_def at_top_mirror filtermap_filtermap .. lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \ (LIM x at_top. f (-x::real) :> F)" unfolding filterlim_def at_bot_mirror filtermap_filtermap .. lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \ (LIM x F. - (f x) :: real :> at_bot)" using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\x. - f x" F] by auto lemma tendsto_at_botI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_bot sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_bot" unfolding filterlim_at_bot_mirror proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" thus "(\n. f (-X n)) \ y" by (intro *) (auto simp: filterlim_uminus_at_top) qed lemma filterlim_at_infinity_imp_filterlim_at_top: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x > 0) F" shows "filterlim f at_top F" proof - from assms(2) have *: "eventually (\x. norm (f x) = f x) F" by eventually_elim simp from assms(1) show ?thesis unfolding filterlim_at_infinity_conv_norm_at_top by (subst (asm) filterlim_cong[OF refl refl *]) qed lemma filterlim_at_infinity_imp_filterlim_at_bot: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x < 0) F" shows "filterlim f at_bot F" proof - from assms(2) have *: "eventually (\x. norm (f x) = -f x) F" by eventually_elim simp from assms(1) have "filterlim (\x. - f x) at_top F" unfolding filterlim_at_infinity_conv_norm_at_top by (subst (asm) filterlim_cong[OF refl refl *]) thus ?thesis by (simp add: filterlim_uminus_at_top) qed lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \ (LIM x F. - (f x) :: real :> at_top)" unfolding filterlim_uminus_at_top by simp lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" unfolding filterlim_at_top_gt[where c=0] eventually_at_filter proof safe fix Z :: real assume [arith]: "0 < Z" then have "eventually (\x. x < inverse Z) (nhds 0)" by (auto simp: eventually_nhds_metric dist_real_def intro!: exI[of _ "\inverse Z\"]) then show "eventually (\x. x \ 0 \ x \ {0<..} \ Z \ inverse x) (nhds 0)" by (auto elim!: eventually_mono simp: inverse_eq_divide field_simps) qed lemma tendsto_inverse_0: fixes x :: "_ \ 'a::real_normed_div_algebra" shows "(inverse \ (0::'a)) at_infinity" unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity proof safe fix r :: real assume "0 < r" show "\b. \x. b \ norm x \ norm (inverse x :: 'a) < r" proof (intro exI[of _ "inverse (r / 2)"] allI impI) fix x :: 'a from \0 < r\ have "0 < inverse (r / 2)" by simp also assume *: "inverse (r / 2) \ norm x" finally show "norm (inverse x) < r" using * \0 < r\ by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) qed qed lemma tendsto_add_filterlim_at_infinity: fixes c :: "'b::real_normed_vector" and F :: "'a filter" assumes "(f \ c) F" and "filterlim g at_infinity F" shows "filterlim (\x. f x + g x) at_infinity F" proof (subst filterlim_at_infinity[OF order_refl], safe) fix r :: real assume r: "r > 0" from assms(1) have "((\x. norm (f x)) \ norm c) F" by (rule tendsto_norm) then have "eventually (\x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all moreover from r have "r + norm c + 1 > 0" by (intro add_pos_nonneg) simp_all with assms(2) have "eventually (\x. norm (g x) \ r + norm c + 1) F" unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all ultimately show "eventually (\x. norm (f x + g x) \ r) F" proof eventually_elim fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \ norm (g x)" from A B have "r \ norm (g x) - norm (f x)" by simp also have "norm (g x) - norm (f x) \ norm (g x + f x)" by (rule norm_diff_ineq) finally show "r \ norm (f x + g x)" by (simp add: add_ac) qed qed lemma tendsto_add_filterlim_at_infinity': fixes c :: "'b::real_normed_vector" and F :: "'a filter" assumes "filterlim f at_infinity F" and "(g \ c) F" shows "filterlim (\x. f x + g x) at_infinity F" by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+ lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)" unfolding filterlim_at by (auto simp: eventually_at_top_dense) (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) lemma filterlim_inverse_at_top: "(f \ (0 :: real)) F \ eventually (\x. 0 < f x) F \ LIM x F. inverse (f x) :> at_top" by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) (simp add: filterlim_def eventually_filtermap eventually_mono at_within_def le_principal) lemma filterlim_inverse_at_bot_neg: "LIM x (at_left (0::real)). inverse x :> at_bot" by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) lemma filterlim_inverse_at_bot: "(f \ (0 :: real)) F \ eventually (\x. f x < 0) F \ LIM x F. inverse (f x) :> at_bot" unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" by (intro filtermap_fun_inverse[symmetric, where g=inverse]) (auto intro: filterlim_inverse_at_top_right filterlim_inverse_at_right_top) lemma eventually_at_right_to_top: "eventually P (at_right (0::real)) \ eventually (\x. P (inverse x)) at_top" unfolding at_right_to_top eventually_filtermap .. lemma filterlim_at_right_to_top: "filterlim f F (at_right (0::real)) \ (LIM x at_top. f (inverse x) :> F)" unfolding filterlim_def at_right_to_top filtermap_filtermap .. lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. lemma eventually_at_top_to_right: "eventually P at_top \ eventually (\x. P (inverse x)) (at_right (0::real))" unfolding at_top_to_right eventually_filtermap .. lemma filterlim_at_top_to_right: "filterlim f F at_top \ (LIM x (at_right (0::real)). f (inverse x) :> F)" unfolding filterlim_def at_top_to_right filtermap_filtermap .. lemma filterlim_inverse_at_infinity: fixes x :: "_ \ 'a::{real_normed_div_algebra, division_ring}" shows "filterlim inverse at_infinity (at (0::'a))" unfolding filterlim_at_infinity[OF order_refl] proof safe fix r :: real assume "0 < r" then show "eventually (\x::'a. r \ norm (inverse x)) (at 0)" unfolding eventually_at norm_inverse by (intro exI[of _ "inverse r"]) (auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) qed lemma filterlim_inverse_at_iff: fixes g :: "'a \ 'b::{real_normed_div_algebra, division_ring}" shows "(LIM x F. inverse (g x) :> at 0) \ (LIM x F. g x :> at_infinity)" unfolding filterlim_def filtermap_filtermap[symmetric] proof assume "filtermap g F \ at_infinity" then have "filtermap inverse (filtermap g F) \ filtermap inverse at_infinity" by (rule filtermap_mono) also have "\ \ at 0" using tendsto_inverse_0[where 'a='b] by (auto intro!: exI[of _ 1] simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity) finally show "filtermap inverse (filtermap g F) \ at 0" . next assume "filtermap inverse (filtermap g F) \ at 0" then have "filtermap inverse (filtermap inverse (filtermap g F)) \ filtermap inverse (at 0)" by (rule filtermap_mono) with filterlim_inverse_at_infinity show "filtermap g F \ at_infinity" by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) qed lemma tendsto_mult_filterlim_at_infinity: fixes c :: "'a::real_normed_field" assumes "(f \ c) F" "c \ 0" assumes "filterlim g at_infinity F" shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ inverse c * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) then have "filterlim (\x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" unfolding filterlim_at using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) then show ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all qed lemma filterlim_power_int_neg_at_infinity: fixes f :: "_ \ 'a::{real_normed_div_algebra, division_ring}" assumes "n < 0" and lim: "(f \ 0) F" and ev: "eventually (\x. f x \ 0) F" shows "filterlim (\x. f x powi n) at_infinity F" proof - have lim': "((\x. f x ^ nat (- n)) \ 0) F" by (rule tendsto_eq_intros lim)+ (use \n < 0\ in auto) have ev': "eventually (\x. f x ^ nat (-n) \ 0) F" using ev by eventually_elim (use \n < 0\ in auto) have "filterlim (\x. inverse (f x ^ nat (-n))) at_infinity F" by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) (use lim' ev' in \auto simp: filterlim_at\) thus ?thesis using \n < 0\ by (simp add: power_int_def power_inverse) qed lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \ ((\x. inverse (f x) :: real) \ 0) F" by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff) lemma filterlim_inverse_at_top_iff: "eventually (\x. 0 < f x) F \ (LIM x F. inverse (f x) :> at_top) \ (f \ (0 :: real)) F" by (auto dest: tendsto_inverse_0_at_top filterlim_inverse_at_top) lemma filterlim_at_top_iff_inverse_0: "eventually (\x. 0 < f x) F \ (LIM x F. f x :> at_top) \ ((inverse \ f) \ (0 :: real)) F" using filterlim_inverse_at_top_iff [of "inverse \ f"] by auto lemma real_tendsto_divide_at_top: fixes c::"real" assumes "(f \ c) F" assumes "filterlim g at_top F" shows "((\x. f x / g x) \ 0) F" by (auto simp: divide_inverse_commute intro!: tendsto_mult[THEN tendsto_eq_rhs] tendsto_inverse_0_at_top assms) lemma mult_nat_left_at_top: "c > 0 \ filterlim (\x. c * x) at_top sequentially" for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) lemma mult_nat_right_at_top: "c > 0 \ filterlim (\x. x * c) at_top sequentially" for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) lemma filterlim_times_pos: "LIM x F1. c * f x :> at_right l" if "filterlim f (at_right p) F1" "0 < c" "l = c * p" for c::"'a::{linordered_field, linorder_topology}" unfolding filterlim_iff proof safe fix P assume "\\<^sub>F x in at_right l. P x" then obtain d where "c * p < d" "\y. y > c * p \ y < d \ P y" unfolding \l = _ \ eventually_at_right_field by auto then have "\\<^sub>F a in at_right p. P (c * a)" by (auto simp: eventually_at_right_field \0 < c\ field_simps intro!: exI[where x="d/c"]) from that(1)[unfolded filterlim_iff, rule_format, OF this] show "\\<^sub>F x in F1. P (c * f x)" . qed lemma filtermap_nhds_times: "c \ 0 \ filtermap (times c) (nhds a) = nhds (c * a)" for a c :: "'a::real_normed_field" by (rule filtermap_fun_inverse[where g="\x. inverse c * x"]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_times_pos_at_right: fixes c::"'a::{linordered_field, linorder_topology}" assumes "c > 0" shows "filtermap (times c) (at_right p) = at_right (c * p)" using assms by (intro filtermap_fun_inverse[where g="\x. inverse c * x"]) (auto intro!: filterlim_ident filterlim_times_pos) lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity" proof (rule antisym) have "(inverse \ (0::'a)) at_infinity" by (fact tendsto_inverse_0) then show "filtermap inverse at_infinity \ at (0::'a)" using filterlim_def filterlim_ident filterlim_inverse_at_iff by fastforce next have "filtermap inverse (filtermap inverse (at (0::'a))) \ filtermap inverse at_infinity" using filterlim_inverse_at_infinity unfolding filterlim_def by (rule filtermap_mono) then show "at (0::'a) \ filtermap inverse at_infinity" by (simp add: filtermap_ident filtermap_filtermap) qed lemma lim_at_infinity_0: fixes l :: "'a::{real_normed_field,field}" shows "(f \ l) at_infinity \ ((f \ inverse) \ l) (at (0::'a))" by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) lemma lim_zero_infinity: fixes l :: "'a::{real_normed_field,field}" shows "((\x. f(1 / x)) \ l) (at (0::'a)) \ (f \ l) at_infinity" by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) text \ We only show rules for multiplication and addition when the functions are either against a real value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}. \ lemma filterlim_tendsto_pos_mult_at_top: assumes f: "(f \ c) F" and c: "0 < c" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f \0 < c\ have "eventually (\x. c / 2 < f x) F" by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono simp: dist_real_def abs_real_def split: if_split_asm) moreover from g have "eventually (\x. (Z / c * 2) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim case (elim x) with \0 < Z\ \0 < c\ have "c / 2 * (Z / c * 2) \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) with \0 < c\ show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_at_top: assumes f: "LIM x F. f x :> at_top" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 1 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim case (elim x) with \0 < Z\ have "1 * Z \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) then show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_tendsto_pos: assumes f: "(f \ c) F" and c: "0 < c" and g: "LIM x F. g x :> at_top" shows "LIM x F. (g x * f x:: real) :> at_top" by (auto simp: mult.commute intro!: filterlim_tendsto_pos_mult_at_top f c g) lemma filterlim_tendsto_pos_mult_at_bot: fixes c :: real assumes "(f \ c) F" "0 < c" "filterlim g at_bot F" shows "LIM x F. f x * g x :> at_bot" using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\x. - g x"] assms(3) unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_neg_mult_at_bot: fixes c :: real assumes c: "(f \ c) F" "c < 0" and g: "filterlim g at_top F" shows "LIM x F. f x * g x :> at_bot" using c filterlim_tendsto_pos_mult_at_top[of "\x. - f x" "- c" F, OF _ _ g] unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp lemma filterlim_cmult_at_bot_at_top: assumes "filterlim (h :: _ \ real) at_top F" "c \ 0" "G = (if c > 0 then at_top else at_bot)" shows "filterlim (\x. c * h x) G F" using assms filterlim_tendsto_pos_mult_at_top[OF tendsto_const[of c], of h F] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const[of c], of h F] by simp lemma filterlim_pow_at_top: fixes f :: "'a \ real" assumes "0 < n" and f: "LIM x F. f x :> at_top" shows "LIM x F. (f x)^n :: real :> at_top" using \0 < n\ proof (induct n) case 0 then show ?case by simp next case (Suc n) with f show ?case by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top) qed lemma filterlim_pow_at_bot_even: fixes f :: "real \ real" shows "0 < n \ LIM x F. f x :> at_bot \ even n \ LIM x F. (f x)^n :> at_top" using filterlim_pow_at_top[of n "\x. - f x" F] by (simp add: filterlim_uminus_at_top) lemma filterlim_pow_at_bot_odd: fixes f :: "real \ real" shows "0 < n \ LIM x F. f x :> at_bot \ odd n \ LIM x F. (f x)^n :> at_bot" using filterlim_pow_at_top[of n "\x. - f x" F] by (simp add: filterlim_uminus_at_bot) lemma filterlim_power_at_infinity [tendsto_intros]: fixes F and f :: "'a \ 'b :: real_normed_div_algebra" assumes "filterlim f at_infinity F" "n > 0" shows "filterlim (\x. f x ^ n) at_infinity F" by (rule filterlim_norm_at_top_imp_at_infinity) (auto simp: norm_power intro!: filterlim_pow_at_top assms intro: filterlim_at_infinity_imp_norm_at_top) lemma filterlim_tendsto_add_at_top: assumes f: "(f \ c) F" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. c - 1 < f x) F" by (auto dest!: tendstoD[where e=1] elim!: eventually_mono simp: dist_real_def) moreover from g have "eventually (\x. Z - (c - 1) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma LIM_at_top_divide: fixes f g :: "'a \ real" assumes f: "(f \ a) F" "0 < a" and g: "(g \ 0) F" "eventually (\x. 0 < g x) F" shows "LIM x F. f x / g x :> at_top" unfolding divide_inverse by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) lemma filterlim_at_top_add_at_top: assumes f: "LIM x F. f x :> at_top" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 0 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma tendsto_divide_0: fixes f :: "_ \ 'a::{real_normed_div_algebra, division_ring}" assumes f: "(f \ c) F" and g: "LIM x F. g x :> at_infinity" shows "((\x. f x / g x) \ 0) F" using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse) lemma linear_plus_1_le_power: fixes x :: real assumes x: "0 \ x" shows "real n * x + 1 \ (x + 1) ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) from x have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" by (simp add: field_simps) also have "\ \ (x + 1)^Suc n" using Suc x by (simp add: mult_left_mono) finally show ?case . qed lemma filterlim_realpow_sequentially_gt1: fixes x :: "'a :: real_normed_div_algebra" assumes x[arith]: "1 < norm x" shows "LIM n sequentially. x ^ n :> at_infinity" proof (intro filterlim_at_infinity[THEN iffD2] allI impI) fix y :: real assume "0 < y" obtain N :: nat where "y < real N * (norm x - 1)" by (meson diff_gt_0_iff_gt reals_Archimedean3 x) also have "\ \ real N * (norm x - 1) + 1" by simp also have "\ \ (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp also have "\ = norm x ^ N" by simp finally have "\n\N. y \ norm x ^ n" by (metis order_less_le_trans power_increasing order_less_imp_le x) then show "eventually (\n. y \ norm (x ^ n)) sequentially" unfolding eventually_sequentially by (auto simp: norm_power) qed simp lemma filterlim_divide_at_infinity: fixes f g :: "'a \ 'a :: real_normed_field" assumes "filterlim f (nhds c) F" "filterlim g (at 0) F" "c \ 0" shows "filterlim (\x. f x / g x) at_infinity F" proof - have "filterlim (\x. f x * inverse (g x)) at_infinity F" by (intro tendsto_mult_filterlim_at_infinity[OF assms(1,3)] filterlim_compose [OF filterlim_inverse_at_infinity assms(2)]) thus ?thesis by (simp add: field_simps) qed subsection \Floor and Ceiling\ lemma eventually_floor_less: fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. of_int (floor l) < f x" by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l) lemma eventually_less_ceiling: fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. f x < of_int (ceiling l)" by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le) lemma eventually_floor_eq: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. floor (f x) = floor l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma eventually_ceiling_eq: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. ceiling (f x) = ceiling l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma tendsto_of_int_floor: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" and "l \ \" shows "((\x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \ of_int (floor l)) F" using eventually_floor_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma tendsto_of_int_ceiling: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" and "l \ \" shows "((\x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \ of_int (ceiling l)) F" using eventually_ceiling_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma continuous_on_of_int_floor: "continuous_on (UNIV - \::'a::{order_topology, floor_ceiling} set) (\x. of_int (floor x)::'b::{ring_1, topological_space})" unfolding continuous_on_def by (auto intro!: tendsto_of_int_floor) lemma continuous_on_of_int_ceiling: "continuous_on (UNIV - \::'a::{order_topology, floor_ceiling} set) (\x. of_int (ceiling x)::'b::{ring_1, topological_space})" unfolding continuous_on_def by (auto intro!: tendsto_of_int_ceiling) subsection \Limits of Sequences\ lemma [trans]: "X = Y \ Y \ z \ X \ z" by simp lemma LIMSEQ_iff: fixes L :: "'a::real_normed_vector" shows "(X \ L) = (\r>0. \no. \n \ no. norm (X n - L) < r)" unfolding lim_sequentially dist_norm .. lemma LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. norm (X n - L) < r) \ X \ L" for L :: "'a::real_normed_vector" by (simp add: LIMSEQ_iff) lemma LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. norm (X n - L) < r" for L :: "'a::real_normed_vector" by (simp add: LIMSEQ_iff) lemma LIMSEQ_linear: "X \ x \ l > 0 \ (\ n. X (n * l)) \ x" unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute) text \Transformation of limit.\ lemma Lim_transform: "(g \ a) F \ ((\x. f x - g x) \ 0) F \ (f \ a) F" for a b :: "'a::real_normed_vector" using tendsto_add [of g a F "\x. f x - g x" 0] by simp lemma Lim_transform2: "(f \ a) F \ ((\x. f x - g x) \ 0) F \ (g \ a) F" for a b :: "'a::real_normed_vector" by (erule Lim_transform) (simp add: tendsto_minus_cancel) proposition Lim_transform_eq: "((\x. f x - g x) \ 0) F \ (f \ a) F \ (g \ a) F" for a :: "'a::real_normed_vector" using Lim_transform Lim_transform2 by blast lemma Lim_transform_eventually: "\(f \ l) F; eventually (\x. f x = g x) F\ \ (g \ l) F" using eventually_elim2 by (fastforce simp add: tendsto_def) lemma Lim_transform_within: assumes "(f \ l) (at x within S)" and "0 < d" and "\x'. x'\S \ 0 < dist x' x \ dist x' x < d \ f x' = g x'" shows "(g \ l) (at x within S)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at x within S)" using assms by (auto simp: eventually_at) show "(f \ l) (at x within S)" by fact qed lemma filterlim_transform_within: assumes "filterlim g G (at x within S)" assumes "G \ F" "0x'. x' \ S \ 0 < dist x' x \ dist x' x < d \ f x' = g x') " shows "filterlim f F (at x within S)" using assms apply (elim filterlim_mono_eventually) unfolding eventually_at by auto text \Common case assuming being away from some crucial point like 0.\ lemma Lim_transform_away_within: fixes a b :: "'a::t1_space" assumes "a \ b" and "\x\S. x \ a \ x \ b \ f x = g x" and "(f \ l) (at a within S)" shows "(g \ l) (at a within S)" proof (rule Lim_transform_eventually) show "(f \ l) (at a within S)" by fact show "eventually (\x. f x = g x) (at a within S)" unfolding eventually_at_topological by (rule exI [where x="- {b}"]) (simp add: open_Compl assms) qed lemma Lim_transform_away_at: fixes a b :: "'a::t1_space" assumes ab: "a \ b" and fg: "\x. x \ a \ x \ b \ f x = g x" and fl: "(f \ l) (at a)" shows "(g \ l) (at a)" using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp text \Alternatively, within an open set.\ lemma Lim_transform_within_open: assumes "(f \ l) (at a within T)" and "open s" and "a \ s" and "\x. x\s \ x \ a \ f x = g x" shows "(g \ l) (at a within T)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at a within T)" unfolding eventually_at_topological using assms by auto show "(f \ l) (at a within T)" by fact qed text \A congruence rule allowing us to transform limits assuming not at point.\ lemma Lim_cong_within: assumes "a = b" and "x = y" and "S = T" and "\x. x \ b \ x \ T \ f x = g x" shows "(f \ x) (at a within S) \ (g \ y) (at b within T)" unfolding tendsto_def eventually_at_topological using assms by simp text \An unbounded sequence's inverse tends to 0.\ lemma LIMSEQ_inverse_zero: assumes "\r::real. \N. \n\N. r < X n" shows "(\n. inverse (X n)) \ 0" apply (rule filterlim_compose[OF tendsto_inverse_0]) by (metis assms eventually_at_top_linorderI filterlim_at_top_dense filterlim_at_top_imp_at_infinity) text \The sequence \<^term>\1/n\ tends to 0 as \<^term>\n\ tends to infinity.\ lemma LIMSEQ_inverse_real_of_nat: "(\n. inverse (real (Suc n))) \ 0" by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity) text \ The sequence \<^term>\r + 1/n\ tends to \<^term>\r\ as \<^term>\n\ tends to infinity is now easily proved. \ lemma LIMSEQ_inverse_real_of_nat_add: "(\n. r + inverse (real (Suc n))) \ r" using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus: "(\n. r + -inverse (real (Suc n))) \ r" using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(\n. r * (1 + - inverse (real (Suc n)))) \ r" using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] by auto lemma lim_inverse_n: "((\n. inverse(of_nat n)) \ (0::'a::real_normed_field)) sequentially" using lim_1_over_n by (simp add: inverse_eq_divide) lemma lim_inverse_n': "((\n. 1 / n) \ 0) sequentially" using lim_inverse_n by (simp add: inverse_eq_divide) lemma LIMSEQ_Suc_n_over_n: "(\n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps) have "(\n. 1 + inverse (of_nat n) :: 'a) \ 1 + 0" by (intro tendsto_add tendsto_const lim_inverse_n) then show "(\n. 1 + inverse (of_nat n) :: 'a) \ 1" by simp qed lemma LIMSEQ_n_over_Suc_n: "(\n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. inverse (of_nat (Suc n) / of_nat n :: 'a) = of_nat n / of_nat (Suc n)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps del: of_nat_Suc) have "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ inverse 1" by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all then show "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ 1" by simp qed subsection \Convergence on sequences\ lemma convergent_cong: assumes "eventually (\x. f x = g x) sequentially" shows "convergent f \ convergent g" unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl) lemma convergent_Suc_iff: "convergent (\n. f (Suc n)) \ convergent f" by (auto simp: convergent_def filterlim_sequentially_Suc) lemma convergent_ignore_initial_segment: "convergent (\n. f (n + m)) = convergent f" proof (induct m arbitrary: f) case 0 then show ?case by simp next case (Suc m) have "convergent (\n. f (n + Suc m)) \ convergent (\n. f (Suc n + m))" by simp also have "\ \ convergent (\n. f (n + m))" by (rule convergent_Suc_iff) also have "\ \ convergent f" by (rule Suc) finally show ?case . qed lemma convergent_add: fixes X Y :: "nat \ 'a::topological_monoid_add" assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n + Y n)" using assms unfolding convergent_def by (blast intro: tendsto_add) lemma convergent_sum: fixes X :: "'a \ nat \ 'b::topological_comm_monoid_add" shows "(\i. i \ A \ convergent (\n. X i n)) \ convergent (\n. \i\A. X i n)" by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add) lemma (in bounded_linear) convergent: assumes "convergent (\n. X n)" shows "convergent (\n. f (X n))" using assms unfolding convergent_def by (blast intro: tendsto) lemma (in bounded_bilinear) convergent: assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n ** Y n)" using assms unfolding convergent_def by (blast intro: tendsto) lemma convergent_minus_iff: fixes X :: "nat \ 'a::topological_group_add" shows "convergent X \ convergent (\n. - X n)" unfolding convergent_def by (force dest: tendsto_minus) lemma convergent_diff: fixes X Y :: "nat \ 'a::topological_group_add" assumes "convergent (\n. X n)" assumes "convergent (\n. Y n)" shows "convergent (\n. X n - Y n)" using assms unfolding convergent_def by (blast intro: tendsto_diff) lemma convergent_norm: assumes "convergent f" shows "convergent (\n. norm (f n))" proof - from assms have "f \ lim f" by (simp add: convergent_LIMSEQ_iff) then have "(\n. norm (f n)) \ norm (lim f)" by (rule tendsto_norm) then show ?thesis by (auto simp: convergent_def) qed lemma convergent_of_real: "convergent f \ convergent (\n. of_real (f n) :: 'a::real_normed_algebra_1)" unfolding convergent_def by (blast intro!: tendsto_of_real) lemma convergent_add_const_iff: "convergent (\n. c + f n :: 'a::topological_ab_group_add) \ convergent f" proof assume "convergent (\n. c + f n)" from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp next assume "convergent f" from convergent_add[OF convergent_const[of c] this] show "convergent (\n. c + f n)" by simp qed lemma convergent_add_const_right_iff: "convergent (\n. f n + c :: 'a::topological_ab_group_add) \ convergent f" using convergent_add_const_iff[of c f] by (simp add: add_ac) lemma convergent_diff_const_right_iff: "convergent (\n. f n - c :: 'a::topological_ab_group_add) \ convergent f" using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac) lemma convergent_mult: fixes X Y :: "nat \ 'a::topological_semigroup_mult" assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n * Y n)" using assms unfolding convergent_def by (blast intro: tendsto_mult) lemma convergent_mult_const_iff: assumes "c \ 0" shows "convergent (\n. c * f n :: 'a::{field,topological_semigroup_mult}) \ convergent f" proof assume "convergent (\n. c * f n)" from assms convergent_mult[OF this convergent_const[of "inverse c"]] show "convergent f" by (simp add: field_simps) next assume "convergent f" from convergent_mult[OF convergent_const[of c] this] show "convergent (\n. c * f n)" by simp qed lemma convergent_mult_const_right_iff: fixes c :: "'a::{field,topological_semigroup_mult}" assumes "c \ 0" shows "convergent (\n. f n * c) \ convergent f" using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac) lemma convergent_imp_Bseq: "convergent f \ Bseq f" by (simp add: Cauchy_Bseq convergent_Cauchy) text \A monotone sequence converges to its least upper bound.\ lemma LIMSEQ_incseq_SUP: fixes X :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology}" assumes u: "bdd_above (range X)" and X: "incseq X" shows "X \ (SUP i. X i)" by (rule order_tendstoI) (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) lemma LIMSEQ_decseq_INF: fixes X :: "nat \ 'a::{conditionally_complete_linorder, linorder_topology}" assumes u: "bdd_below (range X)" and X: "decseq X" shows "X \ (INF i. X i)" by (rule order_tendstoI) (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) text \Main monotonicity theorem.\ lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent X" for X :: "nat \ real" by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below) lemma Bseq_mono_convergent: "Bseq X \ (\m n. m \ n \ X m \ X n) \ convergent X" for X :: "nat \ real" by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def) lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \ convergent f \ Bseq f" for f :: "nat \ real" using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast lemma Bseq_monoseq_convergent'_inc: fixes f :: "nat \ real" shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Bseq_monoseq_convergent'_dec: fixes f :: "nat \ real" shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Cauchy_iff: "Cauchy X \ (\e>0. \M. \m\M. \n\M. norm (X m - X n) < e)" for X :: "nat \ 'a::real_normed_vector" unfolding Cauchy_def dist_norm .. lemma CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e) \ Cauchy X" for X :: "nat \ 'a::real_normed_vector" by (simp add: Cauchy_iff) lemma CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e" for X :: "nat \ 'a::real_normed_vector" by (simp add: Cauchy_iff) lemma incseq_convergent: fixes X :: "nat \ real" assumes "incseq X" and "\i. X i \ B" obtains L where "X \ L" "\i. X i \ L" proof atomize_elim from incseq_bounded[OF assms] \incseq X\ Bseq_monoseq_convergent[of X] obtain L where "X \ L" by (auto simp: convergent_def monoseq_def incseq_def) with \incseq X\ show "\L. X \ L \ (\i. X i \ L)" by (auto intro!: exI[of _ L] incseq_le) qed lemma decseq_convergent: fixes X :: "nat \ real" assumes "decseq X" and "\i. B \ X i" obtains L where "X \ L" "\i. L \ X i" proof atomize_elim from decseq_bounded[OF assms] \decseq X\ Bseq_monoseq_convergent[of X] obtain L where "X \ L" by (auto simp: convergent_def monoseq_def decseq_def) with \decseq X\ show "\L. X \ L \ (\i. L \ X i)" by (auto intro!: exI[of _ L] decseq_ge) qed lemma monoseq_convergent: fixes X :: "nat \ real" assumes X: "monoseq X" and B: "\i. \X i\ \ B" obtains L where "X \ L" using X unfolding monoseq_iff proof assume "incseq X" show thesis using abs_le_D1 [OF B] incseq_convergent [OF \incseq X\] that by meson next assume "decseq X" show thesis using decseq_convergent [OF \decseq X\] that by (metis B abs_le_iff add.inverse_inverse neg_le_iff_le) qed subsection \More about @{term filterlim} (thanks to Wenda Li)\ lemma filterlim_at_infinity_times: fixes f :: "'a \ 'b::real_normed_field" assumes "filterlim f at_infinity F" "filterlim g at_infinity F" shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ 0 * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) then have "filterlim (\x. inverse (f x) * inverse (g x)) (at 0) F" unfolding filterlim_at using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) then show ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all qed lemma filterlim_at_top_at_bot[elim]: fixes f::"'a \ 'b::unbounded_dense_linorder" and F::"'a filter" assumes top:"filterlim f at_top F" and bot: "filterlim f at_bot F" and "F\bot" shows False proof - obtain c::'b where True by auto have "\\<^sub>F x in F. c < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < c" using bot unfolding filterlim_at_bot_dense by auto ultimately have "\\<^sub>F x in F. c < f x \ f x < c" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_top_nhds[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_top F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain c'::'b where "c'>c" using gt_ex by blast have "\\<^sub>F x in F. c' < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < c'" using order_tendstoD[OF tendsto,of c'] \c'>c\ by auto ultimately have "\\<^sub>F x in F. c' < f x \ f x < c'" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_bot_nhds[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_bot F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain c'::'b where "c'\<^sub>F x in F. c' > f x" using top unfolding filterlim_at_bot_dense by auto moreover have "\\<^sub>F x in F. f x > c'" using order_tendstoD[OF tendsto,of c'] \c' by auto ultimately have "\\<^sub>F x in F. c' < f x \ f x < c'" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma eventually_times_inverse_1: fixes f::"'a \ 'b::{field,t2_space}" assumes "(f \ c) F" "c\0" shows "\\<^sub>F x in F. inverse (f x) * f x = 1" by (smt (verit) assms eventually_mono mult.commute right_inverse tendsto_imp_eventually_ne) lemma filterlim_at_infinity_divide_iff: fixes f::"'a \ 'b::real_normed_field" assumes "(f \ c) F" "c\0" shows "(LIM x F. f x / g x :> at_infinity) \ (LIM x F. g x :> at 0)" proof assume "LIM x F. f x / g x :> at_infinity" then have "LIM x F. inverse (f x) * (f x / g x) :> at_infinity" using assms tendsto_inverse tendsto_mult_filterlim_at_infinity by fastforce then have "LIM x F. inverse (g x) :> at_infinity" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono simp add:field_simps) then show "filterlim g (at 0) F" using filterlim_inverse_at_iff[symmetric] by force next assume "filterlim g (at 0) F" then have "filterlim (\x. inverse (g x)) at_infinity F" using filterlim_compose filterlim_inverse_at_infinity by blast then have "LIM x F. f x * inverse (g x) :> at_infinity" using tendsto_mult_filterlim_at_infinity[OF assms, of "\x. inverse(g x)"] by simp then show "LIM x F. f x / g x :> at_infinity" by (simp add: divide_inverse) qed lemma filterlim_tendsto_pos_mult_at_top_iff: fixes f::"'a \ real" assumes "(f \ c) F" and "0 < c" shows "(LIM x F. (f x * g x) :> at_top) \ (LIM x F. g x :> at_top)" proof assume "filterlim g at_top F" then show "LIM x F. f x * g x :> at_top" using filterlim_tendsto_pos_mult_at_top[OF assms] by auto next assume asm:"LIM x F. f x * g x :> at_top" have "((\x. inverse (f x)) \ inverse c) F" using tendsto_inverse[OF assms(1)] \0 by auto moreover have "inverse c >0" using assms(2) by auto ultimately have "LIM x F. inverse (f x) * (f x * g x) :> at_top" using filterlim_tendsto_pos_mult_at_top[OF _ _ asm,of "\x. inverse (f x)" "inverse c"] by auto then show "LIM x F. g x :> at_top" apply (elim filterlim_mono_eventually) apply simp_all[2] using eventually_times_inverse_1[OF assms(1)] \c>0\ eventually_mono by fastforce qed lemma filterlim_tendsto_pos_mult_at_bot_iff: fixes c :: real assumes "(f \ c) F" "0 < c" shows "(LIM x F. f x * g x :> at_bot) \ filterlim g at_bot F" using filterlim_tendsto_pos_mult_at_top_iff[OF assms(1,2), of "\x. - g x"] unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_neg_mult_at_top_iff: fixes f::"'a \ real" assumes "(f \ c) F" and "c < 0" shows "(LIM x F. (f x * g x) :> at_top) \ (LIM x F. g x :> at_bot)" proof - have "(LIM x F. f x * g x :> at_top) = (LIM x F. - g x :> at_top)" apply (rule filterlim_tendsto_pos_mult_at_top_iff[of "\x. - f x" "-c" F "\x. - g x", simplified]) using assms by (auto intro: tendsto_intros ) also have "... = (LIM x F. g x :> at_bot)" using filterlim_uminus_at_bot[symmetric] by auto finally show ?thesis . qed lemma filterlim_tendsto_neg_mult_at_bot_iff: fixes c :: real assumes "(f \ c) F" "0 > c" shows "(LIM x F. f x * g x :> at_bot) \ filterlim g at_top F" using filterlim_tendsto_neg_mult_at_top_iff[OF assms(1,2), of "\x. - g x"] unfolding filterlim_uminus_at_top by simp subsection \Power Sequences\ lemma Bseq_realpow: "0 \ x \ x \ 1 \ Bseq (\n. x ^ n)" for x :: real by (metis decseq_bounded decseq_def power_decreasing zero_le_power) lemma monoseq_realpow: "0 \ x \ x \ 1 \ monoseq (\n. x ^ n)" for x :: real using monoseq_def power_decreasing by blast lemma convergent_realpow: "0 \ x \ x \ 1 \ convergent (\n. x ^ n)" for x :: real by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) lemma LIMSEQ_inverse_realpow_zero: "1 < x \ (\n. inverse (x ^ n)) \ 0" for x :: real by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp lemma LIMSEQ_realpow_zero: fixes x :: real assumes "0 \ x" "x < 1" shows "(\n. x ^ n) \ 0" proof (cases "x = 0") case False with \0 \ x\ have "1 < inverse x" using \x < 1\ by (simp add: one_less_inverse) then have "(\n. inverse (inverse x ^ n)) \ 0" by (rule LIMSEQ_inverse_realpow_zero) then show ?thesis by (simp add: power_inverse) next case True show ?thesis by (rule LIMSEQ_imp_Suc) (simp add: True) qed lemma LIMSEQ_power_zero [tendsto_intros]: "norm x < 1 \ (\n. x ^ n) \ 0" for x :: "'a::real_normed_algebra_1" apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) by (simp add: Zfun_le norm_power_ineq tendsto_Zfun_iff) lemma LIMSEQ_divide_realpow_zero: "1 < x \ (\n. a / (x ^ n) :: real) \ 0" by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp lemma tendsto_power_zero: fixes x::"'a::real_normed_algebra_1" assumes "filterlim f at_top F" assumes "norm x < 1" shows "((\y. x ^ (f y)) \ 0) F" proof (rule tendstoI) fix e::real assume "0 < e" from tendstoD[OF LIMSEQ_power_zero[OF \norm x < 1\] \0 < e\] have "\\<^sub>F xa in sequentially. norm (x ^ xa) < e" by simp then obtain N where N: "norm (x ^ n) < e" if "n \ N" for n by (auto simp: eventually_sequentially) have "\\<^sub>F i in F. f i \ N" using \filterlim f sequentially F\ by (simp add: filterlim_at_top) then show "\\<^sub>F i in F. dist (x ^ f i) 0 < e" by eventually_elim (auto simp: N) qed text \Limit of \<^term>\c^n\ for \<^term>\\c\ < 1\.\ lemma LIMSEQ_abs_realpow_zero: "\c\ < 1 \ (\n. \c\ ^ n :: real) \ 0" by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) lemma LIMSEQ_abs_realpow_zero2: "\c\ < 1 \ (\n. c ^ n :: real) \ 0" by (rule LIMSEQ_power_zero) simp subsection \Limits of Functions\ lemma LIM_eq: "f \a\ L = (\r>0. \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r)" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_def dist_norm) lemma LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r) \ f \a\ L" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_eq) lemma LIM_D: "f \a\ L \ 0 < r \ \s>0.\x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_eq) lemma LIM_offset: "f \a\ L \ (\x. f (x + k)) \(a - k)\ L" for a :: "'a::real_normed_vector" by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap) lemma LIM_offset_zero: "f \a\ L \ (\h. f (a + h)) \0\ L" for a :: "'a::real_normed_vector" by (drule LIM_offset [where k = a]) (simp add: add.commute) lemma LIM_offset_zero_cancel: "(\h. f (a + h)) \0\ L \ f \a\ L" for a :: "'a::real_normed_vector" by (drule LIM_offset [where k = "- a"]) simp lemma LIM_offset_zero_iff: "NO_MATCH 0 a \ f \a\ L \ (\h. f (a + h)) \0\ L" for f :: "'a :: real_normed_vector \ _" using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto lemma tendsto_offset_zero_iff: fixes f :: "'a :: real_normed_vector \ _" assumes " NO_MATCH 0 a" "a \ S" "open S" shows "(f \ L) (at a within S) \ ((\h. f (a + h)) \ L) (at 0)" using assms by (simp add: tendsto_within_open_NO_MATCH LIM_offset_zero_iff) lemma LIM_zero: "(f \ l) F \ ((\x. f x - l) \ 0) F" for f :: "'a \ 'b::real_normed_vector" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_cancel: fixes f :: "'a \ 'b::real_normed_vector" shows "((\x. f x - l) \ 0) F \ (f \ l) F" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_iff: "((\x. f x - l) \ 0) F = (f \ l) F" for f :: "'a \ 'b::real_normed_vector" unfolding tendsto_iff dist_norm by simp lemma LIM_imp_LIM: fixes f :: "'a::topological_space \ 'b::real_normed_vector" fixes g :: "'a::topological_space \ 'c::real_normed_vector" assumes f: "f \a\ l" and le: "\x. x \ a \ norm (g x - m) \ norm (f x - l)" shows "g \a\ m" by (rule metric_LIM_imp_LIM [OF f]) (simp add: dist_norm le) lemma LIM_equal2: fixes f g :: "'a::real_normed_vector \ 'b::topological_space" assumes "0 < R" and "\x. x \ a \ norm (x - a) < R \ f x = g x" shows "g \a\ l \ f \a\ l" by (rule metric_LIM_equal2 [OF _ assms]) (simp_all add: dist_norm) lemma LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ b" shows "(\x. g (f x)) \a\ c" by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" assumes f: "f \a\ 0" and 1: "\x. x \ a \ 0 \ g x" and 2: "\x. x \ a \ g x \ f x" shows "g \a\ 0" proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) fix x assume x: "x \ a" with 1 have "norm (g x - 0) = g x" by simp also have "g x \ f x" by (rule 2 [OF x]) also have "f x \ \f x\" by (rule abs_ge_self) also have "\f x\ = norm (f x - 0)" by simp finally show "norm (g x - 0) \ norm (f x - 0)" . qed subsection \Continuity\ lemma LIM_isCont_iff: "(f \a\ f a) = ((\h. f (a + h)) \0\ f a)" for f :: "'a::real_normed_vector \ 'b::topological_space" by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) lemma isCont_iff: "isCont f x = (\h. f (x + h)) \0\ f x" for f :: "'a::real_normed_vector \ 'b::topological_space" by (simp add: isCont_def LIM_isCont_iff) lemma isCont_LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" by (rule LIM_compose2 [OF f g inj]) lemma isCont_norm [simp]: "isCont f a \ isCont (\x. norm (f x)) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_norm) lemma isCont_rabs [simp]: "isCont f a \ isCont (\x. \f x\) a" for f :: "'a::t2_space \ real" by (fact continuous_rabs) lemma isCont_add [simp]: "isCont f a \ isCont g a \ isCont (\x. f x + g x) a" for f :: "'a::t2_space \ 'b::topological_monoid_add" by (fact continuous_add) lemma isCont_minus [simp]: "isCont f a \ isCont (\x. - f x) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_minus) lemma isCont_diff [simp]: "isCont f a \ isCont g a \ isCont (\x. f x - g x) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_diff) lemma isCont_mult [simp]: "isCont f a \ isCont g a \ isCont (\x. f x * g x) a" for f g :: "'a::t2_space \ 'b::real_normed_algebra" by (fact continuous_mult) lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" by (fact continuous) lemma (in bounded_bilinear) isCont: "isCont f a \ isCont g a \ isCont (\x. f x ** g x) a" by (fact continuous) lemmas isCont_scaleR [simp] = bounded_bilinear.isCont [OF bounded_bilinear_scaleR] lemmas isCont_of_real [simp] = bounded_linear.isCont [OF bounded_linear_of_real] lemma isCont_power [simp]: "isCont f a \ isCont (\x. f x ^ n) a" for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" by (fact continuous_power) lemma isCont_sum [simp]: "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" for f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" by (auto intro: continuous_sum) subsection \Uniform Continuity\ lemma uniformly_continuous_on_def: fixes f :: "'a::metric_space \ 'b::metric_space" shows "uniformly_continuous_on s f \ (\e>0. \d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding uniformly_continuous_on_uniformity uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric) abbreviation isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where "isUCont f \ uniformly_continuous_on UNIV f" lemma isUCont_def: "isUCont f \ (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" by (auto simp: uniformly_continuous_on_def dist_commute) lemma isUCont_isCont: "isUCont f \ isCont f x" by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at) lemma uniformly_continuous_on_Cauchy: fixes f :: "'a::metric_space \ 'b::metric_space" assumes "uniformly_continuous_on S f" "Cauchy X" "\n. X n \ S" shows "Cauchy (\n. f (X n))" using assms unfolding uniformly_continuous_on_def by (meson Cauchy_def) lemma isUCont_Cauchy: "isUCont f \ Cauchy X \ Cauchy (\n. f (X n))" by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all -lemma uniformly_continuous_imp_Cauchy_continuous: - fixes f :: "'a::metric_space \ 'b::metric_space" - shows "\uniformly_continuous_on S f; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" - by (simp add: uniformly_continuous_on_def Cauchy_def) meson - lemma (in bounded_linear) isUCont: "isUCont f" unfolding isUCont_def dist_norm proof (intro allI impI) fix r :: real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "norm (f x) \ norm x * K" for x using pos_bounded by blast show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" proof (rule exI, safe) from r K show "0 < r / K" by simp next fix x y :: 'a assume xy: "norm (x - y) < r / K" have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) also have "\ \ norm (x - y) * K" by (rule norm_le) also from K xy have "\ < r" by (simp only: pos_less_divide_eq) finally show "norm (f x - f y) < r" . qed qed lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" by (rule isUCont [THEN isUCont_Cauchy]) lemma LIM_less_bound: fixes f :: "real \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" shows "0 \ f x" proof (rule tendsto_lowerbound) show "(f \ f x) (at_left x)" using \isCont f x\ by (simp add: filterlim_at_split isCont_def) show "eventually (\x. 0 \ f x) (at_left x)" using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"]) qed simp subsection \Nested Intervals and Bisection -- Needed for Compactness\ lemma nested_sequence_unique: assumes "\n. f n \ f (Suc n)" "\n. g (Suc n) \ g n" "\n. f n \ g n" "(\n. f n - g n) \ 0" shows "\l::real. ((\n. f n \ l) \ f \ l) \ ((\n. l \ g n) \ g \ l)" proof - have "incseq f" unfolding incseq_Suc_iff by fact have "decseq g" unfolding decseq_Suc_iff by fact have "f n \ g 0" for n proof - from \decseq g\ have "g n \ g 0" by (rule decseqD) simp with \\n. f n \ g n\[THEN spec, of n] show ?thesis by auto qed then obtain u where "f \ u" "\i. f i \ u" using incseq_convergent[OF \incseq f\] by auto moreover have "f 0 \ g n" for n proof - from \incseq f\ have "f 0 \ f n" by (rule incseqD) simp with \\n. f n \ g n\[THEN spec, of n] show ?thesis by simp qed then obtain l where "g \ l" "\i. l \ g i" using decseq_convergent[OF \decseq g\] by auto moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \f \ u\ \g \ l\]] ultimately show ?thesis by auto qed lemma Bolzano[consumes 1, case_names trans local]: fixes P :: "real \ real \ bool" assumes [arith]: "a \ b" and trans: "\a b c. P a b \ P b c \ a \ b \ b \ c \ P a c" and local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" shows "P a b" proof - define bisect where "bisect \ \(x,y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2)" define l u where "l n \ fst ((bisect^^n)(a,b))" and "u n \ snd ((bisect^^n)(a,b))" for n have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" by (simp_all add: l_def u_def bisect_def split: prod.split) have [simp]: "l n \ u n" for n by (induct n) auto have "\x. ((\n. l n \ x) \ l \ x) \ ((\n. x \ u n) \ u \ x)" proof (safe intro!: nested_sequence_unique) show "l n \ l (Suc n)" "u (Suc n) \ u n" for n by (induct n) auto next have "l n - u n = (a - b) / 2^n" for n by (induct n) (auto simp: field_simps) then show "(\n. l n - u n) \ 0" by (simp add: LIMSEQ_divide_realpow_zero) qed fact then obtain x where x: "\n. l n \ x" "\n. x \ u n" and "l \ x" "u \ x" by auto obtain d where "0 < d" and d: "a \ x \ x \ b \ b - a < d \ P a b" for a b using \l 0 \ x\ \x \ u 0\ local[of x] by auto show "P a b" proof (rule ccontr) assume "\ P a b" have "\ P (l n) (u n)" for n proof (induct n) case 0 then show ?case by (simp add: \\ P a b\) next case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto qed moreover { have "eventually (\n. x - d / 2 < l n) sequentially" using \0 < d\ \l \ x\ by (intro order_tendstoD[of _ x]) auto moreover have "eventually (\n. u n < x + d / 2) sequentially" using \0 < d\ \u \ x\ by (intro order_tendstoD[of _ x]) auto ultimately have "eventually (\n. P (l n) (u n)) sequentially" proof eventually_elim case (elim n) from add_strict_mono[OF this] have "u n - l n < d" by simp with x show "P (l n) (u n)" by (rule d) qed } ultimately show False by simp qed qed lemma compact_Icc[simp, intro]: "compact {a .. b::real}" proof (cases "a \ b", rule compactI) fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" define T where "T = {a .. b}" from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" proof (induct rule: Bolzano) case (trans a b c) then have *: "{a..c} = {a..b} \ {b..c}" by auto with trans obtain C1 C2 where "C1\C" "finite C1" "{a..b} \ \C1" "C2\C" "finite C2" "{b..c} \ \C2" by auto with trans show ?case unfolding * by (intro exI[of _ "C1 \ C2"]) auto next case (local x) with C have "x \ \C" by auto with C(2) obtain c where "x \ c" "open c" "c \ C" by auto then obtain e where "0 < e" "{x - e <..< x + e} \ c" by (auto simp: open_dist dist_real_def subset_eq Ball_def abs_less_iff) with \c \ C\ show ?case by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto qed qed simp lemma continuous_image_closed_interval: fixes a b and f :: "real \ real" defines "S \ {a..b}" assumes "a \ b" and f: "continuous_on S f" shows "\c d. f`S = {c..d} \ c \ d" proof - have S: "compact S" "S \ {}" using \a \ b\ by (auto simp: S_def) obtain c where "c \ S" "\d\S. f d \ f c" using continuous_attains_sup[OF S f] by auto moreover obtain d where "d \ S" "\c\S. f d \ f c" using continuous_attains_inf[OF S f] by auto moreover have "connected (f`S)" using connected_continuous_image[OF f] connected_Icc by (auto simp: S_def) ultimately have "f ` S = {f d .. f c} \ f d \ f c" by (auto simp: connected_iff_interval) then show ?thesis by auto qed lemma open_Collect_positive: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on s f" shows "\A. open A \ A \ s = {x\s. 0 < f x}" using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"] by (auto simp: Int_def field_simps) lemma open_Collect_less_Int: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" shows "\A. open A \ A \ s = {x\s. f x < g x}" using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) subsection \Boundedness of continuous functions\ text\By bisection, function continuous on closed interval is bounded above\ lemma isCont_eq_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_sup[of "{a..b}" f] by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_eq_Lb: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_inf[of "{a..b}" f] by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_bounded: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" using isCont_eq_Ub[of a b f] by auto lemma isCont_has_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" using isCont_eq_Ub[of a b f] by auto lemma isCont_Lb_Ub: fixes f :: "real \ real" assumes "a \ b" "\x. a \ x \ x \ b \ isCont f x" shows "\L M. (\x. a \ x \ x \ b \ L \ f x \ f x \ M) \ (\y. L \ y \ y \ M \ (\x. a \ x \ x \ b \ (f x = y)))" proof - obtain M where M: "a \ M" "M \ b" "\x. a \ x \ x \ b \ f x \ f M" using isCont_eq_Ub[OF assms] by auto obtain L where L: "a \ L" "L \ b" "\x. a \ x \ x \ b \ f L \ f x" using isCont_eq_Lb[OF assms] by auto have "(\x. a \ x \ x \ b \ f L \ f x \ f x \ f M)" using M L by simp moreover have "(\y. f L \ y \ y \ f M \ (\x\a. x \ b \ f x = y))" proof (cases "L \ M") case True then show ?thesis using IVT[of f L _ M] M L assms by (metis order.trans) next case False then show ?thesis using IVT2[of f L _ M] by (metis L(2) M(1) assms(2) le_cases order.trans) qed ultimately show ?thesis by blast qed text \Continuity of inverse function.\ lemma isCont_inverse_function: fixes f g :: "real \ real" assumes d: "0 < d" and inj: "\z. \z-x\ \ d \ g (f z) = z" and cont: "\z. \z-x\ \ d \ isCont f z" shows "isCont g (f x)" proof - let ?A = "f (x - d)" let ?B = "f (x + d)" let ?D = "{x - d..x + d}" have f: "continuous_on ?D f" using cont by (intro continuous_at_imp_continuous_on ballI) auto then have g: "continuous_on (f`?D) g" using inj by (intro continuous_on_inv) auto from d f have "{min ?A ?B <..< max ?A ?B} \ f ` ?D" by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" by (rule continuous_on_subset) moreover have "(?A < f x \ f x < ?B) \ (?B < f x \ f x < ?A)" using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto then have "f x \ {min ?A ?B <..< max ?A ?B}" by auto ultimately show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma isCont_inverse_function2: fixes f g :: "real \ real" shows "\a < x; x < b; \z. \a \ z; z \ b\ \ g (f z) = z; \z. \a \ z; z \ b\ \ isCont f z\ \ isCont g (f x)" apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"]) apply (simp_all add: abs_le_iff) done text \Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110.\ lemma LIM_fun_gt_zero: "f \c\ l \ 0 < l \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ 0 < f x)" for f :: "real \ real" by (force simp: dest: LIM_D) lemma LIM_fun_less_zero: "f \c\ l \ l < 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x < 0)" for f :: "real \ real" by (drule LIM_D [where r="-l"]) force+ lemma LIM_fun_not_zero: "f \c\ l \ l \ 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x \ 0)" for f :: "real \ real" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff) lemma Lim_topological: "(f \ l) net \ trivial_limit net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" unfolding tendsto_def trivial_limit_eq by auto lemma eventually_within_Un: "eventually P (at x within (s \ t)) \ eventually P (at x within s) \ eventually P (at x within t)" unfolding eventually_at_filter by (auto elim!: eventually_rev_mp) lemma Lim_within_Un: "(f \ l) (at x within (s \ t)) \ (f \ l) (at x within s) \ (f \ l) (at x within t)" unfolding tendsto_def by (auto simp: eventually_within_Un) end