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,2634 +1,3348 @@ 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 + +locale Submetric = Metric_space + fixes A assumes subset: "A \ M" -sublocale submetric \ sub: Metric_space A d +sublocale Submetric \ sub: Metric_space A d by (simp add: subset subspace) -context submetric +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) +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) + +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) 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) + 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: +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 +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" + 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) + 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) + by (meson Submetric.subset subspace) interpret MA: Metric_space A d - by (meson A submetric.subset subspace) + by (meson A Submetric.subset subspace) interpret MB: Metric_space B d - by (meson B submetric.subset subspace) + 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" + 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 sub: "\A. A \ \ \ Submetric M d A" and comp: "\A. A \ \ \ Metric_space.mcomplete A d" - shows "submetric M d (\\)" "Metric_space.mcomplete (\\) d" + shows "Submetric M d (\\)" "Metric_space.mcomplete (\\) d" proof - - show "submetric M d (\\)" - using assms unfolding submetric_def submetric_axioms_def + 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) + 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) + 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) + 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') + 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) + 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) + 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" + 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) + 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\) + 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: +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: +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: +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" + 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 +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) + + + +end + diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy +++ b/src/HOL/Analysis/Abstract_Topology.thy @@ -1,5090 +1,5102 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ theory Abstract_Topology imports Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" begin subsection \General notion of a topology as a value\ definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ (\S T. L S \ L T \ L (S \ T)) \ (\\. (\K\\. L K) \ L (\\))" typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast lemma istopology_openin[iff]: "istopology(openin U)" using openin[of U] by blast lemma istopology_open[iff]: "istopology open" by (auto simp: istopology_def) lemma topology_inverse' [simp]: "istopology U \ openin (topology U) = U" using topology_inverse[unfolded mem_Collect_eq] . lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" by (metis istopology_openin topology_inverse') lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" proof assume "T1 = T2" then show "\S. openin T1 S \ openin T2 S" by simp next assume H: "\S. openin T1 S \ openin T2 S" then have "openin T1 = openin T2" by (simp add: fun_eq_iff) then have "topology (openin T1) = topology (openin T2)" by simp then show "T1 = T2" unfolding openin_inverse . qed text\The "universe": the union of all sets in the topology.\ definition "topspace T = \{S. openin T S}" subsubsection \Main properties of open sets\ proposition openin_clauses: fixes U :: "'a topology" shows "openin U {}" "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def by auto lemma openin_subset: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (rule openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" by (rule openin_clauses) lemma openin_Union[intro]: "(\S. S \ K \ openin U S) \ openin U (\K)" using openin_clauses by blast lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (force simp: openin_Union topspace_def) lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by auto next assume H: ?rhs let ?t = "\{T. openin U T \ T \ S}" have "openin U ?t" by (force simp: openin_Union) also have "?t = S" using H by auto finally show "openin U S" . qed lemma openin_INT [intro]: assumes "finite I" "\i. i \ I \ openin T (U i)" shows "openin T ((\i \ I. U i) \ topspace T)" using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) lemma openin_INT2 [intro]: assumes "finite I" "I \ {}" "\i. i \ I \ openin T (U i)" shows "openin T (\i \ I. U i)" proof - have "(\i \ I. U i) \ topspace T" using \I \ {}\ openin_subset[OF assms(3)] by auto then show ?thesis using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) qed lemma openin_Inter [intro]: assumes "finite \" "\ \ {}" "\X. X \ \ \ openin T X" shows "openin T (\\)" by (metis (full_types) assms openin_INT2 image_ident) lemma openin_Int_Inter: assumes "finite \" "openin T U" "\X. X \ \ \ openin T X" shows "openin T (U \ \\)" using openin_Inter [of "insert U \"] assms by auto subsubsection \Closed sets\ definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" by (simp add: closedin_def) lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" by (auto simp: Diff_Un closedin_def) lemma Diff_Inter[intro]: "A - \S = \{A - s|s. s\S}" by auto lemma closedin_Union: assumes "finite S" "\T. T \ S \ closedin U T" shows "closedin U (\S)" using assms by induction auto lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S. S \K \ closedin U S" shows "closedin U (\K)" using Ke Kc unfolding closedin_def Diff_Inter by auto lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" using assms by blast lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" using closedin_Inter[of "{S,T}" U] by auto lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" by (metis Diff_subset closedin_def double_diff equalityD1 openin_subset) lemma topology_finer_closedin: "topspace X = topspace Y \ (\S. openin Y S \ openin X S) \ (\S. closedin Y S \ closedin X S)" by (metis closedin_def openin_closedin_eq) lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" by (simp add: openin_closedin_eq) lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" by (metis Int_Diff cT closedin_def inf.orderE oS openin_Int openin_subset) lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" by (metis Int_Diff cT closedin_Int closedin_subset inf.orderE oS openin_closedin_eq) lemma all_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) lemma all_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) lemma ex_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) lemma ex_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) subsection\The discrete topology\ definition discrete_topology where "discrete_topology U \ topology (\S. S \ U)" lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \ S \ U" proof - have "istopology (\S. S \ U)" by (auto simp: istopology_def) then show ?thesis by (simp add: discrete_topology_def topology_inverse') qed lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U" by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym) lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \ S \ U" by (simp add: closedin_def) lemma discrete_topology_unique: "discrete_topology U = X \ topspace X = U \ (\x \ U. openin X {x})" (is "?lhs = ?rhs") proof assume R: ?rhs then have "openin X S" if "S \ U" for S using openin_subopen subsetD that by fastforce then show ?lhs by (metis R openin_discrete_topology openin_subset topology_eq) qed auto lemma discrete_topology_unique_alt: "discrete_topology U = X \ topspace X \ U \ (\x \ U. openin X {x})" using openin_subset by (auto simp: discrete_topology_unique) lemma subtopology_eq_discrete_topology_empty: "X = discrete_topology {} \ topspace X = {}" using discrete_topology_unique [of "{}" X] by auto lemma subtopology_eq_discrete_topology_sing: "X = discrete_topology {a} \ topspace X = {a}" by (metis discrete_topology_unique openin_topspace singletonD) subsection \Subspace topology\ definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" (is "istopology ?L") proof - have "?L {}" by blast { fix A B assume A: "?L A" and B: "?L B" from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast have "A \ B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ then have "?L (A \ B)" by blast } moreover { fix K assume K: "K \ Collect ?L" have th0: "Collect ?L = (\S. S \ V) ` Collect (openin U)" by blast from K[unfolded th0 subset_image_iff] obtain Sk where Sk: "Sk \ Collect (openin U)" "K = (\S. S \ V) ` Sk" by blast have "\K = (\Sk) \ V" using Sk by auto moreover have "openin U (\Sk)" using Sk by (auto simp: subset_eq) ultimately have "?L (\K)" by blast } ultimately show ?thesis unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" unfolding subtopology_def topology_inverse'[OF istopology_subtopology] by auto lemma subset_openin_subtopology: "\openin X S; S \ T\ \ openin (subtopology X T) S" by (metis inf.orderE openin_subtopology) lemma openin_subtopology_Int: "openin X S \ openin (subtopology X T) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_Int2: "openin X T \ openin (subtopology X S) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_diff_closed: "\S \ topspace X; closedin X T\ \ openin (subtopology X S) (S - T)" unfolding closedin_def openin_subtopology by (rule_tac x="topspace X - T" in exI) auto lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)" by (force simp: relative_to_def openin_subtopology) lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \ V" by (auto simp: topspace_def openin_subtopology) lemma topspace_subtopology_subset: "S \ topspace X \ topspace(subtopology X S) = S" by (simp add: inf.absorb_iff2) lemma closedin_subtopology: "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) lemma closedin_subtopology_Int_closed: "closedin X T \ closedin (subtopology X S) (S \ T)" using closedin_subtopology inf_commute by blast lemma closedin_subset_topspace: "\closedin X S; S \ T\ \ closedin (subtopology X T) S" using closedin_subtopology by fastforce lemma closedin_relative_to: "(closedin X relative_to S) = closedin (subtopology X S)" by (force simp: relative_to_def closedin_subtopology) lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) lemma subtopology_subtopology: "subtopology (subtopology X S) T = subtopology X (S \ T)" proof - have eq: "\T'. (\S'. T' = S' \ T \ (\T. openin X T \ S' = T \ S)) = (\Sa. T' = Sa \ (S \ T) \ openin X Sa)" by (metis inf_assoc) have "subtopology (subtopology X S) T = topology (\Ta. \Sa. Ta = Sa \ T \ openin (subtopology X S) Sa)" by (simp add: subtopology_def) also have "\ = subtopology X (S \ T)" by (simp add: openin_subtopology eq) (simp add: subtopology_def) finally show ?thesis . qed lemma openin_subtopology_alt: "openin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (openin X)" by (simp add: image_iff inf_commute openin_subtopology) lemma closedin_subtopology_alt: "closedin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (closedin X)" by (simp add: image_iff inf_commute closedin_subtopology) lemma subtopology_superset: assumes UV: "topspace U \ V" shows "subtopology U V = U" proof - { fix S have "openin U S" if "openin U T" "S = T \ V" for T by (metis Int_subset_iff assms inf.orderE openin_subset that) then have "(\T. openin U T \ S = T \ V) \ openin U S" by (metis assms inf.orderE inf_assoc openin_subset) } then show ?thesis unfolding topology_eq openin_subtopology by blast qed lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" by (simp add: subtopology_superset) lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" by (simp add: subtopology_superset) lemma subtopology_restrict: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) lemma openin_subtopology_empty: "openin (subtopology U {}) S \ S = {}" by (metis Int_empty_right openin_empty openin_subtopology) lemma closedin_subtopology_empty: "closedin (subtopology U {}) S \ S = {}" by (metis Int_empty_right closedin_empty closedin_subtopology) lemma closedin_subtopology_refl [simp]: "closedin (subtopology U X) X \ X \ topspace U" by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) lemma closedin_topspace_empty: "topspace T = {} \ (closedin T S \ S = {})" by (simp add: closedin_def) lemma open_in_topspace_empty: "topspace X = {} \ openin X S \ S = {}" by (simp add: openin_closedin_eq) lemma openin_imp_subset: "openin (subtopology U S) T \ T \ S" by (metis Int_iff openin_subtopology subsetI) lemma closedin_imp_subset: "closedin (subtopology U S) T \ T \ S" by (simp add: closedin_def) lemma openin_open_subtopology: "openin X S \ openin (subtopology X S) T \ openin X T \ T \ S" by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology) lemma closedin_closed_subtopology: "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) lemma closedin_trans_full: "\closedin (subtopology X U) S; closedin X U\ \ closedin X S" using closedin_closed_subtopology by blast lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" by (simp add: openin_subtopology) blast lemma closedin_subtopology_Un: "\closedin (subtopology X T) S; closedin (subtopology X U) S\ \ closedin (subtopology X (T \ U)) S" by (simp add: closedin_subtopology) blast lemma openin_trans_full: "\openin (subtopology X U) S; openin X U\ \ openin X S" by (simp add: openin_open_subtopology) subsection \The canonical topology from the underlying type class\ abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" where "top_of_set \ subtopology (topology open)" lemma open_openin: "open S \ openin euclidean S" by simp declare open_openin [symmetric, simp] lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" by (simp) lemma closed_closedin: "closed S \ closedin euclidean S" by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) declare closed_closedin [symmetric, simp] lemma openin_subtopology_self [simp]: "openin (top_of_set S) S" by (metis openin_topspace topspace_euclidean_subtopology) subsubsection\The most basic facts about the usual topology and metric on R\ abbreviation euclideanreal :: "real topology" where "euclideanreal \ topology open" subsection \Basic "localization" results are handy for connectedness.\ lemma openin_open: "openin (top_of_set U) S \ (\T. open T \ (S = U \ T))" by (auto simp: openin_subtopology) lemma openin_Int_open: "\openin (top_of_set U) S; open T\ \ openin (top_of_set U) (S \ T)" by (metis open_Int Int_assoc openin_open) lemma openin_open_Int[intro]: "open S \ openin (top_of_set U) (U \ S)" by (auto simp: openin_open) lemma open_openin_trans[trans]: "open S \ open T \ T \ S \ openin (top_of_set S) T" by (metis Int_absorb1 openin_open_Int) lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_closed: "closedin (top_of_set U) S \ (\T. closed T \ S = U \ T)" by (simp add: closedin_subtopology Int_ac) lemma closedin_closed_Int: "closed S \ closedin (top_of_set U) (U \ S)" by (metis closedin_closed) lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: "\closedin (top_of_set U) V; T \ U; S = V \ T\ \ closedin (top_of_set T) S" by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (top_of_set T) S" by (simp add: finite_imp_closed closed_subset) lemma closedin_singleton [simp]: fixes a :: "'a::t1_space" shows "closedin (top_of_set U) {a} \ a \ U" using closedin_subset by (force intro: closed_subset) lemma openin_euclidean_subtopology_iff: fixes S U :: "'a::metric_space set" shows "openin (top_of_set U) S \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs unfolding openin_open open_dist by blast next define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp apply (rule_tac x="d - dist x a" in exI) by (metis add_0_left dist_commute dist_triangle_lt less_diff_eq) assume ?rhs then have 2: "S = U \ T" unfolding T_def by auto (metis dist_self) from 1 2 show ?lhs unfolding openin_open open_dist by fast qed lemma connected_openin: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_def openin_open disjoint_iff_not_equal by blast lemma connected_openin_eq: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_openin by (metis (no_types, lifting) Un_subset_iff openin_imp_subset subset_antisym) lemma connected_closedin: "connected S \ (\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp add: connected_closed closedin_closed) next assume R: ?rhs then show ?lhs proof (clarsimp simp add: connected_closed closedin_closed) fix A B assume s_sub: "S \ A \ B" "B \ S \ {}" and disj: "A \ B \ S = {}" and cl: "closed A" "closed B" have "S - A = B \ S" using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto then show "A \ S = {}" by (metis Int_Diff_Un Int_Diff_disjoint R cl closedin_closed_Int dual_order.refl inf_commute s_sub(2)) qed qed lemma connected_closedin_eq: "connected S \ \(\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_closedin by (metis Un_subset_iff closedin_imp_subset subset_antisym) text \These "transitivity" results are handy too\ lemma openin_trans[trans]: "openin (top_of_set T) S \ openin (top_of_set U) T \ openin (top_of_set U) S" by (metis openin_Int_open openin_open) lemma openin_open_trans: "openin (top_of_set T) S \ open T \ open S" by (auto simp: openin_open intro: openin_trans) lemma closedin_trans[trans]: "closedin (top_of_set T) S \ closedin (top_of_set U) T \ closedin (top_of_set U) S" by (auto simp: closedin_closed closed_Inter Int_assoc) lemma closedin_closed_trans: "closedin (top_of_set T) S \ closed T \ closed S" by (auto simp: closedin_closed intro: closedin_trans) lemma openin_subtopology_Int_subset: "\openin (top_of_set u) (u \ S); v \ u\ \ openin (top_of_set v) (v \ S)" by (auto simp: openin_subtopology) lemma openin_open_eq: "open s \ (openin (top_of_set s) t \ open t \ t \ s)" using open_subset openin_open_trans openin_subset by fastforce subsection\Derived set (set of limit points)\ definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" lemma derived_set_of_restrict [simp]: "X derived_set_of (topspace X \ S) = X derived_set_of S" by (simp add: derived_set_of_def) (metis openin_subset subset_iff) lemma in_derived_set_of: "x \ X derived_set_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))" by (simp add: derived_set_of_def) lemma derived_set_of_subset_topspace: "X derived_set_of S \ topspace X" by (auto simp add: derived_set_of_def) lemma derived_set_of_subtopology: "(subtopology X U) derived_set_of S = U \ (X derived_set_of (U \ S))" by (simp add: derived_set_of_def openin_subtopology) blast lemma derived_set_of_subset_subtopology: "(subtopology X S) derived_set_of T \ S" by (simp add: derived_set_of_subtopology) lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}" by (auto simp: derived_set_of_def) lemma derived_set_of_mono: "S \ T \ X derived_set_of S \ X derived_set_of T" unfolding derived_set_of_def by blast lemma derived_set_of_Un: "X derived_set_of (S \ T) = X derived_set_of S \ X derived_set_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (clarsimp simp: in_derived_set_of) (metis IntE IntI openin_Int) show "?rhs \ ?lhs" by (simp add: derived_set_of_mono) qed lemma derived_set_of_Union: "finite \ \ X derived_set_of (\\) = (\S \ \. X derived_set_of S)" proof (induction \ rule: finite_induct) case (insert S \) then show ?case by (simp add: derived_set_of_Un) qed auto lemma derived_set_of_topspace: "X derived_set_of (topspace X) = {x \ topspace X. \ openin X {x}}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: in_derived_set_of) show "?rhs \ ?lhs" by (clarsimp simp: in_derived_set_of) (metis openin_closedin_eq openin_subopen singletonD subset_eq) qed lemma discrete_topology_unique_derived_set: "discrete_topology U = X \ topspace X = U \ X derived_set_of U = {}" by (auto simp: discrete_topology_unique derived_set_of_topspace) lemma subtopology_eq_discrete_topology_eq: "subtopology X U = discrete_topology U \ U \ topspace X \ U \ X derived_set_of U = {}" using discrete_topology_unique_derived_set [of U "subtopology X U"] by (auto simp: eq_commute derived_set_of_subtopology) lemma subtopology_eq_discrete_topology: "S \ topspace X \ S \ X derived_set_of S = {} \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq) lemma subtopology_eq_discrete_topology_gen: assumes "S \ X derived_set_of S = {}" shows "subtopology X S = discrete_topology(topspace X \ S)" proof - have "subtopology X S = subtopology X (topspace X \ S)" by (simp add: subtopology_restrict) then show ?thesis using assms by (simp add: inf.assoc subtopology_eq_discrete_topology_eq) qed lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \ S)" proof - have "(\T. \Sa. T = Sa \ S \ Sa \ U) = (\Sa. Sa \ U \ Sa \ S)" by force then show ?thesis by (simp add: subtopology_def) (simp add: discrete_topology_def) qed lemma openin_Int_derived_set_of_subset: "openin X S \ S \ X derived_set_of T \ X derived_set_of (S \ T)" by (auto simp: derived_set_of_def) lemma openin_Int_derived_set_of_eq: assumes "openin X S" shows "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_derived_set_of_subset) show "?rhs \ ?lhs" by (metis derived_set_of_mono inf_commute inf_le1 inf_mono order_refl) qed subsection\ Closure with respect to a topological space\ definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" unfolding closure_of_def using openin_subset by blast lemma in_closure_of: "x \ X closure_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y. y \ S \ y \ T))" by (auto simp: closure_of_def) lemma closure_of: "X closure_of S = topspace X \ (S \ X derived_set_of S)" by (fastforce simp: in_closure_of in_derived_set_of) lemma closure_of_alt: "X closure_of S = topspace X \ S \ X derived_set_of S" using derived_set_of_subset_topspace [of X S] unfolding closure_of_def in_derived_set_of by safe (auto simp: in_derived_set_of) lemma derived_set_of_subset_closure_of: "X derived_set_of S \ X closure_of S" by (fastforce simp: closure_of_def in_derived_set_of) lemma closure_of_subtopology: "(subtopology X U) closure_of S = U \ (X closure_of (U \ S))" unfolding closure_of_def topspace_subtopology openin_subtopology by safe (metis (full_types) IntI Int_iff inf.commute)+ lemma closure_of_empty [simp]: "X closure_of {} = {}" by (simp add: closure_of_alt) lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X" by (simp add: closure_of) lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X" by (simp add: closure_of) lemma closure_of_subset_topspace: "X closure_of S \ topspace X" by (simp add: closure_of) lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \ S" by (simp add: closure_of_subtopology) lemma closure_of_mono: "S \ T \ X closure_of S \ X closure_of T" by (fastforce simp add: closure_of_def) lemma closure_of_subtopology_subset: "(subtopology X U) closure_of S \ (X closure_of S)" unfolding closure_of_subtopology by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2) lemma closure_of_subtopology_mono: "T \ U \ (subtopology X T) closure_of S \ (subtopology X U) closure_of S" unfolding closure_of_subtopology by auto (meson closure_of_mono inf_mono subset_iff) lemma closure_of_Un [simp]: "X closure_of (S \ T) = X closure_of S \ X closure_of T" by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1) lemma closure_of_Union: "finite \ \ X closure_of (\\) = (\S \ \. X closure_of S)" by (induction \ rule: finite_induct) auto lemma closure_of_subset: "S \ topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_Int: "topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_eq: "S \ topspace X \ X closure_of S \ S \ closedin X S" proof - have "openin X (topspace X - S)" if "\x. \x \ topspace X; \T. x \ T \ openin X T \ S \ T \ {}\ \ x \ S" apply (subst openin_subopen) by (metis Diff_iff Diff_mono Diff_triv inf.commute openin_subset order_refl that) then show ?thesis by (auto simp: closedin_def closure_of_def disjoint_iff_not_equal) qed lemma closure_of_eq: "X closure_of S = S \ closedin X S" by (metis closure_of_subset closure_of_subset_eq closure_of_subset_topspace subset_antisym) lemma closedin_contains_derived_set: "closedin X S \ X derived_set_of S \ S \ S \ topspace X" proof (intro iffI conjI) show "closedin X S \ X derived_set_of S \ S" using closure_of_eq derived_set_of_subset_closure_of by fastforce show "closedin X S \ S \ topspace X" using closedin_subset by blast show "X derived_set_of S \ S \ S \ topspace X \ closedin X S" by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE) qed lemma derived_set_subset_gen: "X derived_set_of S \ S \ closedin X (topspace X \ S)" by (simp add: closedin_contains_derived_set derived_set_of_subset_topspace) lemma derived_set_subset: "S \ topspace X \ (X derived_set_of S \ S \ closedin X S)" by (simp add: closedin_contains_derived_set) lemma closedin_derived_set: "closedin (subtopology X T) S \ S \ topspace X \ S \ T \ (\x. x \ X derived_set_of S \ x \ T \ x \ S)" by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1) lemma closedin_Int_closure_of: "closedin (subtopology X S) T \ S \ X closure_of T = T" by (metis Int_left_absorb closure_of_eq closure_of_subtopology) lemma closure_of_closedin: "closedin X S \ X closure_of S = S" by (simp add: closure_of_eq) lemma closure_of_eq_diff: "X closure_of S = topspace X - \{T. openin X T \ disjnt S T}" by (auto simp: closure_of_def disjnt_iff) lemma closedin_closure_of [simp]: "closedin X (X closure_of S)" unfolding closure_of_eq_diff by blast lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S" by (simp add: closure_of_eq) lemma closure_of_hull: assumes "S \ topspace X" shows "X closure_of S = (closedin X) hull S" by (metis assms closedin_closure_of closure_of_eq closure_of_mono closure_of_subset hull_unique) lemma closure_of_minimal: "\S \ T; closedin X T\ \ (X closure_of S) \ T" by (metis closure_of_eq closure_of_mono) lemma closure_of_minimal_eq: "\S \ topspace X; closedin X T\ \ (X closure_of S) \ T \ S \ T" by (meson closure_of_minimal closure_of_subset subset_trans) lemma closure_of_unique: "\S \ T; closedin X T; \T'. \S \ T'; closedin X T'\ \ T \ T'\ \ X closure_of S = T" by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans) lemma closure_of_eq_empty_gen: "X closure_of S = {} \ disjnt (topspace X) S" unfolding disjnt_def closure_of_restrict [where S=S] using closure_of by fastforce lemma closure_of_eq_empty: "S \ topspace X \ X closure_of S = {} \ S = {}" using closure_of_subset by fastforce lemma openin_Int_closure_of_subset: assumes "openin X S" shows "S \ X closure_of T \ X closure_of (S \ T)" proof - have "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" by (meson assms openin_Int_derived_set_of_eq) moreover have "S \ (S \ T) = S \ T" by fastforce ultimately show ?thesis by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1) qed lemma closure_of_openin_Int_closure_of: assumes "openin X S" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" proof show "X closure_of (S \ X closure_of T) \ X closure_of (S \ T)" by (simp add: assms closure_of_minimal openin_Int_closure_of_subset) next show "X closure_of (S \ T) \ X closure_of (S \ X closure_of T)" by (metis Int_subset_iff assms closure_of_alt closure_of_mono inf_mono openin_subset subset_refl sup.coboundedI1) qed lemma openin_Int_closure_of_eq: assumes "openin X S" shows "S \ X closure_of T = S \ X closure_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_closure_of_subset) show "?rhs \ ?lhs" by (metis closure_of_mono inf_commute inf_le1 inf_mono order_refl) qed lemma openin_Int_closure_of_eq_empty: assumes "openin X S" shows "S \ X closure_of T = {} \ S \ T = {}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding disjoint_iff by (meson assms in_closure_of in_mono openin_subset) show "?rhs \ ?lhs" by (simp add: assms openin_Int_closure_of_eq) qed lemma closure_of_openin_Int_superset: "openin X S \ S \ X closure_of T \ X closure_of (S \ T) = X closure_of S" by (metis closure_of_openin_Int_closure_of inf.orderE) lemma closure_of_openin_subtopology_Int_closure_of: assumes S: "openin (subtopology X U) S" and "T \ U" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" (is "?lhs = ?rhs") proof obtain S0 where S0: "openin X S0" "S = S0 \ U" using assms by (auto simp: openin_subtopology) then show "?lhs \ ?rhs" proof - have "S0 \ X closure_of T = S0 \ X closure_of (S0 \ T)" by (meson S0(1) openin_Int_closure_of_eq) moreover have "S0 \ T = S0 \ U \ T" using \T \ U\ by fastforce ultimately have "S \ X closure_of T \ X closure_of (S \ T)" using S0(2) by auto then show ?thesis by (meson closedin_closure_of closure_of_minimal) qed next show "?rhs \ ?lhs" proof - have "T \ S \ T \ X derived_set_of T" by force then show ?thesis by (smt (verit, del_insts) Int_iff in_closure_of inf.orderE openin_subset subsetI) qed qed lemma closure_of_subtopology_open: "openin X U \ S \ U \ (subtopology X U) closure_of S = U \ X closure_of S" by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq) lemma discrete_topology_closure_of: "(discrete_topology U) closure_of S = U \ S" by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl) text\ Interior with respect to a topological space. \ definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: "X interior_of S = X interior_of (topspace X \ S)" using openin_subset by (auto simp: interior_of_def) lemma interior_of_eq: "(X interior_of S = S) \ openin X S" unfolding interior_of_def using openin_subopen by blast lemma interior_of_openin: "openin X S \ X interior_of S = S" by (simp add: interior_of_eq) lemma interior_of_empty [simp]: "X interior_of {} = {}" by (simp add: interior_of_eq) lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X" by (simp add: interior_of_eq) lemma openin_interior_of [simp]: "openin X (X interior_of S)" unfolding interior_of_def using openin_subopen by fastforce lemma interior_of_interior_of [simp]: "X interior_of X interior_of S = X interior_of S" by (simp add: interior_of_eq) lemma interior_of_subset: "X interior_of S \ S" by (auto simp: interior_of_def) lemma interior_of_subset_closure_of: "X interior_of S \ X closure_of S" by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset) lemma subset_interior_of_eq: "S \ X interior_of S \ openin X S" by (metis interior_of_eq interior_of_subset subset_antisym) lemma interior_of_mono: "S \ T \ X interior_of S \ X interior_of T" by (auto simp: interior_of_def) lemma interior_of_maximal: "\T \ S; openin X T\ \ T \ X interior_of S" by (auto simp: interior_of_def) lemma interior_of_maximal_eq: "openin X T \ T \ X interior_of S \ T \ S" by (meson interior_of_maximal interior_of_subset order_trans) lemma interior_of_unique: "\T \ S; openin X T; \T'. \T' \ S; openin X T'\ \ T' \ T\ \ X interior_of S = T" by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym) lemma interior_of_subset_topspace: "X interior_of S \ topspace X" by (simp add: openin_subset) lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \ S" by (meson openin_imp_subset openin_interior_of) lemma interior_of_Int: "X interior_of (S \ T) = X interior_of S \ X interior_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: interior_of_mono) show "?rhs \ ?lhs" by (meson inf_mono interior_of_maximal interior_of_subset openin_Int openin_interior_of) qed lemma interior_of_Inter_subset: "X interior_of (\\) \ (\S \ \. X interior_of S)" by (simp add: INT_greatest Inf_lower interior_of_mono) lemma union_interior_of_subset: "X interior_of S \ X interior_of T \ X interior_of (S \ T)" by (simp add: interior_of_mono) lemma interior_of_eq_empty: "X interior_of S = {} \ (\T. openin X T \ T \ S \ T = {})" by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of) lemma interior_of_eq_empty_alt: "X interior_of S = {} \ (\T. openin X T \ T \ {} \ T - S \ {})" by (auto simp: interior_of_eq_empty) lemma interior_of_Union_openin_subsets: "\{T. openin X T \ T \ S} = X interior_of S" by (rule interior_of_unique [symmetric]) auto lemma interior_of_complement: "X interior_of (topspace X - S) = topspace X - X closure_of S" by (auto simp: interior_of_def closure_of_def) lemma interior_of_closure_of: "X interior_of S = topspace X - X closure_of (topspace X - S)" unfolding interior_of_complement [symmetric] by (metis Diff_Diff_Int interior_of_restrict) lemma closure_of_interior_of: "X closure_of S = topspace X - X interior_of (topspace X - S)" by (simp add: interior_of_complement Diff_Diff_Int closure_of) lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S" unfolding interior_of_def closure_of_def by (blast dest: openin_subset) lemma interior_of_eq_empty_complement: "X interior_of S = {} \ X closure_of (topspace X - S) = topspace X" using interior_of_subset_topspace [of X S] closure_of_complement by fastforce lemma closure_of_eq_topspace: "X closure_of S = topspace X \ X interior_of (topspace X - S) = {}" using closure_of_subset_topspace [of X S] interior_of_complement by fastforce lemma interior_of_subtopology_subset: "U \ X interior_of S \ (subtopology X U) interior_of S" by (auto simp: interior_of_def openin_subtopology) lemma interior_of_subtopology_subsets: "T \ U \ T \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology) lemma interior_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets) lemma interior_of_subtopology_open: assumes "openin X U" shows "(subtopology X U) interior_of S = U \ X interior_of S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (meson assms interior_of_maximal interior_of_subset le_infI openin_interior_of openin_open_subtopology) show "?rhs \ ?lhs" by (simp add: interior_of_subtopology_subset) qed lemma dense_intersects_open: "X closure_of S = topspace X \ (\T. openin X T \ T \ {} \ S \ T \ {})" proof - have "X closure_of S = topspace X \ (topspace X - X interior_of (topspace X - S) = topspace X)" by (simp add: closure_of_interior_of) also have "\ \ X interior_of (topspace X - S) = {}" by (simp add: closure_of_complement interior_of_eq_empty_complement) also have "\ \ (\T. openin X T \ T \ {} \ S \ T \ {})" unfolding interior_of_eq_empty_alt using openin_subset by fastforce finally show ?thesis . qed lemma interior_of_closedin_union_empty_interior_of: assumes "closedin X S" and disj: "X interior_of T = {}" shows "X interior_of (S \ T) = X interior_of S" proof - have "X closure_of (topspace X - T) = topspace X" by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of) then show ?thesis unfolding interior_of_closure_of by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset) qed lemma interior_of_union_eq_empty: "closedin X S \ (X interior_of (S \ T) = {} \ X interior_of S = {} \ X interior_of T = {})" by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset) lemma discrete_topology_interior_of [simp]: "(discrete_topology U) interior_of S = U \ S" by (simp add: interior_of_restrict [of _ S] interior_of_eq) subsection \Frontier with respect to topological space \ definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: "X frontier_of S = X closure_of S \ X closure_of (topspace X - S)" by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of) lemma interior_of_union_frontier_of [simp]: "X interior_of S \ X frontier_of S = X closure_of S" by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym) lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \ S)" by (metis closure_of_restrict frontier_of_def interior_of_restrict) lemma closedin_frontier_of: "closedin X (X frontier_of S)" by (simp add: closedin_Int frontier_of_closures) lemma frontier_of_subset_topspace: "X frontier_of S \ topspace X" by (simp add: closedin_frontier_of closedin_subset) lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \ S" by (metis (no_types) closedin_derived_set closedin_frontier_of) lemma frontier_of_subtopology_subset: "U \ (subtopology X U) frontier_of S \ (X frontier_of S)" proof - have "U \ X interior_of S - subtopology X U interior_of S = {}" by (simp add: interior_of_subtopology_subset) moreover have "X closure_of S \ subtopology X U closure_of S = subtopology X U closure_of S" by (meson closure_of_subtopology_subset inf.absorb_iff2) ultimately show ?thesis unfolding frontier_of_def by blast qed lemma frontier_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X T) frontier_of S \ (subtopology X U) frontier_of S" by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono) lemma clopenin_eq_frontier_of: "closedin X S \ openin X S \ S \ topspace X \ X frontier_of S = {}" proof (cases "S \ topspace X") case True then show ?thesis by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right) next case False then show ?thesis by (simp add: frontier_of_closures openin_closedin_eq) qed lemma frontier_of_eq_empty: "S \ topspace X \ (X frontier_of S = {} \ closedin X S \ openin X S)" by (simp add: clopenin_eq_frontier_of) lemma frontier_of_openin: "openin X S \ X frontier_of S = X closure_of S - S" by (metis (no_types) frontier_of_def interior_of_eq) lemma frontier_of_openin_straddle_Int: assumes "openin X U" "U \ X frontier_of S \ {}" shows "U \ S \ {}" "U - S \ {}" proof - have "U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}" using assms by (simp add: frontier_of_closures) then show "U \ S \ {}" using assms openin_Int_closure_of_eq_empty by fastforce show "U - S \ {}" proof - have "\A. X closure_of (A - S) \ U \ {}" using \U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}\ by blast then have "\ U \ S" by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty) then show ?thesis by blast qed qed lemma frontier_of_subset_closedin: "closedin X S \ (X frontier_of S) \ S" using closure_of_eq frontier_of_def by fastforce lemma frontier_of_empty [simp]: "X frontier_of {} = {}" by (simp add: frontier_of_def) lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}" by (simp add: frontier_of_def) lemma frontier_of_subset_eq: assumes "S \ topspace X" shows "(X frontier_of S) \ S \ closedin X S" proof show "X frontier_of S \ S \ closedin X S" by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff) show "closedin X S \ X frontier_of S \ S" by (simp add: frontier_of_subset_closedin) qed lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S" by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute) lemma frontier_of_disjoint_eq: assumes "S \ topspace X" shows "((X frontier_of S) \ S = {} \ openin X S)" proof assume "X frontier_of S \ S = {}" then have "closedin X (topspace X - S)" using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce then show "openin X S" using assms by (simp add: openin_closedin) next show "openin X S \ X frontier_of S \ S = {}" by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute) qed lemma frontier_of_disjoint_eq_alt: "S \ (topspace X - X frontier_of S) \ openin X S" proof (cases "S \ topspace X") case True show ?thesis using True frontier_of_disjoint_eq by auto next case False then show ?thesis by (meson Diff_subset openin_subset subset_trans) qed lemma frontier_of_Int: "X frontier_of (S \ T) = X closure_of (S \ T) \ (X frontier_of S \ X frontier_of T)" proof - have *: "U \ S \ U \ T \ U \ (S \ A \ T \ B) = U \ (A \ B)" for U S T A B :: "'a set" by blast show ?thesis by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un) qed lemma frontier_of_Int_subset: "X frontier_of (S \ T) \ X frontier_of S \ X frontier_of T" by (simp add: frontier_of_Int) lemma frontier_of_Int_closedin: assumes "closedin X S" "closedin X T" shows "X frontier_of(S \ T) = X frontier_of S \ T \ S \ X frontier_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (force simp add: frontier_of_Int closedin_Int closure_of_closedin) show "?rhs \ ?lhs" using assms frontier_of_subset_closedin by (auto simp add: frontier_of_Int closedin_Int closure_of_closedin) qed lemma frontier_of_Un_subset: "X frontier_of(S \ T) \ X frontier_of S \ X frontier_of T" by (metis Diff_Un frontier_of_Int_subset frontier_of_complement) lemma frontier_of_Union_subset: "finite \ \ X frontier_of (\\) \ (\T \ \. X frontier_of T)" proof (induction \ rule: finite_induct) case (insert A \) then show ?case using frontier_of_Un_subset by fastforce qed simp lemma frontier_of_frontier_of_subset: "X frontier_of (X frontier_of S) \ X frontier_of S" by (simp add: closedin_frontier_of frontier_of_subset_closedin) lemma frontier_of_subtopology_open: "openin X U \ (subtopology X U) frontier_of S = U \ X frontier_of S" by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open) lemma discrete_topology_frontier_of [simp]: "(discrete_topology U) frontier_of S = {}" by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures) lemma openin_subset_topspace_eq: assumes "disjnt S (X frontier_of U)" shows "openin (subtopology X U) S \ openin X S \ S \ U" proof assume S: "openin (subtopology X U) S" show "openin X S \ S \ U" proof show "S \ U" using S openin_imp_subset by blast have "disjnt S (X frontier_of (topspace X \ U))" by (metis assms frontier_of_restrict) moreover have "openin (subtopology X (topspace X \ U)) S" by (simp add: S subtopology_restrict) moreover have "openin X S" if dis: "disjnt S (X frontier_of U)" and ope: "openin (subtopology X U) S" and "U \ topspace X" for S U proof - obtain T where T: "openin X T" "S = T \ U" using ope by (auto simp: openin_subtopology) have "T \ U = T \ X interior_of U" using that T interior_of_subset in_closure_of by (fastforce simp: disjnt_iff frontier_of_def) then show ?thesis using \S = T \ U\ \openin X T\ by auto qed ultimately show "openin X S" by blast qed qed (metis inf.absorb_iff1 openin_subtopology_Int) subsection\Locally finite collections\ definition locally_finite_in where "locally_finite_in X \ \ (\\ \ topspace X) \ (\x \ topspace X. \V. openin X V \ x \ V \ finite {U \ \. U \ V \ {}})" lemma finite_imp_locally_finite_in: "\finite \; \\ \ topspace X\ \ locally_finite_in X \" by (auto simp: locally_finite_in_def) lemma locally_finite_in_subset: assumes "locally_finite_in X \" "\ \ \" shows "locally_finite_in X \" proof - have "finite (\ \ {U. U \ V \ {}}) \ finite (\ \ {U. U \ V \ {}})" for V by (meson \\ \ \\ finite_subset inf_le1 inf_le2 le_inf_iff subset_trans) then show ?thesis using assms unfolding locally_finite_in_def Int_def by fastforce qed lemma locally_finite_in_refinement: assumes \: "locally_finite_in X \" and f: "\S. S \ \ \ f S \ S" shows "locally_finite_in X (f ` \)" proof - show ?thesis unfolding locally_finite_in_def proof safe fix x assume "x \ topspace X" then obtain V where "openin X V" "x \ V" "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast moreover have "{U \ \. f U \ V \ {}} \ {U \ \. U \ V \ {}}" for V using f by blast ultimately have "finite {U \ \. f U \ V \ {}}" using finite_subset by blast moreover have "f ` {U \ \. f U \ V \ {}} = {U \ f ` \. U \ V \ {}}" by blast ultimately have "finite {U \ f ` \. U \ V \ {}}" by (metis (no_types, lifting) finite_imageI) then show "\V. openin X V \ x \ V \ finite {U \ f ` \. U \ V \ {}}" using \openin X V\ \x \ V\ by blast next show "\x xa. \xa \ \; x \ f xa\ \ x \ topspace X" by (meson Sup_upper \ f locally_finite_in_def subset_iff) qed qed lemma locally_finite_in_subtopology: assumes \: "locally_finite_in X \" "\\ \ S" shows "locally_finite_in (subtopology X S) \" unfolding locally_finite_in_def proof safe fix x assume x: "x \ topspace (subtopology X S)" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def topspace_subtopology by blast show "\V. openin (subtopology X S) V \ x \ V \ finite {U \ \. U \ V \ {}}" proof (intro exI conjI) show "openin (subtopology X S) (S \ V)" by (simp add: \openin X V\ openin_subtopology_Int2) have "{U \ \. U \ (S \ V) \ {}} \ {U \ \. U \ V \ {}}" by auto with fin show "finite {U \ \. U \ (S \ V) \ {}}" using finite_subset by auto show "x \ S \ V" using x \x \ V\ by (simp) qed next show "\x A. \x \ A; A \ \\ \ x \ topspace (subtopology X S)" using assms unfolding locally_finite_in_def topspace_subtopology by blast qed lemma closedin_locally_finite_Union: assumes clo: "\S. S \ \ \ closedin X S" and \: "locally_finite_in X \" shows "closedin X (\\)" using \ unfolding locally_finite_in_def closedin_def proof clarify show "openin X (topspace X - \\)" proof (subst openin_subopen, clarify) fix x assume "x \ topspace X" and "x \ \\" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast let ?T = "V - \{S \ \. S \ V \ {}}" show "\T. openin X T \ x \ T \ T \ topspace X - \\" proof (intro exI conjI) show "openin X ?T" by (metis (no_types, lifting) fin \openin X V\ clo closedin_Union mem_Collect_eq openin_diff) show "x \ ?T" using \x \ \\\ \x \ V\ by auto show "?T \ topspace X - \\" using \openin X V\ openin_subset by auto qed qed qed lemma locally_finite_in_closure: assumes \: "locally_finite_in X \" shows "locally_finite_in X ((\S. X closure_of S) ` \)" using \ unfolding locally_finite_in_def proof (intro conjI; clarsimp) fix x A assume "x \ X closure_of A" then show "x \ topspace X" by (meson in_closure_of) next fix x assume "x \ topspace X" and "\\ \ topspace X" then obtain V where V: "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast have eq: "{y \ f ` \. Q y} = f ` {x. x \ \ \ Q(f x)}" for f and Q :: "'a set \ bool" by blast have eq2: "{A \ \. X closure_of A \ V \ {}} = {A \ \. A \ V \ {}}" using openin_Int_closure_of_eq_empty V by blast have "finite {U \ (closure_of) X ` \. U \ V \ {}}" by (simp add: eq eq2 fin) with V show "\V. openin X V \ x \ V \ finite {U \ (closure_of) X ` \. U \ V \ {}}" by blast qed lemma closedin_Union_locally_finite_closure: "locally_finite_in X \ \ closedin X (\((\S. X closure_of S) ` \))" by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" by (simp add: SUP_le_iff Sup_upper closure_of_mono) lemma closure_of_locally_finite_Union: assumes "locally_finite_in X \" shows "X closure_of (\\) = \((\S. X closure_of S) ` \)" proof (rule closure_of_unique) show "\ \ \ \ ((closure_of) X ` \)" using assms by (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def) show "closedin X (\ ((closure_of) X ` \))" using assms by (simp add: closedin_Union_locally_finite_closure) show "\T'. \\ \ \ T'; closedin X T'\ \ \ ((closure_of) X ` \) \ T'" by (simp add: Sup_le_iff closure_of_minimal) qed subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ definition continuous_map where "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" lemma continuous_map: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" by (auto simp: continuous_map_def) lemma continuous_map_image_subset_topspace: "continuous_map X Y f \ f ` (topspace X) \ topspace Y" by (auto simp: continuous_map_def) lemma continuous_map_on_empty: "topspace X = {} \ continuous_map X Y f" by (auto simp: continuous_map_def) lemma continuous_map_closedin: "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" proof - have "(\U. openin Y U \ openin X {x \ topspace X. f x \ U}) = (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" if "\x. x \ topspace X \ f x \ topspace Y" proof - have eq: "{x \ topspace X. f x \ topspace Y \ f x \ C} = (topspace X - {x \ topspace X. f x \ C})" for C using that by blast show ?thesis proof (intro iffI allI impI) fix C assume "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" by (auto simp add: closedin_def eq) next fix U assume "\C. closedin Y C \ closedin X {x \ topspace X. f x \ C}" and "openin Y U" then show "openin X {x \ topspace X. f x \ U}" by (auto simp add: openin_closedin_eq eq) qed qed then show ?thesis by (auto simp: continuous_map_def) qed lemma openin_continuous_map_preimage: "\continuous_map X Y f; openin Y U\ \ openin X {x \ topspace X. f x \ U}" by (simp add: continuous_map_def) lemma closedin_continuous_map_preimage: "\continuous_map X Y f; closedin Y C\ \ closedin X {x \ topspace X. f x \ C}" by (simp add: continuous_map_closedin) lemma openin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "openin X U" "openin Y V" shows "openin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) openin_closedin_eq by fastforce show ?thesis unfolding eq using assms openin_continuous_map_preimage by fastforce qed lemma closedin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "closedin X U" "closedin Y V" shows "closedin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) closedin_def by fastforce show ?thesis unfolding eq using assms closedin_continuous_map_preimage by fastforce qed lemma continuous_map_image_closure_subset: assumes "continuous_map X Y f" shows "f ` (X closure_of S) \ Y closure_of f ` S" proof - have *: "f ` (topspace X) \ topspace Y" by (meson assms continuous_map) have "X closure_of T \ {x \ X closure_of T. f x \ Y closure_of (f ` T)}" if "T \ topspace X" for T proof (rule closure_of_minimal) show "T \ {x \ X closure_of T. f x \ Y closure_of f ` T}" using closure_of_subset * that by (fastforce simp: in_closure_of) next show "closedin X {x \ X closure_of T. f x \ Y closure_of f ` T}" using assms closedin_continuous_map_preimage_gen by fastforce qed then show ?thesis by (smt (verit, ccfv_threshold) assms continuous_map image_eqI image_subset_iff in_closure_of mem_Collect_eq) qed lemma continuous_map_subset_aux1: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_image_closure_subset by blast lemma continuous_map_subset_aux2: assumes "\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S" shows "continuous_map X Y f" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "f x \ topspace Y" using assms closure_of_subset_topspace by fastforce next fix C assume "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" proof (clarsimp simp flip: closure_of_subset_eq, intro conjI) fix x assume x: "x \ X closure_of {x \ topspace X. f x \ C}" and "C \ topspace Y" and "Y closure_of C \ C" show "x \ topspace X" by (meson x in_closure_of) have "{a \ topspace X. f a \ C} \ topspace X" by simp moreover have "Y closure_of f ` {a \ topspace X. f a \ C} \ C" by (simp add: \closedin Y C\ closure_of_minimal image_subset_iff) ultimately show "f x \ C" using x assms by blast qed qed lemma continuous_map_eq_image_closure_subset: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_alt: "continuous_map X Y f \ (\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_gen: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis lemma continuous_map_closure_preimage_subset: "continuous_map X Y f \ X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" unfolding continuous_map_closedin by (rule closure_of_minimal) (use in_closure_of in \fastforce+\) lemma continuous_map_frontier_frontier_preimage_subset: assumes "continuous_map X Y f" shows "X frontier_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y frontier_of T}" proof - have eq: "topspace X - {x \ topspace X. f x \ T} = {x \ topspace X. f x \ topspace Y - T}" using assms unfolding continuous_map_def by blast have "X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" by (simp add: assms continuous_map_closure_preimage_subset) moreover have "X closure_of (topspace X - {x \ topspace X. f x \ T}) \ {x \ topspace X. f x \ Y closure_of (topspace Y - T)}" using continuous_map_closure_preimage_subset [OF assms] eq by presburger ultimately show ?thesis by (auto simp: frontier_of_closures) qed lemma topology_finer_continuous_id: assumes "topspace X = topspace Y" shows "(\S. openin X S \ openin Y S) \ continuous_map Y X id" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding continuous_map_def using assms openin_subopen openin_subset by fastforce show "?rhs \ ?lhs" unfolding continuous_map_def using assms openin_subopen topspace_def by fastforce qed lemma continuous_map_const [simp]: "continuous_map X Y (\x. C) \ topspace X = {} \ C \ topspace Y" proof (cases "topspace X = {}") case False show ?thesis proof (cases "C \ topspace Y") case True with openin_subopen show ?thesis by (auto simp: continuous_map_def) next case False then show ?thesis unfolding continuous_map_def by fastforce qed qed (auto simp: continuous_map_on_empty) declare continuous_map_const [THEN iffD2, continuous_intros] lemma continuous_map_compose [continuous_intros]: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" shows "continuous_map X X'' (g \ f)" unfolding continuous_map_def proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "(g \ f) x \ topspace X''" using assms unfolding continuous_map_def by force next fix U assume "openin X'' U" have eq: "{x \ topspace X. (g \ f) x \ U} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U}}" by auto (meson f continuous_map_def) show "openin X {x \ topspace X. (g \ f) x \ U}" unfolding eq using assms unfolding continuous_map_def using \openin X'' U\ by blast qed lemma continuous_map_eq: assumes "continuous_map X X' f" and "\x. x \ topspace X \ f x = g x" shows "continuous_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms by (simp add: continuous_map_def eq) qed lemma restrict_continuous_map [simp]: "topspace X \ S \ continuous_map X X' (restrict f S) \ continuous_map X X' f" by (auto simp: elim!: continuous_map_eq) lemma continuous_map_in_subtopology: "continuous_map X (subtopology X' S) f \ continuous_map X X' f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof - have "\A. f ` (X closure_of A) \ subtopology X' S closure_of f ` A" by (meson L continuous_map_image_closure_subset) then show ?thesis by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset order.trans) qed next assume R: ?rhs then have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. f x \ U \ f x \ S}" for U by auto show ?lhs using R unfolding continuous_map by (auto simp: openin_subtopology eq) qed lemma continuous_map_from_subtopology: "continuous_map X X' f \ continuous_map (subtopology X S) X' f" by (auto simp: continuous_map openin_subtopology) lemma continuous_map_into_fulltopology: "continuous_map X (subtopology X' T) f \ continuous_map X X' f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_into_subtopology: "\continuous_map X X' f; f ` topspace X \ T\ \ continuous_map X (subtopology X' T) f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_from_subtopology_mono: "\continuous_map (subtopology X T) X' f; S \ T\ \ continuous_map (subtopology X S) X' f" by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology) lemma continuous_map_from_discrete_topology [simp]: "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g" by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant) lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g" by (metis continuous_map_iff_continuous subtopology_UNIV) lemma continuous_map_openin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X (topspace X \ f -` U))" by (auto simp: continuous_map_def vimage_def Int_def) lemma continuous_map_closedin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. closedin Y U \ closedin X (topspace X \ f -` U))" by (auto simp: continuous_map_closedin vimage_def Int_def) lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt" by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt) lemma continuous_map_sqrt [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. sqrt(f x))" by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply) lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id" unfolding continuous_map_def using openin_subopen topspace_def by fastforce declare continuous_map_id [unfolded id_def, simp, continuous_intros] lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id" by (simp add: continuous_map_from_subtopology) declare continuous_map_id_subt [unfolded id_def, simp] lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ topspace T2)" by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute) lemma continuous_map_open [intro]: "continuous_map T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" unfolding continuous_map_alt by auto lemma continuous_map_preimage_topspace [intro]: assumes "continuous_map T1 T2 f" shows "f-`(topspace T2) \ topspace T1 = topspace T1" using assms unfolding continuous_map_def by auto subsection\Open and closed maps (not a priori assumed continuous)\ definition open_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "open_map X1 X2 f \ \U. openin X1 U \ openin X2 (f ` U)" definition closed_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "closed_map X1 X2 f \ \U. closedin X1 U \ closedin X2 (f ` U)" lemma open_map_imp_subset_topspace: "open_map X1 X2 f \ f ` (topspace X1) \ topspace X2" unfolding open_map_def by (simp add: openin_subset) lemma open_map_on_empty: "topspace X = {} \ open_map X Y f" by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset) lemma closed_map_on_empty: "topspace X = {} \ closed_map X Y f" by (simp add: closed_map_def closedin_topspace_empty) lemma closed_map_const: "closed_map X Y (\x. c) \ topspace X = {} \ closedin Y {c}" by (metis closed_map_def closed_map_on_empty closedin_empty closedin_topspace image_constant_conv) lemma open_map_imp_subset: "\open_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" by (meson order_trans open_map_imp_subset_topspace subset_image_iff) lemma topology_finer_open_id: "(\S. openin X S \ openin X' S) \ open_map X X' id" unfolding open_map_def by auto lemma open_map_id: "open_map X X id" unfolding open_map_def by auto lemma open_map_eq: "\open_map X X' f; \x. x \ topspace X \ f x = g x\ \ open_map X X' g" unfolding open_map_def by (metis image_cong openin_subset subset_iff) lemma open_map_inclusion_eq: "open_map (subtopology X S) X id \ openin X (topspace X \ S)" by (metis openin_topspace openin_trans_full subtopology_restrict topology_finer_open_id topspace_subtopology) lemma open_map_inclusion: "openin X S \ open_map (subtopology X S) X id" by (simp add: open_map_inclusion_eq openin_Int) lemma open_map_compose: "\open_map X X' f; open_map X' X'' g\ \ open_map X X'' (g \ f)" by (metis (no_types, lifting) image_comp open_map_def) lemma closed_map_imp_subset_topspace: "closed_map X1 X2 f \ f ` (topspace X1) \ topspace X2" by (simp add: closed_map_def closedin_subset) lemma closed_map_imp_subset: "\closed_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" using closed_map_imp_subset_topspace by blast lemma topology_finer_closed_id: "(\S. closedin X S \ closedin X' S) \ closed_map X X' id" by (simp add: closed_map_def) lemma closed_map_id: "closed_map X X id" by (simp add: closed_map_def) lemma closed_map_eq: "\closed_map X X' f; \x. x \ topspace X \ f x = g x\ \ closed_map X X' g" unfolding closed_map_def by (metis image_cong closedin_subset subset_iff) lemma closed_map_compose: "\closed_map X X' f; closed_map X' X'' g\ \ closed_map X X'' (g \ f)" by (metis (no_types, lifting) closed_map_def image_comp) lemma closed_map_inclusion_eq: "closed_map (subtopology X S) X id \ closedin X (topspace X \ S)" proof - have *: "closedin X (T \ S)" if "closedin X (S \ topspace X)" "closedin X T" for T by (smt (verit, best) closedin_Int closure_of_subset_eq inf_sup_aci le_iff_inf that) then show ?thesis by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *) qed lemma closed_map_inclusion: "closedin X S \ closed_map (subtopology X S) X id" by (simp add: closed_map_inclusion_eq closedin_Int) lemma open_map_into_subtopology: "\open_map X X' f; f ` topspace X \ S\ \ open_map X (subtopology X' S) f" unfolding open_map_def openin_subtopology using openin_subset by fastforce lemma closed_map_into_subtopology: "\closed_map X X' f; f ` topspace X \ S\ \ closed_map X (subtopology X' S) f" unfolding closed_map_def closedin_subtopology using closedin_subset by fastforce lemma open_map_into_discrete_topology: "open_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding open_map_def openin_discrete_topology using openin_subset by blast lemma closed_map_into_discrete_topology: "closed_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast lemma bijective_open_imp_closed_map: "\open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ closed_map X X' f" unfolding open_map_def closed_map_def closedin_def by auto (metis Diff_subset inj_on_image_set_diff) lemma bijective_closed_imp_open_map: "\closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ open_map X X' f" unfolding closed_map_def open_map_def openin_closedin_eq by auto (metis Diff_subset inj_on_image_set_diff) lemma open_map_from_subtopology: "\open_map X X' f; openin X U\ \ open_map (subtopology X U) X' f" unfolding open_map_def openin_subtopology_alt by blast lemma closed_map_from_subtopology: "\closed_map X X' f; closedin X U\ \ closed_map (subtopology X U) X' f" unfolding closed_map_def closedin_subtopology_alt by blast lemma open_map_restriction: assumes f: "open_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "open_map (subtopology X U) (subtopology X' V) f" unfolding open_map_def proof clarsimp fix W assume "openin (subtopology X U) W" then obtain T where "openin X T" "W = T \ U" by (meson openin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding open_map_def openin_closedin_eq by auto then show "openin (subtopology X' V) (f ` W)" by (metis \openin X T\ f open_map_def openin_subtopology_Int) qed lemma closed_map_restriction: assumes f: "closed_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "closed_map (subtopology X U) (subtopology X' V) f" unfolding closed_map_def proof clarsimp fix W assume "closedin (subtopology X U) W" then obtain T where "closedin X T" "W = T \ U" by (meson closedin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding closed_map_def closedin_def by auto then show "closedin (subtopology X' V) (f ` W)" by (metis \closedin X T\ closed_map_def closedin_subtopology f) qed lemma closed_map_closure_of_image: "closed_map X Y f \ f ` topspace X \ topspace Y \ (\S. S \ topspace X \ Y closure_of (f ` S) \ image f (X closure_of S))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (simp add: closed_map_def closed_map_imp_subset_topspace closure_of_minimal closure_of_subset image_mono) next assume ?rhs then show ?lhs by (metis closed_map_def closed_map_into_discrete_topology closure_of_eq closure_of_subset_eq topspace_discrete_topology) qed lemma open_map_interior_of_image_subset: "open_map X Y f \ (\S. image f (X interior_of S) \ Y interior_of (f ` S))" by (metis image_mono interior_of_eq interior_of_maximal interior_of_subset open_map_def openin_interior_of subset_antisym) lemma open_map_interior_of_image_subset_alt: "open_map X Y f \ (\S\topspace X. f ` (X interior_of S) \ Y interior_of f ` S)" by (metis interior_of_eq open_map_def open_map_interior_of_image_subset openin_subset subset_interior_of_eq) lemma open_map_interior_of_image_subset_gen: "open_map X Y f \ (f ` topspace X \ topspace Y \ (\S. f ` (X interior_of S) \ Y interior_of f ` S))" by (meson open_map_imp_subset_topspace open_map_interior_of_image_subset) lemma open_map_preimage_neighbourhood: "open_map X Y f \ (f ` topspace X \ topspace Y \ (\U T. closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` topspace X \ topspace Y" by (simp add: \open_map X Y f\ open_map_imp_subset_topspace) next fix U T assume UT: "closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" have "closedin Y (topspace Y - f ` (topspace X - U))" by (meson UT L open_map_def openin_closedin_eq openin_diff openin_topspace) with UT show "\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" using image_iff by auto qed next assume R: ?rhs show ?lhs unfolding open_map_def proof (intro strip) fix U assume "openin X U" have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" by blast then obtain V where V: "closedin Y V" and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" using R \openin X U\ by (meson Diff_subset openin_closedin_eq) then have "f ` U \ topspace Y - V" using R \openin X U\ openin_subset by fastforce with sub have "f ` U = topspace Y - V" by auto then show "openin Y (f ` U)" using V(1) by auto qed qed lemma closed_map_preimage_neighbourhood: "closed_map X Y f \ image f (topspace X) \ topspace Y \ (\U T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` topspace X \ topspace Y" by (simp add: L closed_map_imp_subset_topspace) next fix U T assume UT: "openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" then have "openin Y (topspace Y - f ` (topspace X - U))" by (meson L closed_map_def closedin_def closedin_diff closedin_topspace) then show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" using UT image_iff by auto qed next assume R: ?rhs show ?lhs unfolding closed_map_def proof (intro strip) fix U assume "closedin X U" have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" by blast then obtain V where V: "openin Y V" and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" using R Diff_subset \closedin X U\ unfolding closedin_def by (smt (verit, ccfv_threshold) Collect_mem_eq Collect_mono_iff) then have "f ` U \ topspace Y - V" using R \closedin X U\ closedin_subset by fastforce with sub have "f ` U = topspace Y - V" by auto with V show "closedin Y (f ` U)" by auto qed qed lemma closed_map_fibre_neighbourhood: "closed_map X Y f \ f ` (topspace X) \ topspace Y \ (\U y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" unfolding closed_map_preimage_neighbourhood proof (intro conj_cong refl all_cong1) fix U assume "f ` topspace X \ topspace Y" show "(\T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)) = (\y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" (is "(\T. ?P T) \ (\y. ?Q y)") proof assume L [rule_format]: "\T. ?P T" show "\y. ?Q y" proof fix y show "?Q y" using L [of "{y}"] by blast qed next assume R: "\y. ?Q y" show "\T. ?P T" proof (cases "openin X U") case True note [[unify_search_bound=3]] obtain V where V: "\y. \y \ topspace Y; {x \ topspace X. f x = y} \ U\ \ openin Y (V y) \ y \ V y \ {x \ topspace X. f x \ V y} \ U" using R by (simp add: True) meson show ?thesis proof clarify fix T assume "openin X U" and "T \ topspace Y" and "{x \ topspace X. f x \ T} \ U" with V show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" by (rule_tac x="\y\T. V y" in exI) fastforce qed qed auto qed qed lemma open_map_in_subtopology: "openin Y S \ (open_map X (subtopology Y S) f \ open_map X Y f \ f ` (topspace X) \ S)" by (metis le_inf_iff open_map_def open_map_imp_subset_topspace open_map_into_subtopology openin_trans_full topspace_subtopology) lemma open_map_from_open_subtopology: "\openin Y S; open_map X (subtopology Y S) f\ \ open_map X Y f" using open_map_in_subtopology by blast lemma closed_map_in_subtopology: "closedin Y S \ closed_map X (subtopology Y S) f \ (closed_map X Y f \ f ` topspace X \ S)" by (metis closed_map_def closed_map_imp_subset_topspace closed_map_into_subtopology closedin_closed_subtopology closedin_subset topspace_subtopology_subset) lemma closed_map_from_closed_subtopology: "\closedin Y S; closed_map X (subtopology Y S) f\ \ closed_map X Y f" using closed_map_in_subtopology by blast lemma closed_map_from_composition_left: assumes cmf: "closed_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "closed_map Y Z g" unfolding closed_map_def proof (intro strip) fix U assume "closedin Y U" then have "closedin X {x \ topspace X. f x \ U}" using contf closedin_continuous_map_preimage by blast then have "closedin Z ((g \ f) ` {x \ topspace X. f x \ U})" using cmf closed_map_def by blast moreover have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" by (smt (verit, ccfv_SIG) \closedin Y U\ closedin_subset fim image_iff subsetD) then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto ultimately show "closedin Z (g ` U)" by metis qed text \identical proof as the above\ lemma open_map_from_composition_left: assumes cmf: "open_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "open_map Y Z g" unfolding open_map_def proof (intro strip) fix U assume "openin Y U" then have "openin X {x \ topspace X. f x \ U}" using contf openin_continuous_map_preimage by blast then have "openin Z ((g \ f) ` {x \ topspace X. f x \ U})" using cmf open_map_def by blast moreover have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" by (smt (verit, ccfv_SIG) \openin Y U\ openin_subset fim image_iff subsetD) then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto ultimately show "openin Z (g ` U)" by metis qed lemma closed_map_from_composition_right: assumes cmf: "closed_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" shows "closed_map X Y f" unfolding closed_map_def proof (intro strip) fix C assume "closedin X C" have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" using \closedin X C\ assms closedin_subset inj_onD by fastforce then have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" using \closedin X C\ assms(2) closedin_subset by fastforce moreover have "closedin Z ((g \ f) ` C)" using \closedin X C\ cmf closed_map_def by blast ultimately show "closedin Y (f ` C)" using assms(3) closedin_continuous_map_preimage by fastforce qed text \identical proof as the above\ lemma open_map_from_composition_right: assumes "open_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" shows "open_map X Y f" unfolding open_map_def proof (intro strip) fix C assume "openin X C" have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" using \openin X C\ assms openin_subset inj_onD by fastforce then have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" using \openin X C\ assms(2) openin_subset by fastforce moreover have "openin Z ((g \ f) ` C)" using \openin X C\ assms(1) open_map_def by blast ultimately show "openin Y (f ` C)" using assms(3) openin_continuous_map_preimage by fastforce qed subsection\Quotient maps\ definition quotient_map where "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (openin X {x. x \ topspace X \ f x \ U} \ openin X' U))" lemma quotient_map_eq: assumes "quotient_map X X' f" "\x. x \ topspace X \ f x = g x" shows "quotient_map X X' g" by (smt (verit) Collect_cong assms image_cong quotient_map_def) lemma quotient_map_compose: assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g" shows "quotient_map X X'' (g \ f)" unfolding quotient_map_def proof (intro conjI allI impI) show "(g \ f) ` topspace X = topspace X''" using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def) next fix U'' assume U'': "U'' \ topspace X''" define U' where "U' \ {y \ topspace X'. g y \ U''}" have "U' \ topspace X'" by (auto simp add: U'_def) then have U': "openin X {x \ topspace X. f x \ U'} = openin X' U'" using assms unfolding quotient_map_def by simp have "{x \ topspace X. f x \ topspace X' \ g (f x) \ U''} = {x \ topspace X. (g \ f) x \ U''}" using f quotient_map_def by fastforce then show "openin X {x \ topspace X. (g \ f) x \ U''} = openin X'' U''" by (smt (verit, best) Collect_cong U' U'_def U'' g mem_Collect_eq quotient_map_def) qed lemma quotient_map_from_composition: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \ f)" shows "quotient_map X' X'' g" unfolding quotient_map_def proof (intro conjI allI impI) show "g ` topspace X' = topspace X''" using assms unfolding continuous_map_def quotient_map_def by fastforce next fix U'' :: "'c set" assume U'': "U'' \ topspace X''" have eq: "{x \ topspace X. g (f x) \ U''} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U''}}" using continuous_map_def f by fastforce show "openin X' {x \ topspace X'. g x \ U''} = openin X'' U''" using assms unfolding continuous_map_def quotient_map_def by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq) qed lemma quotient_imp_continuous_map: "quotient_map X X' f \ continuous_map X X' f" by (simp add: continuous_map openin_subset quotient_map_def) lemma quotient_imp_surjective_map: "quotient_map X X' f \ f ` (topspace X) = topspace X'" by (simp add: quotient_map_def) lemma quotient_map_closedin: "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (closedin X {x. x \ topspace X \ f x \ U} \ closedin X' U))" proof - have eq: "(topspace X - {x \ topspace X. f x \ U'}) = {x \ topspace X. f x \ topspace X' \ f x \ U'}" if "f ` topspace X = topspace X'" "U' \ topspace X'" for U' using that by auto have "(\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U) = (\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U)" if "f ` topspace X = topspace X'" proof (rule iffI; intro allI impI subsetI) fix U' assume *[rule_format]: "\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U" and U': "U' \ topspace X'" show "closedin X {x \ topspace X. f x \ U'} = closedin X' U'" using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that]) next fix U' :: "'b set" assume *[rule_format]: "\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U" and U': "U' \ topspace X'" show "openin X {x \ topspace X. f x \ U'} = openin X' U'" using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that]) qed then show ?thesis unfolding quotient_map_def by force qed lemma continuous_open_imp_quotient_map: assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - { fix U assume U: "U \ topspace X'" and "openin X {x \ topspace X. f x \ U}" then have ope: "openin X' (f ` {x \ topspace X. f x \ U})" using om unfolding open_map_def by blast then have "openin X' U" using U feq by (subst openin_subopen) force } moreover have "openin X {x \ topspace X. f x \ U}" if "U \ topspace X'" and "openin X' U" for U using that assms unfolding continuous_map_def by blast ultimately show ?thesis unfolding quotient_map_def using assms by blast qed lemma continuous_closed_imp_quotient_map: assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - have "f ` {x \ topspace X. f x \ U} = U" if "U \ topspace X'" for U using that feq by auto with assms show ?thesis unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto qed lemma continuous_open_quotient_map: "\continuous_map X X' f; open_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_open_imp_quotient_map quotient_map_def) lemma continuous_closed_quotient_map: "\continuous_map X X' f; closed_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_closed_imp_quotient_map quotient_map_def) lemma injective_quotient_map: assumes "inj_on f (topspace X)" shows "quotient_map X X' f \ continuous_map X X' f \ open_map X X' f \ closed_map X X' f \ f ` (topspace X) = topspace X'" (is "?lhs = ?rhs") proof assume L: ?lhs have om: "open_map X X' f" proof (clarsimp simp add: open_map_def) fix U assume "openin X U" then have "U \ topspace X" by (simp add: openin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "openin X' (f ` U)" using L unfolding quotient_map_def by (metis (no_types, lifting) Collect_cong \openin X U\ image_mono) qed then have "closed_map X X' f" by (simp add: L assms bijective_open_imp_closed_map quotient_imp_surjective_map) then show ?rhs using L om by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map) next assume ?rhs then show ?lhs by (simp add: continuous_closed_imp_quotient_map) qed lemma continuous_compose_quotient_map: assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \ f)" shows "continuous_map X' X'' g" unfolding quotient_map_def continuous_map_def proof (intro conjI ballI allI impI) show "\x'. x' \ topspace X' \ g x' \ topspace X''" using assms unfolding quotient_map_def by (metis (no_types, opaque_lifting) continuous_map_image_subset_topspace image_comp image_subset_iff) next fix U'' :: "'c set" assume U'': "openin X'' U''" have "f ` topspace X = topspace X'" by (simp add: f quotient_imp_surjective_map) then have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U} = {x \ topspace X. g (f x) \ U}" for U by auto have "openin X {x \ topspace X. f x \ topspace X' \ g (f x) \ U''}" unfolding eq using U'' g openin_continuous_map_preimage by fastforce then have *: "openin X {x \ topspace X. f x \ {x \ topspace X'. g x \ U''}}" by auto show "openin X' {x \ topspace X'. g x \ U''}" using f unfolding quotient_map_def by (metis (no_types) Collect_subset *) qed lemma continuous_compose_quotient_map_eq: "quotient_map X X' f \ continuous_map X X'' (g \ f) \ continuous_map X' X'' g" using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast lemma quotient_map_compose_eq: "quotient_map X X' f \ quotient_map X X'' (g \ f) \ quotient_map X' X'' g" by (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_compose quotient_map_from_composition) lemma quotient_map_restriction: assumes quo: "quotient_map X Y f" and U: "{x \ topspace X. f x \ V} = U" and disj: "openin Y V \ closedin Y V" shows "quotient_map (subtopology X U) (subtopology Y V) f" using disj proof assume V: "openin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: openin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" using quo unfolding quotient_map_def by auto have "openin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_def proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "openin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = openin (subtopology Y V) Y'" using U V Y \openin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq) qed next assume V: "closedin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: closedin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin X {x \ topspace X. f x \ U} = closedin Y U" using quo unfolding quotient_map_closedin by auto have "closedin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_closedin proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "closedin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = closedin (subtopology Y V) Y'" using U V Y \closedin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def) qed qed lemma quotient_map_saturated_open: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. openin X U \ {x \ topspace X. f x \ f ` U} \ U \ openin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin Y U = openin X {x \ topspace X. f x \ U}" unfolding quotient_map_def by auto show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "openin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim openin_subset by fastforce+ show "openin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then have YX: "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \openin X U; {x \ topspace X. f x \ f ` U} \ U\ \ openin Y (f ` U)" by (auto simp: quotient_map_def continuous_map_def) show ?lhs proof (simp add: quotient_map_def fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "openin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "openin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "openin Y U" show "openin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed lemma quotient_map_saturated_closed: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. closedin X U \ {x \ topspace X. f x \ f ` U} \ U \ closedin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin Y U = closedin X {x \ topspace X. f x \ U}" by (simp add: quotient_map_closedin) show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "closedin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim closedin_subset by fastforce+ show "closedin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then obtain YX: "\U. closedin Y U \ closedin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \closedin X U; {x \ topspace X. f x \ f ` U} \ U\ \ closedin Y (f ` U)" by (simp add: continuous_map_closedin) show ?lhs proof (simp add: quotient_map_closedin fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "closedin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "closedin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "closedin Y U" show "closedin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed lemma quotient_map_onto_image: assumes "f ` topspace X \ topspace Y" and U: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" shows "quotient_map X (subtopology Y (f ` topspace X)) f" unfolding quotient_map_def topspace_subtopology proof (intro conjI strip) fix U assume "U \ topspace Y \ f ` topspace X" with U have "openin X {x \ topspace X. f x \ U} \ \x. openin Y x \ U = f ` topspace X \ x" by fastforce moreover have "\x. openin Y x \ U = f ` topspace X \ x \ openin X {x \ topspace X. f x \ U}" by (metis (mono_tags, lifting) Collect_cong IntE IntI U image_eqI openin_subset) ultimately show "openin X {x \ topspace X. f x \ U} = openin (subtopology Y (f ` topspace X)) U" by (force simp: openin_subtopology_alt image_iff) qed (use assms in auto) lemma quotient_map_lift_exists: assumes f: "quotient_map X Y f" and h: "continuous_map X Z h" and "\x y. \x \ topspace X; y \ topspace X; f x = f y\ \ h x = h y" obtains g where "continuous_map Y Z g" "g ` topspace Y = h ` topspace X" "\x. x \ topspace X \ g(f x) = h x" proof - obtain g where g: "\x. x \ topspace X \ h x = g(f x)" using function_factors_left_gen[of "\x. x \ topspace X" f h] assms by blast show ?thesis proof show "g ` topspace Y = h ` topspace X" using f g by (force dest!: quotient_imp_surjective_map) show "continuous_map Y Z g" by (smt (verit) f g h continuous_compose_quotient_map_eq continuous_map_eq o_def) qed (simp add: g) qed subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" where "separatedin X S T \ S \ topspace X \ T \ topspace X \ S \ X closure_of T = {} \ T \ X closure_of S = {}" lemma separatedin_empty [simp]: "separatedin X S {} \ S \ topspace X" "separatedin X {} S \ S \ topspace X" by (simp_all add: separatedin_def) lemma separatedin_refl [simp]: "separatedin X S S \ S = {}" by (metis closure_of_subset empty_subsetI inf.orderE separatedin_def) lemma separatedin_sym: "separatedin X S T \ separatedin X T S" by (auto simp: separatedin_def) lemma separatedin_imp_disjoint: "separatedin X S T \ disjnt S T" by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def) lemma separatedin_mono: "\separatedin X S T; S' \ S; T' \ T\ \ separatedin X S' T'" unfolding separatedin_def using closure_of_mono by blast lemma separatedin_open_sets: "\openin X S; openin X T\ \ separatedin X S T \ disjnt S T" unfolding disjnt_def separatedin_def by (auto simp: openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closed_sets: "\closedin X S; closedin X T\ \ separatedin X S T \ disjnt S T" unfolding closure_of_eq disjnt_def separatedin_def by (metis closedin_def closure_of_eq inf_commute) lemma separatedin_subtopology: "separatedin (subtopology X U) S T \ S \ U \ T \ U \ separatedin X S T" by (auto simp: separatedin_def closure_of_subtopology Int_ac disjoint_iff elim!: inf.orderE) lemma separatedin_discrete_topology: "separatedin (discrete_topology U) S T \ S \ U \ T \ U \ disjnt S T" by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology) lemma separated_eq_distinguishable: "separatedin X {x} {y} \ x \ topspace X \ y \ topspace X \ (\U. openin X U \ x \ U \ (y \ U)) \ (\v. openin X v \ y \ v \ (x \ v))" by (force simp: separatedin_def closure_of_def) lemma separatedin_Un [simp]: "separatedin X S (T \ U) \ separatedin X S T \ separatedin X S U" "separatedin X (S \ T) U \ separatedin X S U \ separatedin X T U" by (auto simp: separatedin_def) lemma separatedin_Union: "finite \ \ separatedin X S (\\) \ S \ topspace X \ (\T \ \. separatedin X S T)" "finite \ \ separatedin X (\\) S \ (\T \ \. separatedin X S T) \ S \ topspace X" by (auto simp: separatedin_def closure_of_Union) lemma separatedin_openin_diff: "\openin X S; openin X T\ \ separatedin X (S - T) (T - S)" unfolding separatedin_def by (metis Diff_Int_distrib2 Diff_disjoint Diff_empty Diff_mono empty_Diff empty_subsetI openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closedin_diff: assumes "closedin X S" "closedin X T" shows "separatedin X (S - T) (T - S)" proof - have "S - T \ topspace X" "T - S \ topspace X" using assms closedin_subset by auto with assms show ?thesis by (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2) qed lemma separation_closedin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ closedin (subtopology X (S \ T)) S \ closedin (subtopology X (S \ T)) T" by (auto simp add: separatedin_def closedin_Int_closure_of disjnt_iff dest: closure_of_subset) lemma separation_openin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ openin (subtopology X (S \ T)) S \ openin (subtopology X (S \ T)) T" unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) lemma separatedin_full: "S \ T = topspace X \ separatedin X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" by (metis separatedin_open_sets separation_closedin_Un_gen separation_openin_Un_gen subtopology_topspace) subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ definition homeomorphic_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "homeomorphic_map X Y f \ quotient_map X Y f \ inj_on f (topspace X)" definition homeomorphic_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "homeomorphic_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" lemma homeomorphic_map_eq: "\homeomorphic_map X Y f; \x. x \ topspace X \ f x = g x\ \ homeomorphic_map X Y g" by (meson homeomorphic_map_def inj_on_cong quotient_map_eq) lemma homeomorphic_maps_eq: "\homeomorphic_maps X Y f g; \x. x \ topspace X \ f x = f' x; \y. y \ topspace Y \ g y = g' y\ \ homeomorphic_maps X Y f' g'" unfolding homeomorphic_maps_def by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff) lemma homeomorphic_maps_sym: "homeomorphic_maps X Y f g \ homeomorphic_maps Y X g f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_maps_id: "homeomorphic_maps X Y id id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have "topspace X = topspace Y" by (auto simp: homeomorphic_maps_def continuous_map_def) with L show ?rhs unfolding homeomorphic_maps_def by (metis topology_finer_continuous_id topology_eq) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by auto qed lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have eq: "topspace X = topspace Y" by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def) then have "\S. openin X S \ openin Y S" by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id) then show ?rhs using L unfolding homeomorphic_map_def by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id) next assume ?rhs then show ?lhs unfolding homeomorphic_map_def by (simp add: closed_map_id continuous_closed_imp_quotient_map) qed lemma homeomorphic_map_compose: assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g" shows "homeomorphic_map X X'' (g \ f)" proof - have "inj_on g (f ` topspace X)" by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map) then show ?thesis using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq) qed lemma homeomorphic_maps_compose: "homeomorphic_maps X Y f h \ homeomorphic_maps Y X'' g k \ homeomorphic_maps X X'' (g \ f) (h \ k)" unfolding homeomorphic_maps_def by (auto simp: continuous_map_compose; simp add: continuous_map_def) lemma homeomorphic_eq_everything_map: "homeomorphic_map X Y f \ continuous_map X Y f \ open_map X Y f \ closed_map X Y f \ f ` (topspace X) = topspace Y \ inj_on f (topspace X)" unfolding homeomorphic_map_def by (force simp: injective_quotient_map intro: injective_quotient_map) lemma homeomorphic_imp_continuous_map: "homeomorphic_map X Y f \ continuous_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_open_map: "homeomorphic_map X Y f \ open_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_closed_map: "homeomorphic_map X Y f \ closed_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_surjective_map: "homeomorphic_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_injective_map: "homeomorphic_map X Y f \ inj_on f (topspace X)" by (simp add: homeomorphic_eq_everything_map) lemma bijective_open_imp_homeomorphic_map: "\continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map) lemma bijective_closed_imp_homeomorphic_map: "\continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_closed_quotient_map homeomorphic_map_def) lemma open_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "open_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y open_map_def continuous_map_def eq) qed lemma closed_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "closed_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "closedin X U" for U using closedin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y closed_map_def continuous_map_closedin eq) qed lemma homeomorphic_maps_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f \ homeomorphic_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "continuous_map Y X g" "\x\topspace X. g (f x) = x" "\x'\topspace Y. f (g x') = x'" by (auto simp: homeomorphic_maps_def) show ?rhs proof (intro conjI bijective_open_imp_homeomorphic_map L) show "open_map X Y f" using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def) show "open_map Y X g" using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def) show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X" using L by (force simp: continuous_map_closedin)+ show "inj_on f (topspace X)" "inj_on g (topspace Y)" using L unfolding inj_on_def by metis+ qed next assume ?rhs then show ?lhs by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map) qed lemma homeomorphic_maps_imp_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f" using homeomorphic_maps_map by blast lemma homeomorphic_map_maps: "homeomorphic_map X Y f \ (\g. homeomorphic_maps X Y f g)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f" "f ` (topspace X) = topspace Y" "inj_on f (topspace X)" by (auto simp: homeomorphic_eq_everything_map) have X: "\x. x \ topspace X \ f x \ topspace Y \ inv_into (topspace X) f (f x) = x" using L by auto have Y: "\y. y \ topspace Y \ inv_into (topspace X) f y \ topspace X \ f (inv_into (topspace X) f y) = y" by (simp add: L f_inv_into_f inv_into_into) have "homeomorphic_maps X Y f (inv_into (topspace X) f)" unfolding homeomorphic_maps_def proof (intro conjI L) show "continuous_map Y X (inv_into (topspace X) f)" by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f]) next show "\x\topspace X. inv_into (topspace X) f (f x) = x" "\y\topspace Y. f (inv_into (topspace X) f y) = y" using X Y by auto qed then show ?rhs by metis next assume ?rhs then show ?lhs using homeomorphic_maps_map by blast qed lemma homeomorphic_maps_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_maps X X f f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_map_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_map X X f" using homeomorphic_maps_involution homeomorphic_maps_map by blast lemma homeomorphic_map_openness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "openin Y (f ` U) \ openin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "openin X U \ openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast show "openin Y (f ` U) = openin X U" proof assume L: "openin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "openin X U" by (metis L homeomorphic_imp_open_map open_map_def g) next assume "openin X U" then show "openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast qed qed lemma homeomorphic_map_closedness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "closedin Y (f ` U) \ closedin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "closedin X U \ closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast show "closedin Y (f ` U) = closedin X U" proof assume L: "closedin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "closedin X U" by (metis L homeomorphic_imp_closed_map closed_map_def g) next assume "closedin X U" then show "closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast qed qed lemma homeomorphic_map_openness_eq: "homeomorphic_map X Y f \ openin X U \ U \ topspace X \ openin Y (f ` U)" by (meson homeomorphic_map_openness openin_closedin_eq) lemma homeomorphic_map_closedness_eq: "homeomorphic_map X Y f \ closedin X U \ U \ topspace X \ closedin Y (f ` U)" by (meson closedin_subset homeomorphic_map_closedness) lemma all_openin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. openin Y V \ P V) \ (\U. openin X U \ P(f ` U))" by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff) lemma all_closedin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. closedin Y V \ P V) \ (\U. closedin X U \ P(f ` U))" (is "?lhs = ?rhs") by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff) lemma homeomorphic_map_derived_set_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)" proof - have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)" using hom by (auto simp: homeomorphic_eq_everything_map) have iff: "(\T. x \ T \ openin X T \ (\y. y \ x \ y \ S \ y \ T)) = (\T. T \ topspace Y \ f x \ T \ openin Y T \ (\y. y \ f x \ y \ f ` S \ y \ T))" if "x \ topspace X" for x proof - have \
: "(x \ T \ openin X T) = (T \ topspace X \ f x \ f ` T \ openin Y (f ` T))" for T by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that) moreover have "(\y. y \ x \ y \ S \ y \ T) = (\y. y \ f x \ y \ f ` S \ y \ f ` T)" (is "?lhs = ?rhs") if "T \ topspace X \ f x \ f ` T \ openin Y (f ` T)" for T by (smt (verit, del_insts) S \x \ topspace X\ image_iff inj inj_on_def subsetD that) ultimately show ?thesis by (auto simp flip: fim simp: all_subset_image) qed have *: "\T = f ` S; \x. x \ S \ P x \ Q(f x)\ \ {y. y \ T \ Q y} = f ` {x \ S. P x}" for T S P Q by auto show ?thesis unfolding derived_set_of_def by (rule *) (use fim iff openin_subset in force)+ qed lemma homeomorphic_map_closure_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y closure_of (f ` S) = f ` (X closure_of S)" unfolding closure_of using homeomorphic_imp_surjective_map [OF hom] S by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms]) lemma homeomorphic_map_interior_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y interior_of (f ` S) = f ` (X interior_of S)" proof - { fix y assume "y \ topspace Y" and "y \ Y closure_of (topspace Y - f ` S)" then have "y \ f ` (topspace X - X closure_of (topspace X - S))" using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom] by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) } moreover { fix x assume "x \ topspace X" then have "f x \ topspace Y" using hom homeomorphic_imp_surjective_map by blast } moreover { fix x assume "x \ topspace X" and "x \ X closure_of (topspace X - S)" and "f x \ Y closure_of (topspace Y - f ` S)" then have "False" using homeomorphic_map_closure_of [OF hom] hom unfolding homeomorphic_eq_everything_map by (metis Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff inj_on_image_set_diff) } ultimately show ?thesis by (auto simp: interior_of_closure_of) qed lemma homeomorphic_map_frontier_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y frontier_of (f ` S) = f ` (X frontier_of S)" unfolding frontier_of_def proof (intro equalityI subsetI DiffI) fix y assume "y \ Y closure_of f ` S - Y interior_of f ` S" then show "y \ f ` (X closure_of S - X interior_of S)" using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce next fix y assume "y \ f ` (X closure_of S - X interior_of S)" then show "y \ Y closure_of f ` S" using S hom homeomorphic_map_closure_of by fastforce next fix x assume "x \ f ` (X closure_of S - X interior_of S)" then obtain y where y: "x = f y" "y \ X closure_of S" "y \ X interior_of S" by blast then show "x \ Y interior_of f ` S" using S hom homeomorphic_map_interior_of y(1) unfolding homeomorphic_map_def by (smt (verit, ccfv_SIG) in_closure_of inj_on_image_mem_iff interior_of_subset_topspace) qed lemma homeomorphic_maps_subtopologies: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_maps_subtopologies_alt: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) \ T; g ` (topspace Y \ T) \ S\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_map_subtopologies: "\homeomorphic_map X Y f; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies) lemma homeomorphic_map_subtopologies_alt: assumes hom: "homeomorphic_map X Y f" and S: "\x. \x \ topspace X; f x \ topspace Y\ \ f x \ T \ x \ S" shows "homeomorphic_map (subtopology X S) (subtopology Y T) f" proof - have "homeomorphic_maps (subtopology X S) (subtopology Y T) f g" if "homeomorphic_maps X Y f g" for g proof (rule homeomorphic_maps_subtopologies [OF that]) have "f ` (topspace X \ S) \ topspace Y \ T" using S hom homeomorphic_imp_surjective_map by fastforce then show "f ` (topspace X \ S) = topspace Y \ T" using that unfolding homeomorphic_maps_def continuous_map_def by (smt (verit, del_insts) Int_iff S image_iff subsetI subset_antisym) qed then show ?thesis using hom by (meson homeomorphic_map_maps) qed subsection\Relation of homeomorphism between topological spaces\ definition homeomorphic_space (infixr "homeomorphic'_space" 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl: "X homeomorphic_space X" by (meson homeomorphic_maps_id homeomorphic_space_def) lemma homeomorphic_space_sym: "X homeomorphic_space Y \ Y homeomorphic_space X" unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym) lemma homeomorphic_space_trans [trans]: "\X1 homeomorphic_space X2; X2 homeomorphic_space X3\ \ X1 homeomorphic_space X3" unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose) lemma homeomorphic_space: "X homeomorphic_space Y \ (\f. homeomorphic_map X Y f)" by (simp add: homeomorphic_map_maps homeomorphic_space_def) lemma homeomorphic_maps_imp_homeomorphic_space: "homeomorphic_maps X Y f g \ X homeomorphic_space Y" unfolding homeomorphic_space_def by metis lemma homeomorphic_map_imp_homeomorphic_space: "homeomorphic_map X Y f \ X homeomorphic_space Y" unfolding homeomorphic_map_maps using homeomorphic_space_def by blast lemma homeomorphic_empty_space: "X homeomorphic_space Y \ topspace X = {} \ topspace Y = {}" by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty) lemma homeomorphic_empty_space_eq: assumes "topspace X = {}" shows "X homeomorphic_space Y \ topspace Y = {}" unfolding homeomorphic_maps_def homeomorphic_space_def by (metis assms continuous_map_on_empty continuous_map_closedin ex_in_conv) subsection\Connected topological spaces\ definition connected_space :: "'a topology \ bool" where "connected_space X \ \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where "connectedin X S \ S \ topspace X \ connected_space (subtopology X S)" lemma connected_spaceD: "\connected_space X; openin X U; openin X V; topspace X \ U \ V; U \ V = {}; U \ {}; V \ {}\ \ False" by (auto simp: connected_space_def) lemma connectedin_subset_topspace: "connectedin X S \ S \ topspace X" by (simp add: connectedin_def) lemma connectedin_topspace: "connectedin X (topspace X) \ connected_space X" by (simp add: connectedin_def) lemma connected_space_subtopology: "connectedin X S \ connected_space (subtopology X S)" by (simp add: connectedin_def) lemma connectedin_subtopology: "connectedin (subtopology X S) T \ connectedin X T \ T \ S" by (force simp: connectedin_def subtopology_subtopology inf_absorb2) lemma connected_space_eq: "connected_space X \ (\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_space_def by (metis openin_Un openin_subset subset_antisym) lemma connected_space_closedin: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then have "\E1 E2. \openin X E1; E1 \ E2 = {}; topspace X \ E1 \ E2; openin X E2\ \ E1 = {} \ E2 = {}" by (simp add: connected_space_def) then show ?rhs unfolding connected_space_def by (metis disjnt_def separatedin_closed_sets separation_openin_Un_gen subtopology_superset) next assume R: ?rhs then show ?lhs unfolding connected_space_def by (metis Diff_triv Int_commute separatedin_openin_diff separation_closedin_Un_gen subtopology_superset) qed lemma connected_space_closedin_eq: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" by (metis closedin_Un closedin_def connected_space_closedin subset_antisym) lemma connected_space_clopen_in: "connected_space X \ (\T. openin X T \ closedin X T \ T = {} \ T = topspace X)" proof - have eq: "openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ P \ E2 = topspace X - E1 \ openin X E1 \ openin X E2 \ P" for E1 E2 P using openin_subset by blast show ?thesis unfolding connected_space_eq eq closedin_def by (auto simp: openin_closedin_eq) qed lemma connectedin: "connectedin X S \ S \ topspace X \ (\E1 E2. openin X E1 \ openin X E2 \ S \ E1 \ E2 \ E1 \ E2 \ S = {} \ E1 \ S \ {} \ E2 \ S \ {})" (is "?lhs = ?rhs") proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedinD: "\connectedin X S; openin X E1; openin X E2; S \ E1 \ E2; E1 \ E2 \ S = {}; E1 \ S \ {}; E2 \ S \ {}\ \ False" by (meson connectedin) lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedin_empty [simp]: "connectedin X {}" by (simp add: connectedin) lemma connected_space_topspace_empty: "topspace X = {} \ connected_space X" using connectedin_topspace by fastforce lemma connectedin_sing [simp]: "connectedin X {a} \ a \ topspace X" by (simp add: connectedin) lemma connectedin_absolute [simp]: "connectedin (subtopology X S) S \ connectedin X S" by (simp add: connectedin_subtopology) lemma connectedin_Union: assumes \: "\S. S \ \ \ connectedin X S" and ne: "\\ \ {}" shows "connectedin X (\\)" proof - have "\\ \ topspace X" using \ by (simp add: Union_least connectedin_def) moreover have False if "openin X E1" "openin X E2" and cover: "\\ \ E1 \ E2" and disj: "E1 \ E2 \ \\ = {}" and overlap1: "E1 \ \\ \ {}" and overlap2: "E2 \ \\ \ {}" for E1 E2 proof - have disjS: "E1 \ E2 \ S = {}" if "S \ \" for S using Diff_triv that disj by auto have coverS: "S \ E1 \ E2" if "S \ \" for S using that cover by blast have "\ \ {}" using overlap1 by blast obtain a where a: "\U. U \ \ \ a \ U" using ne by force with \\ \ {}\ have "a \ \\" by blast then consider "a \ E1" | "a \ E2" using \\\ \ E1 \ E2\ by auto then show False proof cases case 1 then obtain b S where "b \ E2" "b \ S" "S \ \" using overlap2 by blast then show ?thesis using "1" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) next case 2 then obtain b S where "b \ E1" "b \ S" "S \ \" using overlap1 by blast then show ?thesis using "2" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) qed qed ultimately show ?thesis unfolding connectedin by blast qed lemma connectedin_Un: "\connectedin X S; connectedin X T; S \ T \ {}\ \ connectedin X (S \ T)" using connectedin_Union [of "{S,T}"] by auto lemma connected_space_subconnected: "connected_space X \ (\x \ topspace X. \y \ topspace X. \S. connectedin X S \ x \ S \ y \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connectedin_topspace by blast next assume R [rule_format]: ?rhs have False if "openin X U" "openin X V" and disj: "U \ V = {}" and cover: "topspace X \ U \ V" and "U \ {}" "V \ {}" for U V proof - obtain u v where "u \ U" "v \ V" using \U \ {}\ \V \ {}\ by auto then obtain T where "u \ T" "v \ T" and T: "connectedin X T" using R [of u v] that by (meson \openin X U\ \openin X V\ subsetD openin_subset) then show False using that unfolding connectedin by (metis IntI \u \ U\ \v \ V\ empty_iff inf_bot_left subset_trans) qed then show ?lhs by (auto simp: connected_space_def) qed lemma connectedin_intermediate_closure_of: assumes "connectedin X S" "S \ T" "T \ X closure_of S" shows "connectedin X T" proof - have S: "S \ topspace X" and T: "T \ topspace X" using assms by (meson closure_of_subset_topspace dual_order.trans)+ have \
: "\E1 E2. \openin X E1; openin X E2; E1 \ S = {} \ E2 \ S = {}\ \ E1 \ T = {} \ E2 \ T = {}" using assms unfolding disjoint_iff by (meson in_closure_of subsetD) then show ?thesis using assms unfolding connectedin closure_of_subset_topspace S T by (metis Int_empty_right T dual_order.trans inf.orderE inf_left_commute) qed lemma connectedin_closure_of: "connectedin X S \ connectedin X (X closure_of S)" by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl) lemma connectedin_separation: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ C1 \ X closure_of C2 = {} \ C2 \ X closure_of C1 = {})" unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology apply (intro conj_cong refl arg_cong [where f=Not]) apply (intro ex_cong1 iffI, blast) using closure_of_subset_Int by force lemma connectedin_eq_not_separated: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" unfolding separatedin_def by (metis connectedin_separation sup.boundedE) lemma connectedin_eq_not_separated_subset: "connectedin X S \ S \ topspace X \ (\C1 C2. S \ C1 \ C2 \ S \ C1 \ {} \ S \ C2 \ {} \ separatedin X C1 C2)" proof - have "\C1 C2. S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" if "\C1 C2. C1 \ C2 = S \ C1 = {} \ C2 = {} \ \ separatedin X C1 C2" proof (intro allI) fix C1 C2 show "S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" using that [of "S \ C1" "S \ C2"] by (auto simp: separatedin_mono) qed then show ?thesis by (metis Un_Int_eq(1) Un_Int_eq(2) connectedin_eq_not_separated order_refl) qed lemma connected_space_eq_not_separated: "connected_space X \ (\C1 C2. C1 \ C2 = topspace X \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (simp add: connectedin_eq_not_separated flip: connectedin_topspace) lemma connected_space_eq_not_separated_subset: "connected_space X \ (\C1 C2. topspace X \ C1 \ C2 \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (metis connected_space_eq_not_separated le_sup_iff separatedin_def subset_antisym) lemma connectedin_subset_separated_union: "\connectedin X C; separatedin X S T; C \ S \ T\ \ C \ S \ C \ T" unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: assumes "connectedin X S" "connectedin X T" "\separatedin X S T" shows "connectedin X (S \ T)" proof - have "\C1 C2. \T \ C1 \ C2; S \ C1 \ C2\ \ S \ C1 = {} \ T \ C1 = {} \ S \ C2 = {} \ T \ C2 = {} \ \ separatedin X C1 C2" using assms unfolding connectedin_eq_not_separated_subset by (metis (no_types, lifting) assms connectedin_subset_separated_union inf.orderE separatedin_empty(1) separatedin_mono separatedin_sym) then show ?thesis unfolding connectedin_eq_not_separated_subset by (simp add: assms connectedin_subset_topspace Int_Un_distrib2) qed lemma connected_space_closures: "connected_space X \ (\e1 e2. e1 \ e2 = topspace X \ X closure_of e1 \ X closure_of e2 = {} \ e1 \ {} \ e2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding connected_space_closedin_eq by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace) next assume ?rhs then show ?lhs unfolding connected_space_closedin_eq by (metis closure_of_eq) qed lemma connectedin_Int_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) moreover have "S - (topspace X \ T) \ {}" using assms(3) by blast moreover have "S \ topspace X \ T \ {}" using assms connectedin by fastforce moreover have False if "S \ T \ {}" "S - T \ {}" "T \ topspace X" "S \ X frontier_of T = {}" for T proof - have null: "S \ (X closure_of T - X interior_of T) = {}" using that unfolding frontier_of_def by blast have "X interior_of T \ (topspace X - X closure_of T) \ S = {}" by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty) moreover have "S \ X interior_of T \ (topspace X - X closure_of T)" using that \S \ topspace X\ null by auto moreover have "S \ X interior_of T \ {}" using closure_of_subset that(1) that(3) null by fastforce ultimately have "S \ X interior_of (topspace X - T) = {}" by (metis "*" inf_commute interior_of_complement openin_interior_of) then have "topspace (subtopology X S) \ X interior_of T = S" using \S \ topspace X\ interior_of_complement null by fastforce then show ?thesis using that by (metis Diff_eq_empty_iff inf_le2 interior_of_subset subset_trans) qed ultimately show ?thesis by (metis Int_lower1 frontier_of_restrict inf_assoc) qed lemma connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and "connectedin X S" shows "connectedin Y (f ` S)" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) show ?thesis unfolding connectedin connected_space_def proof (intro conjI notI; clarify) show "f x \ topspace Y" if "x \ S" for x using \S \ topspace X\ continuous_map_image_subset_topspace f that by blast next fix U V let ?U = "{x \ topspace X. f x \ U}" let ?V = "{x \ topspace X. f x \ V}" assume UV: "openin Y U" "openin Y V" "f ` S \ U \ V" "U \ V \ f ` S = {}" "U \ f ` S \ {}" "V \ f ` S \ {}" then have 1: "?U \ ?V \ S = {}" by auto have 2: "openin X ?U" "openin X ?V" using \openin Y U\ \openin Y V\ continuous_map f by fastforce+ show "False" using * [of ?U ?V] UV \S \ topspace X\ by (auto simp: 1 2) qed qed lemma connected_space_quotient_map_image: "\quotient_map X X' q; connected_space X\ \ connected_space X'" by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) lemma homeomorphic_connected_space: "X homeomorphic_space Y \ connected_space X \ connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis connected_space_subconnected connectedin_continuous_map_image connectedin_topspace continuous_map_image_subset_topspace image_eqI image_subset_iff) lemma homeomorphic_map_connectedness: assumes f: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "connectedin Y (f ` U) \ connectedin X U" proof - have 1: "f ` U \ topspace Y \ U \ topspace X" using U f homeomorphic_imp_surjective_map by blast moreover have "connected_space (subtopology Y (f ` U)) \ connected_space (subtopology X U)" proof (rule homeomorphic_connected_space) have "f ` U \ topspace Y" by (simp add: U 1) then have "topspace Y \ f ` U = f ` U" by (simp add: subset_antisym) then show "subtopology Y (f ` U) homeomorphic_space subtopology X U" by (metis U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym inf.absorb_iff2) qed ultimately show ?thesis by (auto simp: connectedin_def) qed lemma homeomorphic_map_connectedness_eq: "homeomorphic_map X Y f \ connectedin X U \ U \ topspace X \ connectedin Y (f ` U)" using homeomorphic_map_connectedness connectedin_subset_topspace by metis lemma connectedin_discrete_topology: "connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" proof (cases "S \ U") case True show ?thesis proof (cases "S = {}") case False moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" proof show "connectedin (discrete_topology U) S \ \a. S = {a}" using False connectedin_Int_frontier_of insert_Diff by fastforce qed (use True in auto) ultimately show ?thesis by auto qed simp next case False then show ?thesis by (simp add: connectedin_def) qed lemma connected_space_discrete_topology: "connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology) subsection\Compact sets\ definition compactin where "compactin X S \ S \ topspace X \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\))" definition compact_space where "compact_space X \ compactin X (topspace X)" lemma compact_space_alt: "compact_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ topspace X \ \\))" by (simp add: compact_space_def compactin_def) lemma compact_space: "compact_space X \ (\\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. finite \ \ \ \ \ \ \\ = topspace X))" unfolding compact_space_alt using openin_subset by fastforce lemma compactinD: "\compactin X S; \U. U \ \ \ openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" by (auto simp: compactin_def) lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: "compactin (subtopology X S) S \ compactin X S" proof - have eq: "(\U \ \. \Y. openin X Y \ U = Y \ S) \ \ \ (\Y. Y \ S) ` {y. openin X y}" for \ by auto show ?thesis by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) qed lemma compactin_subspace: "compactin X S \ S \ topspace X \ compact_space (subtopology X S)" unfolding compact_space_def topspace_subtopology by (metis compactin_absolute compactin_def inf.absorb2) lemma compact_space_subtopology: "compactin X S \ compact_space (subtopology X S)" by (simp add: compactin_subspace) lemma compactin_subtopology: "compactin (subtopology X S) T \ compactin X T \ T \ S" by (metis compactin_subspace inf.absorb_iff2 le_inf_iff subtopology_subtopology topspace_subtopology) lemma compactin_subset_topspace: "compactin X S \ S \ topspace X" by (simp add: compactin_subspace) lemma compactin_contractive: "\compactin X' S; topspace X' = topspace X; \U. openin X U \ openin X' U\ \ compactin X S" by (simp add: compactin_def) lemma finite_imp_compactin: "\S \ topspace X; finite S\ \ compactin X S" by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology) lemma compactin_empty [iff]: "compactin X {}" by (simp add: finite_imp_compactin) lemma compact_space_topspace_empty: "topspace X = {} \ compact_space X" by (simp add: compact_space_def) lemma finite_imp_compactin_eq: "finite S \ (compactin X S \ S \ topspace X)" using compactin_subset_topspace finite_imp_compactin by blast lemma compactin_sing [simp]: "compactin X {a} \ a \ topspace X" by (simp add: finite_imp_compactin_eq) lemma closed_compactin: assumes XK: "compactin X K" and "C \ K" and XC: "closedin X C" shows "compactin X C" unfolding compactin_def proof (intro conjI allI impI) show "C \ topspace X" by (simp add: XC closedin_subset) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ C \ \\" have "(\U\insert (topspace X - C) \. openin X U)" using XC \ by blast moreover have "K \ \(insert (topspace X - C) \)" using \ XK compactin_subset_topspace by fastforce ultimately obtain \ where "finite \" "\ \ insert (topspace X - C) \" "K \ \\" using assms unfolding compactin_def by metis moreover have "openin X (topspace X - C)" using XC by auto ultimately show "\\. finite \ \ \ \ \ \ C \ \\" using \C \ K\ by (rule_tac x="\ - {topspace X - C}" in exI) auto qed +lemma closed_compactin_Inter: + "\compactin X K; K \ \; \K. K \ \ \ closedin X K\ \ compactin X (\\)" + by (metis Inf_lower closed_compactin closedin_Inter empty_iff) + +lemma closedin_subtopology_Int_subset: + "\closedin (subtopology X U) (U \ S); V \ U\ \ closedin (subtopology X V) (V \ S)" + by (smt (verit, ccfv_SIG) closedin_subtopology inf.left_commute inf.orderE inf_commute) + +lemma closedin_subtopology_Int_closedin: + "\closedin (subtopology X U) S; closedin X T\ \ closedin (subtopology X U) (S \ T)" + by (smt (verit, best) closedin_Int closedin_subtopology inf_assoc inf_commute) + lemma closedin_compact_space: "\compact_space X; closedin X S\ \ compactin X S" by (simp add: closed_compactin closedin_subset compact_space_def) lemma compact_Int_closedin: assumes "compactin X S" "closedin X T" shows "compactin X (S \ T)" proof - have "compactin (subtopology X S) (S \ T)" by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute) then show ?thesis by (simp add: compactin_subtopology) qed lemma closed_Int_compactin: "\closedin X S; compactin X T\ \ compactin X (S \ T)" by (metis compact_Int_closedin inf_commute) lemma compactin_Un: assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \ T)" unfolding compactin_def proof (intro conjI allI impI) show "S \ T \ topspace X" using assms by (auto simp: compactin_def) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ S \ T \ \\" with S obtain \ where \: "finite \" "\ \ \" "S \ \\" unfolding compactin_def by (meson sup.bounded_iff) obtain \ where "finite \" "\ \ \" "T \ \\" using \ T unfolding compactin_def by (meson sup.bounded_iff) with \ show "\\. finite \ \ \ \ \ \ S \ T \ \\" by (rule_tac x="\ \ \" in exI) auto qed lemma compactin_Union: "\finite \; \S. S \ \ \ compactin X S\ \ compactin X (\\)" by (induction rule: finite_induct) (simp_all add: compactin_Un) lemma compactin_subtopology_imp_compact: assumes "compactin (subtopology X S) K" shows "compactin X K" using assms proof (clarsimp simp add: compactin_def) fix \ define \ where "\ \ (\U. U \ S) ` \" assume "K \ topspace X" and "K \ S" and "\x\\. openin X x" and "K \ \\" then have "\V \ \. openin (subtopology X S) V" "K \ \\" unfolding \_def by (auto simp: openin_subtopology) moreover assume "\\. (\x\\. openin (subtopology X S) x) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson then have \: "\U. U \ \ \ V = U \ S" if "V \ \" for V unfolding \_def using that by blast let ?\ = "(\F. @U. U \ \ \ F = U \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using someI_ex [OF \] by blast show "K \ \?\" proof clarsimp fix x assume "x \ K" then show "\V \ \. x \ (SOME U. U \ \ \ V = U \ S)" using \K \ \\\ someI_ex [OF \] by (metis (no_types, lifting) IntD1 Union_iff subsetCE) qed qed qed lemma compact_imp_compactin_subtopology: assumes "compactin X K" "K \ S" shows "compactin (subtopology X S) K" using assms proof (clarsimp simp add: compactin_def) fix \ :: "'a set set" define \ where "\ \ {V. openin X V \ (\U \ \. U = V \ S)}" assume "K \ S" and "K \ topspace X" and "\U\\. openin (subtopology X S) U" and "K \ \\" then have "\V \ \. openin X V" "K \ \\" unfolding \_def by (fastforce simp: subset_eq openin_subtopology)+ moreover assume "\\. (\U\\. openin X U) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson let ?\ = "(\F. F \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using \_def \\ \ \\ by blast show "K \ \?\" using \K \ \\\ assms(2) by auto qed qed proposition compact_space_fip: "compact_space X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" (is "_ = ?rhs") proof (cases "topspace X = {}") case True then show ?thesis unfolding compact_space_def by (metis Sup_bot_conv(1) closedin_topspace_empty compactin_empty finite.emptyI finite_UnionD order_refl) next case False show ?thesis proof safe fix \ :: "'a set set" assume * [rule_format]: "\\. finite \ \ \ \ \ \ \\ \ {}" define \ where "\ \ (\S. topspace X - S) ` \" assume clo: "\C\\. closedin X C" and [simp]: "\\ = {}" then have "\V \ \. openin X V" "topspace X \ \\" by (auto simp: \_def) moreover assume [unfolded compact_space_alt, rule_format, of \]: "compact_space X" ultimately obtain \ where \: "finite \" "\ \ \" "topspace X \ topspace X - \\" by (auto simp: ex_finite_subset_image \_def) moreover have "\ \ {}" using \ \topspace X \ {}\ by blast ultimately show "False" using * [of \] by auto (metis Diff_iff Inter_iff clo closedin_def subsetD) next assume R [rule_format]: ?rhs show "compact_space X" unfolding compact_space_alt proof clarify fix \ :: "'a set set" define \ where "\ \ (\S. topspace X - S) ` \" assume "\C\\. openin X C" and "topspace X \ \\" with \topspace X \ {}\ have *: "\V \ \. closedin X V" "\ \ {}" by (auto simp: \_def) show "\\. finite \ \ \ \ \ \ topspace X \ \\" proof (rule ccontr; simp) assume "\\\\. finite \ \ \ topspace X \ \\" then have "\\. finite \ \ \ \ \ \ \\ \ {}" by (simp add: \_def all_finite_subset_image) with \topspace X \ \\\ show False using R [of \] * by (simp add: \_def) qed qed qed qed corollary compactin_fip: "compactin X S \ S \ topspace X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof (cases "S = {}") case False show ?thesis proof (cases "S \ topspace X") case True then have "compactin X S \ (\\. \ \ (\T. S \ T) ` {T. closedin X T} \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL) also have "\ = (\\\Collect (closedin X). (\\. finite \ \ \ \ (\) S ` \ \ \\ \ {}) \ \ ((\) S ` \) \ {})" by (simp add: all_subset_image) also have "\ = (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof - have eq: "((\\. finite \ \ \ \ \ \ \ ((\) S ` \) \ {}) \ \ ((\) S ` \) \ {}) \ ((\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" for \ by simp (use \S \ {}\ in blast) show ?thesis unfolding imp_conjL [symmetric] all_finite_subset_image eq by blast qed finally show ?thesis using True by simp qed (simp add: compactin_subspace) qed force corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" and ne: "\n. C n \ {}" and dec: "decseq C" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" have "closedin X A" if "A \ ?\" for A using that clo by auto moreover have "(\n\K. \m \ n. C m) \ {}" if "finite K" for K proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast with dec have "C n \ (\n\K. \m \ n. C m)" unfolding decseq_def by blast with ne [of n] show ?thesis by blast qed ultimately show ?thesis using \compact_space X\ [unfolded compact_space_fip, rule_format, of ?\] by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps) qed lemma compactin_discrete_topology: "compactin (discrete_topology X) S \ S \ X \ finite S" (is "?lhs = ?rhs") proof (intro iffI conjI) assume L: ?lhs then show "S \ X" by (auto simp: compactin_def) have *: "\\. Ball \ (openin (discrete_topology X)) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\)" using L by (auto simp: compactin_def) show "finite S" using * [of "(\x. {x}) ` X"] \S \ X\ by clarsimp (metis UN_singleton finite_subset_image infinite_super) next assume ?rhs then show ?lhs by (simp add: finite_imp_compactin) qed lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \ finite X" by (simp add: compactin_discrete_topology compact_space_def) lemma compact_space_imp_Bolzano_Weierstrass: assumes "compact_space X" "infinite S" "S \ topspace X" shows "X derived_set_of S \ {}" proof assume X: "X derived_set_of S = {}" then have "closedin X S" by (simp add: closedin_contains_derived_set assms) then have "compactin X S" by (rule closedin_compact_space [OF \compact_space X\]) with X show False by (metis \infinite S\ compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq) qed lemma compactin_imp_Bolzano_Weierstrass: "\compactin X S; infinite T \ T \ S\ \ S \ X derived_set_of T \ {}" using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"] by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2) lemma compact_closure_of_imp_Bolzano_Weierstrass: "\compactin X (X closure_of S); infinite T; T \ S; T \ topspace X\ \ X derived_set_of T \ {}" using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce lemma discrete_compactin_eq_finite: "S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" by (meson compactin_imp_Bolzano_Weierstrass finite_imp_compactin_eq order_refl) lemma discrete_compact_space_eq_finite: "X derived_set_of (topspace X) = {} \ (compact_space X \ finite(topspace X))" by (metis compact_space_discrete_topology discrete_topology_unique_derived_set) lemma image_compactin: assumes cpt: "compactin X S" and cont: "continuous_map X Y f" shows "compactin Y (f ` S)" unfolding compactin_def proof (intro conjI allI impI) show "f ` S \ topspace Y" using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast next fix \ :: "'b set set" assume \: "Ball \ (openin Y) \ f ` S \ \\" define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "S \ topspace X" and *: "\\. \\U\\. openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using cpt by (auto simp: compactin_def) obtain \ where \: "finite \" "\ \ \" "S \ \\" proof - have 1: "\U\\. openin X U" unfolding \_def using \ cont[unfolded continuous_map] by blast have 2: "S \ \\" unfolding \_def using compactin_subset_topspace cpt \ by fastforce show thesis using * [OF 1 2] that by metis qed have "\v \ \. \U. U \ \ \ v = {x \ topspace X. f x \ U}" using \_def by blast then obtain U where U: "\v \ \. U v \ \ \ v = {x \ topspace X. f x \ U v}" by metis show "\\. finite \ \ \ \ \ \ f ` S \ \\" proof (intro conjI exI) show "finite (U ` \)" by (simp add: \finite \\) next show "U ` \ \ \" using \\ \ \\ U by auto next show "f ` S \ \ (U ` \)" using \(2-3) U UnionE subset_eq U by fastforce qed qed lemma homeomorphic_compact_space: assumes "X homeomorphic_space Y" shows "compact_space X \ compact_space Y" using homeomorphic_space_sym by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin) lemma homeomorphic_map_compactness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "compactin Y (f ` U) \ compactin X U" proof - have "f ` U \ topspace Y" using hom U homeomorphic_imp_surjective_map by blast moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f" using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies) then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)" using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast ultimately show ?thesis by (simp add: compactin_subspace U) qed lemma homeomorphic_map_compactness_eq: "homeomorphic_map X Y f \ compactin X U \ U \ topspace X \ compactin Y (f ` U)" by (meson compactin_subset_topspace homeomorphic_map_compactness) subsection\Embedding maps\ definition embedding_map where "embedding_map X Y f \ homeomorphic_map X (subtopology Y (f ` (topspace X))) f" lemma embedding_map_eq: "\embedding_map X Y f; \x. x \ topspace X \ f x = g x\ \ embedding_map X Y g" unfolding embedding_map_def by (metis homeomorphic_map_eq image_cong) lemma embedding_map_compose: assumes "embedding_map X X' f" "embedding_map X' X'' g" shows "embedding_map X X'' (g \ f)" proof - have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g" using assms by (auto simp: embedding_map_def) then obtain C where "g ` topspace X' \ C = (g \ f) ` topspace X" by (metis homeomorphic_imp_surjective_map image_comp image_mono inf.absorb_iff2 topspace_subtopology) then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \ f) ` topspace X)) g" by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology) then show ?thesis unfolding embedding_map_def using hm(1) homeomorphic_map_compose by blast qed lemma surjective_embedding_map: "embedding_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (force simp: embedding_map_def homeomorphic_eq_everything_map) lemma embedding_map_in_subtopology: "embedding_map X (subtopology Y S) f \ embedding_map X Y f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding embedding_map_def by (metis continuous_map_in_subtopology homeomorphic_imp_continuous_map inf_absorb2 subtopology_subtopology) qed (simp add: embedding_map_def inf.absorb_iff2 subtopology_subtopology) lemma injective_open_imp_embedding_map: "\continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: continuous_map_in_subtopology continuous_open_quotient_map eq_iff homeomorphic_map_def open_map_imp_subset open_map_into_subtopology) lemma injective_closed_imp_embedding_map: "\continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: closed_map_imp_subset closed_map_into_subtopology continuous_closed_quotient_map continuous_map_in_subtopology dual_order.eq_iff homeomorphic_map_def) lemma embedding_map_imp_homeomorphic_space: "embedding_map X Y f \ X homeomorphic_space (subtopology Y (f ` (topspace X)))" unfolding embedding_map_def using homeomorphic_space by blast lemma embedding_imp_closed_map: "\embedding_map X Y f; closedin Y (f ` topspace X)\ \ closed_map X Y f" unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) lemma embedding_imp_closed_map_eq: "embedding_map X Y f \ (closed_map X Y f \ closedin Y (f ` topspace X))" using closed_map_def embedding_imp_closed_map by blast subsection\Retraction and section maps\ definition retraction_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "retraction_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace Y. f(g x) = x)" definition section_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "section_map X Y f \ \g. retraction_maps Y X g f" definition retraction_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "retraction_map X Y f \ \g. retraction_maps X Y f g" lemma retraction_maps_eq: "\retraction_maps X Y f g; \x. x \ topspace X \ f x = f' x; \x. x \ topspace Y \ g x = g' x\ \ retraction_maps X Y f' g'" unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq) lemma section_map_eq: "\section_map X Y f; \x. x \ topspace X \ f x = g x\ \ section_map X Y g" unfolding section_map_def using retraction_maps_eq by blast lemma retraction_map_eq: "\retraction_map X Y f; \x. x \ topspace X \ f x = g x\ \ retraction_map X Y g" unfolding retraction_map_def using retraction_maps_eq by blast lemma homeomorphic_imp_retraction_maps: "homeomorphic_maps X Y f g \ retraction_maps X Y f g" by (simp add: homeomorphic_maps_def retraction_maps_def) lemma section_and_retraction_eq_homeomorphic_map: "section_map X Y f \ retraction_map X Y f \ homeomorphic_map X Y f" (is "?lhs = ?rhs") proof assume ?lhs then obtain g where "homeomorphic_maps X Y f g" unfolding homeomorphic_maps_def retraction_map_def section_map_def by (smt (verit, best) continuous_map_def retraction_maps_def) then show ?rhs using homeomorphic_map_maps by blast next assume ?rhs then show ?lhs unfolding retraction_map_def section_map_def by (meson homeomorphic_imp_retraction_maps homeomorphic_map_maps homeomorphic_maps_sym) qed lemma section_imp_embedding_map: "section_map X Y f \ embedding_map X Y f" unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology) lemma retraction_imp_quotient_map: assumes "retraction_map X Y f" shows "quotient_map X Y f" unfolding quotient_map_def proof (intro conjI subsetI allI impI) show "f ` topspace X = topspace Y" using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def) next fix U assume U: "U \ topspace Y" have "openin Y U" if "\x\topspace Y. g x \ topspace X" "\x\topspace Y. f (g x) = x" "openin Y {x \ topspace Y. g x \ {x \ topspace X. f x \ U}}" for g using openin_subopen U that by fastforce then show "openin X {x \ topspace X. f x \ U} = openin Y U" using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def) qed lemma retraction_maps_compose: "\retraction_maps X Y f f'; retraction_maps Y Z g g'\ \ retraction_maps X Z (g \ f) (f' \ g')" by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def) lemma retraction_map_compose: "\retraction_map X Y f; retraction_map Y Z g\ \ retraction_map X Z (g \ f)" by (meson retraction_map_def retraction_maps_compose) lemma section_map_compose: "\section_map X Y f; section_map Y Z g\ \ section_map X Z (g \ f)" by (meson retraction_maps_compose section_map_def) lemma surjective_section_eq_homeomorphic_map: "section_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map) lemma surjective_retraction_or_section_map: "f ` (topspace X) = topspace Y \ retraction_map X Y f \ section_map X Y f \ retraction_map X Y f" using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce lemma retraction_imp_surjective_map: "retraction_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map) lemma section_imp_injective_map: "\section_map X Y f; x \ topspace X; y \ topspace X\ \ f x = f y \ x = y" by (metis (mono_tags, opaque_lifting) retraction_maps_def section_map_def) lemma retraction_maps_to_retract_maps: "retraction_maps X Y r s \ retraction_maps X (subtopology X (s ` (topspace Y))) (s \ r) id" unfolding retraction_maps_def by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology) subsection \Continuity\ lemma continuous_on_open: "continuous_on S f \ (\T. openin (top_of_set (f ` S)) T \ openin (top_of_set S) (S \ f -` T))" unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_closed: "continuous_on S f \ (\T. closedin (top_of_set (f ` S)) T \ closedin (top_of_set S) (S \ f -` T))" unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_imp_closedin: assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T" shows "closedin (top_of_set S) (S \ f -` T)" using assms continuous_on_closed by blast lemma continuous_map_subtopology_eu [simp]: "continuous_map (top_of_set S) (subtopology euclidean T) h \ continuous_on S h \ h ` S \ T" by (simp add: continuous_map_in_subtopology) lemma continuous_map_euclidean_top_of_set: assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f" shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "openin (top_of_set (f ` S)) (T \ f ` S)" using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto then show ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "closedin (top_of_set (f ` S)) (T \ f ` S)" using closedin_closed_Int[of T "f ` S", OF assms(2)] by (simp add: Int_commute) then show ?thesis using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_openin_preimage_eq: "continuous_on S f \ (\T. open T \ openin (top_of_set S) (S \ f -` T))" by (metis Int_commute continuous_on_open_invariant open_openin openin_subtopology) lemma continuous_closedin_preimage_eq: "continuous_on S f \ (\T. closed T \ closedin (top_of_set S) (S \ f -` T))" by (metis Int_commute closedin_closed continuous_on_closed_invariant) lemma continuous_open_preimage: assumes contf: "continuous_on S f" and "open S" "open T" shows "open (S \ f -` T)" proof- obtain U where "open U" "(S \ f -` T) = S \ U" using continuous_openin_preimage_gen[OF contf \open T\] unfolding openin_open by auto then show ?thesis using open_Int[of S U, OF \open S\] by auto qed lemma continuous_closed_preimage: assumes contf: "continuous_on S f" and "closed S" "closed T" shows "closed (S \ f -` T)" proof- obtain U where "closed U" "(S \ f -` T) = S \ U" using continuous_closedin_preimage[OF contf \closed T\] unfolding closedin_closed by auto then show ?thesis using closed_Int[of S U, OF \closed S\] by auto qed lemma continuous_open_vimage: "open S \ (\x. continuous (at x) f) \ open (f -` S)" by (metis continuous_on_eq_continuous_within open_vimage) lemma continuous_closed_vimage: "closed S \ (\x. continuous (at x) f) \ closed (f -` S)" by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: assumes "(x, y) \ U" "openin (top_of_set (S \ T)) U" obtains V W where "openin (top_of_set S) V" "x \ V" "openin (top_of_set T) W" "y \ W" "(V \ W) \ U" proof - from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" by (auto simp: openin_open) from open_prod_elim[OF \open E\ \(x, y) \ E\] obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" by blast show ?thesis proof show "openin (top_of_set S) (E1 \ S)" "openin (top_of_set T) (E2 \ T)" using \open E1\ \open E2\ by (auto simp: openin_open) show "x \ E1 \ S" "y \ E2 \ T" using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto show "(E1 \ S) \ (E2 \ T) \ U" using \E1 \ E2 \ E\ \U = _\ by auto qed qed lemma closedin_Times: "closedin (top_of_set S) S' \ closedin (top_of_set T) T' \ closedin (top_of_set (S \ T)) (S' \ T')" unfolding closedin_closed using closed_Times by blast lemma openin_Times: "openin (top_of_set S) S' \ openin (top_of_set T) T' \ openin (top_of_set (S \ T)) (S' \ T')" unfolding openin_open using open_Times by blast lemma openin_Times_eq: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (top_of_set (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (top_of_set S) S' \ openin (top_of_set T) T'" (is "?lhs = ?rhs") proof (cases "S' = {} \ T' = {}") case True then show ?thesis by auto next case False then obtain x y where "x \ S'" "y \ T'" by blast show ?thesis proof assume ?lhs have "openin (top_of_set S) S'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set S) U \ x \ U \ U \ S'" if "x \ S'" for x using that \y \ T'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed moreover have "openin (top_of_set T) T'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set T) U \ y \ U \ U \ T'" if "y \ T'" for y using that \x \ S'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed ultimately show ?rhs by simp next assume ?rhs with False show ?lhs by (simp add: openin_Times) qed qed lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" and "openin (top_of_set T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - have "\\<^sub>F x in at a within T. x \ T \ x \ a" by (simp add: eventually_at_filter) moreover from \openin _ _\ obtain U where "open U" "S = T \ U" by (auto simp: openin_open) then have "a \ U" using \a \ S\ by auto from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] have "\\<^sub>F x in at a within T. x \ U" by auto ultimately have "\\<^sub>F x in at a within T. f x = g x" by eventually_elim (auto simp: \S = _\ eq) with f show ?thesis by (rule Lim_transform_eventually) qed lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (top_of_set T) U \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (clarsimp simp add: continuous_openin_preimage_eq openin_open) (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) then show "openin (top_of_set S) (S \ f -` U)" by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: "\continuous_on S f; f ` S \ T; openin (top_of_set T) U\ \ openin (top_of_set S) (S \ f -` U)" by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) (S \ f -` U))" proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast then show ?thesis unfolding continuous_on_open_gen [OF assms] by (metis closedin_def inf.cobounded1 openin_closedin_eq topspace_euclidean_subtopology) qed lemma continuous_closedin_preimage_gen: assumes "continuous_on S f" "f ` S \ T" "closedin (top_of_set T) U" shows "closedin (top_of_set S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: assumes "continuous (at a within T) f" and "openin (top_of_set T) S" "a \ S" and eq: "\x. x \ S \ f x = g x" shows "continuous (at a within T) g" using assms by (simp add: Lim_transform_within_openin continuous_within) subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ the empty set. However, it is convenient to have, and is never a problem in proofs, so I prefer to write it down explicitly. We do not require \UNIV\ to be an open set, as this will not be the case in applications. (We are thinking of a topology on a subset of \UNIV\, the remaining part of \UNIV\ being irrelevant.)\ inductive generate_topology_on for S where Empty: "generate_topology_on S {}" | Int: "generate_topology_on S a \ generate_topology_on S b \ generate_topology_on S (a \ b)" | UN: "(\k. k \ K \ generate_topology_on S k) \ generate_topology_on S (\K)" | Basis: "s \ S \ generate_topology_on S s" lemma istopology_generate_topology_on: "istopology (generate_topology_on S)" unfolding istopology_def by (auto intro: generate_topology_on.intros) text \The basic property of the topology generated by a set \S\ is that it is the smallest topology containing all the elements of \S\:\ lemma generate_topology_on_coarsest: assumes T: "istopology T" "\s. s \ S \ T s" and gen: "generate_topology_on S s0" shows "T s0" using gen by (induct rule: generate_topology_on.induct) (use T in \auto simp: istopology_def\) abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ topology (generate_topology_on S)" lemma openin_topology_generated_by_iff: "openin (topology_generated_by S) s \ generate_topology_on S s" using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp lemma openin_topology_generated_by: "openin (topology_generated_by S) s \ generate_topology_on S s" using openin_topology_generated_by_iff by auto lemma topology_generated_by_topspace [simp]: "topspace (topology_generated_by S) = (\S)" proof { fix s assume "openin (topology_generated_by S) s" then have "generate_topology_on S s" by (rule openin_topology_generated_by) then have "s \ (\S)" by (induct, auto) } then show "topspace (topology_generated_by S) \ (\S)" unfolding topspace_def by auto next have "generate_topology_on S (\S)" using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp then show "(\S) \ topspace (topology_generated_by S)" unfolding topspace_def using openin_topology_generated_by_iff by auto qed lemma topology_generated_by_Basis: "s \ S \ openin (topology_generated_by S) s" by (simp add: Basis openin_topology_generated_by_iff) lemma generate_topology_on_Inter: "\finite \; \K. K \ \ \ generate_topology_on \ K; \ \ {}\ \ generate_topology_on \ (\\)" by (induction \ rule: finite_induct; force intro: generate_topology_on.intros) subsection\Topology bases and sub-bases\ lemma istopology_base_alt: "istopology (arbitrary union_of P) \ (\S T. (arbitrary union_of P) S \ (arbitrary union_of P) T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union) lemma istopology_base_eq: "istopology (arbitrary union_of P) \ (\S T. P S \ P T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_base_alt arbitrary_union_of_Int_eq) lemma istopology_base: "(\S T. \P S; P T\ \ P(S \ T)) \ istopology (arbitrary union_of P)" by (simp add: arbitrary_def istopology_base_eq union_of_inc) lemma openin_topology_base_unique: "openin X = arbitrary union_of P \ (\V. P V \ openin X V) \ (\U x. openin X U \ x \ U \ (\V. P V \ x \ V \ V \ U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: union_of_def arbitrary_def) next assume R: ?rhs then have *: "\\\Collect P. \\ = S" if "openin X S" for S using that by (rule_tac x="{V. P V \ V \ S}" in exI) fastforce from R show ?lhs by (fastforce simp add: union_of_def arbitrary_def intro: *) qed lemma topology_base_unique: assumes "\S. P S \ openin X S" "\U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U" shows "topology (arbitrary union_of P) = X" proof - have "X = topology (openin X)" by (simp add: openin_inverse) also from assms have "openin X = arbitrary union_of P" by (subst openin_topology_base_unique) auto finally show ?thesis .. qed lemma topology_bases_eq_aux: "\(arbitrary union_of P) S; \U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U\ \ (arbitrary union_of Q) S" by (metis arbitrary_union_of_alt arbitrary_union_of_idempot) lemma topology_bases_eq: "\\U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U; \V x. \Q V; x \ V\ \ \U. P U \ x \ U \ U \ V\ \ topology (arbitrary union_of P) = topology (arbitrary union_of Q)" by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux) lemma istopology_subbase: "istopology (arbitrary union_of (finite intersection_of P relative_to S))" by (simp add: finite_intersection_of_Int istopology_base relative_to_Int) lemma openin_subbase: "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S \ (arbitrary union_of (finite intersection_of B relative_to U)) S" by (simp add: istopology_subbase topology_inverse') lemma topspace_subbase [simp]: "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _") proof show "?lhs \ U" by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset) show "U \ ?lhs" by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase openin_subset relative_to_inc subset_UNIV topology_inverse') qed lemma minimal_topology_subbase: assumes X: "\S. P S \ openin X S" and "openin X U" and S: "openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S" shows "openin X S" proof - have "(arbitrary union_of (finite intersection_of P relative_to U)) S" using S openin_subbase by blast with X \openin X U\ show ?thesis by (force simp add: union_of_def intersection_of_def relative_to_def intro: openin_Int_Inter) qed lemma istopology_subbase_UNIV: "istopology (arbitrary union_of (finite intersection_of P))" by (simp add: istopology_base finite_intersection_of_Int) lemma generate_topology_on_eq: "generate_topology_on S = arbitrary union_of finite' intersection_of (\x. x \ S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix A assume "?lhs A" then show "?rhs A" proof induction case (Int a b) then show ?case by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base) next case (UN K) then show ?case by (simp add: arbitrary_union_of_Union) next case (Basis s) then show ?case by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset) qed auto next fix A assume "?rhs A" then obtain \ where \: "\T. T \ \ \ \\. finite' \ \ \ \ S \ \\ = T" and eq: "A = \\" unfolding union_of_def intersection_of_def by auto show "?lhs A" unfolding eq proof (rule generate_topology_on.UN) fix T assume "T \ \" with \ obtain \ where "finite' \" "\ \ S" "\\ = T" by blast have "generate_topology_on S (\\)" proof (rule generate_topology_on_Inter) show "finite \" "\ \ {}" by (auto simp: \finite' \\) show "\K. K \ \ \ generate_topology_on S K" by (metis \\ \ S\ generate_topology_on.simps subset_iff) qed then show "generate_topology_on S T" using \\\ = T\ by blast qed qed lemma continuous_on_generated_topo_iff: "continuous_map T1 (topology_generated_by S) f \ ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" unfolding continuous_map_alt topology_generated_by_topspace proof (auto simp add: topology_generated_by_Basis) assume H: "\U. U \ S \ openin T1 (f -` U \ topspace T1)" fix U assume "openin (topology_generated_by S) U" then have "generate_topology_on S U" by (rule openin_topology_generated_by) then show "openin T1 (f -` U \ topspace T1)" proof (induct) fix a b assume H: "openin T1 (f -` a \ topspace T1)" "openin T1 (f -` b \ topspace T1)" have "f -` (a \ b) \ topspace T1 = (f-`a \ topspace T1) \ (f-`b \ topspace T1)" by auto then show "openin T1 (f -` (a \ b) \ topspace T1)" using H by auto next fix K assume H: "openin T1 (f -` k \ topspace T1)" if "k\ K" for k define L where "L = {f -` k \ topspace T1|k. k \ K}" have *: "openin T1 l" if "l \L" for l using that H unfolding L_def by auto have "openin T1 (\L)" using openin_Union[OF *] by simp moreover have "(\L) = (f -` \K \ topspace T1)" unfolding L_def by auto ultimately show "openin T1 (f -` \K \ topspace T1)" by simp qed (auto simp add: H) qed lemma continuous_on_generated_topo: assumes "\U. U \S \ openin T1 (f-`U \ topspace(T1))" "f`(topspace T1) \ (\ S)" shows "continuous_map T1 (topology_generated_by S) f" using assms continuous_on_generated_topo_iff by blast subsection\Continuity via bases/subbases, hence upper and lower semicontinuity\ lemma continuous_map_into_topology_base: assumes P: "openin Y = arbitrary union_of P" and f: "\x. x \ topspace X \ f x \ topspace Y" and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" shows "continuous_map X Y f" proof - have *: "\\. (\t. t \ \ \ P t) \ openin X {x \ topspace X. \U\\. f x \ U}" by (smt (verit) Ball_Collect ope mem_Collect_eq openin_subopen) show ?thesis using P by (auto simp: continuous_map_def arbitrary_def union_of_def intro!: f *) qed lemma continuous_map_into_topology_base_eq: assumes P: "openin Y = arbitrary union_of P" shows "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" (is "?lhs=?rhs") proof assume L: ?lhs then have "\x. x \ topspace X \ f x \ topspace Y" by (meson continuous_map_def) moreover have "\U. P U \ openin X {x \ topspace X. f x \ U}" using L assms continuous_map openin_topology_base_unique by fastforce ultimately show ?rhs by auto qed (simp add: assms continuous_map_into_topology_base) lemma continuous_map_into_topology_subbase: fixes U P defines "Y \ topology(arbitrary union_of (finite intersection_of P relative_to U))" assumes f: "\x. x \ topspace X \ f x \ topspace Y" and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" shows "continuous_map X Y f" proof (intro continuous_map_into_topology_base) show "openin Y = arbitrary union_of (finite intersection_of P relative_to U)" unfolding Y_def using istopology_subbase topology_inverse' by blast show "openin X {x \ topspace X. f x \ V}" if \
: "(finite intersection_of P relative_to U) V" for V proof - define finv where "finv \ \V. {x \ topspace X. f x \ V}" obtain \ where \: "finite \" "\V. V \ \ \ P V" "{x \ topspace X. f x \ V} = (\V \ insert U \. finv V)" using \
by (fastforce simp: finv_def intersection_of_def relative_to_def) show ?thesis unfolding \ proof (intro openin_Inter ope) have U: "U = topspace Y" unfolding Y_def using topspace_subbase by fastforce fix V assume V: "V \ finv ` insert U \" with U f have "openin X {x \ topspace X. f x \ U}" by (auto simp: openin_subopen [of X "Collect _"]) then show "openin X V" using V \(2) ope by (fastforce simp: finv_def) qed (use \finite \\ in auto) qed qed (use f in auto) lemma continuous_map_into_topology_subbase_eq: assumes "Y = topology(arbitrary union_of (finite intersection_of P relative_to U))" shows "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "\x. x \ topspace X \ f x \ topspace Y" using L continuous_map_def by fastforce fix V assume "P V" have "U = topspace Y" unfolding assms using topspace_subbase by fastforce then have eq: "{x \ topspace X. f x \ V} = {x \ topspace X. f x \ U \ V}" using L by (auto simp: continuous_map) have "openin Y (U \ V)" unfolding assms openin_subbase by (meson \P V\ arbitrary_union_of_inc finite_intersection_of_inc relative_to_inc) show "openin X {x \ topspace X. f x \ V}" using L \openin Y (U \ V)\ continuous_map eq by fastforce qed next show "?rhs \ ?lhs" unfolding assms by (intro continuous_map_into_topology_subbase) auto qed lemma subbase_subtopology_euclidean: fixes U :: "'a::order_topology set" shows "topology (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ range lessThan) relative_to U)) = subtopology euclidean U" proof - have "\V. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ x \ V \ V \ W" if "open W" "x \ W" for W and x::'a using \open W\ [unfolded open_generated_order] \x \ W\ proof (induct rule: generate_topology.induct) case UNIV then show ?case using finite_intersection_of_empty by blast next case (Int a b) then show ?case by (meson Int_iff finite_intersection_of_Int inf_mono) next case (UN K) then show ?case by (meson Union_iff subset_iff) next case (Basis s) then show ?case by (metis (no_types, lifting) Un_iff finite_intersection_of_inc order_refl) qed moreover have "\V::'a set. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ open V" by (force simp: intersection_of_def subset_iff) ultimately have *: "openin (euclidean::'a topology) = (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)))" by (smt (verit, best) openin_topology_base_unique open_openin) show ?thesis unfolding subtopology_def arbitrary_union_of_relative_to [symmetric] apply (simp add: relative_to_def flip: *) by (metis Int_commute) qed lemma continuous_map_upper_lower_semicontinuous_lt_gen: fixes U :: "'a::order_topology set" shows "continuous_map X (subtopology euclidean U) f \ (\x \ topspace X. f x \ U) \ (\a. openin X {x \ topspace X. f x > a}) \ (\a. openin X {x \ topspace X. f x < a})" by (auto simp: continuous_map_into_topology_subbase_eq [OF subbase_subtopology_euclidean [symmetric, of U]] greaterThan_def lessThan_def image_iff simp flip: all_simps) lemma continuous_map_upper_lower_semicontinuous_lt: fixes f :: "'a \ 'b::order_topology" shows "continuous_map X euclidean f \ (\a. openin X {x \ topspace X. f x > a}) \ (\a. openin X {x \ topspace X. f x < a})" using continuous_map_upper_lower_semicontinuous_lt_gen [where U=UNIV] by auto lemma Int_Collect_imp_eq: "A \ {x. x\A \ P x} = {x\A. P x}" by blast lemma continuous_map_upper_lower_semicontinuous_le_gen: shows "continuous_map X (subtopology euclideanreal U) f \ (\x \ topspace X. f x \ U) \ (\a. closedin X {x \ topspace X. f x \ a}) \ (\a. closedin X {x \ topspace X. f x \ a})" unfolding continuous_map_upper_lower_semicontinuous_lt_gen by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) lemma continuous_map_upper_lower_semicontinuous_le: "continuous_map X euclideanreal f \ (\a. closedin X {x \ topspace X. f x \ a}) \ (\a. closedin X {x \ topspace X. f x \ a})" using continuous_map_upper_lower_semicontinuous_le_gen [where U=UNIV] by auto lemma continuous_map_upper_lower_semicontinuous_lte_gen: "continuous_map X (subtopology euclideanreal U) f \ (\x \ topspace X. f x \ U) \ (\a. openin X {x \ topspace X. f x < a}) \ (\a. closedin X {x \ topspace X. f x \ a})" unfolding continuous_map_upper_lower_semicontinuous_lt_gen by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) lemma continuous_map_upper_lower_semicontinuous_lte: "continuous_map X euclideanreal f \ (\a. openin X {x \ topspace X. f x < a}) \ (\a. closedin X {x \ topspace X. f x \ a})" using continuous_map_upper_lower_semicontinuous_lte_gen [where U=UNIV] by auto subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as we will need it to define the strong operator topology on the space of continuous linear operators, by pulling back the product topology on the space of all functions.\ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: "istopology (\S. \U. openin T U \ S = f-`U \ A)" unfolding istopology_def proof (auto) fix K assume "\S\K. \U. openin T U \ S = f -` U \ A" then have "\U. \S\K. openin T (U S) \ S = f-`(U S) \ A" by (rule bchoice) then obtain U where U: "\S\K. openin T (U S) \ S = f-`(U S) \ A" by blast define V where "V = (\S\K. U S)" have "openin T V" "\K = f -` V \ A" unfolding V_def using U by auto then show "\V. openin T V \ \K = f -` V \ A" by auto qed lemma openin_pullback_topology: "openin (pullback_topology A f T) S \ (\U. openin T U \ S = f-`U \ A)" unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto lemma topspace_pullback_topology: "topspace (pullback_topology A f T) = f-`(topspace T) \ A" by (auto simp add: topspace_def openin_pullback_topology) proposition continuous_map_pullback [intro]: assumes "continuous_map T1 T2 g" shows "continuous_map (pullback_topology A f T1) T2 (g o f)" unfolding continuous_map_alt proof (auto) fix U::"'b set" assume "openin T2 U" then have "openin T1 (g-`U \ topspace T1)" using assms unfolding continuous_map_alt by auto have "(g o f)-`U \ topspace (pullback_topology A f T1) = (g o f)-`U \ A \ f-`(topspace T1)" unfolding topspace_pullback_topology by auto also have "... = f-`(g-`U \ topspace T1) \ A " by auto also have "openin (pullback_topology A f T1) (...)" unfolding openin_pullback_topology using \openin T1 (g-`U \ topspace T1)\ by auto finally show "openin (pullback_topology A f T1) ((g \ f) -` U \ topspace (pullback_topology A f T1))" by auto next fix x assume "x \ topspace (pullback_topology A f T1)" then have "f x \ topspace T1" unfolding topspace_pullback_topology by auto then show "g (f x) \ topspace T2" using assms unfolding continuous_map_def by auto qed proposition continuous_map_pullback' [intro]: assumes "continuous_map T1 T2 (f o g)" "topspace T1 \ g-`A" shows "continuous_map T1 (pullback_topology A f T2) g" unfolding continuous_map_alt proof (auto) fix U assume "openin (pullback_topology A f T2) U" then have "\V. openin T2 V \ U = f-`V \ A" unfolding openin_pullback_topology by auto then obtain V where "openin T2 V" "U = f-`V \ A" by blast then have "g -` U \ topspace T1 = g-`(f-`V \ A) \ topspace T1" by blast also have "... = (f o g)-`V \ (g-`A \ topspace T1)" by auto also have "... = (f o g)-`V \ topspace T1" using assms(2) by auto also have "openin T1 (...)" using assms(1) \openin T2 V\ by auto finally show "openin T1 (g -` U \ topspace T1)" by simp next fix x assume "x \ topspace T1" have "(f o g) x \ topspace T2" using assms(1) \x \ topspace T1\ unfolding continuous_map_def by auto then have "g x \ f-`(topspace T2)" unfolding comp_def by blast moreover have "g x \ A" using assms(2) \x \ topspace T1\ by blast ultimately show "g x \ topspace (pullback_topology A f T2)" unfolding topspace_pullback_topology by blast qed subsection\Proper maps (not a priori assumed continuous) \ definition proper_map where "proper_map X Y f \ closed_map X Y f \ (\y \ topspace Y. compactin X {x \ topspace X. f x = y})" lemma proper_imp_closed_map: "proper_map X Y f \ closed_map X Y f" by (simp add: proper_map_def) lemma proper_map_imp_subset_topspace: "proper_map X Y f \ f ` (topspace X) \ topspace Y" by (simp add: closed_map_imp_subset_topspace proper_map_def) lemma proper_map_restriction: assumes "proper_map X Y f" "{x \ topspace X. f x \ V} = U" shows "proper_map (subtopology X U) (subtopology Y V) f" proof - have [simp]: "{x \ topspace X. f x \ V \ f x = y} = {x \ topspace X. f x = y}" if "y \ V" for y using that by auto show ?thesis using assms unfolding proper_map_def using closed_map_restriction by (force simp: compactin_subtopology) qed lemma closed_injective_imp_proper_map: assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)" shows "proper_map X Y f" unfolding proper_map_def proof (clarsimp simp: f) show "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y using inj_on_eq_iff [OF inj] that proof - have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using inj_on_eq_iff [OF inj] by auto then show ?thesis using that by (metis (no_types, lifting) compactin_empty compactin_sing) qed qed lemma injective_imp_proper_eq_closed_map: "inj_on f (topspace X) \ (proper_map X Y f \ closed_map X Y f)" using closed_injective_imp_proper_map proper_imp_closed_map by blast lemma homeomorphic_imp_proper_map: "homeomorphic_map X Y f \ proper_map X Y f" by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map) lemma compactin_proper_map_preimage: assumes f: "proper_map X Y f" and "compactin Y K" shows "compactin X {x. x \ topspace X \ f x \ K}" proof - have "f ` (topspace X) \ topspace Y" by (simp add: f proper_map_imp_subset_topspace) have *: "\y. y \ topspace Y \ compactin X {x \ topspace X. f x = y}" using f by (auto simp: proper_map_def) show ?thesis unfolding compactin_def proof clarsimp show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x \ K} \ \\" if \: "\U\\. openin X U" and sub: "{x \ topspace X. f x \ K} \ \\" for \ proof - have "\y \ K. \\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" proof fix y assume "y \ K" then have "compactin X {x \ topspace X. f x = y}" by (metis "*" \compactin Y K\ compactin_subspace subsetD) with \y \ K\ show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" unfolding compactin_def using \ sub by fastforce qed then obtain \ where \: "\y. y \ K \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by (metis (full_types)) define F where "F \ \y. topspace Y - f ` (topspace X - \(\ y))" have "\\. finite \ \ \ \ F ` K \ K \ \\" proof (rule compactinD [OF \compactin Y K\]) have "\x. x \ K \ closedin Y (f ` (topspace X - \(\ x)))" using f unfolding proper_map_def closed_map_def by (meson \ \ openin_Union openin_closedin_eq subsetD) then show "openin Y U" if "U \ F ` K" for U using that by (auto simp: F_def) show "K \ \(F ` K)" using \ \compactin Y K\ unfolding F_def compactin_def by fastforce qed then obtain J where "finite J" "J \ K" and J: "K \ \(F ` J)" by (auto simp: ex_finite_subset_image) show ?thesis unfolding F_def proof (intro exI conjI) show "finite (\(\ ` J))" using \ \J \ K\ \finite J\ by blast show "\(\ ` J) \ \" using \ \J \ K\ by blast show "{x \ topspace X. f x \ K} \ \(\(\ ` J))" using J \J \ K\ unfolding F_def by auto qed qed qed qed lemma compact_space_proper_map_preimage: assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y" shows "compact_space X" proof - have eq: "topspace X = {x \ topspace X. f x \ topspace Y}" using fim by blast moreover have "compactin Y (topspace Y)" using \compact_space Y\ compact_space_def by auto ultimately show ?thesis unfolding compact_space_def using eq f compactin_proper_map_preimage by fastforce qed lemma proper_map_alt: "proper_map X Y f \ closed_map X Y f \ (\K. compactin Y K \ compactin X {x. x \ topspace X \ f x \ K})" proof (intro iffI conjI allI impI) show "compactin X {x \ topspace X. f x \ K}" if "proper_map X Y f" and "compactin Y K" for K using that by (simp add: compactin_proper_map_preimage) show "proper_map X Y f" if f: "closed_map X Y f \ (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" proof - have "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "compactin X {x \ topspace X. f x \ {y}}" using f compactin_sing that by fastforce then show ?thesis by auto qed with f show ?thesis by (auto simp: proper_map_def) qed qed (simp add: proper_imp_closed_map) lemma proper_map_on_empty: "topspace X = {} \ proper_map X Y f" by (auto simp: proper_map_def closed_map_on_empty) lemma proper_map_id [simp]: "proper_map X X id" proof (clarsimp simp: proper_map_alt closed_map_id) fix K assume K: "compactin X K" then have "{a \ topspace X. a \ K} = K" by (simp add: compactin_subspace subset_antisym subset_iff) then show "compactin X {a \ topspace X. a \ K}" using K by auto qed lemma proper_map_compose: assumes "proper_map X Y f" "proper_map Y Z g" shows "proper_map X Z (g \ f)" proof - have "closed_map X Y f" and f: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" and "closed_map Y Z g" and g: "\K. compactin Z K \ compactin Y {x \ topspace Y. g x \ K}" using assms by (auto simp: proper_map_alt) show ?thesis unfolding proper_map_alt proof (intro conjI allI impI) show "closed_map X Z (g \ f)" using \closed_map X Y f\ \closed_map Y Z g\ closed_map_compose by blast have "{x \ topspace X. g (f x) \ K} = {x \ topspace X. f x \ {b \ topspace Y. g b \ K}}" for K using \closed_map X Y f\ closed_map_imp_subset_topspace by blast then show "compactin X {x \ topspace X. (g \ f) x \ K}" if "compactin Z K" for K using f [OF g [OF that]] by auto qed qed lemma proper_map_const: "proper_map X Y (\x. c) \ compact_space X \ (topspace X = {} \ closedin Y {c})" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: compact_space_topspace_empty proper_map_on_empty) next case False have *: "compactin X {x \ topspace X. c = y}" if "compact_space X" for y using that unfolding compact_space_def by (metis (mono_tags, lifting) compactin_empty empty_subsetI mem_Collect_eq subsetI subset_antisym) then show ?thesis using closed_compactin closedin_subset by (force simp: False proper_map_def closed_map_const compact_space_def) qed lemma proper_map_inclusion: "S \ topspace X \ proper_map (subtopology X S) X id \ closedin X S \ (\k. compactin X k \ compactin X (S \ k))" by (metis closed_Int_compactin closed_map_inclusion_eq inf.absorb_iff2 inj_on_id injective_imp_proper_eq_closed_map) lemma proper_map_into_subtopology: "\proper_map X Y f; f ` topspace X \ C\ \ proper_map X (subtopology Y C) f" by (simp add: closed_map_into_subtopology compactin_subtopology proper_map_alt) lemma proper_map_from_composition_left: assumes gf: "proper_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "proper_map Y Z g" unfolding proper_map_def proof (intro strip conjI) show "closed_map Y Z g" by (meson closed_map_from_composition_left gf contf fim proper_imp_closed_map) fix z assume "z \ topspace Z" have eq: "{y \ topspace Y. g y = z} = f ` {x \ topspace X. (g \ f) x = z}" using fim by force show "compactin Y {x \ topspace Y. g x = z}" unfolding eq proof (rule image_compactin [OF _ contf]) show "compactin X {x \ topspace X. (g \ f) x = z}" using \z \ topspace Z\ gf proper_map_def by fastforce qed qed lemma proper_map_from_composition_right_inj: assumes gf: "proper_map X Z (g \ f)" and fim: "f ` topspace X \ topspace Y" and contf: "continuous_map Y Z g" and inj: "inj_on g (topspace Y)" shows "proper_map X Y f" unfolding proper_map_def proof (intro strip conjI) show "closed_map X Y f" by (meson closed_map_from_composition_right assms proper_imp_closed_map) fix y assume "y \ topspace Y" with fim inj have eq: "{x \ topspace X. f x = y} = {x \ topspace X. (g \ f) x = g y}" by (auto simp: image_subset_iff inj_onD) show "compactin X {x \ topspace X. f x = y}" unfolding eq by (smt (verit) Collect_cong \y \ topspace Y\ contf continuous_map_closedin gf proper_map_def) qed subsection\Perfect maps (proper, continuous and surjective)\ definition perfect_map where "perfect_map X Y f \ continuous_map X Y f \ proper_map X Y f \ f ` (topspace X) = topspace Y" lemma homeomorphic_imp_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f" by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def) lemma perfect_imp_quotient_map: "perfect_map X Y f \ quotient_map X Y f" by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def) lemma homeomorphic_eq_injective_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f \ inj_on f (topspace X)" using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast lemma perfect_injective_eq_homeomorphic_map: "perfect_map X Y f \ inj_on f (topspace X) \ homeomorphic_map X Y f" by (simp add: homeomorphic_eq_injective_perfect_map) lemma perfect_map_id [simp]: "perfect_map X X id" by (simp add: homeomorphic_imp_perfect_map) lemma perfect_map_compose: "\perfect_map X Y f; perfect_map Y Z g\ \ perfect_map X Z (g \ f)" by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def) lemma perfect_imp_continuous_map: "perfect_map X Y f \ continuous_map X Y f" using perfect_map_def by blast lemma perfect_imp_closed_map: "perfect_map X Y f \ closed_map X Y f" by (simp add: perfect_map_def proper_map_def) lemma perfect_imp_proper_map: "perfect_map X Y f \ proper_map X Y f" by (simp add: perfect_map_def) lemma perfect_imp_surjective_map: "perfect_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: perfect_map_def) lemma perfect_map_from_composition_left: assumes "perfect_map X Z (g \ f)" and "continuous_map X Y f" and "continuous_map Y Z g" and "f ` topspace X = topspace Y" shows "perfect_map Y Z g" using assms unfolding perfect_map_def by (metis image_comp proper_map_from_composition_left) end diff --git a/src/HOL/Analysis/Analysis.thy b/src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy +++ b/src/HOL/Analysis/Analysis.thy @@ -1,58 +1,59 @@ theory Analysis imports (* Linear Algebra *) Convex Determinants (* Topology *) FSigma Sum_Topology Abstract_Topological_Spaces Abstract_Metric_Spaces + Urysohn Connected Abstract_Limits Isolated (* Functional Analysis *) Elementary_Normed_Spaces Norm_Arith (* Vector Analysis *) Convex_Euclidean_Space Operator_Norm (* Unsorted *) Line_Segment Derivative Cartesian_Euclidean_Space Weierstrass_Theorems (* Measure and Integration Theory *) Ball_Volume Integral_Test Improper_Integral Equivalence_Measurable_On_Borel Lebesgue_Integral_Substitution Embed_Measure Complete_Measure Radon_Nikodym Fashoda_Theorem Cross3 Homeomorphism Bounded_Continuous_Function Abstract_Topology Product_Topology Lindelof_Spaces Infinite_Products Infinite_Sum Infinite_Set_Sum Polytope Jordan_Curve Poly_Roots Generalised_Binomial_Theorem Gamma_Function Change_Of_Vars Multivariate_Analysis Simplex_Content FPS_Convergence Smooth_Paths Abstract_Euclidean_Space Function_Metric begin end diff --git a/src/HOL/Analysis/Homotopy.thy b/src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy +++ b/src/HOL/Analysis/Homotopy.thy @@ -1,5379 +1,5403 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Homotopy of Maps\ theory Homotopy imports Path_Connected Product_Topology Uncountable_Sets begin definition\<^marker>\tag important\ homotopic_with where "homotopic_with P X Y f g \ (\h. continuous_map (prod_topology (top_of_set {0..1::real}) X) Y h \ (\x. h(0, x) = f x) \ (\x. h(1, x) = g x) \ (\t \ {0..1}. P(\x. h(t,x))))" text\\p\, \q\ are functions \X \ Y\, and the property \P\ restricts all intermediate maps. We often just want to require that \P\ fixes some subset, but to include the case of a loop homotopy, it is convenient to have a general property \P\.\ abbreviation homotopic_with_canon :: "[('a::topological_space \ 'b::topological_space) \ bool, 'a set, 'b set, 'a \ 'b, 'a \ 'b] \ bool" where "homotopic_with_canon P S T p q \ homotopic_with P (top_of_set S) (top_of_set T) p q" lemma split_01: "{0..1::real} = {0..1/2} \ {1/2..1}" by force lemma split_01_prod: "{0..1::real} \ X = ({0..1/2} \ X) \ ({1/2..1} \ X)" by force lemma image_Pair_const: "(\x. (x, c)) ` A = A \ {c}" by auto lemma fst_o_paired [simp]: "fst \ (\(x,y). (f x y, g x y)) = (\(x,y). f x y)" by auto lemma snd_o_paired [simp]: "snd \ (\(x,y). (f x y, g x y)) = (\(x,y). g x y)" by auto lemma continuous_on_o_Pair: "\continuous_on (T \ X) h; t \ T\ \ continuous_on X (h \ Pair t)" by (fast intro: continuous_intros elim!: continuous_on_subset) lemma continuous_map_o_Pair: assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \ topspace X" shows "continuous_map Y Z (h \ Pair t)" by (intro continuous_map_compose [OF _ h] continuous_intros; simp add: t) subsection\<^marker>\tag unimportant\\Trivial properties\ text \We often want to just localize the ending function equality or whatever.\ text\<^marker>\tag important\ \%whitespace\ proposition homotopic_with: assumes "\h k. (\x. x \ topspace X \ h x = k x) \ (P h \ P k)" shows "homotopic_with P X Y p q \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ (\x \ topspace X. h(0,x) = p x) \ (\x \ topspace X. h(1,x) = q x) \ (\t \ {0..1}. P(\x. h(t, x))))" unfolding homotopic_with_def apply (rule iffI, blast, clarify) apply (rule_tac x="\(u,v). if v \ topspace X then h(u,v) else if u = 0 then p v else q v" in exI) apply simp by (smt (verit, best) SigmaE assms case_prod_conv continuous_map_eq topspace_prod_topology) lemma homotopic_with_mono: assumes hom: "homotopic_with P X Y f g" and Q: "\h. \continuous_map X Y h; P h\ \ Q h" shows "homotopic_with Q X Y f g" using hom unfolding homotopic_with_def by (force simp: o_def dest: continuous_map_o_Pair intro: Q) lemma homotopic_with_imp_continuous_maps: assumes "homotopic_with P X Y f g" shows "continuous_map X Y f \ continuous_map X Y g" proof - obtain h :: "real \ 'a \ 'b" where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) Y h" and h: "\x. h (0, x) = f x" "\x. h (1, x) = g x" using assms by (auto simp: homotopic_with_def) have *: "t \ {0..1} \ continuous_map X Y (h \ (\x. (t,x)))" for t by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise) show ?thesis using h *[of 0] *[of 1] by (simp add: continuous_map_eq) qed lemma homotopic_with_imp_continuous: assumes "homotopic_with_canon P X Y f g" shows "continuous_on X f \ continuous_on X g" by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_property: assumes "homotopic_with P X Y f g" shows "P f \ P g" proof obtain h where h: "\x. h(0, x) = f x" "\x. h(1, x) = g x" and P: "\t. t \ {0..1::real} \ P(\x. h(t,x))" using assms by (force simp: homotopic_with_def) show "P f" "P g" using P [of 0] P [of 1] by (force simp: h)+ qed lemma homotopic_with_equal: assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\x. x \ topspace X \ f x = g x" shows "homotopic_with P X Y f g" unfolding homotopic_with_def proof (intro exI conjI allI ballI) let ?h = "\(t::real,x). if t = 1 then g x else f x" show "continuous_map (prod_topology (top_of_set {0..1}) X) Y ?h" proof (rule continuous_map_eq) show "continuous_map (prod_topology (top_of_set {0..1}) X) Y (f \ snd)" by (simp add: contf continuous_map_of_snd) qed (auto simp: fg) show "P (\x. ?h (t, x))" if "t \ {0..1}" for t by (cases "t = 1") (simp_all add: assms) qed auto lemma homotopic_with_imp_subset1: "homotopic_with_canon P X Y f g \ f ` X \ Y" by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_subset2: "homotopic_with_canon P X Y f g \ g ` X \ Y" by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_subset_left: "\homotopic_with_canon P X Y f g; Z \ X\ \ homotopic_with_canon P Z Y f g" unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) lemma homotopic_with_subset_right: "\homotopic_with_canon P X Y f g; Y \ Z\ \ homotopic_with_canon P X Z f g" unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) subsection\Homotopy with P is an equivalence relation\ text \(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\ lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \ continuous_map X Y f \ P f" by (metis homotopic_with_equal homotopic_with_imp_continuous_maps homotopic_with_imp_property) lemma homotopic_with_symD: assumes "homotopic_with P X Y f g" shows "homotopic_with P X Y g f" proof - let ?I01 = "subtopology euclideanreal {0..1}" let ?j = "\y. (1 - fst y, snd y)" have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j" by (intro continuous_intros; simp add: continuous_map_subtopology_fst prod_topology_subtopology) have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j" proof - have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \ topspace X)) ?j" by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff) then show ?thesis by (simp add: prod_topology_subtopology(1)) qed show ?thesis using assms apply (clarsimp simp: homotopic_with_def) subgoal for h by (rule_tac x="h \ (\y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *]) done qed lemma homotopic_with_sym: "homotopic_with P X Y f g \ homotopic_with P X Y g f" by (metis homotopic_with_symD) proposition homotopic_with_trans: assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h" shows "homotopic_with P X Y f h" proof - let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X" obtain k1 k2 where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2" and k12: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" "\x. k1 (0, x) = f x" "\x. k2 (1, x) = h x" and P: "\t\{0..1}. P (\x. k1 (t, x))" "\t\{0..1}. P (\x. k2 (t, x))" using assms by (auto simp: homotopic_with_def) define k where "k \ \y. if fst y \ 1/2 then (k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y else (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v by (simp add: k12 that) show ?thesis unfolding homotopic_with_def proof (intro exI conjI) show "continuous_map ?X01 Y k" unfolding k_def proof (rule continuous_map_cases_le) show fst: "continuous_map ?X01 euclideanreal fst" using continuous_map_fst continuous_map_in_subtopology by blast show "continuous_map ?X01 euclideanreal (\x. 1/2)" by simp show "continuous_map (subtopology ?X01 {y \ topspace ?X01. fst y \ 1/2}) Y (k1 \ (\x. (2 *\<^sub>R fst x, snd x)))" apply (intro fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ by (force simp: prod_topology_subtopology) show "continuous_map (subtopology ?X01 {y \ topspace ?X01. 1/2 \ fst y}) Y (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x)))" apply (intro fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ by (force simp: prod_topology_subtopology) show "(k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y = (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" if "y \ topspace ?X01" and "fst y = 1/2" for y using that by (simp add: keq) qed show "\x. k (0, x) = f x" by (simp add: k12 k_def) show "\x. k (1, x) = h x" by (simp add: k12 k_def) show "\t\{0..1}. P (\x. k (t, x))" proof fix t show "t\{0..1} \ P (\x. k (t, x))" by (cases "t \ 1/2") (auto simp: k_def P) qed qed qed lemma homotopic_with_id2: "(\x. x \ topspace X \ g (f x) = x) \ homotopic_with (\x. True) X X (g \ f) id" by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD) subsection\Continuity lemmas\ lemma homotopic_with_compose_continuous_map_left: "\homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \j. p j \ q(h \ j)\ \ homotopic_with q X1 X3 (h \ f) (h \ g)" unfolding homotopic_with_def apply clarify subgoal for k by (rule_tac x="h \ k" in exI) (rule conjI continuous_map_compose | simp add: o_def)+ done lemma homotopic_with_compose_continuous_map_right: assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h" and q: "\j. p j \ q(j \ h)" shows "homotopic_with q X1 X3 (f \ h) (g \ h)" proof - obtain k where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k" and k: "\x. k (0, x) = f x" "\x. k (1, x) = g x" and p: "\t. t\{0..1} \ p (\x. k (t, x))" using hom unfolding homotopic_with_def by blast have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \ snd)" by (rule continuous_map_compose [OF continuous_map_snd conth]) let ?h = "k \ (\(t,x). (t,h x))" show ?thesis unfolding homotopic_with_def proof (intro exI conjI allI ballI) have "continuous_map (prod_topology (top_of_set {0..1}) X1) (prod_topology (top_of_set {0..1::real}) X2) (\(t, x). (t, h x))" by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd) then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h" by (intro conjI continuous_map_compose [OF _ contk]) show "q (\x. ?h (t, x))" if "t \ {0..1}" for t using q [OF p [OF that]] by (simp add: o_def) qed (auto simp: k) qed corollary homotopic_compose: assumes "homotopic_with (\x. True) X Y f f'" "homotopic_with (\x. True) Y Z g g'" shows "homotopic_with (\x. True) X Z (g \ f) (g' \ f')" by (metis assms homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right homotopic_with_imp_continuous_maps homotopic_with_trans) proposition homotopic_with_compose_continuous_right: "\homotopic_with_canon (\f. p (f \ h)) X Y f g; continuous_on W h; h ` W \ X\ \ homotopic_with_canon p W Y (f \ h) (g \ h)" by (simp add: homotopic_with_compose_continuous_map_right) proposition homotopic_with_compose_continuous_left: "\homotopic_with_canon (\f. p (h \ f)) X Y f g; continuous_on Y h; h ` Y \ Z\ \ homotopic_with_canon p X Z (h \ f) (h \ g)" by (simp add: homotopic_with_compose_continuous_map_left) lemma homotopic_from_subtopology: "homotopic_with P X X' f g \ homotopic_with P (subtopology X S) X' f g" by (metis continuous_map_id_subt homotopic_with_compose_continuous_map_right o_id) lemma homotopic_on_emptyI: assumes "topspace X = {}" "P f" "P g" shows "homotopic_with P X X' f g" by (metis assms continuous_map_on_empty empty_iff homotopic_with_equal) lemma homotopic_on_empty: "topspace X = {} \ (homotopic_with P X X' f g \ P f \ P g)" using homotopic_on_emptyI homotopic_with_imp_property by metis lemma homotopic_with_canon_on_empty [simp]: "homotopic_with_canon (\x. True) {} t f g" by (auto intro: homotopic_with_equal) lemma homotopic_constant_maps: "homotopic_with (\x. True) X X' (\x. a) (\x. b) \ topspace X = {} \ path_component_of X' a b" (is "?lhs = ?rhs") proof (cases "topspace X = {}") case False then obtain c where c: "c \ topspace X" by blast have "\g. continuous_map (top_of_set {0..1::real}) X' g \ g 0 = a \ g 1 = b" if "x \ topspace X" and hom: "homotopic_with (\x. True) X X' (\x. a) (\x. b)" for x proof - obtain h :: "real \ 'a \ 'b" where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h" and h: "\x. h (0, x) = a" "\x. h (1, x) = b" using hom by (auto simp: homotopic_with_def) have cont: "continuous_map (top_of_set {0..1}) X' (h \ (\t. (t, c)))" by (rule continuous_map_compose [OF _ conth] continuous_intros c | simp)+ then show ?thesis by (force simp: h) qed moreover have "homotopic_with (\x. True) X X' (\x. g 0) (\x. g 1)" if "x \ topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g" for x and g :: "real \ 'b" unfolding homotopic_with_def by (force intro!: continuous_map_compose continuous_intros c that) ultimately show ?thesis using False by (auto simp: path_component_of_def pathin_def) qed (simp add: homotopic_on_empty) proposition homotopic_with_eq: assumes h: "homotopic_with P X Y f g" and f': "\x. x \ topspace X \ f' x = f x" and g': "\x. x \ topspace X \ g' x = g x" and P: "(\h k. (\x. x \ topspace X \ h x = k x) \ P h \ P k)" shows "homotopic_with P X Y f' g'" by (smt (verit, ccfv_SIG) assms homotopic_with) lemma homotopic_with_prod_topology: assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'" and r: "\i j. \p i; q j\ \ r(\(x,y). (i x, j y))" shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2) (\z. (f(fst z),g(snd z))) (\z. (f'(fst z), g'(snd z)))" proof - obtain h where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h" and h0: "\x. h (0, x) = f x" and h1: "\x. h (1, x) = f' x" and p: "\t. \0 \ t; t \ 1\ \ p (\x. h (t,x))" using assms unfolding homotopic_with_def by auto obtain k where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k" and k0: "\x. k (0, x) = g x" and k1: "\x. k (1, x) = g' x" and q: "\t. \0 \ t; t \ 1\ \ q (\x. k (t,x))" using assms unfolding homotopic_with_def by auto let ?hk = "\(t,x,y). (h(t,x), k(t,y))" show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2)) (prod_topology Y1 Y2) ?hk" unfolding continuous_map_pairwise case_prod_unfold by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def] continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def] continuous_map_compose [OF _ h, unfolded o_def] continuous_map_compose [OF _ k, unfolded o_def])+ next fix x show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))" by (auto simp: case_prod_beta h0 k0 h1 k1) qed (auto simp: p q r) qed lemma homotopic_with_product_topology: assumes ht: "\i. i \ I \ homotopic_with (p i) (X i) (Y i) (f i) (g i)" and pq: "\h. (\i. i \ I \ p i (h i)) \ q(\x. (\i\I. h i (x i)))" shows "homotopic_with q (product_topology X I) (product_topology Y I) (\z. (\i\I. (f i) (z i))) (\z. (\i\I. (g i) (z i)))" proof - obtain h where h: "\i. i \ I \ continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)" and h0: "\i x. i \ I \ h i (0, x) = f i x" and h1: "\i x. i \ I \ h i (1, x) = g i x" and p: "\i t. \i \ I; t \ {0..1}\ \ p i (\x. h i (t,x))" using ht unfolding homotopic_with_def by metis show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) let ?h = "\(t,z). \i\I. h i (t,z i)" have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (Y i) (\x. h i (fst x, snd x i))" if "i \ I" for i proof - have \
: "continuous_map (prod_topology (top_of_set {0..1}) (product_topology X I)) (X i) (\x. snd x i)" using continuous_map_componentwise continuous_map_snd that by fastforce show ?thesis unfolding continuous_map_pairwise case_prod_unfold by (intro conjI that \
continuous_intros continuous_map_compose [OF _ h, unfolded o_def]) qed then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (product_topology Y I) ?h" by (auto simp: continuous_map_componentwise case_prod_beta) show "?h (0, x) = (\i\I. f i (x i))" "?h (1, x) = (\i\I. g i (x i))" for x by (auto simp: case_prod_beta h0 h1) show "\t\{0..1}. q (\x. ?h (t, x))" by (force intro: p pq) qed qed text\Homotopic triviality implicitly incorporates path-connectedness.\ lemma homotopic_triviality: shows "(\f g. continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T \ homotopic_with_canon (\x. True) S T f g) \ (S = {} \ path_connected T) \ (\f. continuous_on S f \ f ` S \ T \ (\c. homotopic_with_canon (\x. True) S T f (\x. c)))" (is "?lhs = ?rhs") proof (cases "S = {} \ T = {}") case True then show ?thesis by (auto simp: homotopic_on_emptyI) next case False show ?thesis proof assume LHS [rule_format]: ?lhs have pab: "path_component T a b" if "a \ T" "b \ T" for a b proof - have "homotopic_with_canon (\x. True) S T (\x. a) (\x. b)" by (simp add: LHS image_subset_iff that) then show ?thesis using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] by auto qed moreover have "\c. homotopic_with_canon (\x. True) S T f (\x. c)" if "continuous_on S f" "f ` S \ T" for f using False LHS continuous_on_const that by blast ultimately show ?rhs by (simp add: path_connected_component) next assume RHS: ?rhs with False have T: "path_connected T" by blast show ?lhs proof clarify fix f g assume "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" obtain c d where c: "homotopic_with_canon (\x. True) S T f (\x. c)" and d: "homotopic_with_canon (\x. True) S T g (\x. d)" using False \continuous_on S f\ \f ` S \ T\ RHS \continuous_on S g\ \g ` S \ T\ by blast with T have "path_component T c d" by (metis False ex_in_conv homotopic_with_imp_subset2 image_subset_iff path_connected_component) then have "homotopic_with_canon (\x. True) S T (\x. c) (\x. d)" by (simp add: homotopic_constant_maps) with c d show "homotopic_with_canon (\x. True) S T f g" by (meson homotopic_with_symD homotopic_with_trans) qed qed qed subsection\Homotopy of paths, maintaining the same endpoints\ definition\<^marker>\tag important\ homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" where "homotopic_paths S p q \ homotopic_with_canon (\r. pathstart r = pathstart p \ pathfinish r = pathfinish p) {0..1} S p q" lemma homotopic_paths: "homotopic_paths S p q \ (\h. continuous_on ({0..1} \ {0..1}) h \ h ` ({0..1} \ {0..1}) \ S \ (\x \ {0..1}. h(0,x) = p x) \ (\x \ {0..1}. h(1,x) = q x) \ (\t \ {0..1::real}. pathstart(h \ Pair t) = pathstart p \ pathfinish(h \ Pair t) = pathfinish p))" by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) proposition homotopic_paths_imp_pathstart: "homotopic_paths S p q \ pathstart p = pathstart q" by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) proposition homotopic_paths_imp_pathfinish: "homotopic_paths S p q \ pathfinish p = pathfinish q" by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) lemma homotopic_paths_imp_path: "homotopic_paths S p q \ path p \ path q" using homotopic_paths_def homotopic_with_imp_continuous_maps path_def continuous_map_subtopology_eu by blast lemma homotopic_paths_imp_subset: "homotopic_paths S p q \ path_image p \ S \ path_image q \ S" by (metis (mono_tags) continuous_map_subtopology_eu homotopic_paths_def homotopic_with_imp_continuous_maps path_image_def) proposition homotopic_paths_refl [simp]: "homotopic_paths S p p \ path p \ path_image p \ S" by (simp add: homotopic_paths_def path_def path_image_def) proposition homotopic_paths_sym: "homotopic_paths S p q \ homotopic_paths S q p" by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) proposition homotopic_paths_sym_eq: "homotopic_paths S p q \ homotopic_paths S q p" by (metis homotopic_paths_sym) proposition homotopic_paths_trans [trans]: assumes "homotopic_paths S p q" "homotopic_paths S q r" shows "homotopic_paths S p r" using assms homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart unfolding homotopic_paths_def by (smt (verit, ccfv_SIG) homotopic_with_mono homotopic_with_trans) proposition homotopic_paths_eq: "\path p; path_image p \ S; \t. t \ {0..1} \ p t = q t\ \ homotopic_paths S p q" by (smt (verit, best) homotopic_paths homotopic_paths_refl) proposition homotopic_paths_reparametrize: assumes "path p" and pips: "path_image p \ S" and contf: "continuous_on {0..1} f" and f01:"f ` {0..1} \ {0..1}" and [simp]: "f(0) = 0" "f(1) = 1" and q: "\t. t \ {0..1} \ q(t) = p(f t)" shows "homotopic_paths S p q" proof - have contp: "continuous_on {0..1} p" by (metis \path p\ path_def) then have "continuous_on {0..1} (p \ f)" using contf continuous_on_compose continuous_on_subset f01 by blast then have "path q" by (simp add: path_def) (metis q continuous_on_cong) have piqs: "path_image q \ S" by (metis (no_types, opaque_lifting) pips f01 image_subset_iff path_image_def q) have fb0: "\a b. \0 \ a; a \ 1; 0 \ b; b \ 1\ \ 0 \ (1 - a) * f b + a * b" using f01 by force have fb1: "\0 \ a; a \ 1; 0 \ b; b \ 1\ \ (1 - a) * f b + a * b \ 1" for a b using f01 [THEN subsetD, of "f b"] by (simp add: convex_bound_le) have "homotopic_paths S q p" proof (rule homotopic_paths_trans) show "homotopic_paths S q (p \ f)" using q by (force intro: homotopic_paths_eq [OF \path q\ piqs]) next show "homotopic_paths S (p \ f) p" using pips [unfolded path_image_def] apply (simp add: homotopic_paths_def homotopic_with_def) apply (rule_tac x="p \ (\y. (1 - (fst y)) *\<^sub>R ((f \ snd) y) + (fst y) *\<^sub>R snd y)" in exI) apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ by (auto simp: fb0 fb1 pathstart_def pathfinish_def) qed then show ?thesis by (simp add: homotopic_paths_sym) qed lemma homotopic_paths_subset: "\homotopic_paths S p q; S \ t\ \ homotopic_paths t p q" unfolding homotopic_paths by fast text\ A slightly ad-hoc but useful lemma in constructing homotopies.\ lemma continuous_on_homotopic_join_lemma: fixes q :: "[real,real] \ 'a::topological_space" assumes p: "continuous_on ({0..1} \ {0..1}) (\y. p (fst y) (snd y))" (is "continuous_on ?A ?p") and q: "continuous_on ({0..1} \ {0..1}) (\y. q (fst y) (snd y))" (is "continuous_on ?A ?q") and pf: "\t. t \ {0..1} \ pathfinish(p t) = pathstart(q t)" shows "continuous_on ({0..1} \ {0..1}) (\y. (p(fst y) +++ q(fst y)) (snd y))" proof - have \
: "(\t. p (fst t) (2 * snd t)) = ?p \ (\y. (fst y, 2 * snd y))" "(\t. q (fst t) (2 * snd t - 1)) = ?q \ (\y. (fst y, 2 * snd y - 1))" by force+ show ?thesis unfolding joinpaths_def proof (rule continuous_on_cases_le) show "continuous_on {y \ ?A. snd y \ 1/2} (\t. p (fst t) (2 * snd t))" "continuous_on {y \ ?A. 1/2 \ snd y} (\t. q (fst t) (2 * snd t - 1))" "continuous_on ?A snd" unfolding \
by (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ qed (use pf in \auto simp: mult.commute pathstart_def pathfinish_def\) qed text\ Congruence properties of homotopy w.r.t. path-combining operations.\ lemma homotopic_paths_reversepath_D: assumes "homotopic_paths S p q" shows "homotopic_paths S (reversepath p) (reversepath q)" using assms apply (simp add: homotopic_paths_def homotopic_with_def, clarify) apply (rule_tac x="h \ (\x. (fst x, 1 - snd x))" in exI) apply (rule conjI continuous_intros)+ apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) done proposition homotopic_paths_reversepath: "homotopic_paths S (reversepath p) (reversepath q) \ homotopic_paths S p q" using homotopic_paths_reversepath_D by force proposition homotopic_paths_join: "\homotopic_paths S p p'; homotopic_paths S q q'; pathfinish p = pathstart q\ \ homotopic_paths S (p +++ q) (p' +++ q')" apply (clarsimp simp: homotopic_paths_def homotopic_with_def) apply (rename_tac k1 k2) apply (rule_tac x="(\y. ((k1 \ Pair (fst y)) +++ (k2 \ Pair (fst y))) (snd y))" in exI) apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def) done proposition homotopic_paths_continuous_image: "\homotopic_paths S f g; continuous_on S h; h ` S \ t\ \ homotopic_paths t (h \ f) (h \ g)" unfolding homotopic_paths_def by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose) subsection\Group properties for homotopy of paths\ text\<^marker>\tag important\\So taking equivalence classes under homotopy would give the fundamental group\ proposition homotopic_paths_rid: assumes "path p" "path_image p \ S" shows "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) p" proof - have \
: "continuous_on {0..1} (\t::real. if t \ 1/2 then 2 *\<^sub>R t else 1)" unfolding split_01 by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ show ?thesis apply (rule homotopic_paths_sym) using assms unfolding pathfinish_def joinpaths_def by (intro \
continuous_on_cases continuous_intros homotopic_paths_reparametrize [where f = "\t. if t \ 1/2 then 2 *\<^sub>R t else 1"]; force) qed proposition homotopic_paths_lid: "\path p; path_image p \ S\ \ homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p) p" using homotopic_paths_rid [of "reversepath p" S] by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath pathfinish_reversepath reversepath_joinpaths reversepath_linepath) proposition homotopic_paths_assoc: "\path p; path_image p \ S; path q; path_image q \ S; path r; path_image r \ S; pathfinish p = pathstart q; pathfinish q = pathstart r\ \ homotopic_paths S (p +++ (q +++ r)) ((p +++ q) +++ r)" apply (subst homotopic_paths_sym) apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1/2 then inverse 2 *\<^sub>R t else if t \ 3 / 4 then t - (1 / 4) else 2 *\<^sub>R t - 1"]) apply (simp_all del: le_divide_eq_numeral1 add: subset_path_image_join) apply (rule continuous_on_cases_1 continuous_intros | auto simp: joinpaths_def)+ done proposition homotopic_paths_rinv: assumes "path p" "path_image p \ S" shows "homotopic_paths S (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" proof - have p: "continuous_on {0..1} p" using assms by (auto simp: path_def) let ?A = "{0..1} \ {0..1}" have "continuous_on ?A (\x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" unfolding joinpaths_def subpath_def reversepath_def path_def add_0_right diff_0_right proof (rule continuous_on_cases_le) show "continuous_on {x \ ?A. snd x \ 1/2} (\t. p (fst t * (2 * snd t)))" "continuous_on {x \ ?A. 1/2 \ snd x} (\t. p (fst t * (1 - (2 * snd t - 1))))" "continuous_on ?A snd" by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp: mult_le_one)+ qed (auto simp: algebra_simps) then show ?thesis using assms apply (subst homotopic_paths_sym_eq) unfolding homotopic_paths_def homotopic_with_def apply (rule_tac x="(\y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) apply (force simp: mult_le_one path_defs joinpaths_def subpath_def reversepath_def) done qed proposition homotopic_paths_linv: assumes "path p" "path_image p \ S" shows "homotopic_paths S (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" using homotopic_paths_rinv [of "reversepath p" S] assms by simp subsection\Homotopy of loops without requiring preservation of endpoints\ definition\<^marker>\tag important\ homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where "homotopic_loops S p q \ homotopic_with_canon (\r. pathfinish r = pathstart r) {0..1} S p q" lemma homotopic_loops: "homotopic_loops S p q \ (\h. continuous_on ({0..1::real} \ {0..1}) h \ image h ({0..1} \ {0..1}) \ S \ (\x \ {0..1}. h(0,x) = p x) \ (\x \ {0..1}. h(1,x) = q x) \ (\t \ {0..1}. pathfinish(h \ Pair t) = pathstart(h \ Pair t)))" by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) proposition homotopic_loops_imp_loop: "homotopic_loops S p q \ pathfinish p = pathstart p \ pathfinish q = pathstart q" using homotopic_with_imp_property homotopic_loops_def by blast proposition homotopic_loops_imp_path: "homotopic_loops S p q \ path p \ path q" unfolding homotopic_loops_def path_def using homotopic_with_imp_continuous_maps continuous_map_subtopology_eu by blast proposition homotopic_loops_imp_subset: "homotopic_loops S p q \ path_image p \ S \ path_image q \ S" unfolding homotopic_loops_def path_image_def by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) proposition homotopic_loops_refl: "homotopic_loops S p p \ path p \ path_image p \ S \ pathfinish p = pathstart p" by (simp add: homotopic_loops_def path_image_def path_def) proposition homotopic_loops_sym: "homotopic_loops S p q \ homotopic_loops S q p" by (simp add: homotopic_loops_def homotopic_with_sym) proposition homotopic_loops_sym_eq: "homotopic_loops S p q \ homotopic_loops S q p" by (metis homotopic_loops_sym) proposition homotopic_loops_trans: "\homotopic_loops S p q; homotopic_loops S q r\ \ homotopic_loops S p r" unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) proposition homotopic_loops_subset: "\homotopic_loops S p q; S \ t\ \ homotopic_loops t p q" by (fastforce simp add: homotopic_loops) proposition homotopic_loops_eq: "\path p; path_image p \ S; pathfinish p = pathstart p; \t. t \ {0..1} \ p(t) = q(t)\ \ homotopic_loops S p q" unfolding homotopic_loops_def path_image_def path_def pathstart_def pathfinish_def by (auto intro: homotopic_with_eq [OF homotopic_with_refl [where f = p, THEN iffD2]]) proposition homotopic_loops_continuous_image: "\homotopic_loops S f g; continuous_on S h; h ` S \ t\ \ homotopic_loops t (h \ f) (h \ g)" unfolding homotopic_loops_def by (simp add: homotopic_with_compose_continuous_map_left pathfinish_def pathstart_def) subsection\Relations between the two variants of homotopy\ proposition homotopic_paths_imp_homotopic_loops: "\homotopic_paths S p q; pathfinish p = pathstart p; pathfinish q = pathstart p\ \ homotopic_loops S p q" by (auto simp: homotopic_with_def homotopic_paths_def homotopic_loops_def) proposition homotopic_loops_imp_homotopic_paths_null: assumes "homotopic_loops S p (linepath a a)" shows "homotopic_paths S p (linepath (pathstart p) (pathstart p))" proof - have "path p" by (metis assms homotopic_loops_imp_path) have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) have pip: "path_image p \ S" by (metis assms homotopic_loops_imp_subset) let ?A = "{0..1::real} \ {0..1::real}" obtain h where conth: "continuous_on ?A h" and hs: "h ` ?A \ S" and h0[simp]: "\x. x \ {0..1} \ h(0,x) = p x" and h1[simp]: "\x. x \ {0..1} \ h(1,x) = a" and ends: "\t. t \ {0..1} \ pathfinish (h \ Pair t) = pathstart (h \ Pair t)" using assms by (auto simp: homotopic_loops homotopic_with) have conth0: "path (\u. h (u, 0))" unfolding path_def proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) show "continuous_on ((\x. (x, 0)) ` {0..1}) h" by (force intro: continuous_on_subset [OF conth]) qed (force intro: continuous_intros) have pih0: "path_image (\u. h (u, 0)) \ S" using hs by (force simp: path_image_def) have c1: "continuous_on ?A (\x. h (fst x * snd x, 0))" proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) show "continuous_on ((\x. (fst x * snd x, 0)) ` ?A) h" by (force simp: mult_le_one intro: continuous_on_subset [OF conth]) qed (force intro: continuous_intros)+ have c2: "continuous_on ?A (\x. h (fst x - fst x * snd x, 0))" proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) show "continuous_on ((\x. (fst x - fst x * snd x, 0)) ` ?A) h" by (auto simp: algebra_simps add_increasing2 mult_left_le intro: continuous_on_subset [OF conth]) qed (force intro: continuous_intros) have [simp]: "\t. \0 \ t \ t \ 1\ \ h (t, 1) = h (t, 0)" using ends by (simp add: pathfinish_def pathstart_def) have adhoc_le: "c * 4 \ 1 + c * (d * 4)" if "\ d * 4 \ 3" "0 \ c" "c \ 1" for c d::real proof - have "c * 3 \ c * (d * 4)" using that less_eq_real_def by auto with \c \ 1\ show ?thesis by fastforce qed have *: "\p x. \path p \ path(reversepath p); path_image p \ S \ path_image(reversepath p) \ S; pathfinish p = pathstart(linepath a a +++ reversepath p) \ pathstart(reversepath p) = a \ pathstart p = x\ \ homotopic_paths S (p +++ linepath a a +++ reversepath p) (linepath x x)" by (metis homotopic_paths_lid homotopic_paths_join homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) have 1: "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" using \path p\ homotopic_paths_rid homotopic_paths_sym pip by blast moreover have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" S] by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_sym pathstart_join) moreover have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0)))" unfolding homotopic_paths_def homotopic_with_def proof (intro exI strip conjI) let ?h = "\y. (subpath 0 (fst y) (\u. h (u, 0)) +++ (\u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\u. h (u, 0))) (snd y)" have "continuous_on ?A ?h" by (intro continuous_on_homotopic_join_lemma; simp add: path_defs joinpaths_def subpath_def conth c1 c2) moreover have "?h ` ?A \ S" unfolding joinpaths_def subpath_def by (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le) ultimately show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set S) ?h" by (simp add: subpath_reversepath) qed (use ploop in \simp_all add: reversepath_def path_defs joinpaths_def o_def subpath_def conth c1 c2\) moreover have "homotopic_paths S ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0))) (linepath (pathstart p) (pathstart p))" by (rule *; simp add: pih0 pathstart_def pathfinish_def conth0; simp add: reversepath_def joinpaths_def) ultimately show ?thesis by (blast intro: homotopic_paths_trans) qed proposition homotopic_loops_conjugate: fixes S :: "'a::real_normed_vector set" assumes "path p" "path q" and pip: "path_image p \ S" and piq: "path_image q \ S" and pq: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" shows "homotopic_loops S (p +++ q +++ reversepath p) q" proof - have contp: "continuous_on {0..1} p" using \path p\ [unfolded path_def] by blast have contq: "continuous_on {0..1} q" using \path q\ [unfolded path_def] by blast let ?A = "{0..1::real} \ {0..1::real}" have c1: "continuous_on ?A (\x. p ((1 - fst x) * snd x + fst x))" proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) show "continuous_on ((\x. (1 - fst x) * snd x + fst x) ` ?A) p" by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) qed (force intro: continuous_intros) have c2: "continuous_on ?A (\x. p ((fst x - 1) * snd x + 1))" proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) show "continuous_on ((\x. (fst x - 1) * snd x + 1) ` ?A) p" by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_left_le_one_le) qed (force intro: continuous_intros) have ps1: "\a b. \b * 2 \ 1; 0 \ b; 0 \ a; a \ 1\ \ p ((1 - a) * (2 * b) + a) \ S" using sum_le_prod1 by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) have ps2: "\a b. \\ 4 * b \ 3; b \ 1; 0 \ a; a \ 1\ \ p ((a - 1) * (4 * b - 3) + 1) \ S" apply (rule pip [unfolded path_image_def, THEN subsetD]) apply (rule image_eqI, blast) apply (simp add: algebra_simps) by (metis add_mono affine_ineq linear mult.commute mult.left_neutral mult_right_mono add.commute zero_le_numeral) have qs: "\a b. \4 * b \ 3; \ b * 2 \ 1\ \ q (4 * b - 2) \ S" using path_image_def piq by fastforce have "homotopic_loops S (p +++ q +++ reversepath p) (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" unfolding homotopic_loops_def homotopic_with_def proof (intro exI strip conjI) let ?h = "(\y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" have "continuous_on ?A (\y. q (snd y))" by (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) then have "continuous_on ?A ?h" using pq qloop by (intro continuous_on_homotopic_join_lemma) (auto simp: path_defs joinpaths_def subpath_def c1 c2) then show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set S) ?h" by (auto simp: joinpaths_def subpath_def ps1 ps2 qs) show "?h (1,x) = (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) x" for x using pq by (simp add: pathfinish_def subpath_refl) qed (auto simp: subpath_reversepath) moreover have "homotopic_loops S (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" proof - have "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q) q" using \path q\ homotopic_paths_lid qloop piq by auto hence 1: "\f. homotopic_paths S f q \ \ homotopic_paths S f (linepath (pathfinish q) (pathfinish q) +++ q)" using homotopic_paths_trans by blast hence "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" by (smt (verit, best) \path q\ homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_lid homotopic_paths_rid homotopic_paths_trans pathstart_join piq qloop) thus ?thesis by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) qed ultimately show ?thesis by (blast intro: homotopic_loops_trans) qed lemma homotopic_paths_loop_parts: assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" shows "homotopic_paths S p q" proof - have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp then have "path p" using \path q\ homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast show ?thesis proof (cases "pathfinish p = pathfinish q") case True obtain pipq: "path_image p \ S" "path_image q \ S" by (metis Un_subset_iff paths \path p\ \path q\ homotopic_loops_imp_subset homotopic_paths_imp_path loops path_image_join path_image_reversepath path_imp_reversepath path_join_eq) have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" using \path p\ \path_image p \ S\ homotopic_paths_rid homotopic_paths_sym by blast moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" by (simp add: True \path p\ \path q\ pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" by (simp add: True \path p\ \path q\ homotopic_paths_assoc pipq) moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" by (simp add: \path q\ homotopic_paths_join paths pipq) ultimately show ?thesis by (metis \path q\ homotopic_paths_imp_path homotopic_paths_lid homotopic_paths_trans path_join_path_ends pathfinish_linepath pipq(2)) next case False then show ?thesis using \path q\ homotopic_loops_imp_path loops path_join_path_ends by fastforce qed qed subsection\<^marker>\tag unimportant\\Homotopy of "nearby" function, paths and loops\ lemma homotopic_with_linear: fixes f g :: "_ \ 'b::real_normed_vector" assumes contf: "continuous_on S f" and contg:"continuous_on S g" and sub: "\x. x \ S \ closed_segment (f x) (g x) \ t" shows "homotopic_with_canon (\z. True) S t f g" unfolding homotopic_with_def apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) using sub closed_segment_def by (fastforce intro: continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]) lemma homotopic_paths_linear: fixes g h :: "real \ 'a::real_normed_vector" assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" "\t. t \ {0..1} \ closed_segment (g t) (h t) \ S" shows "homotopic_paths S g h" using assms unfolding path_def apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R (g \ snd) y + (fst y) *\<^sub>R (h \ snd) y)" in exI) apply (intro conjI subsetI continuous_intros; force) done lemma homotopic_loops_linear: fixes g h :: "real \ 'a::real_normed_vector" assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" "\t x. t \ {0..1} \ closed_segment (g t) (h t) \ S" shows "homotopic_loops S g h" using assms unfolding path_defs homotopic_loops_def homotopic_with_def apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) by (force simp: closed_segment_def intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) lemma homotopic_paths_nearby_explicit: assumes \
: "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" and no: "\t x. \t \ {0..1}; x \ S\ \ norm(h t - g t) < norm(g t - x)" shows "homotopic_paths S g h" using homotopic_paths_linear [OF \
] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) lemma homotopic_loops_nearby_explicit: assumes \
: "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" and no: "\t x. \t \ {0..1}; x \ S\ \ norm(h t - g t) < norm(g t - x)" shows "homotopic_loops S g h" using homotopic_loops_linear [OF \
] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) lemma homotopic_nearby_paths: fixes g h :: "real \ 'a::euclidean_space" assumes "path g" "open S" "path_image g \ S" shows "\e. 0 < e \ (\h. path h \ pathstart h = pathstart g \ pathfinish h = pathfinish g \ (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_paths S g h)" proof - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - S \ e \ dist x y" using separate_compact_closed [of "path_image g" "-S"] assms by force show ?thesis using e [unfolded dist_norm] \e > 0\ by (fastforce simp: path_image_def intro!: homotopic_paths_nearby_explicit assms exI) qed lemma homotopic_nearby_loops: fixes g h :: "real \ 'a::euclidean_space" assumes "path g" "open S" "path_image g \ S" "pathfinish g = pathstart g" shows "\e. 0 < e \ (\h. path h \ pathfinish h = pathstart h \ (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_loops S g h)" proof - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - S \ e \ dist x y" using separate_compact_closed [of "path_image g" "-S"] assms by force show ?thesis using e [unfolded dist_norm] \e > 0\ by (fastforce simp: path_image_def intro!: homotopic_loops_nearby_explicit assms exI) qed subsection\ Homotopy and subpaths\ lemma homotopic_join_subpaths1: assumes "path g" and pag: "path_image g \ S" and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" "u \ v" "v \ w" shows "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" proof - have 1: "t * 2 \ 1 \ u + t * (v * 2) \ v + t * (u * 2)" for t using affine_ineq \u \ v\ by fastforce have 2: "t * 2 > 1 \ u + (2*t - 1) * v \ v + (2*t - 1) * w" for t by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \u \ v\ \v \ w\) have t2: "\t::real. t*2 = 1 \ t = 1/2" by auto have "homotopic_paths (path_image g) (subpath u v g +++ subpath v w g) (subpath u w g)" proof (cases "w = u") case True then show ?thesis by (metis \path g\ homotopic_paths_rinv path_image_subpath_subset path_subpath pathstart_subpath reversepath_subpath subpath_refl u v) next case False let ?f = "\t. if t \ 1/2 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))" show ?thesis proof (rule homotopic_paths_sym [OF homotopic_paths_reparametrize [where f = ?f]]) show "path (subpath u w g)" using assms(1) path_subpath u w(1) by blast show "path_image (subpath u w g) \ path_image g" by (meson path_image_subpath_subset u w(1)) show "continuous_on {0..1} ?f" unfolding split_01 by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def dest!: t2)+ show "?f ` {0..1} \ {0..1}" using False assms by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2) show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \ {0..1}" for t using assms unfolding joinpaths_def subpath_def by (auto simp: divide_simps add.commute mult.commute mult.left_commute) qed (use False in auto) qed then show ?thesis by (rule homotopic_paths_subset [OF _ pag]) qed lemma homotopic_join_subpaths2: assumes "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" shows "homotopic_paths S (subpath w v g +++ subpath v u g) (subpath w u g)" by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) lemma homotopic_join_subpaths3: assumes hom: "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" and "path g" and pag: "path_image g \ S" and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" shows "homotopic_paths S (subpath v w g +++ subpath w u g) (subpath v u g)" proof - obtain wvg: "path (subpath w v g)" "path_image (subpath w v g) \ S" and wug: "path (subpath w u g)" "path_image (subpath w u g) \ S" and vug: "path (subpath v u g)" "path_image (subpath v u g) \ S" by (meson \path g\ pag path_image_subpath_subset path_subpath subset_trans u v w) have "homotopic_paths S (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" by (simp add: hom homotopic_paths_join homotopic_paths_sym wvg) also have "homotopic_paths S ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)" using wvg vug \path g\ by (metis homotopic_paths_assoc homotopic_paths_sym path_image_subpath_commute path_subpath pathfinish_subpath pathstart_subpath u v w) also have "homotopic_paths S (subpath u v g +++ subpath v w g +++ subpath w v g) (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" using wvg vug \path g\ by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl path_image_subpath_commute path_subpath pathfinish_subpath pathstart_join pathstart_subpath reversepath_subpath u v) also have "homotopic_paths S (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" using vug \path g\ by (metis homotopic_paths_rid path_image_subpath_commute path_subpath u v) finally have "homotopic_paths S (subpath u w g +++ subpath w v g) (subpath u v g)" . then show ?thesis using homotopic_join_subpaths2 by blast qed proposition homotopic_join_subpaths: "\path g; path_image g \ S; u \ {0..1}; v \ {0..1}; w \ {0..1}\ \ homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" by (smt (verit, del_insts) homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3) text\Relating homotopy of trivial loops to path-connectedness.\ lemma path_component_imp_homotopic_points: assumes "path_component S a b" shows "homotopic_loops S (linepath a a) (linepath b b)" proof - obtain g :: "real \ 'a" where g: "continuous_on {0..1} g" "g ` {0..1} \ S" "g 0 = a" "g 1 = b" using assms by (auto simp: path_defs) then have "continuous_on ({0..1} \ {0..1}) (g \ fst)" by (fastforce intro!: continuous_intros)+ with g show ?thesis by (auto simp: homotopic_loops_def homotopic_with_def path_defs image_subset_iff) qed lemma homotopic_loops_imp_path_component_value: "\homotopic_loops S p q; 0 \ t; t \ 1\ \ path_component S (p t) (q t)" apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) apply (rule_tac x="h \ (\u. (u, t))" in exI) apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) done lemma homotopic_points_eq_path_component: "homotopic_loops S (linepath a a) (linepath b b) \ path_component S a b" using homotopic_loops_imp_path_component_value path_component_imp_homotopic_points by fastforce lemma path_connected_eq_homotopic_points: "path_connected S \ (\a b. a \ S \ b \ S \ homotopic_loops S (linepath a a) (linepath b b))" by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) subsection\Simply connected sets\ text\<^marker>\tag important\\defined as "all loops are homotopic (as loops)\ definition\<^marker>\tag important\ simply_connected where "simply_connected S \ \p q. path p \ pathfinish p = pathstart p \ path_image p \ S \ path q \ pathfinish q = pathstart q \ path_image q \ S \ homotopic_loops S p q" lemma simply_connected_empty [iff]: "simply_connected {}" by (simp add: simply_connected_def) lemma simply_connected_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S" by (simp add: simply_connected_def path_connected_eq_homotopic_points) lemma simply_connected_imp_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ connected S" by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) lemma simply_connected_eq_contractible_loop_any: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ (\p a. path p \ path_image p \ S \ pathfinish p = pathstart p \ a \ S \ homotopic_loops S p (linepath a a))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding simply_connected_def by force next assume ?rhs then show ?lhs unfolding simply_connected_def by (metis pathfinish_in_path_image subsetD homotopic_loops_trans homotopic_loops_sym) qed lemma simply_connected_eq_contractible_loop_some: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ (\a. a \ S \ homotopic_loops S p (linepath a a)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected) next assume ?rhs then show ?lhs by (meson homotopic_loops_trans path_connected_eq_homotopic_points simply_connected_eq_contractible_loop_any) qed lemma simply_connected_eq_contractible_loop_all: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ S = {} \ (\a \ S. \p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a))" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by force next case False then obtain a where "a \ S" by blast then show ?thesis by (meson False homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) qed lemma simply_connected_eq_contractible_path: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_paths S p (linepath (pathstart p) (pathstart p)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding simply_connected_imp_path_connected by (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) next assume ?rhs then show ?lhs using homotopic_paths_imp_homotopic_loops simply_connected_eq_contractible_loop_some by fastforce qed lemma simply_connected_eq_homotopic_paths: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p q. path p \ path_image p \ S \ path q \ path_image q \ S \ pathstart q = pathstart p \ pathfinish q = pathfinish p \ homotopic_paths S p q)" (is "?lhs = ?rhs") proof assume ?lhs then have pc: "path_connected S" and *: "\p. \path p; path_image p \ S; pathfinish p = pathstart p\ \ homotopic_paths S p (linepath (pathstart p) (pathstart p))" by (auto simp: simply_connected_eq_contractible_path) have "homotopic_paths S p q" if "path p" "path_image p \ S" "path q" "path_image q \ S" "pathstart q = pathstart p" "pathfinish q = pathfinish p" for p q proof - have "homotopic_paths S p (p +++ reversepath q +++ q)" using that by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym homotopic_paths_trans pathstart_linepath) also have "homotopic_paths S (p +++ reversepath q +++ q) ((p +++ reversepath q) +++ q)" by (simp add: that homotopic_paths_assoc) also have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart q) (pathstart q) +++ q)" using * [of "p +++ reversepath q"] that by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join) also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" using that homotopic_paths_lid by blast finally show ?thesis . qed then show ?rhs by (blast intro: pc *) next assume ?rhs then show ?lhs by (force simp: simply_connected_eq_contractible_path) qed proposition simply_connected_Times: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes S: "simply_connected S" and T: "simply_connected T" shows "simply_connected(S \ T)" proof - have "homotopic_loops (S \ T) p (linepath (a, b) (a, b))" if "path p" "path_image p \ S \ T" "p 1 = p 0" "a \ S" "b \ T" for p a b proof - have "path (fst \ p)" by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \path p\]) moreover have "path_image (fst \ p) \ S" using that by (force simp add: path_image_def) ultimately have p1: "homotopic_loops S (fst \ p) (linepath a a)" using S that by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) have "path (snd \ p)" by (simp add: continuous_on_snd Path_Connected.path_continuous_image [OF \path p\]) moreover have "path_image (snd \ p) \ T" using that by (force simp: path_image_def) ultimately have p2: "homotopic_loops T (snd \ p) (linepath b b)" using T that by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) show ?thesis using p1 p2 unfolding homotopic_loops apply clarify subgoal for h k by (rule_tac x="\z. (h z, k z)" in exI) (force intro: continuous_intros simp: path_defs) done qed with assms show ?thesis by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) qed subsection\Contractible sets\ definition\<^marker>\tag important\ contractible where "contractible S \ \a. homotopic_with_canon (\x. True) S S id (\x. a)" proposition contractible_imp_simply_connected: fixes S :: "_::real_normed_vector set" assumes "contractible S" shows "simply_connected S" proof (cases "S = {}") case True then show ?thesis by force next case False obtain a where a: "homotopic_with_canon (\x. True) S S id (\x. a)" using assms by (force simp: contractible_def) then have "a \ S" by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_in_topspace topspace_euclidean_subtopology) have "\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a)" using a apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) apply (rule_tac x="(h \ (\y. (fst y, (p \ snd) y)))" in exI) apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset) done with \a \ S\ show ?thesis by (auto simp: simply_connected_eq_contractible_loop_all False) qed corollary contractible_imp_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ connected S" by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) lemma contractible_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ path_connected S" by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) lemma nullhomotopic_through_contractible: fixes S :: "_::topological_space set" assumes f: "continuous_on S f" "f ` S \ T" and g: "continuous_on T g" "g ` T \ U" and T: "contractible T" obtains c where "homotopic_with_canon (\h. True) S U (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with_canon (\x. True) T T id (\x. b)" using assms by (force simp: contractible_def) have "homotopic_with_canon (\f. True) T U (g \ id) (g \ (\x. b))" by (metis Abstract_Topology.continuous_map_subtopology_eu b g homotopic_with_compose_continuous_map_left) then have "homotopic_with_canon (\f. True) S U (g \ id \ f) (g \ (\x. b) \ f)" by (simp add: f homotopic_with_compose_continuous_map_right) then show ?thesis by (simp add: comp_def that) qed lemma nullhomotopic_into_contractible: assumes f: "continuous_on S f" "f ` S \ T" and T: "contractible T" obtains c where "homotopic_with_canon (\h. True) S T f (\x. c)" by (rule nullhomotopic_through_contractible [OF f, of id T]) (use assms in auto) lemma nullhomotopic_from_contractible: assumes f: "continuous_on S f" "f ` S \ T" and S: "contractible S" obtains c where "homotopic_with_canon (\h. True) S T f (\x. c)" by (auto simp: comp_def intro: nullhomotopic_through_contractible [OF continuous_on_id _ f S]) lemma homotopic_through_contractible: fixes S :: "_::real_normed_vector set" assumes "continuous_on S f1" "f1 ` S \ T" "continuous_on T g1" "g1 ` T \ U" "continuous_on S f2" "f2 ` S \ T" "continuous_on T g2" "g2 ` T \ U" "contractible T" "path_connected U" shows "homotopic_with_canon (\h. True) S U (g1 \ f1) (g2 \ f2)" proof - obtain c1 where c1: "homotopic_with_canon (\h. True) S U (g1 \ f1) (\x. c1)" by (rule nullhomotopic_through_contractible [of S f1 T g1 U]) (use assms in auto) obtain c2 where c2: "homotopic_with_canon (\h. True) S U (g2 \ f2) (\x. c2)" by (rule nullhomotopic_through_contractible [of S f2 T g2 U]) (use assms in auto) have "S = {} \ (\t. path_connected t \ t \ U \ c2 \ t \ c1 \ t)" proof (cases "S = {}") case True then show ?thesis by force next case False with c1 c2 have "c1 \ U" "c2 \ U" using homotopic_with_imp_continuous_maps by fastforce+ with \path_connected U\ show ?thesis by blast qed then have "homotopic_with_canon (\h. True) S U (\x. c2) (\x. c1)" by (simp add: path_component homotopic_constant_maps) then show ?thesis using c1 c2 homotopic_with_symD homotopic_with_trans by blast qed lemma homotopic_into_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f ` S \ T" and g: "continuous_on S g" "g ` S \ T" and T: "contractible T" shows "homotopic_with_canon (\h. True) S T f g" using homotopic_through_contractible [of S f T id T g id] by (simp add: assms contractible_imp_path_connected) lemma homotopic_from_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f ` S \ T" and g: "continuous_on S g" "g ` S \ T" and "contractible S" "path_connected T" shows "homotopic_with_canon (\h. True) S T f g" using homotopic_through_contractible [of S id S f T id g] by (simp add: assms contractible_imp_path_connected) subsection\Starlike sets\ definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" lemma starlike_UNIV [simp]: "starlike UNIV" by (simp add: starlike_def) lemma convex_imp_starlike: "convex S \ S \ {} \ starlike S" unfolding convex_contains_segment starlike_def by auto lemma starlike_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" shows "starlike T" proof - have "rel_interior S \ {}" by (simp add: assms rel_interior_eq_empty) with ST obtain a where a: "a \ rel_interior S" and "a \ T" by blast have "\x. x \ T \ open_segment a x \ rel_interior S" by (rule rel_interior_closure_convex_segment [OF \convex S\ a]) (use assms in auto) then have "\x\T. a \ T \ open_segment a x \ T" using ST by (blast intro: a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) then show ?thesis unfolding starlike_def using bexI [OF _ \a \ T\] by (simp add: closed_segment_eq_open) qed lemma starlike_imp_contractible_gen: fixes S :: "'a::real_normed_vector set" assumes S: "starlike S" and P: "\a T. \a \ S; 0 \ T; T \ 1\ \ P(\x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" obtains a where "homotopic_with_canon P S S (\x. x) (\x. a)" proof - obtain a where "a \ S" and a: "\x. x \ S \ closed_segment a x \ S" using S by (auto simp: starlike_def) have "\t b. 0 \ t \ t \ 1 \ \u. (1 - t) *\<^sub>R b + t *\<^sub>R a = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" by (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) then have "(\y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \ S) \ S" using a [unfolded closed_segment_def] by force then have "homotopic_with_canon P S S (\x. x) (\x. a)" using \a \ S\ unfolding homotopic_with_def apply (rule_tac x="\y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) apply (force simp add: P intro: continuous_intros) done then show ?thesis using that by blast qed lemma starlike_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "starlike S \ contractible S" using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" by (simp add: starlike_imp_contractible) lemma starlike_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ simply_connected S" by (simp add: contractible_imp_simply_connected starlike_imp_contractible) lemma convex_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "convex S \ simply_connected S" using convex_imp_starlike starlike_imp_simply_connected by blast lemma starlike_imp_path_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ path_connected S" by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) lemma starlike_imp_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ connected S" by (simp add: path_connected_imp_connected starlike_imp_path_connected) lemma is_interval_simply_connected_1: fixes S :: "real set" shows "is_interval S \ simply_connected S" by (meson convex_imp_simply_connected is_interval_connected_1 is_interval_convex_1 simply_connected_imp_connected) lemma contractible_empty [simp]: "contractible {}" by (simp add: contractible_def homotopic_on_emptyI) lemma contractible_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" and TS: "rel_interior S \ T" "T \ closure S" shows "contractible T" by (metis assms closure_eq_empty contractible_empty empty_subsetI starlike_convex_tweak_boundary_points starlike_imp_contractible subset_antisym) lemma convex_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "convex S \ contractible S" using contractible_empty convex_imp_starlike starlike_imp_contractible by blast lemma contractible_sing [simp]: fixes a :: "'a::real_normed_vector" shows "contractible {a}" by (rule convex_imp_contractible [OF convex_singleton]) lemma is_interval_contractible_1: fixes S :: "real set" shows "is_interval S \ contractible S" using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 is_interval_simply_connected_1 by auto lemma contractible_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "contractible S" and T: "contractible T" shows "contractible (S \ T)" proof - obtain a h where conth: "continuous_on ({0..1} \ S) h" and hsub: "h ` ({0..1} \ S) \ S" and [simp]: "\x. x \ S \ h (0, x) = x" and [simp]: "\x. x \ S \ h (1::real, x) = a" using S by (auto simp: contractible_def homotopic_with) obtain b k where contk: "continuous_on ({0..1} \ T) k" and ksub: "k ` ({0..1} \ T) \ T" and [simp]: "\x. x \ T \ k (0, x) = x" and [simp]: "\x. x \ T \ k (1::real, x) = b" using T by (auto simp: contractible_def homotopic_with) show ?thesis apply (simp add: contractible_def homotopic_with) apply (rule exI [where x=a]) apply (rule exI [where x=b]) apply (rule exI [where x = "\z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) using hsub ksub apply (fastforce intro!: continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) done qed subsection\Local versions of topological properties in general\ definition\<^marker>\tag important\ locally :: "('a::topological_space set \ bool) \ 'a set \ bool" where "locally P S \ \w x. openin (top_of_set S) w \ x \ w \ (\u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w)" lemma locallyI: assumes "\w x. \openin (top_of_set S) w; x \ w\ \ \u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w" shows "locally P S" using assms by (force simp: locally_def) lemma locallyE: assumes "locally P S" "openin (top_of_set S) w" "x \ w" obtains u v where "openin (top_of_set S) u" "P v" "x \ u" "u \ v" "v \ w" using assms unfolding locally_def by meson lemma locally_mono: assumes "locally P S" "\T. P T \ Q T" shows "locally Q S" by (metis assms locally_def) lemma locally_open_subset: assumes "locally P S" "openin (top_of_set S) t" shows "locally P t" by (smt (verit, ccfv_SIG) assms order.trans locally_def openin_imp_subset openin_subset_trans openin_trans) lemma locally_diff_closed: "\locally P S; closedin (top_of_set S) t\ \ locally P (S - t)" using locally_open_subset closedin_def by fastforce lemma locally_empty [iff]: "locally P {}" by (simp add: locally_def openin_subtopology) lemma locally_singleton [iff]: fixes a :: "'a::metric_space" shows "locally P {a} \ P {a}" proof - have "\x::real. \ 0 < x \ P {a}" using zero_less_one by blast then show ?thesis unfolding locally_def by (auto simp: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR) qed lemma locally_iff: "locally P S \ (\T x. open T \ x \ S \ T \ (\U. open U \ (\V. P V \ x \ S \ U \ S \ U \ V \ V \ S \ T)))" by (smt (verit) locally_def openin_open) lemma locally_Int: assumes S: "locally P S" and T: "locally P T" and P: "\S T. P S \ P T \ P(S \ T)" shows "locally P (S \ T)" unfolding locally_iff proof clarify fix A x assume "open A" "x \ A" "x \ S" "x \ T" then obtain U1 V1 U2 V2 where "open U1" "P V1" "x \ S \ U1" "S \ U1 \ V1 \ V1 \ S \ A" "open U2" "P V2" "x \ T \ U2" "T \ U2 \ V2 \ V2 \ T \ A" using S T unfolding locally_iff by (meson IntI) then have "S \ T \ (U1 \ U2) \ V1 \ V2" "V1 \ V2 \ S \ T \ A" "x \ S \ T \ (U1 \ U2)" by blast+ moreover have "P (V1 \ V2)" by (simp add: P \P V1\ \P V2\) ultimately show "\U. open U \ (\V. P V \ x \ S \ T \ U \ S \ T \ U \ V \ V \ S \ T \ A)" using \open U1\ \open U2\ by blast qed lemma locally_Times: fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" assumes PS: "locally P S" and QT: "locally Q T" and R: "\S T. P S \ Q T \ R(S \ T)" shows "locally R (S \ T)" unfolding locally_def proof (clarify) fix W x y assume W: "openin (top_of_set (S \ T)) W" and xy: "(x, y) \ W" then obtain U V where "openin (top_of_set S) U" "x \ U" "openin (top_of_set T) V" "y \ V" "U \ V \ W" using Times_in_interior_subtopology by metis then obtain U1 U2 V1 V2 where opeS: "openin (top_of_set S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" and opeT: "openin (top_of_set T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" by (meson PS QT locallyE) then have "openin (top_of_set (S \ T)) (U1 \ V1)" by (simp add: openin_Times) moreover have "R (U2 \ V2)" by (simp add: R opeS opeT) moreover have "U1 \ V1 \ U2 \ V2 \ U2 \ V2 \ W" using opeS opeT \U \ V \ W\ by auto ultimately show "\U V. openin (top_of_set (S \ T)) U \ R V \ (x,y) \ U \ U \ V \ V \ W" using opeS opeT by auto qed proposition homeomorphism_locally_imp: fixes S :: "'a::metric_space set" and T :: "'b::t2_space set" assumes S: "locally P S" and hom: "homeomorphism S T f g" and Q: "\S S'. \P S; homeomorphism S S' f g\ \ Q S'" shows "locally Q T" proof (clarsimp simp: locally_def) fix W y assume "y \ W" and "openin (top_of_set T) W" then obtain A where T: "open A" "W = T \ A" by (force simp: openin_open) then have "W \ T" by auto have f: "\x. x \ S \ g(f x) = x" "f ` S = T" "continuous_on S f" and g: "\y. y \ T \ f(g y) = y" "g ` T = S" "continuous_on T g" using hom by (auto simp: homeomorphism_def) have gw: "g ` W = S \ f -` W" using \W \ T\ g by force have "openin (top_of_set S) (g ` W)" using \openin (top_of_set T) W\ continuous_on_open f gw by auto then obtain U V where osu: "openin (top_of_set S) U" and uv: "P V" "g y \ U" "U \ V" "V \ g ` W" by (metis S \y \ W\ image_eqI locallyE) have "V \ S" using uv by (simp add: gw) have fv: "f ` V = T \ {x. g x \ V}" using \f ` S = T\ f \V \ S\ by auto have contvf: "continuous_on V f" using \V \ S\ continuous_on_subset f(3) by blast have "f ` V \ W" using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto then have contvg: "continuous_on (f ` V) g" using \W \ T\ continuous_on_subset [OF g(3)] by blast have "V \ g ` f ` V" by (metis \V \ S\ hom homeomorphism_def homeomorphism_of_subsets order_refl) then have homv: "homeomorphism V (f ` V) f g" using \V \ S\ f by (auto simp: homeomorphism_def contvf contvg) have "openin (top_of_set (g ` T)) U" using \g ` T = S\ by (simp add: osu) then have "openin (top_of_set T) (T \ g -` U)" using \continuous_on T g\ continuous_on_open [THEN iffD1] by blast moreover have "\V. Q V \ y \ (T \ g -` U) \ (T \ g -` U) \ V \ V \ W" proof (intro exI conjI) show "Q (f ` V)" using Q homv \P V\ by blast show "y \ T \ g -` U" using T(2) \y \ W\ \g y \ U\ by blast show "T \ g -` U \ f ` V" using g(1) image_iff uv(3) by fastforce show "f ` V \ W" using \f ` V \ W\ by blast qed ultimately show "\U. openin (top_of_set T) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" by meson qed lemma homeomorphism_locally: fixes f:: "'a::metric_space \ 'b::metric_space" assumes "homeomorphism S T f g" and "\S T. homeomorphism S T f g \ (P S \ Q T)" shows "locally P S \ locally Q T" by (smt (verit) assms homeomorphism_locally_imp homeomorphism_symD) lemma homeomorphic_locally: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" assumes hom: "S homeomorphic T" and iff: "\X Y. X homeomorphic Y \ (P X \ Q Y)" shows "locally P S \ locally Q T" by (smt (verit, ccfv_SIG) hom homeomorphic_def homeomorphism_locally homeomorphism_locally_imp iff) lemma homeomorphic_local_compactness: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" shows "S homeomorphic T \ locally compact S \ locally compact T" by (simp add: homeomorphic_compactness homeomorphic_locally) lemma locally_translation: fixes P :: "'a :: real_normed_vector set \ bool" shows "(\S. P ((+) a ` S) = P S) \ locally P ((+) a ` S) = locally P S" using homeomorphism_locally [OF homeomorphism_translation] by (metis (full_types) homeomorphism_image2) lemma locally_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" and iff: "\S. P (f ` S) \ Q S" shows "locally P (f ` S) \ locally Q S" using homeomorphism_locally [of "f`S" _ _ f] linear_homeomorphism_image [OF f] by (metis (no_types, lifting) homeomorphism_image2 iff) lemma locally_open_map_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes P: "locally P S" and f: "continuous_on S f" and oo: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" and Q: "\T. \T \ S; P T\ \ Q(f ` T)" shows "locally Q (f ` S)" proof (clarsimp simp: locally_def) fix W y assume oiw: "openin (top_of_set (f ` S)) W" and "y \ W" then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) have oivf: "openin (top_of_set S) (S \ f -` W)" by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) then obtain x where "x \ S" "f x = y" using \W \ f ` S\ \y \ W\ by blast then obtain U V where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" by (metis IntI P \y \ W\ locallyE oivf vimageI) then have "openin (top_of_set (f ` S)) (f ` U)" by (simp add: oo) then show "\X. openin (top_of_set (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" using Q \P V\ \U \ V\ \V \ S \ f -` W\ \f x = y\ \x \ U\ by blast qed subsection\An induction principle for connected sets\ proposition connected_induction: assumes "connected S" and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ Q x \ Q y)" and etc: "a \ S" "b \ S" "P a" "P b" "Q a" shows "Q b" proof - let ?A = "{b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ Q x)}" let ?B = "{b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ \ Q x)}" have "?A \ ?B = {}" by (clarsimp simp: set_eq_iff) (metis (no_types, opaque_lifting) Int_iff opD openin_Int) moreover have "S \ ?A \ ?B" by clarsimp (meson opI) moreover have "openin (top_of_set S) ?A" by (subst openin_subopen, blast) moreover have "openin (top_of_set S) ?B" by (subst openin_subopen, blast) ultimately have "?A = {} \ ?B = {}" by (metis (no_types, lifting) \connected S\ connected_openin) then show ?thesis by clarsimp (meson opI etc) qed lemma connected_equivalence_relation_gen: assumes "connected S" and etc: "a \ S" "b \ S" "P a" "P b" and trans: "\x y z. \R x y; R y z\ \ R x z" and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ R x y)" shows "R a b" proof - have "\a b c. \a \ S; P a; b \ S; c \ S; P b; P c; R a b\ \ R a c" apply (rule connected_induction [OF \connected S\ opD], simp_all) by (meson trans opI) then show ?thesis by (metis etc opI) qed lemma connected_induction_simple: assumes "connected S" and etc: "a \ S" "b \ S" "P a" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y)" shows "P b" by (rule connected_induction [OF \connected S\ _, where P = "\x. True"]) (use opI etc in auto) lemma connected_equivalence_relation: assumes "connected S" and etc: "a \ S" "b \ S" and sym: "\x y. \R x y; x \ S; y \ S\ \ R y x" and trans: "\x y z. \R x y; R y z; x \ S; y \ S; z \ S\ \ R x z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. R a x)" shows "R a b" proof - have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" apply (rule connected_induction_simple [OF \connected S\], simp_all) by (meson local.sym local.trans opI openin_imp_subset subsetCE) then show ?thesis by (metis etc opI) qed lemma locally_constant_imp_constant: assumes "connected S" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. f x = f a)" shows "f constant_on S" proof - have "\x y. x \ S \ y \ S \ f x = f y" apply (rule connected_equivalence_relation [OF \connected S\], simp_all) by (metis opI) then show ?thesis by (metis constant_on_def) qed lemma locally_constant: assumes "connected S" shows "locally (\U. f constant_on U) S \ f constant_on S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (smt (verit, del_insts) assms constant_on_def locally_constant_imp_constant locally_def openin_subtopology_self subset_iff) next assume ?rhs then show ?lhs by (metis constant_on_subset locallyI openin_imp_subset order_refl) qed subsection\Basic properties of local compactness\ proposition locally_compact: fixes S :: "'a :: metric_space set" shows "locally compact S \ (\x \ S. \u v. x \ u \ u \ v \ v \ S \ openin (top_of_set S) u \ compact v)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson locallyE openin_subtopology_self) next assume r [rule_format]: ?rhs have *: "\u v. openin (top_of_set S) u \ compact v \ x \ u \ u \ v \ v \ S \ T" if "open T" "x \ S" "x \ T" for x T proof - obtain u v where uv: "x \ u" "u \ v" "v \ S" "compact v" "openin (top_of_set S) u" using r [OF \x \ S\] by auto obtain e where "e>0" and e: "cball x e \ T" using open_contains_cball \open T\ \x \ T\ by blast show ?thesis apply (rule_tac x="(S \ ball x e) \ u" in exI) apply (rule_tac x="cball x e \ v" in exI) using that \e > 0\ e uv apply auto done qed show ?lhs by (rule locallyI) (metis "*" Int_iff openin_open) qed lemma locally_compactE: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains u v where "\x. x \ S \ x \ u x \ u x \ v x \ v x \ S \ openin (top_of_set S) (u x) \ compact (v x)" using assms unfolding locally_compact by metis lemma locally_compact_alt: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\x \ S. \U. x \ U \ openin (top_of_set S) U \ compact(closure U) \ closure U \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson bounded_subset closure_minimal compact_closure compact_imp_bounded compact_imp_closed dual_order.trans locally_compactE) next assume ?rhs then show ?lhs by (meson closure_subset locally_compact) qed lemma locally_compact_Int_cball: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\x \ S. \e. 0 < e \ closed(cball x e \ S))" (is "?lhs = ?rhs") proof assume L: ?lhs then have "\x U V e. \U \ V; V \ S; compact V; 0 < e; cball x e \ S \ U\ \ closed (cball x e \ S)" by (metis compact_Int compact_cball compact_imp_closed inf.absorb_iff2 inf.assoc inf.orderE) with L show ?rhs by (meson locally_compactE openin_contains_cball) next assume R: ?rhs show ?lhs unfolding locally_compact proof fix x assume "x \ S" then obtain e where "e>0" and "compact (cball x e \ S)" by (metis Int_commute compact_Int_closed compact_cball inf.right_idem R) moreover have "\y\ball x e \ S. \\>0. cball y \ \ S \ ball x e" by (meson Elementary_Metric_Spaces.open_ball IntD1 le_infI1 open_contains_cball_eq) moreover have "openin (top_of_set S) (ball x e \ S)" by (simp add: inf_commute openin_open_Int) ultimately show "\U V. x \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" by (metis Int_iff \0 < e\ \x \ S\ ball_subset_cball centre_in_ball inf_commute inf_le1 inf_mono order_refl) qed qed lemma locally_compact_compact: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\K. K \ S \ compact K \ (\U V. K \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V))" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ S \ x \ u x \ u x \ v x \ v x \ S \ openin (top_of_set S) (u x) \ compact (v x)" by (metis locally_compactE) have *: "\U V. K \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" if "K \ S" "compact K" for K proof - have "\C. (\c\C. openin (top_of_set K) c) \ K \ \C \ \D\C. finite D \ K \ \D" using that by (simp add: compact_eq_openin_cover) moreover have "\c \ (\x. K \ u x) ` K. openin (top_of_set K) c" using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) moreover have "K \ \((\x. K \ u x) ` K)" using that by clarsimp (meson subsetCE uv) ultimately obtain D where "D \ (\x. K \ u x) ` K" "finite D" "K \ \D" by metis then obtain T where T: "T \ K" "finite T" "K \ \((\x. K \ u x) ` T)" by (metis finite_subset_image) have Tuv: "\(u ` T) \ \(v ` T)" using T that by (force dest!: uv) moreover have "openin (top_of_set S) (\ (u ` T))" using T that uv by fastforce moreover obtain "compact (\ (v ` T))" "\ (v ` T) \ S" by (metis T UN_subset_iff \K \ S\ compact_UN subset_iff uv) ultimately show ?thesis using T by auto qed show ?rhs by (blast intro: *) next assume ?rhs then show ?lhs apply (clarsimp simp: locally_compact) apply (drule_tac x="{x}" in spec, simp) done qed lemma open_imp_locally_compact: fixes S :: "'a :: heine_borel set" assumes "open S" shows "locally compact S" proof - have *: "\U V. x \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" if "x \ S" for x proof - obtain e where "e>0" and e: "cball x e \ S" using open_contains_cball assms \x \ S\ by blast have ope: "openin (top_of_set S) (ball x e)" by (meson e open_ball ball_subset_cball dual_order.trans open_subset) show ?thesis by (meson \0 < e\ ball_subset_cball centre_in_ball compact_cball e ope) qed show ?thesis unfolding locally_compact by (blast intro: *) qed lemma closed_imp_locally_compact: fixes S :: "'a :: heine_borel set" assumes "closed S" shows "locally compact S" proof - have *: "\U V. x \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" if "x \ S" for x apply (rule_tac x = "S \ ball x 1" in exI, rule_tac x = "S \ cball x 1" in exI) using \x \ S\ assms by auto show ?thesis unfolding locally_compact by (blast intro: *) qed lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" by (simp add: closed_imp_locally_compact) lemma locally_compact_Int: fixes S :: "'a :: t2_space set" shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" by (simp add: compact_Int locally_Int) lemma locally_compact_closedin: fixes S :: "'a :: heine_borel set" shows "\closedin (top_of_set S) T; locally compact S\ \ locally compact T" unfolding closedin_closed using closed_imp_locally_compact locally_compact_Int by blast lemma locally_compact_delete: fixes S :: "'a :: t1_space set" shows "locally compact S \ locally compact (S - {a})" by (auto simp: openin_delete locally_open_subset) lemma locally_closed: fixes S :: "'a :: heine_borel set" shows "locally closed S \ locally compact S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding locally_def apply (elim all_forward imp_forward asm_rl exE) apply (rule_tac x = "u \ ball x 1" in exI) apply (rule_tac x = "v \ cball x 1" in exI) apply (force intro: openin_trans) done next assume ?rhs then show ?lhs using compact_eq_bounded_closed locally_mono by blast qed lemma locally_compact_openin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" and opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ S" by (meson \x \ S\ opS openin_contains_cball) then have "cball x e2 \ (S \ T) = cball x e2 \ S" by force ultimately have "closed (cball x (min e1 e2) \ (S \ T))" by (metis (no_types, lifting) cball_min_Int closed_Int closed_cball inf_assoc inf_commute) then show ?thesis by (metis \0 < e1\ \0 < e2\ min_def) qed moreover have "\e>0. closed (cball x e \ (S \ T))" if "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ T" by (meson \x \ T\ opT openin_contains_cball) then have "cball x e2 \ (S \ T) = cball x e2 \ T" by force moreover have "closed (cball x e1 \ (cball x e2 \ T))" by (metis closed_Int closed_cball e1 inf_left_commute) ultimately show ?thesis by (rule_tac x="min e1 e2" in exI) (simp add: \0 < e2\ cball_min_Int inf_assoc) qed ultimately show ?thesis by (force simp: locally_compact_Int_cball) qed lemma locally_compact_closedin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" and clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover have "closed (cball x (min e1 e2) \ (S \ T))" by (smt (verit) Int_Un_distrib2 Int_commute cball_min_Int closed_Int closed_Un closed_cball e1 e2 inf_left_commute) ultimately show ?thesis by (rule_tac x="min e1 e2" in exI) linarith qed moreover have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S - T" using clT x by (fastforce simp: openin_contains_cball closedin_def) then have "closed (cball x e2 \ T)" proof - have "{} = T - (T - cball x e2)" using Diff_subset Int_Diff \cball x e2 \ (S \ T) \ S - T\ by auto then show ?thesis by (simp add: Diff_Diff_Int inf_commute) qed with e1 have "closed ((cball x e1 \ cball x e2) \ (S \ T))" apply (simp add: inf_commute inf_sup_distrib2) by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) then have "closed (cball x (min e1 e2) \ (S \ T))" by (simp add: cball_min_Int inf_commute) ultimately show ?thesis using \0 < e2\ by (rule_tac x="min e1 e2" in exI) linarith qed moreover have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S \ T - S" using clS x by (fastforce simp: openin_contains_cball closedin_def) then have "closed (cball x e2 \ S)" by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) with e1 have "closed ((cball x e1 \ cball x e2) \ (S \ T))" apply (simp add: inf_commute inf_sup_distrib2) by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) then have "closed (cball x (min e1 e2) \ (S \ T))" by (auto simp: cball_min_Int) ultimately show ?thesis using \0 < e2\ by (rule_tac x="min e1 e2" in exI) linarith qed ultimately show ?thesis by (auto simp: locally_compact_Int_cball) qed lemma locally_compact_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" by (auto simp: compact_Times locally_Times) lemma locally_compact_compact_subopen: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\K T. K \ S \ compact K \ open T \ K \ T \ (\U V. K \ U \ U \ V \ U \ T \ V \ S \ openin (top_of_set S) U \ compact V))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof clarify fix K :: "'a set" and T :: "'a set" assume "K \ S" and "compact K" and "open T" and "K \ T" obtain U V where "K \ U" "U \ V" "V \ S" "compact V" and ope: "openin (top_of_set S) U" using L unfolding locally_compact_compact by (meson \K \ S\ \compact K\) show "\U V. K \ U \ U \ V \ U \ T \ V \ S \ openin (top_of_set S) U \ compact V" proof (intro exI conjI) show "K \ U \ T" by (simp add: \K \ T\ \K \ U\) show "U \ T \ closure(U \ T)" by (rule closure_subset) show "closure (U \ T) \ S" by (metis \U \ V\ \V \ S\ \compact V\ closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) show "openin (top_of_set S) (U \ T)" by (simp add: \open T\ ope openin_Int_open) show "compact (closure (U \ T))" by (meson Int_lower1 \U \ V\ \compact V\ bounded_subset compact_closure compact_eq_bounded_closed) qed auto qed next assume ?rhs then show ?lhs unfolding locally_compact_compact by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) qed subsection\Sura-Bura's results about compact components of sets\ proposition Sura_Bura_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" and C: "C \ components S" shows "C = \{T. C \ T \ openin (top_of_set S) T \ closedin (top_of_set S) T}" (is "C = \?\") proof obtain x where x: "C = connected_component_set S x" and "x \ S" using C by (auto simp: components_def) have "C \ S" by (simp add: C in_components_subset) have "\?\ \ connected_component_set S x" proof (rule connected_component_maximal) have "x \ C" by (simp add: \x \ S\ x) then show "x \ \?\" by blast have clo: "closed (\?\)" by (simp add: \compact S\ closed_Inter closedin_compact_eq compact_imp_closed) have False if K1: "closedin (top_of_set (\?\)) K1" and K2: "closedin (top_of_set (\?\)) K2" and K12_Int: "K1 \ K2 = {}" and K12_Un: "K1 \ K2 = \?\" and "K1 \ {}" "K2 \ {}" for K1 K2 proof - have "closed K1" "closed K2" using closedin_closed_trans clo K1 K2 by blast+ then obtain V1 V2 where "open V1" "open V2" "K1 \ V1" "K2 \ V2" and V12: "V1 \ V2 = {}" using separation_normal \K1 \ K2 = {}\ by metis have SV12_ne: "(S - (V1 \ V2)) \ (\?\) \ {}" proof (rule compact_imp_fip) show "compact (S - (V1 \ V2))" by (simp add: \open V1\ \open V2\ \compact S\ compact_diff open_Un) show clo\: "closed T" if "T \ ?\" for T using that \compact S\ by (force intro: closedin_closed_trans simp add: compact_imp_closed) show "(S - (V1 \ V2)) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof assume djo: "(S - (V1 \ V2)) \ \\ = {}" obtain D where opeD: "openin (top_of_set S) D" and cloD: "closedin (top_of_set S) D" and "C \ D" and DV12: "D \ V1 \ V2" proof (cases "\ = {}") case True with \C \ S\ djo that show ?thesis by force next case False show ?thesis proof show ope: "openin (top_of_set S) (\\)" using openin_Inter \finite \\ False \ by blast then show "closedin (top_of_set S) (\\)" by (meson clo\ \ closed_Inter closed_subset openin_imp_subset subset_eq) show "C \ \\" using \ by auto show "\\ \ V1 \ V2" using ope djo openin_imp_subset by fastforce qed qed have "connected C" by (simp add: x) have "closed D" using \compact S\ cloD closedin_closed_trans compact_imp_closed by blast have cloV1: "closedin (top_of_set D) (D \ closure V1)" and cloV2: "closedin (top_of_set D) (D \ closure V2)" by (simp_all add: closedin_closed_Int) moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" using \D \ V1 \ V2\ \open V1\ \open V2\ V12 by (auto simp: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) ultimately have cloDV1: "closedin (top_of_set D) (D \ V1)" and cloDV2: "closedin (top_of_set D) (D \ V2)" by metis+ then obtain U1 U2 where "closed U1" "closed U2" and D1: "D \ V1 = D \ U1" and D2: "D \ V2 = D \ U2" by (auto simp: closedin_closed) have "D \ U1 \ C \ {}" proof assume "D \ U1 \ C = {}" then have *: "C \ D \ V2" using D1 DV12 \C \ D\ by auto have 1: "openin (top_of_set S) (D \ V2)" by (simp add: \open V2\ opeD openin_Int_open) have 2: "closedin (top_of_set S) (D \ V2)" using cloD cloDV2 closedin_trans by blast have "\ ?\ \ D \ V2" by (rule Inter_lower) (use * 1 2 in simp) then show False using K1 V12 \K1 \ {}\ \K1 \ V1\ closedin_imp_subset by blast qed moreover have "D \ U2 \ C \ {}" proof assume "D \ U2 \ C = {}" then have *: "C \ D \ V1" using D2 DV12 \C \ D\ by auto have 1: "openin (top_of_set S) (D \ V1)" by (simp add: \open V1\ opeD openin_Int_open) have 2: "closedin (top_of_set S) (D \ V1)" using cloD cloDV1 closedin_trans by blast have "\?\ \ D \ V1" by (rule Inter_lower) (use * 1 2 in simp) then show False using K2 V12 \K2 \ {}\ \K2 \ V2\ closedin_imp_subset by blast qed ultimately show False using \connected C\ [unfolded connected_closed, simplified, rule_format, of concl: "D \ U1" "D \ U2"] using \C \ D\ D1 D2 V12 DV12 \closed U1\ \closed U2\ \closed D\ by blast qed qed show False by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \K1 \ V1\ \K2 \ V2\ disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) qed then show "connected (\?\)" by (auto simp: connected_closedin_eq) show "\?\ \ S" by (fastforce simp: C in_components_subset) qed with x show "\?\ \ C" by simp qed auto corollary Sura_Bura_clopen_subset: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and U: "open U" "C \ U" obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof (rule ccontr) assume "\ thesis" with that have neg: "\K. openin (top_of_set S) K \ compact K \ C \ K \ K \ U" by metis obtain V K where "C \ V" "V \ U" "V \ K" "K \ S" "compact K" and opeSV: "openin (top_of_set S) V" using S U \compact C\ by (meson C in_components_subset locally_compact_compact_subopen) let ?\ = "{T. C \ T \ openin (top_of_set K) T \ compact T \ T \ K}" have CK: "C \ components K" by (meson C \C \ V\ \K \ S\ \V \ K\ components_intermediate_subset subset_trans) with \compact K\ have "C = \{T. C \ T \ openin (top_of_set K) T \ closedin (top_of_set K) T}" by (simp add: Sura_Bura_compact) then have Ceq: "C = \?\" by (simp add: closedin_compact_eq \compact K\) obtain W where "open W" and W: "V = S \ W" using opeSV by (auto simp: openin_open) have "-(U \ W) \ \?\ \ {}" proof (rule closed_imp_fip_compact) show "- (U \ W) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof (cases "\ = {}") case True have False if "U = UNIV" "W = UNIV" proof - have "V = S" by (simp add: W \W = UNIV\) with neg show False using \C \ V\ \K \ S\ \V \ K\ \V \ U\ \compact K\ by auto qed with True show ?thesis by auto next case False show ?thesis proof assume "- (U \ W) \ \\ = {}" then have FUW: "\\ \ U \ W" by blast have "C \ \\" using \ by auto moreover have "compact (\\)" by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \) moreover have "\\ \ K" using False that(2) by fastforce moreover have opeKF: "openin (top_of_set K) (\\)" using False \ \finite \\ by blast then have opeVF: "openin (top_of_set V) (\\)" using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce then have "openin (top_of_set S) (\\)" by (metis opeSV openin_trans) moreover have "\\ \ U" by (meson \V \ U\ opeVF dual_order.trans openin_imp_subset) ultimately show False using neg by blast qed qed qed (use \open W\ \open U\ in auto) with W Ceq \C \ V\ \C \ U\ show False by auto qed corollary Sura_Bura_clopen_subset_alt: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and opeSU: "openin (top_of_set S) U" and "C \ U" obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof - obtain V where "open V" "U = S \ V" using opeSU by (auto simp: openin_open) with \C \ U\ have "C \ V" by auto then show ?thesis using Sura_Bura_clopen_subset [OF S C \compact C\ \open V\] by (metis \U = S \ V\ inf.bounded_iff openin_imp_subset that) qed corollary Sura_Bura: fixes S :: "'a::euclidean_space set" assumes "locally compact S" "C \ components S" "compact C" shows "C = \ {K. C \ K \ compact K \ openin (top_of_set S) K}" (is "C = ?rhs") proof show "?rhs \ C" proof (clarsimp, rule ccontr) fix x assume *: "\X. C \ X \ compact X \ openin (top_of_set S) X \ x \ X" and "x \ C" obtain U V where "open U" "open V" "{x} \ U" "C \ V" "U \ V = {}" using separation_normal [of "{x}" C] by (metis Int_empty_left \x \ C\ \compact C\ closed_empty closed_insert compact_imp_closed insert_disjoint(1)) have "x \ V" using \U \ V = {}\ \{x} \ U\ by blast then show False by (meson "*" Sura_Bura_clopen_subset \C \ V\ \open V\ assms(1) assms(2) assms(3) subsetCE) qed qed blast subsection\Special cases of local connectedness and path connectedness\ lemma locally_connected_1: assumes "\V x. \openin (top_of_set S) V; x \ V\ \ \U. openin (top_of_set S) U \ connected U \ x \ U \ U \ V" shows "locally connected S" by (metis assms locally_def) lemma locally_connected_2: assumes "locally connected S" "openin (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (connected_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. connected v \ x \ u \ u \ v \ v \ w))" and "connected_component t x y" then have "y \ t" and y: "y \ connected_component_set t x" using connected_component_subset by blast+ obtain F where "\x y. (\w. openin ?SS w \ (\u. connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. connected u \ x \ F x y \ F x y \ u \ u \ y))" by moura then obtain G where "\a A. (\U. openin ?SS U \ (\V. connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" by moura then have *: "openin ?SS (F y t) \ connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" using 1 \y \ t\ by presburger have "G y t \ connected_component_set t y" by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) then have "\A. openin ?SS A \ y \ A \ A \ connected_component_set t x" by (metis (no_types) * connected_component_eq dual_order.trans y) } then show ?thesis using assms openin_subopen by (force simp: locally_def) qed lemma locally_connected_3: assumes "\t x. \openin (top_of_set S) t; x \ t\ \ openin (top_of_set S) (connected_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" using assms connected_component_subset by fastforce lemma locally_connected: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_connected_open_connected_component: "locally connected S \ (\t x. openin (top_of_set S) t \ x \ t \ openin (top_of_set S) (connected_component_set t x))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_path_connected_1: assumes "\v x. \openin (top_of_set S) v; x \ v\ \ \u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" shows "locally path_connected S" by (force simp add: locally_def dest: assms) lemma locally_path_connected_2: assumes "locally path_connected S" "openin (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (path_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. path_connected v \ x \ u \ u \ v \ v \ w))" and "path_component t x y" then have "y \ t" and y: "y \ path_component_set t x" using path_component_mem(2) by blast+ obtain F where "\x y. (\w. openin ?SS w \ (\u. path_connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. path_connected u \ x \ F x y \ F x y \ u \ u \ y))" by moura then obtain G where "\a A. (\U. openin ?SS U \ (\V. path_connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ path_connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" by moura then have *: "openin ?SS (F y t) \ path_connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" using 1 \y \ t\ by presburger have "G y t \ path_component_set t y" using * path_component_maximal rev_subsetD by blast then have "\A. openin ?SS A \ y \ A \ A \ path_component_set t x" by (metis "*" \G y t \ path_component_set t y\ dual_order.trans path_component_eq y) } then show ?thesis using assms openin_subopen by (force simp: locally_def) qed lemma locally_path_connected_3: assumes "\t x. \openin (top_of_set S) t; x \ t\ \ openin (top_of_set S) (path_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" proof - have "path_component v x x" by (meson assms(3) path_component_refl) then show ?thesis by (metis assms mem_Collect_eq path_component_subset path_connected_path_component) qed proposition locally_path_connected: "locally path_connected S \ (\V x. openin (top_of_set S) V \ x \ V \ (\U. openin (top_of_set S) U \ path_connected U \ x \ U \ U \ V))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) proposition locally_path_connected_open_path_component: "locally path_connected S \ (\t x. openin (top_of_set S) t \ x \ t \ openin (top_of_set S) (path_component_set t x))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) lemma locally_connected_open_component: "locally connected S \ (\t c. openin (top_of_set S) t \ c \ components t \ openin (top_of_set S) c)" by (metis components_iff locally_connected_open_connected_component) proposition locally_connected_im_kleinen: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\c. connected c \ c \ v \ x \ c \ y \ c))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (fastforce simp add: locally_connected) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ c" if "openin (top_of_set S) t" and c: "c \ components t" and "x \ c" for t c x proof - from that \?rhs\ [rule_format, of t x] obtain u where u: "openin (top_of_set S) u \ x \ u \ u \ t \ (\y. y \ u \ (\c. connected c \ c \ t \ x \ c \ y \ c))" using in_components_subset by auto obtain F :: "'a set \ 'a set \ 'a" where "\x y. (\z. z \ x \ y = connected_component_set x z) = (F x y \ x \ y = connected_component_set x (F x y))" by moura then have F: "F t c \ t \ c = connected_component_set t (F t c)" by (meson components_iff c) obtain G :: "'a set \ 'a set \ 'a" where G: "\x y. (\z. z \ y \ z \ x) = (G x y \ y \ G x y \ x)" by moura have "G c u \ u \ G c u \ c" using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) then show ?thesis using G u by auto qed show ?lhs unfolding locally_connected_open_component by (meson "*" openin_subopen) qed proposition locally_path_connected_im_kleinen: "locally path_connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\p. path p \ path_image p \ v \ pathstart p = x \ pathfinish p = y))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: locally_path_connected path_connected_def) apply (erule all_forward ex_forward imp_forward conjE | simp)+ by (meson dual_order.trans) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ path_component_set u z" if "openin (top_of_set S) u" and "z \ u" and c: "path_component u z x" for u z x proof - have "x \ u" by (meson c path_component_mem(2)) with that \?rhs\ [rule_format, of u x] obtain U where U: "openin (top_of_set S) U \ x \ U \ U \ u \ (\y. y \ U \ (\p. path p \ path_image p \ u \ pathstart p = x \ pathfinish p = y))" by blast show ?thesis by (metis U c mem_Collect_eq path_component_def path_component_eq subsetI) qed show ?lhs unfolding locally_path_connected_open_path_component using "*" openin_subopen by fastforce qed lemma locally_path_connected_imp_locally_connected: "locally path_connected S \ locally connected S" using locally_mono path_connected_imp_connected by blast lemma locally_connected_components: "\locally connected S; c \ components S\ \ locally connected c" by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) lemma locally_path_connected_components: "\locally path_connected S; c \ components S\ \ locally path_connected c" by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) lemma locally_path_connected_connected_component: "locally path_connected S \ locally path_connected (connected_component_set S x)" by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) lemma open_imp_locally_path_connected: fixes S :: "'a :: real_normed_vector set" assumes "open S" shows "locally path_connected S" proof (rule locally_mono) show "locally convex S" using assms unfolding locally_def by (meson open_ball centre_in_ball convex_ball openE open_subset openin_imp_subset openin_open_trans subset_trans) show "\T::'a set. convex T \ path_connected T" using convex_imp_path_connected by blast qed lemma open_imp_locally_connected: fixes S :: "'a :: real_normed_vector set" shows "open S \ locally connected S" by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" by (simp add: open_imp_locally_path_connected) lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" by (simp add: open_imp_locally_connected) lemma openin_connected_component_locally_connected: "locally connected S \ openin (top_of_set S) (connected_component_set S x)" by (metis connected_component_eq_empty locally_connected_2 openin_empty openin_subtopology_self) lemma openin_components_locally_connected: "\locally connected S; c \ components S\ \ openin (top_of_set S) c" using locally_connected_open_component openin_subtopology_self by blast lemma openin_path_component_locally_path_connected: "locally path_connected S \ openin (top_of_set S) (path_component_set S x)" by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) lemma closedin_path_component_locally_path_connected: assumes "locally path_connected S" shows "closedin (top_of_set S) (path_component_set S x)" proof - have "openin (top_of_set S) (\ ({path_component_set S y |y. y \ S} - {path_component_set S x}))" using locally_path_connected_2 assms by fastforce then show ?thesis by (simp add: closedin_def path_component_subset complement_path_component_Union) qed lemma convex_imp_locally_path_connected: fixes S :: "'a:: real_normed_vector set" assumes "convex S" shows "locally path_connected S" proof (clarsimp simp: locally_path_connected) fix V x assume "openin (top_of_set S) V" and "x \ V" then obtain T e where "V = S \ T" "x \ S" "0 < e" "ball x e \ T" by (metis Int_iff openE openin_open) then have "openin (top_of_set S) (S \ ball x e)" "path_connected (S \ ball x e)" by (simp_all add: assms convex_Int convex_imp_path_connected openin_open_Int) then show "\U. openin (top_of_set S) U \ path_connected U \ x \ U \ U \ V" using \0 < e\ \V = S \ T\ \ball x e \ T\ \x \ S\ by auto qed lemma convex_imp_locally_connected: fixes S :: "'a:: real_normed_vector set" shows "convex S \ locally connected S" by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) subsection\Relations between components and path components\ lemma path_component_eq_connected_component: assumes "locally path_connected S" shows "(path_component S x = connected_component S x)" proof (cases "x \ S") case True have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)" proof (rule openin_subset_trans) show "openin (top_of_set S) (path_component_set S x)" by (simp add: True assms locally_path_connected_2) show "connected_component_set S x \ S" by (simp add: connected_component_subset) qed (simp add: path_component_subset_connected_component) moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)" proof (rule closedin_subset_trans [of S]) show "closedin (top_of_set S) (path_component_set S x)" by (simp add: assms closedin_path_component_locally_path_connected) show "connected_component_set S x \ S" by (simp add: connected_component_subset) qed (simp add: path_component_subset_connected_component) ultimately have *: "path_component_set S x = connected_component_set S x" by (metis connected_connected_component connected_clopen True path_component_eq_empty) then show ?thesis by blast next case False then show ?thesis by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) qed lemma path_component_eq_connected_component_set: "locally path_connected S \ (path_component_set S x = connected_component_set S x)" by (simp add: path_component_eq_connected_component) lemma locally_path_connected_path_component: "locally path_connected S \ locally path_connected (path_component_set S x)" using locally_path_connected_connected_component path_component_eq_connected_component by fastforce lemma open_path_connected_component: fixes S :: "'a :: real_normed_vector set" shows "open S \ path_component S x = connected_component S x" by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) lemma open_path_connected_component_set: fixes S :: "'a :: real_normed_vector set" shows "open S \ path_component_set S x = connected_component_set S x" by (simp add: open_path_connected_component) proposition locally_connected_quotient_image: assumes lcS: "locally connected S" and oo: "\T. T \ f ` S \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally connected (f ` S)" proof (clarsimp simp: locally_connected_open_component) fix U C assume opefSU: "openin (top_of_set (f ` S)) U" and "C \ components U" then have "C \ U" "U \ f ` S" by (meson in_components_subset openin_imp_subset)+ then have "openin (top_of_set (f ` S)) C \ openin (top_of_set S) (S \ f -` C)" by (auto simp: oo) moreover have "openin (top_of_set S) (S \ f -` C)" proof (subst openin_subopen, clarify) fix x assume "x \ S" "f x \ C" show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` C)" proof (intro conjI exI) show "openin (top_of_set S) (connected_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set S) (connected_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" using \U \ f ` S\ opefSU lcS locally_connected_2 oo by blast with ** show False by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) qed next show "x \ connected_component_set (S \ f -` U) x" using \C \ U\ \f x \ C\ \x \ S\ by auto next have contf: "continuous_on S f" by (simp add: continuous_on_open oo openin_imp_subset) then have "continuous_on (connected_component_set (S \ f -` U) x) f" by (meson connected_component_subset continuous_on_subset inf.boundedE) then have "connected (f ` connected_component_set (S \ f -` U) x)" by (rule connected_continuous_image [OF _ connected_connected_component]) moreover have "f ` connected_component_set (S \ f -` U) x \ U" using connected_component_in by blast moreover have "C \ f ` connected_component_set (S \ f -` U) x \ {}" using \C \ U\ \f x \ C\ \x \ S\ by fastforce ultimately have fC: "f ` (connected_component_set (S \ f -` U) x) \ C" by (rule components_maximal [OF \C \ components U\]) have cUC: "connected_component_set (S \ f -` U) x \ (S \ f -` C)" using connected_component_subset fC by blast have "connected_component_set (S \ f -` U) x \ connected_component_set (S \ f -` C) x" proof - { assume "x \ connected_component_set (S \ f -` U) x" then have ?thesis using cUC connected_component_idemp connected_component_mono by blast } then show ?thesis using connected_component_eq_empty by auto qed also have "\ \ (S \ f -` C)" by (rule connected_component_subset) finally show "connected_component_set (S \ f -` U) x \ (S \ f -` C)" . qed qed ultimately show "openin (top_of_set (f ` S)) C" by metis qed text\The proof resembles that above but is not identical!\ proposition locally_path_connected_quotient_image: assumes lcS: "locally path_connected S" and oo: "\T. T \ f ` S \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally path_connected (f ` S)" proof (clarsimp simp: locally_path_connected_open_path_component) fix U y assume opefSU: "openin (top_of_set (f ` S)) U" and "y \ U" then have "path_component_set U y \ U" "U \ f ` S" by (meson path_component_subset openin_imp_subset)+ then have "openin (top_of_set (f ` S)) (path_component_set U y) \ openin (top_of_set S) (S \ f -` path_component_set U y)" proof - have "path_component_set U y \ f ` S" using \U \ f ` S\ \path_component_set U y \ U\ by blast then show ?thesis using oo by blast qed moreover have "openin (top_of_set S) (S \ f -` path_component_set U y)" proof (subst openin_subopen, clarify) fix x assume "x \ S" and Uyfx: "path_component U y (f x)" then have "f x \ U" using path_component_mem by blast show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" proof (intro conjI exI) show "openin (top_of_set S) (path_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set S) (path_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" by (metis (no_types, lifting) \U \ f ` S\ opefSU lcS oo locally_path_connected_open_path_component) then show False using ** \path_component_set U y \ U\ \x \ S\ \path_component U y (f x)\ by blast qed next show "x \ path_component_set (S \ f -` U) x" by (simp add: \f x \ U\ \x \ S\ path_component_refl) next have contf: "continuous_on S f" by (simp add: continuous_on_open oo openin_imp_subset) then have "continuous_on (path_component_set (S \ f -` U) x) f" by (meson Int_lower1 continuous_on_subset path_component_subset) then have "path_connected (f ` path_component_set (S \ f -` U) x)" by (simp add: path_connected_continuous_image) moreover have "f ` path_component_set (S \ f -` U) x \ U" using path_component_mem by fastforce moreover have "f x \ f ` path_component_set (S \ f -` U) x" by (force simp: \x \ S\ \f x \ U\ path_component_refl_eq) ultimately have "f ` (path_component_set (S \ f -` U) x) \ path_component_set U (f x)" by (meson path_component_maximal) also have "\ \ path_component_set U y" by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) finally have fC: "f ` (path_component_set (S \ f -` U) x) \ path_component_set U y" . have cUC: "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" using path_component_subset fC by blast have "path_component_set (S \ f -` U) x \ path_component_set (S \ f -` path_component_set U y) x" proof - have "\a. path_component_set (path_component_set (S \ f -` U) x) a \ path_component_set (S \ f -` path_component_set U y) a" using cUC path_component_mono by blast then show ?thesis using path_component_path_component by blast qed also have "\ \ (S \ f -` path_component_set U y)" by (rule path_component_subset) finally show "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" . qed qed ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)" by metis qed subsection\<^marker>\tag unimportant\\Components, continuity, openin, closedin\ lemma continuous_on_components_gen: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "\C. C \ components S \ openin (top_of_set S) C \ continuous_on C f" shows "continuous_on S f" proof (clarsimp simp: continuous_openin_preimage_eq) fix t :: "'b set" assume "open t" have *: "S \ f -` t = (\c \ components S. c \ f -` t)" by auto show "openin (top_of_set S) (S \ f -` t)" unfolding * using \open t\ assms continuous_openin_preimage_gen openin_trans openin_Union by blast qed lemma continuous_on_components: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "locally connected S " "\C. C \ components S \ continuous_on C f" shows "continuous_on S f" proof (rule continuous_on_components_gen) fix C assume "C \ components S" then show "openin (top_of_set S) C \ continuous_on C f" by (simp add: assms openin_components_locally_connected) qed lemma continuous_on_components_eq: "locally connected S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" by (meson continuous_on_components continuous_on_subset in_components_subset) lemma continuous_on_components_open: fixes S :: "'a::real_normed_vector set" assumes "open S " "\c. c \ components S \ continuous_on c f" shows "continuous_on S f" using continuous_on_components open_imp_locally_connected assms by blast lemma continuous_on_components_open_eq: fixes S :: "'a::real_normed_vector set" shows "open S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" using continuous_on_subset in_components_subset by (blast intro: continuous_on_components_open) lemma closedin_union_complement_components: assumes U: "locally connected U" and S: "closedin (top_of_set U) S" and cuS: "c \ components(U - S)" shows "closedin (top_of_set U) (S \ \c)" proof - have di: "(\S T. S \ c \ T \ c' \ disjnt S T) \ disjnt (\ c) (\ c')" for c' by (simp add: disjnt_def) blast have "S \ U" using S closedin_imp_subset by blast moreover have "U - S = \c \ \(components (U - S) - c)" by (metis Diff_partition Union_components Union_Un_distrib assms(3)) moreover have "disjnt (\c) (\(components (U - S) - c))" apply (rule di) by (metis di DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) ultimately have eq: "S \ \c = U - (\(components(U - S) - c))" by (auto simp: disjnt_def) have *: "openin (top_of_set U) (\(components (U - S) - c))" proof (rule openin_Union [OF openin_trans [of "U - S"]]) show "openin (top_of_set (U - S)) T" if "T \ components (U - S) - c" for T using that by (simp add: U S locally_diff_closed openin_components_locally_connected) show "openin (top_of_set U) (U - S)" if "T \ components (U - S) - c" for T using that by (simp add: openin_diff S) qed have "closedin (top_of_set U) (U - \ (components (U - S) - c))" by (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) then have "openin (top_of_set U) (U - (U - \(components (U - S) - c)))" by (simp add: openin_diff) then show ?thesis by (force simp: eq closedin_def) qed lemma closed_union_complement_components: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and c: "c \ components(- S)" shows "closed(S \ \ c)" proof - have "closedin (top_of_set UNIV) (S \ \c)" by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_union_complement_components locally_connected_UNIV subtopology_UNIV) then show ?thesis by simp qed lemma closedin_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes u: "locally connected u" and S: "closedin (top_of_set u) S" and c: " c \ components(u - S)" shows "closedin (top_of_set u) (S \ c)" proof - have "closedin (top_of_set u) (S \ \{c})" using c by (blast intro: closedin_union_complement_components [OF u S]) then show ?thesis by simp qed lemma closed_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and c: " c \ components(-S)" shows "closed (S \ c)" by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component locally_connected_UNIV subtopology_UNIV) subsection\Existence of isometry between subspaces of same dimension\ lemma isometry_subset_subspace: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S \ dim T" obtains f where "linear f" "f ` S \ T" "\x. x \ S \ norm(f x) = norm x" proof - obtain B where "B \ S" and Borth: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" "finite B" "card B = dim S" "span B = S" by (metis orthonormal_basis_subspace [OF S] independent_finite) obtain C where "C \ T" and Corth: "pairwise orthogonal C" and C1:"\x. x \ C \ norm x = 1" and "independent C" "finite C" "card C = dim T" "span C = T" by (metis orthonormal_basis_subspace [OF T] independent_finite) obtain fb where "fb ` B \ C" "inj_on fb B" by (metis \card B = dim S\ \card C = dim T\ \finite B\ \finite C\ card_le_inj d) then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" using Corth unfolding pairwise_def inj_on_def by (blast intro: orthogonal_clauses) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce have "span (f ` B) \ span C" by (metis \fb ` B \ C\ ffb image_cong span_mono) then have "f ` S \ T" unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \fb ` B \ C\ by auto have "norm (f x) = norm x" if "x \ S" for x proof - interpret linear f by fact obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" proof (rule norm_sum_Pythagorean [OF \finite B\]) show "pairwise (\v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) qed also have "\ = norm x ^2" by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) finally show ?thesis by (simp add: norm_eq_sqrt_inner) qed then show ?thesis by (rule that [OF \linear f\ \f ` S \ T\]) qed proposition isometries_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" "\x. x \ S \ norm(f x) = norm x" "\x. x \ T \ norm(g x) = norm x" "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" proof - obtain B where "B \ S" and Borth: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" "finite B" "card B = dim S" "span B = S" by (metis orthonormal_basis_subspace [OF S] independent_finite) obtain C where "C \ T" and Corth: "pairwise orthogonal C" and C1:"\x. x \ C \ norm x = 1" and "independent C" "finite C" "card C = dim T" "span C = T" by (metis orthonormal_basis_subspace [OF T] independent_finite) obtain fb where "bij_betw fb B C" by (metis \finite B\ \finite C\ bij_betw_iff_card \card B = dim S\ \card C = dim T\ d) then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" using Corth unfolding pairwise_def inj_on_def bij_betw_def by (blast intro: orthogonal_clauses) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce interpret f: linear f by fact define gb where "gb \ inv_into B fb" then have pairwise_orth_gb: "pairwise (\v j. orthogonal (gb v) (gb j)) C" using Borth \bij_betw fb B C\ unfolding pairwise_def bij_betw_def by force obtain g where "linear g" and ggb: "\x. x \ C \ g x = gb x" using linear_independent_extend \independent C\ by fastforce interpret g: linear g by fact have "span (f ` B) \ span C" by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb image_cong) then have "f ` S \ T" unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on by fastforce have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \ S" for x proof - obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce have "f x = (\v \ B. f (a v *\<^sub>R v))" using linear_sum [OF \linear f\] x by auto also have "\ = (\v \ B. a v *\<^sub>R f v)" by (simp add: f.sum f.scale) also have "\ = (\v \ B. a v *\<^sub>R fb v)" by (simp add: ffb cong: sum.cong) finally have *: "f x = (\v\B. a v *\<^sub>R fb v)" . then have "(norm (f x))\<^sup>2 = (norm (\v\B. a v *\<^sub>R fb v))\<^sup>2" by simp also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" proof (rule norm_sum_Pythagorean [OF \finite B\]) show "pairwise (\v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) qed also have "\ = (norm x)\<^sup>2" by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) finally show "norm (f x) = norm x" by (simp add: norm_eq_sqrt_inner) have "g (f x) = g (\v\B. a v *\<^sub>R fb v)" by (simp add: *) also have "\ = (\v\B. g (a v *\<^sub>R fb v))" by (simp add: g.sum g.scale) also have "\ = (\v\B. a v *\<^sub>R g (fb v))" by (simp add: g.scale) also have "\ = (\v\B. a v *\<^sub>R v)" proof (rule sum.cong [OF refl]) show "a x *\<^sub>R g (fb x) = a x *\<^sub>R x" if "x \ B" for x using that \bij_betw fb B C\ bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce qed also have "\ = x" using x by blast finally show "g (f x) = x" . qed have [simp]: "\x. x \ C \ norm (gb x) = norm x" by (metis B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on gb_def inv_into_into) have g [simp]: "f (g x) = x" if "x \ T" for x proof - obtain a where x: "x = (\v \ C. a v *\<^sub>R v)" using \finite C\ \span C = T\ \x \ T\ span_finite by fastforce have "g x = (\v \ C. g (a v *\<^sub>R v))" by (simp add: x g.sum) also have "\ = (\v \ C. a v *\<^sub>R g v)" by (simp add: g.scale) also have "\ = (\v \ C. a v *\<^sub>R gb v)" by (simp add: ggb cong: sum.cong) finally have "f (g x) = f (\v\C. a v *\<^sub>R gb v)" by simp also have "\ = (\v\C. f (a v *\<^sub>R gb v))" by (simp add: f.scale f.sum) also have "\ = (\v\C. a v *\<^sub>R f (gb v))" by (simp add: f.scale f.sum) also have "\ = (\v\C. a v *\<^sub>R v)" using \bij_betw fb B C\ by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) also have "\ = x" using x by blast finally show "f (g x) = x" . qed have gim: "g ` T = S" by (metis (full_types) S T \f ` S \ T\ d dim_eq_span dim_image_le f(2) g.linear_axioms image_iff linear_subspace_image span_eq_iff subset_iff) have fim: "f ` S = T" using \g ` T = S\ image_iff by fastforce have [simp]: "norm (g x) = norm x" if "x \ T" for x using fim that by auto show ?thesis by (rule that [OF \linear f\ \linear g\]) (simp_all add: fim gim) qed corollary isometry_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" obtains f where "linear f" "f ` S = T" "\x. x \ S \ norm(f x) = norm x" using isometries_subspaces [OF assms] by metis corollary isomorphisms_UNIV_UNIV: assumes "DIM('M) = DIM('N)" obtains f::"'M::euclidean_space \'N::euclidean_space" and g where "linear f" "linear g" "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g (f x) = x" "\y. f(g y) = y" using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) lemma homeomorphic_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" shows "S homeomorphic T" proof - obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" by (blast intro: isometries_subspaces [OF assms]) then show ?thesis unfolding homeomorphic_def homeomorphism_def apply (rule_tac x=f in exI, rule_tac x=g in exI) apply (auto simp: linear_continuous_on linear_conv_bounded_linear) done qed lemma homeomorphic_affine_sets: assumes "affine S" "affine T" "aff_dim S = aff_dim T" shows "S homeomorphic T" proof (cases "S = {} \ T = {}") case True with assms aff_dim_empty homeomorphic_empty show ?thesis by metis next case False then obtain a b where ab: "a \ S" "b \ T" by auto then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) have "S homeomorphic ((+) (- a) ` S)" by (fact homeomorphic_translation) also have "\ homeomorphic ((+) (- b) ` T)" by (rule homeomorphic_subspaces [OF ss dd]) also have "\ homeomorphic T" using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T]) finally show ?thesis . qed subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ locale\<^marker>\tag important\ Retracts = fixes S h t k assumes conth: "continuous_on S h" and imh: "h ` S = t" and contk: "continuous_on t k" and imk: "k ` t \ S" and idhk: "\y. y \ t \ h(k y) = y" begin lemma homotopically_trivial_retraction_gen: assumes P: "\f. \continuous_on U f; f ` U \ t; Q f\ \ P(k \ f)" and Q: "\f. \continuous_on U f; f ` U \ S; P f\ \ Q(h \ f)" and Qeq: "\h k. (\x. x \ U \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on U f; f ` U \ S; P f; continuous_on U g; g ` U \ S; P g\ \ homotopic_with_canon P U S f g" and contf: "continuous_on U f" and imf: "f ` U \ t" and Qf: "Q f" and contg: "continuous_on U g" and img: "g ` U \ t" and Qg: "Q g" shows "homotopic_with_canon Q U t f g" proof - have "continuous_on U (k \ f)" using contf continuous_on_compose continuous_on_subset contk imf by blast moreover have "(k \ f) ` U \ S" using imf imk by fastforce moreover have "P (k \ f)" by (simp add: P Qf contf imf) moreover have "continuous_on U (k \ g)" using contg continuous_on_compose continuous_on_subset contk img by blast moreover have "(k \ g) ` U \ S" using img imk by fastforce moreover have "P (k \ g)" by (simp add: P Qg contg img) ultimately have "homotopic_with_canon P U S (k \ f) (k \ g)" by (rule hom) then have "homotopic_with_canon Q U t (h \ (k \ f)) (h \ (k \ g))" apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) using Q by (auto simp: conth imh) then show ?thesis proof (rule homotopic_with_eq; simp) show "\h k. (\x. x \ U \ h x = k x) \ Q h = Q k" using Qeq topspace_euclidean_subtopology by blast show "\x. x \ U \ f x = h (k (f x))" "\x. x \ U \ g x = h (k (g x))" using idhk imf img by auto qed qed lemma homotopically_trivial_retraction_null_gen: assumes P: "\f. \continuous_on U f; f ` U \ t; Q f\ \ P(k \ f)" and Q: "\f. \continuous_on U f; f ` U \ S; P f\ \ Q(h \ f)" and Qeq: "\h k. (\x. x \ U \ h x = k x) \ Q h = Q k" and hom: "\f. \continuous_on U f; f ` U \ S; P f\ \ \c. homotopic_with_canon P U S f (\x. c)" and contf: "continuous_on U f" and imf:"f ` U \ t" and Qf: "Q f" obtains c where "homotopic_with_canon Q U t f (\x. c)" proof - have feq: "\x. x \ U \ (h \ (k \ f)) x = f x" using idhk imf by auto have "continuous_on U (k \ f)" using contf continuous_on_compose continuous_on_subset contk imf by blast moreover have "(k \ f) ` U \ S" using imf imk by fastforce moreover have "P (k \ f)" by (simp add: P Qf contf imf) ultimately obtain c where "homotopic_with_canon P U S (k \ f) (\x. c)" by (metis hom) then have "homotopic_with_canon Q U t (h \ (k \ f)) (h \ (\x. c))" apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) using Q by (auto simp: conth imh) then have "homotopic_with_canon Q U t f (\x. h c)" proof (rule homotopic_with_eq) show "\x. x \ topspace (top_of_set U) \ f x = (h \ (k \ f)) x" using feq by auto show "\h k. (\x. x \ topspace (top_of_set U) \ h x = k x) \ Q h = Q k" using Qeq topspace_euclidean_subtopology by blast qed auto then show ?thesis using that by blast qed lemma cohomotopically_trivial_retraction_gen: assumes P: "\f. \continuous_on t f; f ` t \ U; Q f\ \ P(f \ h)" and Q: "\f. \continuous_on S f; f ` S \ U; P f\ \ Q(f \ k)" and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on S f; f ` S \ U; P f; continuous_on S g; g ` S \ U; P g\ \ homotopic_with_canon P S U f g" and contf: "continuous_on t f" and imf: "f ` t \ U" and Qf: "Q f" and contg: "continuous_on t g" and img: "g ` t \ U" and Qg: "Q g" shows "homotopic_with_canon Q t U f g" proof - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto have geq: "\x. x \ t \ (g \ h \ k) x = g x" using idhk img by auto have "continuous_on S (f \ h)" using contf conth continuous_on_compose imh by blast moreover have "(f \ h) ` S \ U" using imf imh by fastforce moreover have "P (f \ h)" by (simp add: P Qf contf imf) moreover have "continuous_on S (g \ h)" using contg continuous_on_compose continuous_on_subset conth imh by blast moreover have "(g \ h) ` S \ U" using img imh by fastforce moreover have "P (g \ h)" by (simp add: P Qg contg img) ultimately have "homotopic_with_canon P S U (f \ h) (g \ h)" by (rule hom) then have "homotopic_with_canon Q t U (f \ h \ k) (g \ h \ k)" apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) using Q by (auto simp: contk imk) then show ?thesis proof (rule homotopic_with_eq) show "f x = (f \ h \ k) x" "g x = (g \ h \ k) x" if "x \ topspace (top_of_set t)" for x using feq geq that by force+ qed (use Qeq topspace_euclidean_subtopology in blast) qed lemma cohomotopically_trivial_retraction_null_gen: assumes P: "\f. \continuous_on t f; f ` t \ U; Q f\ \ P(f \ h)" and Q: "\f. \continuous_on S f; f ` S \ U; P f\ \ Q(f \ k)" and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on S f; f ` S \ U; P f\ \ \c. homotopic_with_canon P S U f (\x. c)" and contf: "continuous_on t f" and imf: "f ` t \ U" and Qf: "Q f" obtains c where "homotopic_with_canon Q t U f (\x. c)" proof - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto have "continuous_on S (f \ h)" using contf conth continuous_on_compose imh by blast moreover have "(f \ h) ` S \ U" using imf imh by fastforce moreover have "P (f \ h)" by (simp add: P Qf contf imf) ultimately obtain c where "homotopic_with_canon P S U (f \ h) (\x. c)" by (metis hom) then have \
: "homotopic_with_canon Q t U (f \ h \ k) ((\x. c) \ k)" proof (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) show "\h. \continuous_map (top_of_set S) (top_of_set U) h; P h\ \ Q (h \ k)" using Q by auto qed (auto simp: contk imk) moreover have "homotopic_with_canon Q t U f (\x. c)" using homotopic_with_eq [OF \
] feq Qeq by fastforce ultimately show ?thesis using that by blast qed end lemma simply_connected_retraction_gen: shows "\simply_connected S; continuous_on S h; h ` S = T; continuous_on T k; k ` T \ S; \y. y \ T \ h(k y) = y\ \ simply_connected T" apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) apply (rule Retracts.homotopically_trivial_retraction_gen [of S h _ k _ "\p. pathfinish p = pathstart p" "\p. pathfinish p = pathstart p"]) apply (simp_all add: Retracts_def pathfinish_def pathstart_def) done lemma homeomorphic_simply_connected: "\S homeomorphic T; simply_connected S\ \ simply_connected T" by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) lemma homeomorphic_simply_connected_eq: "S homeomorphic T \ (simply_connected S \ simply_connected T)" by (metis homeomorphic_simply_connected homeomorphic_sym) subsection\Homotopy equivalence\ subsection\Homotopy equivalence of topological spaces.\ definition\<^marker>\tag important\ homotopy_equivalent_space (infix "homotopy'_equivalent'_space" 50) where "X homotopy_equivalent_space Y \ (\f g. continuous_map X Y f \ continuous_map Y X g \ homotopic_with (\x. True) X X (g \ f) id \ homotopic_with (\x. True) Y Y (f \ g) id)" lemma homeomorphic_imp_homotopy_equivalent_space: "X homeomorphic_space Y \ X homotopy_equivalent_space Y" unfolding homeomorphic_space_def homotopy_equivalent_space_def apply (erule ex_forward)+ by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def) lemma homotopy_equivalent_space_refl: "X homotopy_equivalent_space X" by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl) lemma homotopy_equivalent_space_sym: "X homotopy_equivalent_space Y \ Y homotopy_equivalent_space X" by (meson homotopy_equivalent_space_def) lemma homotopy_eqv_trans [trans]: assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U" shows "X homotopy_equivalent_space U" proof - obtain f1 g1 where f1: "continuous_map X Y f1" and g1: "continuous_map Y X g1" and hom1: "homotopic_with (\x. True) X X (g1 \ f1) id" "homotopic_with (\x. True) Y Y (f1 \ g1) id" using 1 by (auto simp: homotopy_equivalent_space_def) obtain f2 g2 where f2: "continuous_map Y U f2" and g2: "continuous_map U Y g2" and hom2: "homotopic_with (\x. True) Y Y (g2 \ f2) id" "homotopic_with (\x. True) U U (f2 \ g2) id" using 2 by (auto simp: homotopy_equivalent_space_def) have "homotopic_with (\f. True) X Y (g2 \ f2 \ f1) (id \ f1)" using f1 hom2(1) homotopic_with_compose_continuous_map_right by metis then have "homotopic_with (\f. True) X Y (g2 \ (f2 \ f1)) (id \ f1)" by (simp add: o_assoc) then have "homotopic_with (\x. True) X X (g1 \ (g2 \ (f2 \ f1))) (g1 \ (id \ f1))" by (simp add: g1 homotopic_with_compose_continuous_map_left) moreover have "homotopic_with (\x. True) X X (g1 \ id \ f1) id" using hom1 by simp ultimately have SS: "homotopic_with (\x. True) X X (g1 \ g2 \ (f2 \ f1)) id" by (metis comp_assoc homotopic_with_trans id_comp) have "homotopic_with (\f. True) U Y (f1 \ g1 \ g2) (id \ g2)" using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce then have "homotopic_with (\f. True) U Y (f1 \ (g1 \ g2)) (id \ g2)" by (simp add: o_assoc) then have "homotopic_with (\x. True) U U (f2 \ (f1 \ (g1 \ g2))) (f2 \ (id \ g2))" by (simp add: f2 homotopic_with_compose_continuous_map_left) moreover have "homotopic_with (\x. True) U U (f2 \ id \ g2) id" using hom2 by simp ultimately have UU: "homotopic_with (\x. True) U U (f2 \ f1 \ (g1 \ g2)) id" by (simp add: fun.map_comp hom2(2) homotopic_with_trans) show ?thesis unfolding homotopy_equivalent_space_def by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU) qed lemma deformation_retraction_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X (S \ r) id; retraction_maps X Y r S\ \ X homotopy_equivalent_space Y" unfolding homotopy_equivalent_space_def retraction_maps_def using homotopic_with_id2 by fastforce lemma deformation_retract_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X r id; retraction_maps X Y r id\ \ X homotopy_equivalent_space Y" using deformation_retraction_imp_homotopy_equivalent_space by force lemma deformation_retract_of_space: "S \ topspace X \ (\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` (topspace X) \ S)" proof (cases "S \ topspace X") case True moreover have "(\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ (S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` topspace X \ S))" unfolding retract_of_space_def proof safe fix f r assume f: "homotopic_with (\x. True) X X id f" and fS: "f ` topspace X \ S" and r: "continuous_map X (subtopology X S) r" and req: "\x\S. r x = x" show "\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id" proof (intro exI conjI) have "homotopic_with (\x. True) X X f r" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) X X (r \ f) (r \ id)" by (metis continuous_map_into_fulltopology f homotopic_with_compose_continuous_map_left homotopic_with_symD r) show "f x = (r \ f) x" if "x \ topspace X" for x using that fS req by auto qed auto then show "homotopic_with (\x. True) X X id r" by (rule homotopic_with_trans [OF f]) next show "retraction_maps X (subtopology X S) r id" by (simp add: r req retraction_maps_def) qed qed (use True in \auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\) ultimately show ?thesis by simp qed (auto simp: retract_of_space_def retraction_maps_def) subsection\Contractible spaces\ text\The definition (which agrees with "contractible" on subsets of Euclidean space) is a little cryptic because we don't in fact assume that the constant "a" is in the space. This forces the convention that the empty space / set is contractible, avoiding some special cases. \ definition contractible_space where "contractible_space X \ \a. homotopic_with (\x. True) X X id (\x. a)" lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \ contractible S" by (auto simp: contractible_space_def contractible_def) lemma contractible_space_empty: "topspace X = {} \ contractible_space X" unfolding contractible_space_def homotopic_with_def apply (rule_tac x=undefined in exI) apply (rule_tac x="\(t,x). if t = 0 then x else undefined" in exI) apply (auto simp: continuous_map_on_empty) done lemma contractible_space_singleton: "topspace X = {a} \ contractible_space X" unfolding contractible_space_def homotopic_with_def apply (rule_tac x=a in exI) apply (rule_tac x="\(t,x). if t = 0 then x else a" in exI) apply (auto intro: continuous_map_eq [where f = "\z. a"]) done lemma contractible_space_subset_singleton: "topspace X \ {a} \ contractible_space X" by (meson contractible_space_empty contractible_space_singleton subset_singletonD) lemma contractible_space_subtopology_singleton: "contractible_space(subtopology X {a})" by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI) lemma contractible_space: "contractible_space X \ topspace X = {} \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" proof (cases "topspace X = {}") case False then show ?thesis using homotopic_with_imp_continuous_maps by (fastforce simp: contractible_space_def) qed (simp add: contractible_space_empty) lemma contractible_imp_path_connected_space: assumes "contractible_space X" shows "path_connected_space X" proof (cases "topspace X = {}") case False have *: "path_connected_space X" if "a \ topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h" and h: "\x. h (0, x) = x" "\x. h (1, x) = a" for a and h :: "real \ 'a \ 'a" proof - have "path_component_of X b a" if "b \ topspace X" for b unfolding path_component_of_def proof (intro exI conjI) let ?g = "h \ (\x. (x,b))" show "pathin X ?g" unfolding pathin_def proof (rule continuous_map_compose [OF _ conth]) show "continuous_map (top_of_set {0..1}) (prod_topology (top_of_set {0..1}) X) (\x. (x, b))" using that by (auto intro!: continuous_intros) qed qed (use h in auto) then show ?thesis by (metis path_component_of_equiv path_connected_space_iff_path_component) qed show ?thesis using assms False by (auto simp: contractible_space homotopic_with_def *) qed (simp add: path_connected_space_topspace_empty) lemma contractible_imp_connected_space: "contractible_space X \ connected_space X" by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space) lemma contractible_space_alt: "contractible_space X \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" (is "?lhs = ?rhs") proof assume X: ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) show ?rhs proof show "homotopic_with (\x. True) X X id (\x. b)" if "b \ topspace X" for b proof (rule homotopic_with_trans [OF a]) show "homotopic_with (\x. True) X X (\x. a) (\x. b)" using homotopic_constant_maps path_connected_space_imp_path_component_of by (metis (full_types) X a continuous_map_const contractible_imp_path_connected_space homotopic_with_imp_continuous_maps that) qed qed next assume R: ?rhs then show ?lhs unfolding contractible_space_def by (metis equals0I homotopic_on_emptyI) qed lemma compose_const [simp]: "f \ (\x. a) = (\x. f a)" "(\x. a) \ g = (\x. a)" by (simp_all add: o_def) lemma nullhomotopic_through_contractible_space: assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Z (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with (\x. True) Y Y id (\x. b)" using Y by (auto simp: contractible_space_def) show thesis using homotopic_with_compose_continuous_map_right [OF homotopic_with_compose_continuous_map_left [OF b g] f] by (force simp add: that) qed lemma nullhomotopic_into_contractible_space: assumes f: "continuous_map X Y f" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF f _ Y] by (metis continuous_map_id id_comp) lemma nullhomotopic_from_contractible_space: assumes f: "continuous_map X Y f" and X: "contractible_space X" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF _ f X] by (metis comp_id continuous_map_id) lemma homotopy_dominated_contractibility: assumes f: "continuous_map X Y f" and g: "continuous_map Y X g" and hom: "homotopic_with (\x. True) Y Y (f \ g) id" and X: "contractible_space X" shows "contractible_space Y" proof - obtain c where c: "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_from_contractible_space [OF f X] . have "homotopic_with (\x. True) Y Y (f \ g) (\x. c)" using homotopic_with_compose_continuous_map_right [OF c g] by fastforce then have "homotopic_with (\x. True) Y Y id (\x. c)" using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast then show ?thesis unfolding contractible_space_def .. qed lemma homotopy_equivalent_space_contractibility: "X homotopy_equivalent_space Y \ (contractible_space X \ contractible_space Y)" unfolding homotopy_equivalent_space_def by (blast intro: homotopy_dominated_contractibility) lemma homeomorphic_space_contractibility: "X homeomorphic_space Y \ (contractible_space X \ contractible_space Y)" by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) +lemma homotopic_through_contractible_space: + "continuous_map X Y f \ + continuous_map X Y f' \ + continuous_map Y Z g \ + continuous_map Y Z g' \ + contractible_space Y \ path_connected_space Z + \ homotopic_with (\h. True) X Z (g \ f) (g' \ f')" + using nullhomotopic_through_contractible_space [of X Y f Z g] + using nullhomotopic_through_contractible_space [of X Y f' Z g'] + by (metis continuous_map_const homotopic_constant_maps homotopic_with_imp_continuous_maps + homotopic_with_trans path_connected_space_iff_path_component homotopic_with_sym) + +lemma homotopic_from_contractible_space: + "continuous_map X Y f \ continuous_map X Y g \ + contractible_space X \ path_connected_space Y + \ homotopic_with (\x. True) X Y f g" + by (metis comp_id continuous_map_id homotopic_through_contractible_space) + +lemma homotopic_into_contractible_space: + "continuous_map X Y f \ continuous_map X Y g \ + contractible_space Y + \ homotopic_with (\x. True) X Y f g" + by (metis continuous_map_id contractible_imp_path_connected_space homotopic_through_contractible_space id_comp) + lemma contractible_eq_homotopy_equivalent_singleton_subtopology: "contractible_space X \ topspace X = {} \ (\a \ topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs") proof (cases "topspace X = {}") case False show ?thesis proof assume ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) then have "a \ topspace X" by (metis False continuous_map_const homotopic_with_imp_continuous_maps) then have "homotopic_with (\x. True) (subtopology X {a}) (subtopology X {a}) id (\x. a)" using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce then have "X homotopy_equivalent_space subtopology X {a}" unfolding homotopy_equivalent_space_def using \a \ topspace X\ by (metis (full_types) a comp_id continuous_map_const continuous_map_id_subt empty_subsetI homotopic_with_symD id_comp insertI1 insert_subset topspace_subtopology_subset) with \a \ topspace X\ show ?rhs by blast next assume ?rhs then show ?lhs by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility) qed qed (simp add: contractible_space_empty) lemma contractible_space_retraction_map_image: assumes "retraction_map X Y f" and X: "contractible_space X" shows "contractible_space Y" proof - obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\y \ topspace Y. f(g y) = y" using assms by (auto simp: retraction_map_def retraction_maps_def) obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" using X by (auto simp: contractible_space_def) have "homotopic_with (\x. True) Y Y id (\x. f a)" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) Y Y (f \ id \ g) (f \ (\x. a) \ g)" using f g a homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right by metis qed (use fg in auto) then show ?thesis unfolding contractible_space_def by blast qed lemma contractible_space_prod_topology: "contractible_space(prod_topology X Y) \ topspace X = {} \ topspace Y = {} \ contractible_space X \ contractible_space Y" proof (cases "topspace X = {} \ topspace Y = {}") case True then have "topspace (prod_topology X Y) = {}" by simp then show ?thesis by (auto simp: contractible_space_empty) next case False have "contractible_space(prod_topology X Y) \ contractible_space X \ contractible_space Y" proof safe assume XY: "contractible_space (prod_topology X Y)" with False have "retraction_map (prod_topology X Y) X fst" by (auto simp: contractible_space False retraction_map_fst) then show "contractible_space X" by (rule contractible_space_retraction_map_image [OF _ XY]) have "retraction_map (prod_topology X Y) Y snd" using False XY by (auto simp: contractible_space False retraction_map_snd) then show "contractible_space Y" by (rule contractible_space_retraction_map_image [OF _ XY]) next assume "contractible_space X" and "contractible_space Y" with False obtain a b where "a \ topspace X" and a: "homotopic_with (\x. True) X X id (\x. a)" and "b \ topspace Y" and b: "homotopic_with (\x. True) Y Y id (\x. b)" by (auto simp: contractible_space) with False show "contractible_space (prod_topology X Y)" apply (simp add: contractible_space) apply (rule_tac x=a in bexI) apply (rule_tac x=b in bexI) using homotopic_with_prod_topology [OF a b] apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff) apply auto done qed with False show ?thesis by auto qed lemma contractible_space_product_topology: "contractible_space(product_topology X I) \ topspace (product_topology X I) = {} \ (\i \ I. contractible_space(X i))" proof (cases "topspace (product_topology X I) = {}") case False have 1: "contractible_space (X i)" if XI: "contractible_space (product_topology X I)" and "i \ I" for i proof (rule contractible_space_retraction_map_image [OF _ XI]) show "retraction_map (product_topology X I) (X i) (\x. x i)" using False by (simp add: retraction_map_product_projection \i \ I\) qed have 2: "contractible_space (product_topology X I)" if "x \ topspace (product_topology X I)" and cs: "\i\I. contractible_space (X i)" for x :: "'a \ 'b" proof - obtain f where f: "\i. i\I \ homotopic_with (\x. True) (X i) (X i) id (\x. f i)" using cs unfolding contractible_space_def by metis have "homotopic_with (\x. True) (product_topology X I) (product_topology X I) id (\x. restrict f I)" by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto) then show ?thesis by (auto simp: contractible_space_def) qed show ?thesis using False 1 2 by blast qed (simp add: contractible_space_empty) lemma contractible_space_subtopology_euclideanreal [simp]: "contractible_space(subtopology euclideanreal S) \ is_interval S" (is "?lhs = ?rhs") proof assume ?lhs then have "path_connectedin (subtopology euclideanreal S) S" using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute by (simp add: contractible_imp_path_connected) then show ?rhs by (simp add: is_interval_path_connected_1) next assume ?rhs then have "convex S" by (simp add: is_interval_convex_1) show ?lhs proof (cases "S = {}") case False then obtain z where "z \ S" by blast show ?thesis unfolding contractible_space_def homotopic_with_def proof (intro exI conjI allI) note \
= convexD [OF \convex S\, simplified] show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S) (\(t,x). (1 - t) * x + t * z)" using \z \ S\ by (auto simp: case_prod_unfold intro!: continuous_intros \
) qed auto qed (simp add: contractible_space_empty) qed corollary contractible_space_euclideanreal: "contractible_space euclideanreal" proof - have "contractible_space (subtopology euclideanreal UNIV)" using contractible_space_subtopology_euclideanreal by blast then show ?thesis by simp qed abbreviation\<^marker>\tag important\ homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" (infix "homotopy'_eqv" 50) where "S homotopy_eqv T \ top_of_set S homotopy_equivalent_space top_of_set T" lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def by (metis continuous_map_subtopology_eu homotopic_with_id2 openin_imp_subset openin_subtopology_self topspace_euclidean_subtopology) lemma homotopy_eqv_inj_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(f ` S) homotopy_eqv S" by (metis assms homeomorphic_sym homeomorphic_imp_homotopy_eqv linear_homeomorphic_image) lemma homotopy_eqv_translation: fixes S :: "'a::real_normed_vector set" shows "(+) a ` S homotopy_eqv S" using homeomorphic_imp_homotopy_eqv homeomorphic_translation homeomorphic_sym by blast lemma homotopy_eqv_homotopic_triviality_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on U f" "f ` U \ T" and g: "continuous_on U g" "g ` U \ T" and homUS: "\f g. \continuous_on U f; f ` U \ S; continuous_on U g; g ` U \ S\ \ homotopic_with_canon (\x. True) U S f g" shows "homotopic_with_canon (\x. True) U T f g" proof - obtain h k where h: "continuous_on S h" "h ` S \ T" and k: "continuous_on T k" "k ` T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) have "homotopic_with_canon (\f. True) U S (k \ f) (k \ g)" proof (rule homUS) show "continuous_on U (k \ f)" "continuous_on U (k \ g)" using continuous_on_compose continuous_on_subset f g k by blast+ qed (use f g k in \(force simp: o_def)+\ ) then have "homotopic_with_canon (\x. True) U T (h \ (k \ f)) (h \ (k \ g))" by (rule homotopic_with_compose_continuous_left; simp add: h) moreover have "homotopic_with_canon (\x. True) U T (h \ k \ f) (id \ f)" by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom f) moreover have "homotopic_with_canon (\x. True) U T (h \ k \ g) (id \ g)" by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom g) ultimately show "homotopic_with_canon (\x. True) U T f g" unfolding o_assoc by (metis homotopic_with_trans homotopic_with_sym id_comp) qed lemma homotopy_eqv_homotopic_triviality: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f g. continuous_on U f \ f ` U \ S \ continuous_on U g \ g ` U \ S \ homotopic_with_canon (\x. True) U S f g) \ (\f g. continuous_on U f \ f ` U \ T \ continuous_on U g \ g ` U \ T \ homotopic_with_canon (\x. True) U T f g)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis assms homotopy_eqv_homotopic_triviality_imp) next assume ?rhs moreover have "T homotopy_eqv S" using assms homotopy_equivalent_space_sym by blast ultimately show ?lhs by (blast intro: homotopy_eqv_homotopic_triviality_imp) qed lemma homotopy_eqv_cohomotopic_triviality_null_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on T f" "f ` T \ U" and homSU: "\f. \continuous_on S f; f ` S \ U\ \ \c. homotopic_with_canon (\x. True) S U f (\x. c)" obtains c where "homotopic_with_canon (\x. True) T U f (\x. c)" proof - obtain h k where h: "continuous_on S h" "h ` S \ T" and k: "continuous_on T k" "k ` T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) obtain c where "homotopic_with_canon (\x. True) S U (f \ h) (\x. c)" proof (rule exE [OF homSU]) show "continuous_on S (f \ h)" using continuous_on_compose continuous_on_subset f h by blast qed (use f h in force) then have "homotopic_with_canon (\x. True) T U ((f \ h) \ k) ((\x. c) \ k)" by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto) moreover have "homotopic_with_canon (\x. True) T U (f \ id) (f \ (h \ k))" by (rule homotopic_with_compose_continuous_left [where Y=T]) (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis using that homotopic_with_trans by (fastforce simp add: o_def) qed lemma homotopy_eqv_cohomotopic_triviality_null: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f. continuous_on S f \ f ` S \ U \ (\c. homotopic_with_canon (\x. True) S U f (\x. c))) \ (\f. continuous_on T f \ f ` T \ U \ (\c. homotopic_with_canon (\x. True) T U f (\x. c)))" by (rule iffI; metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym) text \Similar to the proof above\ lemma homotopy_eqv_homotopic_triviality_null_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on U f" "f ` U \ T" and homSU: "\f. \continuous_on U f; f ` U \ S\ \ \c. homotopic_with_canon (\x. True) U S f (\x. c)" shows "\c. homotopic_with_canon (\x. True) U T f (\x. c)" proof - obtain h k where h: "continuous_on S h" "h ` S \ T" and k: "continuous_on T k" "k ` T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) obtain c::'a where "homotopic_with_canon (\x. True) U S (k \ f) (\x. c)" proof (rule exE [OF homSU [of "k \ f"]]) show "continuous_on U (k \ f)" using continuous_on_compose continuous_on_subset f k by blast qed (use f k in force) then have "homotopic_with_canon (\x. True) U T (h \ (k \ f)) (h \ (\x. c))" by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto) moreover have "homotopic_with_canon (\x. True) U T (id \ f) ((h \ k) \ f)" by (rule homotopic_with_compose_continuous_right [where X=T]) (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis using homotopic_with_trans by (fastforce simp add: o_def) qed lemma homotopy_eqv_homotopic_triviality_null: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f. continuous_on U f \ f ` U \ S \ (\c. homotopic_with_canon (\x. True) U S f (\x. c))) \ (\f. continuous_on U f \ f ` U \ T \ (\c. homotopic_with_canon (\x. True) U T f (\x. c)))" by (rule iffI; metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym) lemma homotopy_eqv_contractible_sets: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "contractible S" "contractible T" "S = {} \ T = {}" shows "S homotopy_eqv T" proof (cases "S = {}") case True with assms show ?thesis by (simp add: homeomorphic_imp_homotopy_eqv) next case False with assms obtain a b where "a \ S" "b \ T" by auto then show ?thesis unfolding homotopy_equivalent_space_def apply (rule_tac x="\x. b" in exI, rule_tac x="\x. a" in exI) apply (intro assms conjI continuous_on_id' homotopic_into_contractible; force) done qed lemma homotopy_eqv_empty1 [simp]: fixes S :: "'a::real_normed_vector set" shows "S homotopy_eqv ({}::'b::real_normed_vector set) \ S = {}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_map_subtopology_eu empty_iff equalityI homotopy_equivalent_space_def image_subset_iff subsetI) qed (simp add: homotopy_eqv_contractible_sets) lemma homotopy_eqv_empty2 [simp]: fixes S :: "'a::real_normed_vector set" shows "({}::'b::real_normed_vector set) homotopy_eqv S \ S = {}" using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast lemma homotopy_eqv_contractibility: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ (contractible S \ contractible T)" by (meson contractible_space_top_of_set homotopy_equivalent_space_contractibility) lemma homotopy_eqv_sing: fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" shows "S homotopy_eqv {a} \ S \ {} \ contractible S" proof (cases "S = {}") case False then show ?thesis by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets) qed simp lemma homeomorphic_contractible_eq: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homeomorphic T \ (contractible S \ contractible T)" by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) lemma homeomorphic_contractible: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "\contractible S; S homeomorphic T\ \ contractible T" by (metis homeomorphic_contractible_eq) subsection\<^marker>\tag unimportant\\Misc other results\ lemma bounded_connected_Compl_real: fixes S :: "real set" assumes "bounded S" and conn: "connected(- S)" shows "S = {}" proof - obtain a b where "S \ box a b" by (meson assms bounded_subset_box_symmetric) then have "a \ S" "b \ S" by auto then have "\x. a \ x \ x \ b \ x \ - S" by (meson Compl_iff conn connected_iff_interval) then show ?thesis using \S \ box a b\ by auto qed corollary bounded_path_connected_Compl_real: fixes S :: "real set" assumes "bounded S" "path_connected(- S)" shows "S = {}" by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected) lemma bounded_connected_Compl_1: fixes S :: "'a::{euclidean_space} set" assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" shows "S = {}" proof - have "DIM('a) = DIM(real)" by (simp add: "1") then obtain f::"'a \ real" and g where "linear f" "\x. norm(f x) = norm x" and fg: "\x. g(f x) = x" "\y. f(g y) = y" by (rule isomorphisms_UNIV_UNIV) blast with \bounded S\ have "bounded (f ` S)" using bounded_linear_image linear_linear by blast have "bij f" by (metis fg bijI') have "connected (f ` (-S))" using connected_linear_image assms \linear f\ by blast moreover have "f ` (-S) = - (f ` S)" by (simp add: \bij f\ bij_image_Compl_eq) finally have "connected (- (f ` S))" by simp then have "f ` S = {}" using \bounded (f ` S)\ bounded_connected_Compl_real by blast then show ?thesis by blast qed lemma connected_card_eq_iff_nontrivial: fixes S :: "'a::metric_space set" shows "connected S \ uncountable S \ \(\a. S \ {a})" by (metis connected_uncountable finite.emptyI finite.insertI rev_finite_subset singleton_iff subsetI uncountable_infinite) lemma connected_finite_iff_sing: fixes S :: "'a::metric_space set" assumes "connected S" shows "finite S \ S = {} \ (\a. S = {a})" using assms connected_uncountable countable_finite by blast subsection\<^marker>\tag unimportant\\ Some simple positive connection theorems\ proposition path_connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "path_connected(U - S)" proof (clarsimp simp: path_connected_def) fix a b assume "a \ U" "a \ S" "b \ U" "b \ S" let ?m = "midpoint a b" show "\g. path g \ path_image g \ U - S \ pathstart g = a \ pathfinish g = b" proof (cases "a = b") case True then show ?thesis by (metis DiffI \a \ U\ \a \ S\ path_component_def path_component_refl) next case False then have "a \ ?m" "b \ ?m" using midpoint_eq_endpoint by fastforce+ have "?m \ U" using \a \ U\ \b \ U\ \convex U\ convex_contains_segment by force obtain c where "c \ U" and nc_abc: "\ collinear {a,b,c}" by (metis False \a \ U\ \b \ U\ \\ collinear U\ collinear_triples insert_absorb) have ncoll_mca: "\ collinear {?m,c,a}" by (metis (full_types) \a \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have ncoll_mcb: "\ collinear {?m,c,b}" by (metis (full_types) \b \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have "c \ ?m" by (metis collinear_midpoint insert_commute nc_abc) then have "closed_segment ?m c \ U" by (simp add: \c \ U\ \?m \ U\ \convex U\ closed_segment_subset) then obtain z where z: "z \ closed_segment ?m c" and disjS: "(closed_segment a z \ closed_segment z b) \ S = {}" proof - have False if "closed_segment ?m c \ {z. (closed_segment a z \ closed_segment z b) \ S \ {}}" proof - have closb: "closed_segment ?m c \ {z \ closed_segment ?m c. closed_segment a z \ S \ {}} \ {z \ closed_segment ?m c. closed_segment z b \ S \ {}}" using that by blast have *: "countable {z \ closed_segment ?m c. closed_segment z u \ S \ {}}" if "u \ U" "u \ S" and ncoll: "\ collinear {?m, c, u}" for u proof - have **: False if x1: "x1 \ closed_segment ?m c" and x2: "x2 \ closed_segment ?m c" and "x1 \ x2" "x1 \ u" and w: "w \ closed_segment x1 u" "w \ closed_segment x2 u" and "w \ S" for x1 x2 w proof - have "x1 \ affine hull {?m,c}" "x2 \ affine hull {?m,c}" using segment_as_ball x1 x2 by auto then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) have "\ collinear {x1, u, x2}" proof assume "collinear {x1, u, x2}" then have "collinear {?m, c, u}" by (metis (full_types) \c \ ?m\ coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \x1 \ x2\) with ncoll show False .. qed then have "closed_segment x1 u \ closed_segment u x2 = {u}" by (blast intro!: Int_closed_segment) then have "w = u" using closed_segment_commute w by auto show ?thesis using \u \ S\ \w = u\ that(7) by auto qed then have disj: "disjoint ((\z\closed_segment ?m c. {closed_segment z u \ S}))" by (fastforce simp: pairwise_def disjnt_def) have cou: "countable ((\z \ closed_segment ?m c. {closed_segment z u \ S}) - {{}})" apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) apply (rule countable_subset [OF _ \countable S\], auto) done define f where "f \ \X. (THE z. z \ closed_segment ?m c \ X = closed_segment z u \ S)" show ?thesis proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) fix x assume x: "x \ closed_segment ?m c" "closed_segment x u \ S \ {}" show "x \ f ` ((\z\closed_segment ?m c. {closed_segment z u \ S}) - {{}})" proof (rule_tac x="closed_segment x u \ S" in image_eqI) show "x = f (closed_segment x u \ S)" unfolding f_def by (rule the_equality [symmetric]) (use x in \auto dest: **\) qed (use x in auto) qed qed have "uncountable (closed_segment ?m c)" by (metis \c \ ?m\ uncountable_closed_segment) then show False using closb * [OF \a \ U\ \a \ S\ ncoll_mca] * [OF \b \ U\ \b \ S\ ncoll_mcb] by (simp add: closed_segment_commute countable_subset) qed then show ?thesis by (force intro: that) qed show ?thesis proof (intro exI conjI) have "path_image (linepath a z +++ linepath z b) \ U" by (metis \a \ U\ \b \ U\ \closed_segment ?m c \ U\ z \convex U\ closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) with disjS show "path_image (linepath a z +++ linepath z b) \ U - S" by (force simp: path_image_join) qed auto qed qed corollary connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "connected(U - S)" by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) lemma path_connected_punctured_convex: assumes "convex S" and aff: "aff_dim S \ 1" shows "path_connected(S - {a})" proof - consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \ 2" using assms aff_dim_geq [of S] by linarith then show ?thesis proof cases assume "aff_dim S = -1" then show ?thesis by (metis aff_dim_empty empty_Diff path_connected_empty) next assume "aff_dim S = 0" then show ?thesis by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) next assume ge2: "aff_dim S \ 2" then have "\ collinear S" proof (clarsimp simp: collinear_affine_hull) fix u v assume "S \ affine hull {u, v}" then have "aff_dim S \ aff_dim {u, v}" by (metis (no_types) aff_dim_affine_hull aff_dim_subset) with ge2 show False by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) qed moreover have "countable {a}" by simp ultimately show ?thesis by (metis path_connected_convex_diff_countable [OF \convex S\]) qed qed lemma connected_punctured_convex: shows "\convex S; aff_dim S \ 1\ \ connected(S - {a})" using path_connected_imp_connected path_connected_punctured_convex by blast lemma path_connected_complement_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "countable S" shows "path_connected(- S)" proof - have "\ collinear (UNIV::'a set)" using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) then have "path_connected(UNIV - S)" by (simp add: \countable S\ path_connected_convex_diff_countable) then show ?thesis by (simp add: Compl_eq_Diff_UNIV) qed proposition path_connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "path_connected(S - T)" proof (clarsimp simp: path_connected_component) fix x y assume xy: "x \ S" "x \ T" "y \ S" "y \ T" show "path_component (S - T) x y" proof (rule connected_equivalence_relation_gen [OF \connected S\, where P = "\x. x \ T"]) show "\z. z \ U \ z \ T" if opeU: "openin (top_of_set S) U" and "x \ U" for U x proof - have "openin (top_of_set (affine hull S)) U" using opeU ope openin_trans by blast with \x \ U\ obtain r where Usub: "U \ affine hull S" and "r > 0" and subU: "ball x r \ affine hull S \ U" by (auto simp: openin_contains_ball) with \x \ U\ have x: "x \ ball x r \ affine hull S" by auto have "\ S \ {x}" using \\ collinear S\ collinear_subset by blast then obtain x' where "x' \ x" "x' \ S" by blast obtain y where y: "y \ x" "y \ ball x r \ affine hull S" proof show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \ x" using \x' \ x\ \r > 0\ by auto show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \ ball x r \ affine hull S" using \x' \ x\ \r > 0\ \x' \ S\ x by (simp add: dist_norm mem_affine_3_minus hull_inc) qed have "convex (ball x r \ affine hull S)" by (simp add: affine_imp_convex convex_Int) with x y subU have "uncountable U" by (meson countable_subset uncountable_convex) then have "\ U \ T" using \countable T\ countable_subset by blast then show ?thesis by blast qed show "\U. openin (top_of_set S) U \ x \ U \ (\x\U. \y\U. x \ T \ y \ T \ path_component (S - T) x y)" if "x \ S" for x proof - obtain r where Ssub: "S \ affine hull S" and "r > 0" and subS: "ball x r \ affine hull S \ S" using ope \x \ S\ by (auto simp: openin_contains_ball) then have conv: "convex (ball x r \ affine hull S)" by (simp add: affine_imp_convex convex_Int) have "\ aff_dim (affine hull S) \ 1" using \\ collinear S\ collinear_aff_dim by auto then have "\ aff_dim (ball x r \ affine hull S) \ 1" by (metis (no_types, opaque_lifting) aff_dim_convex_Int_open IntI open_ball \0 < r\ aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) then have "\ collinear (ball x r \ affine hull S)" by (simp add: collinear_aff_dim) then have *: "path_connected ((ball x r \ affine hull S) - T)" by (rule path_connected_convex_diff_countable [OF conv _ \countable T\]) have ST: "ball x r \ affine hull S - T \ S - T" using subS by auto show ?thesis proof (intro exI conjI) show "x \ ball x r \ affine hull S" using \x \ S\ \r > 0\ by (simp add: hull_inc) have "openin (top_of_set (affine hull S)) (ball x r \ affine hull S)" by (subst inf.commute) (simp add: openin_Int_open) then show "openin (top_of_set S) (ball x r \ affine hull S)" by (rule openin_subset_trans [OF _ subS Ssub]) qed (use * path_component_trans in \auto simp: path_connected_component path_component_of_subset [OF ST]\) qed qed (use xy path_component_trans in auto) qed corollary connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "connected(S - T)" by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) corollary path_connected_open_diff_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "open S" "connected S" "countable T" shows "path_connected(S - T)" proof (cases "S = {}") case True then show ?thesis by (simp) next case False show ?thesis proof (rule path_connected_openin_diff_countable) show "openin (top_of_set (affine hull S)) S" by (simp add: assms hull_subset open_subset) show "\ collinear S" using assms False by (simp add: collinear_aff_dim aff_dim_open) qed (simp_all add: assms) qed corollary connected_open_diff_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "open S" "connected S" "countable T" shows "connected(S - T)" by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) subsection\<^marker>\tag unimportant\ \Self-homeomorphisms shuffling points about\ subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_moving_points_exists\\ lemma homeomorphism_moving_point_1: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" "f a = u" "\x. x \ sphere a r \ f x = x" proof - have nou: "norm (u - a) < r" and "u \ T" using u by (auto simp: dist_norm norm_minus_commute) then have "0 < r" by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) define f where "f \ \x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a proof - have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" using eq by (simp add: algebra_simps) then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" by (metis diff_divide_distrib) also have "\ \ norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" using norm_triangle_ineq by blast also have "\ = norm y + (norm x - norm y) * (norm u / r)" using yx \r > 0\ by (simp add: field_split_simps) also have "\ < norm y + (norm x - norm y) * 1" proof (subst add_less_cancel_left) show "(norm x - norm y) * (norm u / r) < (norm x - norm y) * 1" proof (rule mult_strict_left_mono) show "norm u / r < 1" using \0 < r\ divide_less_eq_1_pos nou by blast qed (simp add: yx) qed also have "\ = norm x" by simp finally show False by simp qed have "inj f" unfolding f_def proof (clarsimp simp: inj_on_def) fix x y assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" by (auto simp: algebra_simps) show "x=y" proof (cases "norm (x - a) = norm (y - a)") case True then show ?thesis using eq by auto next case False then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" by linarith then have "False" proof cases case 1 show False using * [OF _ nou 1] eq by simp next case 2 with * [OF eq nou] show False by auto qed then show "x=y" .. qed qed then have inj_onf: "inj_on f (cball a r \ T)" using inj_on_Int by fastforce have contf: "continuous_on (cball a r \ T) f" unfolding f_def using \0 < r\ by (intro continuous_intros) blast have fim: "f ` (cball a r \ T) = cball a r \ T" proof have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \ r" if "norm y \ r" "norm u < r" for y u::'a proof - have "norm (y + (1 - norm y / r) *\<^sub>R u) \ norm y + norm((1 - norm y / r) *\<^sub>R u)" using norm_triangle_ineq by blast also have "\ = norm y + abs(1 - norm y / r) * norm u" by simp also have "\ \ r" proof - have "(r - norm u) * (r - norm y) \ 0" using that by auto then have "r * norm u + r * norm y \ r * r + norm u * norm y" by (simp add: algebra_simps) then show ?thesis using that \0 < r\ by (simp add: abs_if field_simps) qed finally show ?thesis . qed have "f ` (cball a r) \ cball a r" using * nou apply (clarsimp simp: dist_norm norm_minus_commute f_def) by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute) moreover have "f ` T \ T" unfolding f_def using \affine T\ \a \ T\ \u \ T\ by (force simp: add.commute mem_affine_3_minus) ultimately show "f ` (cball a r \ T) \ cball a r \ T" by blast next show "cball a r \ T \ f ` (cball a r \ T)" proof (clarsimp simp: dist_norm norm_minus_commute) fix x assume x: "norm (x - a) \ r" and "x \ T" have "\v \ {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \ 1 = 0" by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) then obtain v where "0 \ v" "v \ 1" and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" by auto then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v" by (simp add: field_simps norm_minus_commute) show "x \ f ` (cball a r \ T)" proof (rule image_eqI) show "x = f (x - v *\<^sub>R (u - a))" using \r > 0\ v by (simp add: f_def) (simp add: field_simps) have "x - v *\<^sub>R (u - a) \ cball a r" using \r > 0\\0 \ v\ by (simp add: dist_norm n) moreover have "x - v *\<^sub>R (u - a) \ T" by (simp add: f_def \u \ T\ \x \ T\ assms mem_affine_3_minus2) ultimately show "x - v *\<^sub>R (u - a) \ cball a r \ T" by blast qed qed qed have "compact (cball a r \ T)" by (simp add: affine_closed compact_Int_closed \affine T\) then obtain g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" by (metis homeomorphism_compact [OF _ contf fim inj_onf]) then show thesis apply (rule_tac f=f in that) using \r > 0\ by (simp_all add: f_def dist_norm norm_minus_commute) qed corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_2: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" "f u = v" "\x. \x \ sphere a r; x \ T\ \ f x = x" proof - have "0 < r" by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) obtain f1 g1 where hom1: "homeomorphism (cball a r \ T) (cball a r \ T) f1 g1" and "f1 a = u" and f1: "\x. x \ sphere a r \ f1 x = x" using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ u] by blast obtain f2 g2 where hom2: "homeomorphism (cball a r \ T) (cball a r \ T) f2 g2" and "f2 a = v" and f2: "\x. x \ sphere a r \ f2 x = x" using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ v] by blast show ?thesis proof show "homeomorphism (cball a r \ T) (cball a r \ T) (f2 \ g1) (f1 \ g2)" by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) have "g1 u = a" using \0 < r\ \f1 a = u\ assms hom1 homeomorphism_apply1 by fastforce then show "(f2 \ g1) u = v" by (simp add: \f2 a = v\) show "\x. \x \ sphere a r; x \ T\ \ (f2 \ g1) x = x" using f1 f2 hom1 homeomorphism_apply1 by fastforce qed qed corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_3: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism S S f g" "f u = v" "{x. \ (f x = x \ g x = x)} \ ball a r \ T" proof - obtain f g where hom: "homeomorphism (cball a r \ T) (cball a r \ T) f g" and "f u = v" and fid: "\x. \x \ sphere a r; x \ T\ \ f x = x" using homeomorphism_moving_point_2 [OF \affine T\ \a \ T\ u v] by blast have gid: "\x. \x \ sphere a r; x \ T\ \ g x = x" using fid hom homeomorphism_apply1 by fastforce define ff where "ff \ \x. if x \ ball a r \ T then f x else x" define gg where "gg \ \x. if x \ ball a r \ T then g x else x" show ?thesis proof show "homeomorphism S S ff gg" proof (rule homeomorphismI) have "continuous_on ((cball a r \ T) \ (T - ball a r)) ff" unfolding ff_def using homeomorphism_cont1 [OF hom] by (intro continuous_on_cases) (auto simp: affine_closed \affine T\ fid) then show "continuous_on S ff" by (rule continuous_on_subset) (use ST in auto) have "continuous_on ((cball a r \ T) \ (T - ball a r)) gg" unfolding gg_def using homeomorphism_cont2 [OF hom] by (intro continuous_on_cases) (auto simp: affine_closed \affine T\ gid) then show "continuous_on S gg" by (rule continuous_on_subset) (use ST in auto) show "ff ` S \ S" proof (clarsimp simp: ff_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "f x \ cball a r \ T" using homeomorphism_image1 [OF hom] by force then show "f x \ S" using ST(1) \x \ T\ gid hom homeomorphism_def x by fastforce qed show "gg ` S \ S" proof (clarsimp simp: gg_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "g x \ cball a r \ T" using homeomorphism_image2 [OF hom] by force then have "g x \ ball a r" using homeomorphism_apply2 [OF hom] by (metis Diff_Diff_Int Diff_iff \x \ T\ cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) then show "g x \ S" using ST(1) \g x \ cball a r \ T\ by force qed show "\x. x \ S \ gg (ff x) = x" unfolding ff_def gg_def using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by simp (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) show "\x. x \ S \ ff (gg x) = x" unfolding ff_def gg_def using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by simp (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) qed show "ff u = v" using u by (auto simp: ff_def \f u = v\) show "{x. \ (ff x = x \ gg x = x)} \ ball a r \ T" by (auto simp: ff_def gg_def) qed qed proposition\<^marker>\tag unimportant\ homeomorphism_moving_point: fixes a :: "'a::euclidean_space" assumes ope: "openin (top_of_set (affine hull S)) S" and "S \ T" and TS: "T \ affine hull S" and S: "connected S" "a \ S" "b \ S" obtains f g where "homeomorphism T T f g" "f a = b" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" proof - have 1: "\h k. homeomorphism T T h k \ h (f d) = d \ {x. \ (h x = x \ k x = x)} \ S \ bounded {x. \ (h x = x \ k x = x)}" if "d \ S" "f d \ S" and homfg: "homeomorphism T T f g" and S: "{x. \ (f x = x \ g x = x)} \ S" and bo: "bounded {x. \ (f x = x \ g x = x)}" for d f g proof (intro exI conjI) show homgf: "homeomorphism T T g f" by (metis homeomorphism_symD homfg) then show "g (f d) = d" by (meson \S \ T\ homeomorphism_def subsetD \d \ S\) show "{x. \ (g x = x \ f x = x)} \ S" using S by blast show "bounded {x. \ (g x = x \ f x = x)}" using bo by (simp add: conj_commute) qed have 2: "\f g. homeomorphism T T f g \ f x = f2 (f1 x) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" if "x \ S" "f1 x \ S" "f2 (f1 x) \ S" and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" and sub: "{x. \ (f1 x = x \ g1 x = x)} \ S" "{x. \ (f2 x = x \ g2 x = x)} \ S" and bo: "bounded {x. \ (f1 x = x \ g1 x = x)}" "bounded {x. \ (f2 x = x \ g2 x = x)}" for x f1 f2 g1 g2 proof (intro exI conjI) show homgf: "homeomorphism T T (f2 \ f1) (g1 \ g2)" by (metis homeomorphism_compose hom) then show "(f2 \ f1) x = f2 (f1 x)" by force show "{x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)} \ S" using sub by force have "bounded ({x. \(f1 x = x \ g1 x = x)} \ {x. \(f2 x = x \ g2 x = x)})" using bo by simp then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" by (rule bounded_subset) auto qed have 3: "\U. openin (top_of_set S) U \ d \ U \ (\x\U. \f g. homeomorphism T T f g \ f d = x \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)})" if "d \ S" for d proof - obtain r where "r > 0" and r: "ball d r \ affine hull S \ S" by (metis \d \ S\ ope openin_contains_ball) have *: "\f g. homeomorphism T T f g \ f d = e \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" if "e \ S" "e \ ball d r" for e apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) using r \S \ T\ TS that apply (auto simp: \d \ S\ \0 < r\ hull_inc) using bounded_subset by blast show ?thesis by (rule_tac x="S \ ball d r" in exI) (fastforce simp: openin_open_Int \0 < r\ that intro: *) qed have "\f g. homeomorphism T T f g \ f a = b \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" by (rule connected_equivalence_relation [OF S]; blast intro: 1 2 3) then show ?thesis using that by auto qed lemma homeomorphism_moving_points_exists_gen: assumes K: "finite K" "\i. i \ K \ x i \ S \ y i \ S" "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and "2 \ aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" and "S \ T" "T \ affine hull S" "connected S" shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using assms proof (induction K) case empty then show ?case by (force simp: homeomorphism_ident) next case (insert i K) then have xney: "\j. \j \ K; j \ i\ \ x i \ x j \ y i \ y j" and pw: "pairwise (\i j. x i \ x j \ y i \ y j) K" and "x i \ S" "y i \ S" and xyS: "\i. i \ K \ x i \ S \ y i \ S" by (simp_all add: pairwise_insert) obtain f g where homfg: "homeomorphism T T f g" and feq: "\i. i \ K \ f(x i) = y i" and fg_sub: "{x. \ (f x = x \ g x = x)} \ S" and bo_fg: "bounded {x. \ (f x = x \ g x = x)}" using insert.IH [OF xyS pw] insert.prems by (blast intro: that) then have "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using insert by blast have aff_eq: "affine hull (S - y ` K) = affine hull S" proof (rule affine_hull_Diff [OF ope]) show "finite (y ` K)" by (simp add: insert.hyps(1)) show "y ` K \ S" using \y i \ S\ insert.hyps(2) xney xyS by fastforce qed have f_in_S: "f x \ S" if "x \ S" for x using homfg fg_sub homeomorphism_apply1 \S \ T\ proof - have "(f (f x) \ f x \ g (f x) \ f x) \ f x \ S" by (metis \S \ T\ homfg subsetD homeomorphism_apply1 that) then show ?thesis using fg_sub by force qed obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" and hk_sub: "{x. \ (h x = x \ k x = x)} \ S - y ` K" and bo_hk: "bounded {x. \ (h x = x \ k x = x)}" proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)" by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) show "S - y ` K \ T" using \S \ T\ by auto show "T \ affine hull (S - y ` K)" using insert by (simp add: aff_eq) show "connected (S - y ` K)" proof (rule connected_openin_diff_countable [OF \connected S\ ope]) show "\ collinear S" using collinear_aff_dim \2 \ aff_dim S\ by force show "countable (y ` K)" using countable_finite insert.hyps(1) by blast qed have "\k. \f (x i) = y k; k \ K\ \ False" by (metis feq homfg \x i \ S\ homeomorphism_def \S \ T\ \i \ K\ subsetCE xney xyS) then show "f (x i) \ S - y ` K" by (auto simp: f_in_S \x i \ S\) show "y i \ S - y ` K" using insert.hyps xney by (auto simp: \y i \ S\) qed blast show ?case proof (intro exI conjI) show "homeomorphism T T (h \ f) (g \ k)" using homfg homhk homeomorphism_compose by blast show "\i \ insert i K. (h \ f) (x i) = y i" using feq hk_sub by (auto simp: heq) show "{x. \ ((h \ f) x = x \ (g \ k) x = x)} \ S" using fg_sub hk_sub by force have "bounded ({x. \(f x = x \ g x = x)} \ {x. \(h x = x \ k x = x)})" using bo_fg bo_hk bounded_Un by blast then show "bounded {x. \ ((h \ f) x = x \ (g \ k) x = x)}" by (rule bounded_subset) auto qed qed proposition\<^marker>\tag unimportant\ homeomorphism_moving_points_exists: fixes S :: "'a::euclidean_space set" assumes 2: "2 \ DIM('a)" "open S" "connected S" "S \ T" "finite K" and KS: "\i. i \ K \ x i \ S \ y i \ S" and pw: "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "\i. i \ K \ f(x i) = y i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof (cases "S = {}") case True then show ?thesis using KS homeomorphism_ident that by fastforce next case False then have affS: "affine hull S = UNIV" by (simp add: affine_hull_open \open S\) then have ope: "openin (top_of_set (affine hull S)) S" using \open S\ open_openin by auto have "2 \ DIM('a)" by (rule 2) also have "\ = aff_dim (UNIV :: 'a set)" by simp also have "\ \ aff_dim S" by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) finally have "2 \ aff_dim S" by linarith then show ?thesis using homeomorphism_moving_points_exists_gen [OF \finite K\ KS pw _ ope S] that by fastforce qed subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_grouping_points_exists\\ lemma homeomorphism_grouping_point_1: fixes a::real and c::real assumes "a < b" "c < d" obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" proof - define f where "f \ \x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" have "\g. homeomorphism (cbox a b) (cbox c d) f g" proof (rule homeomorphism_compact) show "continuous_on (cbox a b) f" unfolding f_def by (intro continuous_intros) have "f ` {a..b} = {c..d}" unfolding f_def image_affinity_atLeastAtMost using assms sum_sqs_eq by (auto simp: field_split_simps) then show "f ` cbox a b = cbox c d" by auto show "inj_on f (cbox a b)" unfolding f_def inj_on_def using assms by auto qed auto then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. then show ?thesis proof show "f a = c" by (simp add: f_def) show "f b = d" using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps) qed qed lemma homeomorphism_grouping_point_2: fixes a::real and w::real assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" and "b \ cbox a c" "v \ cbox u w" and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" "\x. x \ cbox a b \ f x = f1 x" "\x. x \ cbox b c \ f x = f2 x" proof - have le: "a \ b" "b \ c" "u \ v" "v \ w" using assms by simp_all then have ac: "cbox a c = cbox a b \ cbox b c" and uw: "cbox u w = cbox u v \ cbox v w" by auto define f where "f \ \x. if x \ b then f1 x else f2 x" have "\g. homeomorphism (cbox a c) (cbox u w) f g" proof (rule homeomorphism_compact) have cf1: "continuous_on (cbox a b) f1" using hom_ab homeomorphism_cont1 by blast have cf2: "continuous_on (cbox b c) f2" using hom_bc homeomorphism_cont1 by blast show "continuous_on (cbox a c) f" unfolding f_def using le eq by (force intro: continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" unfolding f_def using eq by force+ then show "f ` cbox a c = cbox u w" unfolding ac uw image_Un by (metis hom_ab hom_bc homeomorphism_def) have neq12: "f1 x \ f2 y" if x: "a \ x" "x \ b" and y: "b < y" "y \ c" for x y proof - have "f1 x \ cbox u v" by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) moreover have "f2 y \ cbox v w" by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) moreover have "f2 y \ f2 b" by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) ultimately show ?thesis using le eq by simp qed have "inj_on f1 (cbox a b)" by (metis (full_types) hom_ab homeomorphism_def inj_onI) moreover have "inj_on f2 (cbox b c)" by (metis (full_types) hom_bc homeomorphism_def inj_onI) ultimately show "inj_on f (cbox a c)" apply (simp (no_asm) add: inj_on_def) apply (simp add: f_def inj_on_eq_iff) using neq12 by force qed auto then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. then show ?thesis using eq f_def le that by force qed lemma homeomorphism_grouping_point_3: fixes a::real assumes cbox_sub: "cbox c d \ box a b" "cbox u v \ box a b" and box_ne: "box c d \ {}" "box u v \ {}" obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" "\x. x \ cbox c d \ f x \ cbox u v" proof - have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \ {}" using assms by (simp_all add: cbox_sub subset_eq) obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" and f1_eq: "f1 a = a" "f1 c = u" using homeomorphism_grouping_point_1 [OF \a < c\ \a < u\] . obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" and f2_eq: "f2 c = u" "f2 d = v" using homeomorphism_grouping_point_1 [OF \c < d\ \u < v\] . obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" and f3_eq: "f3 d = v" "f3 b = b" using homeomorphism_grouping_point_1 [OF \d < b\ \v < b\] . obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" and f4_eq: "\x. x \ cbox a c \ f4 x = f1 x" "\x. x \ cbox c d \ f4 x = f2 x" using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" and f_eq: "\x. x \ cbox a d \ f x = f4 x" "\x. x \ cbox d b \ f x = f3 x" using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) show ?thesis proof (rule that [OF fg]) show "f x \ cbox u v" if "x \ cbox c d" for x using that f4_eq f_eq homeomorphism_image1 [OF 2] by (metis atLeastAtMost_iff box_real(2) image_eqI less(1) less_eq_real_def order_trans) qed qed lemma homeomorphism_grouping_point_4: fixes T :: "real set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "\x. x \ K \ f x \ U" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof - obtain c d where "box c d \ {}" "cbox c d \ U" proof - obtain u where "u \ U" using \U \ {}\ by blast then obtain e where "e > 0" "cball u e \ U" using \open U\ open_contains_cball by blast then show ?thesis by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) qed have "compact K" by (simp add: \finite K\ finite_imp_compact) obtain a b where "box a b \ {}" "K \ cbox a b" "cbox a b \ S" proof (cases "K = {}") case True then show ?thesis using \box c d \ {}\ \cbox c d \ U\ \U \ S\ that by blast next case False then obtain a b where "a \ K" "b \ K" and a: "\x. x \ K \ a \ x" and b: "\x. x \ K \ x \ b" using compact_attains_inf compact_attains_sup by (metis \compact K\)+ obtain e where "e > 0" "cball b e \ S" using \open S\ open_contains_cball by (metis \b \ K\ \K \ S\ subsetD) show ?thesis proof show "box a (b + e) \ {}" using \0 < e\ \b \ K\ a by force show "K \ cbox a (b + e)" using \0 < e\ a b by fastforce have "a \ S" using \a \ K\ assms(6) by blast have "b + e \ S" using \0 < e\ \cball b e \ S\ by (force simp: dist_norm) show "cbox a (b + e) \ S" using \a \ S\ \b + e \ S\ \connected S\ connected_contains_Icc by auto qed qed obtain w z where "cbox w z \ S" and sub_wz: "cbox a b \ cbox c d \ box w z" proof - have "a \ S" "b \ S" using \box a b \ {}\ \cbox a b \ S\ by auto moreover have "c \ S" "d \ S" using \box c d \ {}\ \cbox c d \ U\ \U \ S\ by force+ ultimately have "min a c \ S" "max b d \ S" by linarith+ then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \ S" "e2 > 0" "cball (max b d) e2 \ S" using \open S\ open_contains_cball by metis then have *: "min a c - e1 \ S" "max b d + e2 \ S" by (auto simp: dist_norm) show ?thesis proof show "cbox (min a c - e1) (max b d+ e2) \ S" using * \connected S\ connected_contains_Icc by auto show "cbox a b \ cbox c d \ box (min a c - e1) (max b d + e2)" using \0 < e1\ \0 < e2\ by auto qed qed then obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" and "f w = w" "f z = z" and fin: "\x. x \ cbox a b \ f x \ cbox c d" using homeomorphism_grouping_point_3 [of a b w z c d] using \box a b \ {}\ \box c d \ {}\ by blast have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" using hom homeomorphism_def by blast+ define f' where "f' \ \x. if x \ cbox w z then f x else x" define g' where "g' \ \x. if x \ cbox w z then g x else x" show ?thesis proof have T: "cbox w z \ (T - box w z) = T" using \cbox w z \ S\ \S \ T\ by auto show "homeomorphism T T f' g'" proof have clo: "closedin (top_of_set (cbox w z \ (T - box w z))) (T - box w z)" by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) have "\x. \w \ x \ x \ z; w < x \ \ x < z\ \ f x = x" using \f w = w\ \f z = z\ by auto moreover have "\x. \w \ x \ x \ z; w < x \ \ x < z\ \ g x = x" using \f w = w\ \f z = z\ hom homeomorphism_apply1 by fastforce ultimately have "continuous_on (cbox w z \ (T - box w z)) f'" "continuous_on (cbox w z \ (T - box w z)) g'" unfolding f'_def g'_def by (intro continuous_on_cases_local contfg continuous_on_id clo; auto simp: closed_subset)+ then show "continuous_on T f'" "continuous_on T g'" by (simp_all only: T) show "f' ` T \ T" unfolding f'_def by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) show "g' ` T \ T" unfolding g'_def by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) show "\x. x \ T \ g' (f' x) = x" unfolding f'_def g'_def using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce show "\y. y \ T \ f' (g' y) = y" unfolding f'_def g'_def using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce qed show "\x. x \ K \ f' x \ U" using fin sub_wz \K \ cbox a b\ \cbox c d \ U\ by (force simp: f'_def) show "{x. \ (f' x = x \ g' x = x)} \ S" using \cbox w z \ S\ by (auto simp: f'_def g'_def) show "bounded {x. \ (f' x = x \ g' x = x)}" proof (rule bounded_subset [of "cbox w z"]) show "bounded (cbox w z)" using bounded_cbox by blast show "{x. \ (f' x = x \ g' x = x)} \ cbox w z" by (auto simp: f'_def g'_def) qed qed qed proposition\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists: fixes S :: "'a::euclidean_space set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ DIM('a)") case True have TS: "T \ affine hull S" using affine_hull_open assms by blast have "infinite U" using \open U\ \U \ {}\ finite_imp_not_open by blast then obtain P where "P \ U" "finite P" "card K = card P" using infinite_arbitrarily_large by metis then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast obtain f g where "homeomorphism T T f g" "\i. i \ K \ f (id i) = \ i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists [OF True \open S\ \connected S\ \S \ T\ \finite K\]) show "\i. i \ K \ id i \ S \ \ i \ S" using \P \ U\ \bij_betw \ K P\ \K \ S\ \U \ S\ bij_betwE by blast show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed (use affine_hull_open assms that in auto) then show ?thesis using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) next case False with DIM_positive have "DIM('a) = 1" by (simp add: dual_order.antisym) then obtain h::"'a \real" and j where "linear h" "linear j" and noh: "\x. norm(h x) = norm x" and noj: "\y. norm(j y) = norm y" and hj: "\x. j(h x) = x" "\y. h(j y) = y" and ranh: "surj h" using isomorphisms_UNIV_UNIV by (metis (mono_tags, opaque_lifting) DIM_real UNIV_eq_I range_eqI) obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" and f: "\x. x \ h ` K \ f x \ h ` U" and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" and bou: "bounded {x. \ (f x = x \ g x = x)}" apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) by (simp_all add: assms image_mono \linear h\ open_surjective_linear_image connected_linear_image ranh) have jf: "j (f (h x)) = x \ f (h x) = h x" for x by (metis hj) have jg: "j (g (h x)) = x \ g (h x) = h x" for x by (metis hj) have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y by (simp_all add: \linear h\ \linear j\ linear_linear linear_continuous_on) show ?thesis proof show "homeomorphism T T (j \ f \ h) (j \ g \ h)" proof show "continuous_on T (j \ f \ h)" "continuous_on T (j \ g \ h)" using hom homeomorphism_def by (blast intro: continuous_on_compose cont_hj)+ show "(j \ f \ h) ` T \ T" "(j \ g \ h) ` T \ T" by auto (metis (mono_tags, opaque_lifting) hj(1) hom homeomorphism_def imageE imageI)+ show "\x. x \ T \ (j \ g \ h) ((j \ f \ h) x) = x" using hj hom homeomorphism_apply1 by fastforce show "\y. y \ T \ (j \ f \ h) ((j \ g \ h) y) = y" using hj hom homeomorphism_apply2 by fastforce qed show "{x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)} \ S" proof (clarsimp simp: jf jg hj) show "f (h x) = h x \ g (h x) \ h x \ x \ S" for x using sub [THEN subsetD, of "h x"] hj by simp (metis imageE) qed have "bounded (j ` {x. (\ (f x = x \ g x = x))})" by (rule bounded_linear_image [OF bou]) (use \linear j\ linear_conv_bounded_linear in auto) moreover have *: "{x. \((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (\ (f x = x \ g x = x))}" using hj by (auto simp: jf jg image_iff, metis+) ultimately show "bounded {x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)}" by metis show "\x. x \ K \ (j \ f \ h) x \ U" using f hj by fastforce qed qed proposition\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists_gen: fixes S :: "'a::euclidean_space set" assumes opeU: "openin (top_of_set S) U" and opeS: "openin (top_of_set (affine hull S)) S" and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ aff_dim S") case True have opeU': "openin (top_of_set (affine hull S)) U" using opeS opeU openin_trans by blast obtain u where "u \ U" "u \ S" using \U \ {}\ opeU openin_imp_subset by fastforce+ have "infinite U" proof (rule infinite_openin [OF opeU \u \ U\]) show "u islimpt S" using True \u \ S\ assms(8) connected_imp_perfect_aff_dim by fastforce qed then obtain P where "P \ U" "finite P" "card K = card P" using infinite_arbitrarily_large by metis then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast have "\f g. homeomorphism T T f g \ (\i \ K. f(id i) = \ i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists_gen [OF \finite K\ _ _ True opeS S]) show "\i. i \ K \ id i \ S \ \ i \ S" by (metis id_apply opeU openin_contains_cball subsetCE \P \ U\ \bij_betw \ K P\ \K \ S\ bij_betwE) show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed then show ?thesis using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) next case False with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith then show ?thesis proof cases assume "aff_dim S = -1" then have "S = {}" using aff_dim_empty by blast then have "False" using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast then show ?thesis .. next assume "aff_dim S = 0" then obtain a where "S = {a}" using aff_dim_eq_0 by blast then have "K \ U" using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast show ?thesis using \K \ U\ by (intro that [of id id]) (auto intro: homeomorphismI) next assume "aff_dim S = 1" then have "affine hull S homeomorphic (UNIV :: real set)" by (auto simp: homeomorphic_affine_sets) then obtain h::"'a\real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" using homeomorphic_def by blast then have h: "\x. x \ affine hull S \ j(h(x)) = x" and j: "\y. j y \ affine hull S \ h(j y) = y" by (auto simp: homeomorphism_def) have connh: "connected (h ` S)" by (meson Topological_Spaces.connected_continuous_image \connected S\ homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) have hUS: "h ` U \ h ` S" by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) have opn: "openin (top_of_set (affine hull S)) U \ open (h ` U)" for U using homeomorphism_imp_open_map [OF homhj] by simp have "open (h ` U)" "open (h ` S)" by (auto intro: opeS opeU openin_trans opn) then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" and f: "\x. x \ h ` K \ f x \ h ` U" and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" and bou: "bounded {x. \ (f x = x \ g x = x)}" apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) using assms by (auto simp: connh hUS) have jf: "\x. x \ affine hull S \ j (f (h x)) = x \ f (h x) = h x" by (metis h j) have jg: "\x. x \ affine hull S \ j (g (h x)) = x \ g (h x) = h x" by (metis h j) have cont_hj: "continuous_on T h" "continuous_on Y j" for Y proof (rule continuous_on_subset [OF _ \T \ affine hull S\]) show "continuous_on (affine hull S) h" using homeomorphism_def homhj by blast qed (meson continuous_on_subset homeomorphism_def homhj top_greatest) define f' where "f' \ \x. if x \ affine hull S then (j \ f \ h) x else x" define g' where "g' \ \x. if x \ affine hull S then (j \ g \ h) x else x" show ?thesis proof show "homeomorphism T T f' g'" proof have "continuous_on T (j \ f \ h)" using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast then show "continuous_on T f'" apply (rule continuous_on_eq) using \T \ affine hull S\ f'_def by auto have "continuous_on T (j \ g \ h)" using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast then show "continuous_on T g'" apply (rule continuous_on_eq) using \T \ affine hull S\ g'_def by auto show "f' ` T \ T" proof (clarsimp simp: f'_def) fix x assume "x \ T" then have "f (h x) \ h ` T" by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) then show "j (f (h x)) \ T" using \T \ affine hull S\ h by auto qed show "g' ` T \ T" proof (clarsimp simp: g'_def) fix x assume "x \ T" then have "g (h x) \ h ` T" by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) then show "j (g (h x)) \ T" using \T \ affine hull S\ h by auto qed show "\x. x \ T \ g' (f' x) = x" using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def) show "\y. y \ T \ f' (g' y) = y" using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def) qed next have \
: "\x y. \x \ affine hull S; h x = h y; y \ S\ \ x \ S" by (metis h hull_inc) show "{x. \ (f' x = x \ g' x = x)} \ S" using sub by (simp add: f'_def g'_def jf jg) (force elim: \
) next have "compact (j ` closure {x. \ (f x = x \ g x = x)})" using bou by (auto simp: compact_continuous_image cont_hj) then have "bounded (j ` {x. \ (f x = x \ g x = x)})" by (rule bounded_closure_image [OF compact_imp_bounded]) moreover have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (\ (f x = x \ g x = x))}" using h j by (auto simp: image_iff; metis) ultimately have "bounded {x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x}" by metis then show "bounded {x. \ (f' x = x \ g' x = x)}" by (simp add: f'_def g'_def Collect_mono bounded_subset) next show "f' x \ U" if "x \ K" for x proof - have "U \ S" using opeU openin_imp_subset by blast then have "j (f (h x)) \ U" using f h hull_subset that by fastforce then show "f' x \ U" using \K \ S\ S f'_def that by auto qed qed qed qed subsection\Nullhomotopic mappings\ text\ A mapping out of a sphere is nullhomotopic iff it extends to the ball. This even works out in the degenerate cases when the radius is \\\ 0, and we also don't need to explicitly assume continuity since it's already implicit in both sides of the equivalence.\ lemma nullhomotopic_from_lemma: assumes contg: "continuous_on (cball a r - {a}) g" and fa: "\e. 0 < e \ \d. 0 < d \ (\x. x \ a \ norm(x - a) < d \ norm(g x - f a) < e)" and r: "\x. x \ cball a r \ x \ a \ f x = g x" shows "continuous_on (cball a r) f" proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) fix x assume x: "dist a x \ r" show "continuous (at x within cball a r) f" proof (cases "x=a") case True then show ?thesis by (metis continuous_within_eps_delta fa dist_norm dist_self r) next case False show ?thesis proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) have "\d>0. \x'\cball a r. dist x' x < d \ dist (g x') (g x) < e" if "e>0" for e proof - obtain d where "d > 0" and d: "\x'. \dist x' a \ r; x' \ a; dist x' x < d\ \ dist (g x') (g x) < e" using contg False x \e>0\ unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that) show ?thesis using \d > 0\ \x \ a\ by (rule_tac x="min d (norm(x - a))" in exI) (auto simp: dist_commute dist_norm [symmetric] intro!: d) qed then show "continuous (at x within cball a r) g" using contg False by (auto simp: continuous_within_eps_delta) show "0 < norm (x - a)" using False by force show "x \ cball a r" by (simp add: x) show "\x'. \x' \ cball a r; dist x' x < norm (x - a)\ \ g x' = f x'" by (metis dist_commute dist_norm less_le r) qed qed qed proposition nullhomotopic_from_sphere_extension: fixes f :: "'M::euclidean_space \ 'a::real_normed_vector" shows "(\c. homotopic_with_canon (\x. True) (sphere a r) S f (\x. c)) \ (\g. continuous_on (cball a r) g \ g ` (cball a r) \ S \ (\x \ sphere a r. g x = f x))" (is "?lhs = ?rhs") proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp add: homotopic_on_emptyI) next case equal show ?thesis proof assume L: ?lhs with equal have [simp]: "f a \ S" using homotopic_with_imp_subset1 by fastforce obtain h:: "real \ 'M \ 'a" where h: "continuous_on ({0..1} \ {a}) h" "h ` ({0..1} \ {a}) \ S" "h (0, a) = f a" using L equal by (auto simp: homotopic_with) then have "continuous_on (cball a r) (\x. h (0, a))" "(\x. h (0, a)) ` cball a r \ S" by (auto simp: equal) then show ?rhs using h(3) local.equal by force next assume ?rhs then show ?lhs using equal continuous_on_const by (force simp add: homotopic_with) qed next case greater let ?P = "continuous_on {x. norm(x - a) = r} f \ f ` {x. norm(x - a) = r} \ S" have ?P if ?lhs using that proof fix c assume c: "homotopic_with_canon (\x. True) (sphere a r) S f (\x. c)" then have contf: "continuous_on (sphere a r) f" by (metis homotopic_with_imp_continuous) moreover have fim: "f ` sphere a r \ S" by (meson continuous_map_subtopology_eu c homotopic_with_imp_continuous_maps) show ?P using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) qed moreover have ?P if ?rhs using that proof fix g assume g: "continuous_on (cball a r) g \ g ` cball a r \ S \ (\xa\sphere a r. g xa = f xa)" then have "f ` {x. norm (x - a) = r} \ S" using sphere_cball [of a r] unfolding image_subset_iff sphere_def by (metis dist_commute dist_norm mem_Collect_eq subset_eq) with g show ?P by (auto simp: dist_norm norm_minus_commute elim!: continuous_on_eq [OF continuous_on_subset]) qed moreover have ?thesis if ?P proof assume ?lhs then obtain c where "homotopic_with_canon (\x. True) (sphere a r) S (\x. c) f" using homotopic_with_sym by blast then obtain h where conth: "continuous_on ({0..1::real} \ sphere a r) h" and him: "h ` ({0..1} \ sphere a r) \ S" and h: "\x. h(0, x) = c" "\x. h(1, x) = f x" by (auto simp: homotopic_with_def) obtain b1::'M where "b1 \ Basis" using SOME_Basis by auto have "c \ h ` ({0..1} \ sphere a r)" proof show "c = h (0, a + r *\<^sub>R b1)" by (simp add: h) show "(0, a + r *\<^sub>R b1) \ {0..1::real} \ sphere a r" using greater \b1 \ Basis\ by (auto simp: dist_norm) qed then have "c \ S" using him by blast have uconth: "uniformly_continuous_on ({0..1::real} \ (sphere a r)) h" by (force intro: compact_Times conth compact_uniformly_continuous) let ?g = "\x. h (norm (x - a)/r, a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" let ?g' = "\x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" show ?rhs proof (intro exI conjI) have "continuous_on (cball a r - {a}) ?g'" using greater by (force simp: dist_norm norm_minus_commute intro: continuous_on_compose2 [OF conth] continuous_intros) then show "continuous_on (cball a r) ?g" proof (rule nullhomotopic_from_lemma) show "\d>0. \x. x \ a \ norm (x - a) < d \ norm (?g' x - ?g a) < e" if "0 < e" for e proof - obtain d where "0 < d" and d: "\x x'. \x \ {0..1} \ sphere a r; x' \ {0..1} \ sphere a r; norm ( x' - x) < d\ \ norm (h x' - h x) < e" using uniformly_continuous_onE [OF uconth \0 < e\] by (auto simp: dist_norm) have *: "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" (is "norm (?ha - ?hb) < e") if "x \ a" "norm (x - a) < r" "norm (x - a) < d * r" for x proof - have "norm (?ha - ?hb) = norm (?ha - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" by (simp add: h) also have "\ < e" using greater \0 < d\ \b1 \ Basis\ that by (intro d) (simp_all add: dist_norm, simp add: field_simps) finally show ?thesis . qed show ?thesis using greater \0 < d\ by (rule_tac x = "min r (d * r)" in exI) (auto simp: *) qed show "\x. x \ cball a r \ x \ a \ ?g x = ?g' x" by auto qed next show "?g ` cball a r \ S" using greater him \c \ S\ by (force simp: h dist_norm norm_minus_commute) next show "\x\sphere a r. ?g x = f x" using greater by (auto simp: h dist_norm norm_minus_commute) qed next assume ?rhs then obtain g where contg: "continuous_on (cball a r) g" and gim: "g ` cball a r \ S" and gf: "\x \ sphere a r. g x = f x" by auto let ?h = "\y. g (a + (fst y) *\<^sub>R (snd y - a))" have "continuous_on ({0..1} \ sphere a r) ?h" proof (rule continuous_on_compose2 [OF contg]) show "continuous_on ({0..1} \ sphere a r) (\x. a + fst x *\<^sub>R (snd x - a))" by (intro continuous_intros) qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) moreover have "?h ` ({0..1} \ sphere a r) \ S" by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) moreover have "\x\sphere a r. ?h (0, x) = g a" "\x\sphere a r. ?h (1, x) = f x" by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) ultimately have "homotopic_with_canon (\x. True) (sphere a r) S (\x. g a) f" by (auto simp: homotopic_with) then show ?lhs using homotopic_with_symD by blast qed ultimately show ?thesis by meson qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Linear_Algebra.thy b/src/HOL/Analysis/Linear_Algebra.thy --- a/src/HOL/Analysis/Linear_Algebra.thy +++ b/src/HOL/Analysis/Linear_Algebra.thy @@ -1,1661 +1,1661 @@ (* Title: HOL/Analysis/Linear_Algebra.thy Author: Amine Chaieb, University of Cambridge *) section \Elementary Linear Algebra on Euclidean Spaces\ theory Linear_Algebra imports Euclidean_Space "HOL-Library.Infinite_Set" begin lemma linear_simps: assumes "bounded_linear f" shows "f (a + b) = f a + f b" "f (a - b) = f a - f b" "f 0 = 0" "f (- a) = - f a" "f (s *\<^sub>R v) = s *\<^sub>R (f v)" proof - interpret f: bounded_linear f by fact show "f (a + b) = f a + f b" by (rule f.add) show "f (a - b) = f a - f b" by (rule f.diff) show "f 0 = 0" by (rule f.zero) show "f (- a) = - f a" by (rule f.neg) show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scale) qed lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x \ (UNIV::'a::finite set)}" using finite finite_image_set by blast lemma substdbasis_expansion_unique: includes inner_syntax assumes d: "d \ Basis" shows "(\i\d. f i *\<^sub>R i) = (x::'a::euclidean_space) \ (\i\Basis. (i \ d \ f i = x \ i) \ (i \ d \ x \ i = 0))" proof - have *: "\x a b P. x * (if P then a else b) = (if P then x * a else x * b)" by auto have **: "finite d" by (auto intro: finite_subset[OF assms]) have ***: "\i. i \ Basis \ (\i\d. f i *\<^sub>R i) \ i = (\x\d. if x = i then f x else 0)" using d by (auto intro!: sum.cong simp: inner_Basis inner_sum_left) show ?thesis unfolding euclidean_eq_iff[where 'a='a] by (auto simp: sum.delta[OF **] ***) qed lemma independent_substdbasis: "d \ Basis \ independent d" by (rule independent_mono[OF independent_Basis]) lemma subset_translation_eq [simp]: fixes a :: "'a::real_vector" shows "(+) a ` s \ (+) a ` t \ s \ t" by auto lemma translate_inj_on: fixes A :: "'a::ab_group_add set" shows "inj_on (\x. a + x) A" unfolding inj_on_def by auto lemma translation_assoc: fixes a b :: "'a::ab_group_add" shows "(\x. b + x) ` ((\x. a + x) ` S) = (\x. (a + b) + x) ` S" by auto lemma translation_invert: fixes a :: "'a::ab_group_add" assumes "(\x. a + x) ` A = (\x. a + x) ` B" shows "A = B" using assms translation_assoc by fastforce lemma translation_galois: fixes a :: "'a::ab_group_add" shows "T = ((\x. a + x) ` S) \ S = ((\x. (- a) + x) ` T)" by (metis add.right_inverse group_cancel.rule0 translation_invert translation_assoc) lemma translation_inverse_subset: assumes "((\x. - a + x) ` V) \ (S :: 'n::ab_group_add set)" shows "V \ ((\x. a + x) ` S)" by (metis assms subset_image_iff translation_galois) subsection\<^marker>\tag unimportant\ \More interesting properties of the norm\ unbundle inner_syntax text\Equality of vectors in terms of \<^term>\(\)\ products.\ lemma linear_componentwise: fixes f:: "'a::euclidean_space \ 'b::real_inner" assumes lf: "linear f" shows "(f x) \ j = (\i\Basis. (x\i) * (f i\j))" (is "?lhs = ?rhs") proof - interpret linear f by fact have "?rhs = (\i\Basis. (x\i) *\<^sub>R (f i))\j" by (simp add: inner_sum_left) then show ?thesis by (simp add: euclidean_representation sum[symmetric] scale[symmetric]) qed lemma vector_eq: "x = y \ x \ x = x \ y \ y \ y = x \ x" by (metis (no_types, opaque_lifting) inner_commute inner_diff_right inner_eq_zero_iff right_minus_eq) lemma norm_triangle_half_r: "norm (y - x1) < e/2 \ norm (y - x2) < e/2 \ norm (x1 - x2) < e" using dist_triangle_half_r unfolding dist_norm[symmetric] by auto lemma norm_triangle_half_l: assumes "norm (x - y) < e/2" and "norm (x' - y) < e/2" shows "norm (x - x') < e" by (metis assms dist_norm dist_triangle_half_l) lemma abs_triangle_half_r: fixes y :: "'a::linordered_field" shows "abs (y - x1) < e/2 \ abs (y - x2) < e/2 \ abs (x1 - x2) < e" by linarith lemma abs_triangle_half_l: fixes y :: "'a::linordered_field" assumes "abs (x - y) < e/2" and "abs (x' - y) < e/2" shows "abs (x - x') < e" using assms by linarith lemma sum_clauses: shows "sum f {} = 0" and "finite S \ sum f (insert x S) = (if x \ S then sum f S else f x + sum f S)" by (auto simp add: insert_absorb) lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = z" and vector_eq_rdot: "(\z. x \ z = y \ z) \ x = y" by (metis inner_commute vector_eq)+ subsection \Substandard Basis\ lemma ex_card: assumes "n \ card A" shows "\S\A. card S = n" by (meson assms obtain_subset_with_card_n) lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" by (auto simp: subspace_def inner_add_left) lemma dim_substandard: assumes d: "d \ Basis" shows "dim {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0} = card d" (is "dim ?A = _") proof (rule dim_unique) from d show "d \ ?A" by (auto simp: inner_Basis) from d show "independent d" by (rule independent_mono [OF independent_Basis]) have "x \ span d" if "\i\Basis. i \ d \ x \ i = 0" for x proof - have "finite d" by (rule finite_subset [OF d finite_Basis]) then have "(\i\d. (x \ i) *\<^sub>R i) \ span d" by (simp add: span_sum span_clauses) also have "(\i\d. (x \ i) *\<^sub>R i) = (\i\Basis. (x \ i) *\<^sub>R i)" by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) finally show "x \ span d" by (simp only: euclidean_representation) qed then show "?A \ span d" by auto qed simp subsection \Orthogonality\ definition\<^marker>\tag important\ (in real_inner) "orthogonal x y \ x \ y = 0" context real_inner begin lemma orthogonal_self: "orthogonal x x \ x = 0" by (simp add: orthogonal_def) lemma orthogonal_clauses: "orthogonal a 0" "orthogonal a x \ orthogonal a (c *\<^sub>R x)" "orthogonal a x \ orthogonal a (- x)" "orthogonal a x \ orthogonal a y \ orthogonal a (x + y)" "orthogonal a x \ orthogonal a y \ orthogonal a (x - y)" "orthogonal 0 a" "orthogonal x a \ orthogonal (c *\<^sub>R x) a" "orthogonal x a \ orthogonal (- x) a" "orthogonal x a \ orthogonal y a \ orthogonal (x + y) a" "orthogonal x a \ orthogonal y a \ orthogonal (x - y) a" unfolding orthogonal_def inner_add inner_diff by auto end lemma orthogonal_commute: "orthogonal x y \ orthogonal y x" by (simp add: orthogonal_def inner_commute) lemma orthogonal_scaleR [simp]: "c \ 0 \ orthogonal (c *\<^sub>R x) = orthogonal x" by (rule ext) (simp add: orthogonal_def) lemma pairwise_ortho_scaleR: "pairwise (\i j. orthogonal (f i) (g j)) B \ pairwise (\i j. orthogonal (a i *\<^sub>R f i) (a j *\<^sub>R g j)) B" by (auto simp: pairwise_def orthogonal_clauses) lemma orthogonal_rvsum: "\finite s; \y. y \ s \ orthogonal x (f y)\ \ orthogonal x (sum f s)" by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) lemma orthogonal_lvsum: "\finite s; \x. x \ s \ orthogonal (f x) y\ \ orthogonal (sum f s) y" by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) lemma norm_add_Pythagorean: assumes "orthogonal a b" shows "(norm (a + b))\<^sup>2 = (norm a)\<^sup>2 + (norm b)\<^sup>2" proof - from assms have "(a - (0 - b)) \ (a - (0 - b)) = a \ a - (0 - b \ b)" by (simp add: algebra_simps orthogonal_def inner_commute) then show ?thesis by (simp add: power2_norm_eq_inner) qed lemma norm_sum_Pythagorean: assumes "finite I" "pairwise (\i j. orthogonal (f i) (f j)) I" shows "(norm (sum f I))\<^sup>2 = (\i\I. (norm (f i))\<^sup>2)" using assms proof (induction I rule: finite_induct) case empty then show ?case by simp next case (insert x I) then have "orthogonal (f x) (sum f I)" by (metis pairwise_insert orthogonal_rvsum) with insert show ?case by (simp add: pairwise_insert norm_add_Pythagorean) qed subsection \Orthogonality of a transformation\ definition\<^marker>\tag important\ "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" lemma\<^marker>\tag unimportant\ orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\v. norm (f v) = norm v)" by (smt (verit, ccfv_threshold) dot_norm linear_add norm_eq_sqrt_inner orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_id [simp]: "orthogonal_transformation (\x. x)" by (simp add: linear_iff orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_orthogonal_transformation: "orthogonal_transformation f \ orthogonal (f x) (f y) \ orthogonal x y" by (simp add: orthogonal_def orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_compose: "\orthogonal_transformation f; orthogonal_transformation g\ \ orthogonal_transformation(f \ g)" by (auto simp: orthogonal_transformation_def linear_compose) lemma\<^marker>\tag unimportant\ orthogonal_transformation_neg: "orthogonal_transformation(\x. -(f x)) \ orthogonal_transformation f" by (auto simp: orthogonal_transformation_def dest: linear_compose_neg) lemma\<^marker>\tag unimportant\ orthogonal_transformation_scaleR: "orthogonal_transformation f \ f (c *\<^sub>R v) = c *\<^sub>R f v" by (simp add: linear_iff orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_linear: "orthogonal_transformation f \ linear f" by (simp add: orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_inj: "orthogonal_transformation f \ inj f" unfolding orthogonal_transformation_def inj_on_def by (metis vector_eq) lemma\<^marker>\tag unimportant\ orthogonal_transformation_surj: "orthogonal_transformation f \ surj f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear) lemma\<^marker>\tag unimportant\ orthogonal_transformation_bij: "orthogonal_transformation f \ bij f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj) lemma\<^marker>\tag unimportant\ orthogonal_transformation_inv: "orthogonal_transformation f \ orthogonal_transformation (inv f)" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (metis (no_types, opaque_lifting) bijection.inv_right bijection_def inj_linear_imp_inv_linear orthogonal_transformation orthogonal_transformation_bij orthogonal_transformation_inj) lemma\<^marker>\tag unimportant\ orthogonal_transformation_norm: "orthogonal_transformation f \ norm (f x) = norm x" by (metis orthogonal_transformation) subsection \Bilinear functions\ definition\<^marker>\tag important\ bilinear :: "('a::real_vector \ 'b::real_vector \ 'c::real_vector) \ bool" where "bilinear f \ (\x. linear (\y. f x y)) \ (\y. linear (\x. f x y))" lemma bilinear_ladd: "bilinear h \ h (x + y) z = h x z + h y z" by (simp add: bilinear_def linear_iff) lemma bilinear_radd: "bilinear h \ h x (y + z) = h x y + h x z" by (simp add: bilinear_def linear_iff) lemma bilinear_times: fixes c::"'a::real_algebra" shows "bilinear (\x y::'a. x*y)" by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) lemma bilinear_lmul: "bilinear h \ h (c *\<^sub>R x) y = c *\<^sub>R h x y" by (simp add: bilinear_def linear_iff) lemma bilinear_rmul: "bilinear h \ h x (c *\<^sub>R y) = c *\<^sub>R h x y" by (simp add: bilinear_def linear_iff) lemma bilinear_lneg: "bilinear h \ h (- x) y = - h x y" by (drule bilinear_lmul [of _ "- 1"]) simp lemma bilinear_rneg: "bilinear h \ h x (- y) = - h x y" by (drule bilinear_rmul [of _ _ "- 1"]) simp lemma (in ab_group_add) eq_add_iff: "x = x + y \ y = 0" using add_left_imp_eq[of x y 0] by auto lemma bilinear_lzero: assumes "bilinear h" shows "h 0 x = 0" using bilinear_ladd [OF assms, of 0 0 x] by (simp add: eq_add_iff field_simps) lemma bilinear_rzero: assumes "bilinear h" shows "h x 0 = 0" using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps) lemma bilinear_lsub: "bilinear h \ h (x - y) z = h x z - h y z" using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg) lemma bilinear_rsub: "bilinear h \ h z (x - y) = h z x - h z y" using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg) lemma bilinear_sum: assumes "bilinear h" shows "h (sum f S) (sum g T) = sum (\(i,j). h (f i) (g j)) (S \ T) " proof - interpret l: linear "\x. h x y" for y using assms by (simp add: bilinear_def) interpret r: linear "\y. h x y" for x using assms by (simp add: bilinear_def) have "h (sum f S) (sum g T) = sum (\x. h (f x) (sum g T)) S" by (simp add: l.sum) also have "\ = sum (\x. sum (\y. h (f x) (g y)) T) S" by (rule sum.cong) (simp_all add: r.sum) finally show ?thesis unfolding sum.cartesian_product . qed subsection \Adjoints\ definition\<^marker>\tag important\ adjoint :: "(('a::real_inner) \ ('b::real_inner)) \ 'b \ 'a" where "adjoint f = (SOME f'. \x y. f x \ y = x \ f' y)" lemma adjoint_unique: assumes "\x y. inner (f x) y = inner x (g y)" shows "adjoint f = g" unfolding adjoint_def proof (rule some_equality) show "\x y. inner (f x) y = inner x (g y)" by (rule assms) next fix h assume "\x y. inner (f x) y = inner x (h y)" then show "h = g" by (metis assms ext vector_eq_ldot) qed text \TODO: The following lemmas about adjoints should hold for any Hilbert space (i.e. complete inner product space). (see \<^url>\https://en.wikipedia.org/wiki/Hermitian_adjoint\) \ lemma adjoint_works: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" proof - interpret linear f by fact have "\y. \w. \x. f x \ y = x \ w" proof (intro allI exI) fix y :: "'m" and x let ?w = "(\i\Basis. (f i \ y) *\<^sub>R i) :: 'n" have "f x \ y = f (\i\Basis. (x \ i) *\<^sub>R i) \ y" by (simp add: euclidean_representation) also have "\ = (\i\Basis. (x \ i) *\<^sub>R f i) \ y" by (simp add: sum scale) finally show "f x \ y = x \ ?w" by (simp add: inner_sum_left inner_sum_right mult.commute) qed then show ?thesis unfolding adjoint_def choice_iff by (intro someI2_ex[where Q="\f'. x \ f' y = f x \ y"]) auto qed lemma adjoint_clauses: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" and "adjoint f y \ x = y \ f x" by (simp_all add: adjoint_works[OF lf] inner_commute) lemma adjoint_linear: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "linear (adjoint f)" by (simp add: lf linear_iff euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m] adjoint_clauses[OF lf] inner_distrib) lemma adjoint_adjoint: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "adjoint (adjoint f) = f" by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) subsection\<^marker>\tag unimportant\ \Euclidean Spaces as Typeclass\ lemma independent_Basis: "independent Basis" by (rule independent_Basis) lemma span_Basis [simp]: "span Basis = UNIV" by (rule span_Basis) lemma in_span_Basis: "x \ span Basis" unfolding span_Basis .. subsection\<^marker>\tag unimportant\ \Linearity and Bilinearity continued\ lemma linear_bounded: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" shows "\B. \x. norm (f x) \ B * norm x" proof interpret linear f by fact let ?B = "\b\Basis. norm (f b)" show "\x. norm (f x) \ ?B * norm x" proof fix x :: 'a let ?g = "\b. (x \ b) *\<^sub>R f b" have "norm (f x) = norm (f (\b\Basis. (x \ b) *\<^sub>R b))" unfolding euclidean_representation .. also have "\ = norm (sum ?g Basis)" by (simp add: sum scale) finally have th0: "norm (f x) = norm (sum ?g Basis)" . have th: "norm (?g i) \ norm (f i) * norm x" if "i \ Basis" for i proof - from Basis_le_norm[OF that, of x] show "norm (?g i) \ norm (f i) * norm x" unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) qed from sum_norm_le[of _ ?g, OF th] show "norm (f x) \ ?B * norm x" by (simp add: sum_distrib_right th0) qed qed lemma linear_conv_bounded_linear: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ bounded_linear f" by (metis mult.commute bounded_linear_axioms.intro bounded_linear_def linear_bounded) lemmas linear_linear = linear_conv_bounded_linear[symmetric] lemma inj_linear_imp_inv_bounded_linear: fixes f::"'a::euclidean_space \ 'a" shows "\bounded_linear f; inj f\ \ bounded_linear (inv f)" by (simp add: inj_linear_imp_inv_linear linear_linear) lemma linear_bounded_pos: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" obtains B where "B > 0" "\x. norm (f x) \ B * norm x" by (metis bounded_linear.pos_bounded lf linear_linear mult.commute) lemma linear_invertible_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "linear g" and gf: "g \ f = id" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" proof - obtain B where "B > 0" and B: "\x. norm (g x) \ B * norm x" using linear_bounded_pos [OF \linear g\] by blast show thesis proof show "0 < 1/B" by (simp add: \B > 0\) show "1/B * norm x \ norm (f x)" for x by (smt (verit, ccfv_SIG) B \0 < B\ gf comp_apply divide_inverse id_apply inverse_eq_divide less_divide_eq mult.commute) qed qed lemma linear_inj_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "inj f" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" using linear_injective_left_inverse [OF assms] linear_invertible_bounded_below_pos assms by blast lemma bounded_linearI': fixes f ::"'a::euclidean_space \ 'b::real_normed_vector" assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>R x) = c *\<^sub>R f x" shows "bounded_linear f" using assms linearI linear_conv_bounded_linear by blast lemma bilinear_bounded: fixes h :: "'m::euclidean_space \ 'n::euclidean_space \ 'k::real_normed_vector" assumes bh: "bilinear h" shows "\B. \x y. norm (h x y) \ B * norm x * norm y" proof (clarify intro!: exI[of _ "\i\Basis. \j\Basis. norm (h i j)"]) fix x :: 'm fix y :: 'n have "norm (h x y) = norm (h (sum (\i. (x \ i) *\<^sub>R i) Basis) (sum (\i. (y \ i) *\<^sub>R i) Basis))" by (simp add: euclidean_representation) also have "\ = norm (sum (\ (i,j). h ((x \ i) *\<^sub>R i) ((y \ j) *\<^sub>R j)) (Basis \ Basis))" unfolding bilinear_sum[OF bh] .. finally have th: "norm (h x y) = \" . have "\i j. \i \ Basis; j \ Basis\ \ \x \ i\ * (\y \ j\ * norm (h i j)) \ norm x * (norm y * norm (h i j))" by (auto simp add: zero_le_mult_iff Basis_le_norm mult_mono) then show "norm (h x y) \ (\i\Basis. \j\Basis. norm (h i j)) * norm x * norm y" unfolding sum_distrib_right th sum.cartesian_product by (clarsimp simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR intro!: sum_norm_le) qed lemma bilinear_conv_bounded_bilinear: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" shows "bilinear h \ bounded_bilinear h" proof assume "bilinear h" show "bounded_bilinear h" proof fix x y z show "h (x + y) z = h x z + h y z" using \bilinear h\ unfolding bilinear_def linear_iff by simp next fix x y z show "h x (y + z) = h x y + h x z" using \bilinear h\ unfolding bilinear_def linear_iff by simp next show "h (scaleR r x) y = scaleR r (h x y)" "h x (scaleR r y) = scaleR r (h x y)" for r x y using \bilinear h\ unfolding bilinear_def linear_iff by simp_all next have "\B. \x y. norm (h x y) \ B * norm x * norm y" using \bilinear h\ by (rule bilinear_bounded) then show "\K. \x y. norm (h x y) \ norm x * norm y * K" by (simp add: ac_simps) qed next assume "bounded_bilinear h" then interpret h: bounded_bilinear h . show "bilinear h" unfolding bilinear_def linear_conv_bounded_linear using h.bounded_linear_left h.bounded_linear_right by simp qed lemma bilinear_bounded_pos: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes bh: "bilinear h" shows "\B > 0. \x y. norm (h x y) \ B * norm x * norm y" by (metis mult.assoc bh bilinear_conv_bounded_bilinear bounded_bilinear.pos_bounded mult.commute) lemma bounded_linear_imp_has_derivative: "bounded_linear f \ (f has_derivative f) net" by (auto simp add: has_derivative_def linear_diff linear_linear linear_def dest: bounded_linear.linear) lemma linear_imp_has_derivative: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ (f has_derivative f) net" by (simp add: bounded_linear_imp_has_derivative linear_conv_bounded_linear) lemma bounded_linear_imp_differentiable: "bounded_linear f \ f differentiable net" using bounded_linear_imp_has_derivative differentiable_def by blast lemma linear_imp_differentiable: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable net" by (metis linear_imp_has_derivative differentiable_def) lemma of_real_differentiable [simp,derivative_intros]: "of_real differentiable F" by (simp add: bounded_linear_imp_differentiable bounded_linear_of_real) subsection\<^marker>\tag unimportant\ \We continue\ lemma independent_bound: fixes S :: "'a::euclidean_space set" shows "independent S \ finite S \ card S \ DIM('a)" by (metis dim_subset_UNIV finiteI_independent dim_span_eq_card_independent) lemmas independent_imp_finite = finiteI_independent corollary\<^marker>\tag unimportant\ independent_card_le: fixes S :: "'a::euclidean_space set" assumes "independent S" shows "card S \ DIM('a)" using assms independent_bound by auto lemma dependent_biggerset: fixes S :: "'a::euclidean_space set" shows "(finite S \ card S > DIM('a)) \ dependent S" by (metis independent_bound not_less) text \Picking an orthogonal replacement for a spanning set.\ lemma vector_sub_project_orthogonal: fixes b x :: "'a::euclidean_space" shows "b \ (x - ((b \ x) / (b \ b)) *\<^sub>R b) = 0" unfolding inner_simps by auto lemma pairwise_orthogonal_insert: assumes "pairwise orthogonal S" and "\y. y \ S \ orthogonal x y" shows "pairwise orthogonal (insert x S)" using assms by (auto simp: pairwise_def orthogonal_commute) lemma basis_orthogonal: fixes B :: "'a::real_inner set" assumes fB: "finite B" shows "\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C" (is " \C. ?P B C") using fB proof (induct rule: finite_induct) case empty then show ?case using pairwise_empty by blast next case (insert a B) note fB = \finite B\ and aB = \a \ B\ from \\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C\ obtain C where C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast let ?a = "a - sum (\x. (x \ a / (x \ x)) *\<^sub>R x) C" let ?C = "insert ?a C" from C(1) have fC: "finite ?C" by simp have cC: "card ?C \ card (insert a B)" using C aB card_insert_if local.insert(1) by fastforce { fix x k have th0: "\(a::'a) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps) have "x - k *\<^sub>R (a - (\x\C. (x \ a / (x \ x)) *\<^sub>R x)) \ span C \ x - k *\<^sub>R a \ span C" unfolding scaleR_right_diff_distrib th0 by (intro span_add_eq span_scale span_sum span_base) } then have SC: "span ?C = span (insert a B)" unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto { fix y assume yC: "y \ C" then have Cy: "C = insert y (C - {y})" by blast have fth: "finite (C - {y})" using C by simp have "y \ 0 \ \x\C - {y}. x \ a * (x \ y) / (x \ x) = 0" using \pairwise orthogonal C\ by (metis Cy DiffE div_0 insertCI mult_zero_right orthogonal_def pairwise_insert) then have "orthogonal ?a y" unfolding orthogonal_def unfolding inner_diff inner_sum_left right_minus_eq unfolding sum.remove [OF \finite C\ \y \ C\] by (auto simp add: sum.neutral inner_commute[of y a]) } with \pairwise orthogonal C\ have CPO: "pairwise orthogonal ?C" by (rule pairwise_orthogonal_insert) from fC cC SC CPO have "?P (insert a B) ?C" by blast then show ?case by blast qed lemma orthogonal_basis_exists: fixes V :: "('a::euclidean_space) set" shows "\B. independent B \ B \ span V \ V \ span B \ (card B = dim V) \ pairwise orthogonal B" proof - from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "card B = dim V" by force from B have fB: "finite B" "card B = dim V" using independent_bound by auto from basis_orthogonal[OF fB(1)] obtain C where C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast from C B have CSV: "C \ span V" by (metis span_superset span_mono subset_trans) from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) from C fB have "card C \ dim V" by simp moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] by simp ultimately have "card C = dim V" using C(1) by simp with C B CSV show ?thesis by (metis SVC card_eq_dim dim_span) qed text \Low-dimensional subset is in a hyperplane (weak orthogonal complement).\ lemma span_not_univ_orthogonal: fixes S :: "'a::euclidean_space set" assumes sU: "span S \ UNIV" shows "\a::'a. a \ 0 \ (\x \ span S. a \ x = 0)" proof - from sU obtain a where a: "a \ span S" by blast from orthogonal_basis_exists obtain B where B: "independent B" "B \ span S" "S \ span B" "card B = dim S" "pairwise orthogonal B" by blast from B have fB: "finite B" "card B = dim S" using independent_bound by auto have sSB: "span S = span B" by (simp add: B span_eq) let ?a = "a - sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B" have "sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B \ span S" by (simp add: sSB span_base span_mul span_sum) with a have a0:"?a \ 0" by auto have "?a \ x = 0" if "x\span B" for x proof (rule span_induct [OF that]) show "subspace {x. ?a \ x = 0}" by (auto simp add: subspace_def inner_add) next { fix x assume x: "x \ B" from x have B': "B = insert x (B - {x})" by blast have fth: "finite (B - {x})" using fB by simp have "(\b\B - {x}. a \ b * (b \ x) / (b \ b)) = 0" if "x \ 0" by (smt (verit) B' B(5) DiffD2 divide_eq_0_iff inner_real_def inner_zero_right insertCI orthogonal_def pairwise_insert sum.neutral) then have "?a \ x = 0" apply (subst B') using fB fth unfolding sum_clauses(2)[OF fth] by (auto simp add: inner_add_left inner_diff_left inner_sum_left) } then show "?a \ x = 0" if "x \ B" for x using that by blast qed with a0 sSB show ?thesis by blast qed lemma span_not_univ_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes SU: "span S \ UNIV" shows "\ a. a \0 \ span S \ {x. a \ x = 0}" using span_not_univ_orthogonal[OF SU] by auto lemma lowdim_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes d: "dim S < DIM('a)" shows "\a::'a. a \ 0 \ span S \ {x. a \ x = 0}" using d dim_eq_full nless_le span_not_univ_subset_hyperplane by blast lemma linear_eq_stdbasis: fixes f :: "'a::euclidean_space \ _" assumes lf: "linear f" and lg: "linear g" and fg: "\b. b \ Basis \ f b = g b" shows "f = g" using linear_eq_on_span[OF lf lg, of Basis] fg by auto text \Similar results for bilinear functions.\ lemma bilinear_eq: assumes bf: "bilinear f" and bg: "bilinear g" and SB: "S \ span B" and TC: "T \ span C" and "x\S" "y\T" and fg: "\x y. \x \ B; y\ C\ \ f x y = g x y" shows "f x y = g x y" proof - let ?P = "{x. \y\ span C. f x y = g x y}" from bf bg have sp: "subspace ?P" unfolding bilinear_def linear_iff subspace_def bf bg by (auto simp add: span_zero bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) have sfg: "\x. x \ B \ subspace {a. f x a = g x a}" by (auto simp: subspace_def bf bg bilinear_rzero bilinear_radd bilinear_rmul) have "\y\ span C. f x y = g x y" if "x \ span B" for x using span_induct [OF that sp] fg sfg span_induct by blast then show ?thesis using SB TC assms by auto qed lemma bilinear_eq_stdbasis: fixes f :: "'a::euclidean_space \ 'b::euclidean_space \ _" assumes bf: "bilinear f" and bg: "bilinear g" and fg: "\i j. i \ Basis \ j \ Basis \ f i j = g i j" shows "f = g" using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis]] fg by blast subsection \Infinity norm\ definition\<^marker>\tag important\ "infnorm (x::'a::euclidean_space) = Sup {\x \ b\ |b. b \ Basis}" lemma infnorm_set_image: fixes x :: "'a::euclidean_space" shows "{\x \ i\ |i. i \ Basis} = (\i. \x \ i\) ` Basis" by blast lemma infnorm_Max: fixes x :: "'a::euclidean_space" shows "infnorm x = Max ((\i. \x \ i\) ` Basis)" by (simp add: infnorm_def infnorm_set_image cSup_eq_Max) lemma infnorm_set_lemma: fixes x :: "'a::euclidean_space" shows "finite {\x \ i\ |i. i \ Basis}" and "{\x \ i\ |i. i \ Basis} \ {}" unfolding infnorm_set_image by auto lemma infnorm_pos_le: fixes x :: "'a::euclidean_space" shows "0 \ infnorm x" by (simp add: infnorm_Max Max_ge_iff ex_in_conv) lemma infnorm_triangle: fixes x :: "'a::euclidean_space" shows "infnorm (x + y) \ infnorm x + infnorm y" proof - have *: "\a b c d :: real. \a\ \ c \ \b\ \ d \ \a + b\ \ c + d" by simp show ?thesis by (auto simp: infnorm_Max inner_add_left intro!: *) qed lemma infnorm_eq_0: fixes x :: "'a::euclidean_space" shows "infnorm x = 0 \ x = 0" proof - have "infnorm x \ 0 \ x = 0" unfolding infnorm_Max by (simp add: euclidean_all_zero_iff) then show ?thesis using infnorm_pos_le[of x] by simp qed lemma infnorm_0: "infnorm 0 = 0" by (simp add: infnorm_eq_0) lemma infnorm_neg: "infnorm (- x) = infnorm x" unfolding infnorm_def by simp lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" by (metis infnorm_neg minus_diff_eq) lemma absdiff_infnorm: "\infnorm x - infnorm y\ \ infnorm (x - y)" by (smt (verit, del_insts) diff_add_cancel infnorm_sub infnorm_triangle) lemma real_abs_infnorm: "\infnorm x\ = infnorm x" using infnorm_pos_le[of x] by arith lemma Basis_le_infnorm: fixes x :: "'a::euclidean_space" shows "b \ Basis \ \x \ b\ \ infnorm x" by (simp add: infnorm_Max) lemma infnorm_mul: "infnorm (a *\<^sub>R x) = \a\ * infnorm x" unfolding infnorm_Max proof (safe intro!: Max_eqI) let ?B = "(\i. \x \ i\) ` Basis" { fix b :: 'a assume "b \ Basis" then show "\a *\<^sub>R x \ b\ \ \a\ * Max ?B" by (simp add: abs_mult mult_left_mono) next from Max_in[of ?B] obtain b where "b \ Basis" "Max ?B = \x \ b\" by (auto simp del: Max_in) then show "\a\ * Max ((\i. \x \ i\) ` Basis) \ (\i. \a *\<^sub>R x \ i\) ` Basis" by (intro image_eqI[where x=b]) (auto simp: abs_mult) } qed simp lemma infnorm_mul_lemma: "infnorm (a *\<^sub>R x) \ \a\ * infnorm x" unfolding infnorm_mul .. lemma infnorm_pos_lt: "infnorm x > 0 \ x \ 0" using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith text \Prove that it differs only up to a bound from Euclidean norm.\ lemma infnorm_le_norm: "infnorm x \ norm x" by (simp add: Basis_le_norm infnorm_Max) lemma norm_le_infnorm: fixes x :: "'a::euclidean_space" shows "norm x \ sqrt DIM('a) * infnorm x" unfolding norm_eq_sqrt_inner id_def -proof (rule real_le_lsqrt[OF inner_ge_zero]) +proof (rule real_le_lsqrt) show "sqrt DIM('a) * infnorm x \ 0" by (simp add: zero_le_mult_iff infnorm_pos_le) have "x \ x \ (\b\Basis. x \ b * (x \ b))" by (metis euclidean_inner order_refl) also have "\ \ DIM('a) * \infnorm x\\<^sup>2" by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm) also have "\ \ (sqrt DIM('a) * infnorm x)\<^sup>2" by (simp add: power_mult_distrib) finally show "x \ x \ (sqrt DIM('a) * infnorm x)\<^sup>2" . qed lemma tendsto_infnorm [tendsto_intros]: assumes "(f \ a) F" shows "((\x. infnorm (f x)) \ infnorm a) F" proof (rule tendsto_compose [OF LIM_I assms]) fix r :: real assume "r > 0" then show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (infnorm x - infnorm a) < r" by (metis real_norm_def le_less_trans absdiff_infnorm infnorm_le_norm) qed text \Equality in Cauchy-Schwarz and triangle inequalities.\ lemma norm_cauchy_schwarz_eq: "x \ y = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x" (is "?lhs \ ?rhs") proof (cases "x=0") case True then show ?thesis by auto next case False from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"] have "?rhs \ (norm y * (norm y * norm x * norm x - norm x * (x \ y)) - norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" using False unfolding inner_simps by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" using False by (simp add: field_simps inner_commute) also have "\ \ ?lhs" using False by auto finally show ?thesis by metis qed lemma norm_cauchy_schwarz_abs_eq: "\x \ y\ = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm x *\<^sub>R y = - norm y *\<^sub>R x" using norm_cauchy_schwarz_eq [symmetric, of x y] using norm_cauchy_schwarz_eq [symmetric, of "-x" y] Cauchy_Schwarz_ineq2 [of x y] by auto lemma norm_triangle_eq: fixes x y :: "'a::real_inner" shows "norm (x + y) = norm x + norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x" proof (cases "x = 0 \ y = 0") case True then show ?thesis by force next case False then have n: "norm x > 0" "norm y > 0" by auto have "norm (x + y) = norm x + norm y \ (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2" by simp also have "\ \ norm x *\<^sub>R y = norm y *\<^sub>R x" by (smt (verit, best) dot_norm inner_real_def inner_simps norm_cauchy_schwarz_eq power2_eq_square) finally show ?thesis . qed lemma dist_triangle_eq: fixes x y z :: "'a::real_inner" shows "dist x z = dist x y + dist y z \ norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" by (metis (no_types, lifting) add_diff_eq diff_add_cancel dist_norm norm_triangle_eq) subsection \Collinearity\ definition\<^marker>\tag important\ collinear :: "'a::real_vector set \ bool" where "collinear S \ (\u. \x \ S. \ y \ S. \c. x - y = c *\<^sub>R u)" lemma collinear_alt: "collinear S \ (\u v. \x \ S. \c. x = u + c *\<^sub>R v)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding collinear_def by (metis add.commute diff_add_cancel) next assume ?rhs then obtain u v where *: "\x. x \ S \ \c. x = u + c *\<^sub>R v" by auto have "\c. x - y = c *\<^sub>R v" if "x \ S" "y \ S" for x y by (metis *[OF \x \ S\] *[OF \y \ S\] scaleR_left.diff add_diff_cancel_left) then show ?lhs using collinear_def by blast qed lemma collinear: fixes S :: "'a::{perfect_space,real_vector} set" shows "collinear S \ (\u. u \ 0 \ (\x \ S. \ y \ S. \c. x - y = c *\<^sub>R u))" proof - have "\v. v \ 0 \ (\x\S. \y\S. \c. x - y = c *\<^sub>R v)" if "\x\S. \y\S. \c. x - y = c *\<^sub>R u" "u=0" for u proof - have "\x\S. \y\S. x = y" using that by auto moreover obtain v::'a where "v \ 0" using UNIV_not_singleton [of 0] by auto ultimately have "\x\S. \y\S. \c. x - y = c *\<^sub>R v" by auto then show ?thesis using \v \ 0\ by blast qed then show ?thesis by (metis collinear_def) qed lemma collinear_subset: "\collinear T; S \ T\ \ collinear S" by (meson collinear_def subsetCE) lemma collinear_empty [iff]: "collinear {}" by (simp add: collinear_def) lemma collinear_sing [iff]: "collinear {x}" by (simp add: collinear_def) lemma collinear_2 [iff]: "collinear {x, y}" by (simp add: collinear_def) (metis minus_diff_eq scaleR_left.minus scaleR_one) lemma collinear_lemma: "collinear {0, x, y} \ x = 0 \ y = 0 \ (\c. y = c *\<^sub>R x)" (is "?lhs \ ?rhs") proof (cases "x = 0 \ y = 0") case True then show ?thesis by (auto simp: insert_commute) next case False show ?thesis proof assume h: "?lhs" then obtain u where u: "\ x\ {0,x,y}. \y\ {0,x,y}. \c. x - y = c *\<^sub>R u" unfolding collinear_def by blast from u[rule_format, of x 0] u[rule_format, of y 0] obtain cx and cy where cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u" by auto from cx cy False have cx0: "cx \ 0" and cy0: "cy \ 0" by auto let ?d = "cy / cx" from cx cy cx0 have "y = ?d *\<^sub>R x" by simp then show ?rhs using False by blast next assume h: "?rhs" then obtain c where c: "y = c *\<^sub>R x" using False by blast show ?lhs apply (simp add: collinear_def c) by (metis (mono_tags, lifting) scaleR_left.minus scaleR_left_diff_distrib scaleR_one) qed qed lemma collinear_iff_Reals: "collinear {0::complex,w,z} \ z/w \ \" proof show "z/w \ \ \ collinear {0,w,z}" by (metis Reals_cases collinear_lemma nonzero_divide_eq_eq scaleR_conv_of_real) qed (auto simp: collinear_lemma scaleR_conv_of_real) lemma collinear_scaleR_iff: "collinear {0, \ *\<^sub>R w, \ *\<^sub>R z} \ collinear {0,w,z} \ \=0 \ \=0" (is "?lhs = ?rhs") proof (cases "\=0 \ \=0") case False then have "(\c. \ *\<^sub>R z = (c * \) *\<^sub>R w) = (\c. z = c *\<^sub>R w)" by (metis mult.commute scaleR_scaleR vector_fraction_eq_iff) then show ?thesis by (auto simp add: collinear_lemma) qed (auto simp: collinear_lemma) lemma norm_cauchy_schwarz_equal: "\x \ y\ = norm x * norm y \ collinear {0, x, y}" proof (cases "x=0") case True then show ?thesis by (auto simp: insert_commute) next case False then have nnz: "norm x \ 0" by auto show ?thesis proof assume "\x \ y\ = norm x * norm y" then show "collinear {0, x, y}" unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (meson eq_vector_fraction_iff nnz) next assume "collinear {0, x, y}" with False show "\x \ y\ = norm x * norm y" unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (auto simp: abs_if) qed qed lemma norm_triangle_eq_imp_collinear: fixes x y :: "'a::real_inner" assumes "norm (x + y) = norm x + norm y" shows "collinear{0,x,y}" using assms norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq by blast subsection\Properties of special hyperplanes\ lemma subspace_hyperplane: "subspace {x. a \ x = 0}" by (simp add: subspace_def inner_right_distrib) lemma subspace_hyperplane2: "subspace {x. x \ a = 0}" by (simp add: inner_commute inner_right_distrib subspace_def) lemma special_hyperplane_span: fixes S :: "'n::euclidean_space set" assumes "k \ Basis" shows "{x. k \ x = 0} = span (Basis - {k})" proof - have *: "x \ span (Basis - {k})" if "k \ x = 0" for x proof - have "x = (\b\Basis. (x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also have "\ = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" by (auto simp: sum.remove [of _ k] inner_commute assms that) finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . then show ?thesis by (simp add: span_finite) qed show ?thesis apply (rule span_subspace [symmetric]) using assms apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) done qed lemma dim_special_hyperplane: fixes k :: "'n::euclidean_space" shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" by (metis Diff_subset card_Diff_singleton indep_card_eq_dim_span independent_substdbasis special_hyperplane_span) proposition dim_hyperplane: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "dim {x. a \ x = 0} = DIM('a) - 1" proof - have span0: "span {x. a \ x = 0} = {x. a \ x = 0}" by (rule span_unique) (auto simp: subspace_hyperplane) then obtain B where "independent B" and Bsub: "B \ {x. a \ x = 0}" and subspB: "{x. a \ x = 0} \ span B" and card0: "(card B = dim {x. a \ x = 0})" and ortho: "pairwise orthogonal B" using orthogonal_basis_exists by metis with assms have "a \ span B" by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0) then have ind: "independent (insert a B)" by (simp add: \independent B\ independent_insert) have "finite B" using \independent B\ independent_bound by blast have "UNIV \ span (insert a B)" proof fix y::'a obtain r z where "y = r *\<^sub>R a + z" "a \ z = 0" by (metis add.commute diff_add_cancel vector_sub_project_orthogonal) then show "y \ span (insert a B)" by (metis (mono_tags, lifting) Bsub add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_eq subspB) qed then have "DIM('a) = dim(insert a B)" by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI) then show ?thesis by (metis One_nat_def \a \ span B\ \finite B\ card0 card_insert_disjoint diff_Suc_Suc diff_zero dim_eq_card_independent ind span_base) qed lemma lowdim_eq_hyperplane: fixes S :: "'a::euclidean_space set" assumes "dim S = DIM('a) - 1" obtains a where "a \ 0" and "span S = {x. a \ x = 0}" proof - obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" by (metis DIM_positive assms diff_less zero_less_one lowdim_subset_hyperplane) then show ?thesis by (metis assms dim_hyperplane dim_span dim_subset subspace_dim_equal subspace_hyperplane subspace_span that) qed lemma dim_eq_hyperplane: fixes S :: "'n::euclidean_space set" shows "dim S = DIM('n) - 1 \ (\a. a \ 0 \ span S = {x. a \ x = 0})" by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) subsection\ Orthogonal bases and Gram-Schmidt process\ lemma pairwise_orthogonal_independent: assumes "pairwise orthogonal S" and "0 \ S" shows "independent S" proof - have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" using assms by (simp add: pairwise_def orthogonal_def) have "False" if "a \ S" and a: "a \ span (S - {a})" for a proof - obtain T U where "T \ S - {a}" "a = (\v\T. U v *\<^sub>R v)" using a by (force simp: span_explicit) then have "a \ a = a \ (\v\T. U v *\<^sub>R v)" by simp also have "\ = 0" apply (simp add: inner_sum_right) by (smt (verit) "0" DiffE \T \ S - {a}\ in_mono insertCI mult_not_zero sum.neutral that(1)) finally show ?thesis using \0 \ S\ \a \ S\ by auto qed then show ?thesis by (force simp: dependent_def) qed lemma pairwise_orthogonal_imp_finite: fixes S :: "'a::euclidean_space set" assumes "pairwise orthogonal S" shows "finite S" by (metis Set.set_insert assms finite_insert independent_bound pairwise_insert pairwise_orthogonal_independent) lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) lemma subspace_orthogonal_to_vectors: "subspace {y. \x \ S. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) lemma orthogonal_to_span: assumes a: "a \ span S" and x: "\y. y \ S \ orthogonal x y" shows "orthogonal x a" by (metis a orthogonal_clauses(1,2,4) span_induct_alt x) proposition Gram_Schmidt_step: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" and x: "x \ span S" shows "orthogonal x (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b))" proof - have "finite S" by (simp add: S pairwise_orthogonal_imp_finite) have "orthogonal (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)) x" if "x \ S" for x proof - have "a \ x = (\y\S. if y = x then y \ a else 0)" by (simp add: \finite S\ inner_commute that) also have "\ = (\b\S. b \ a * (b \ x) / (b \ b))" apply (rule sum.cong [OF refl], simp) by (meson S orthogonal_def pairwise_def that) finally show ?thesis by (simp add: orthogonal_def algebra_simps inner_sum_left) qed then show ?thesis using orthogonal_to_span orthogonal_commute x by blast qed lemma orthogonal_extension_aux: fixes S :: "'a::euclidean_space set" assumes "finite T" "finite S" "pairwise orthogonal S" shows "\U. pairwise orthogonal (S \ U) \ span (S \ U) = span (S \ T)" using assms proof (induction arbitrary: S) case empty then show ?case by simp (metis sup_bot_right) next case (insert a T) have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" using insert by (simp add: pairwise_def orthogonal_def) define a' where "a' = a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)" obtain U where orthU: "pairwise orthogonal (S \ insert a' U)" and spanU: "span (insert a' S \ U) = span (insert a' S \ T)" by (rule exE [OF insert.IH [of "insert a' S"]]) (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) have orthS: "\x. x \ S \ a' \ x = 0" using Gram_Schmidt_step a'_def insert.prems orthogonal_commute orthogonal_def span_base by blast have "span (S \ insert a' U) = span (insert a' (S \ T))" using spanU by simp also have "\ = span (insert a (S \ T))" by (simp add: a'_def span_neg span_sum span_base span_mul eq_span_insert_eq) also have "\ = span (S \ insert a T)" by simp finally show ?case using orthU by blast qed proposition orthogonal_extension: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - obtain B where "finite B" "span B = span T" using basis_subspace_exists [of "span T"] subspace_span by metis with orthogonal_extension_aux [of B S] obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ B)" using assms pairwise_orthogonal_imp_finite by auto with \span B = span T\ show ?thesis by (rule_tac U=U in that) (auto simp: span_Un) qed corollary\<^marker>\tag unimportant\ orthogonal_extension_strong: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - obtain U where U: "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" using orthogonal_extension assms by blast moreover have "pairwise orthogonal (S \ (U - insert 0 S))" by (smt (verit, best) Un_Diff_Int Un_iff U pairwise_def) ultimately show ?thesis by (metis Diff_disjoint Un_Diff_cancel Un_insert_left inf_commute span_insert_0 that) qed subsection\Decomposing a vector into parts in orthogonal subspaces\ text\existence of orthonormal basis for a subspace.\ lemma orthogonal_spanningset_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" "span B = S" by (metis assms basis_orthogonal basis_subspace_exists span_eq) lemma orthogonal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "0 \ B" "B \ S" "pairwise orthogonal B" "independent B" "card B = dim S" "span B = S" by (metis assms dependent_zero orthogonal_basis_exists span_eq span_eq_iff) proposition orthonormal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = S" proof - obtain B where "0 \ B" "B \ S" and orth: "pairwise orthogonal B" and "independent B" "card B = dim S" "span B = S" by (blast intro: orthogonal_basis_subspace [OF assms]) have 1: "(\x. x /\<^sub>R norm x) ` B \ S" using \span B = S\ span_superset span_mul by fastforce have 2: "pairwise orthogonal ((\x. x /\<^sub>R norm x) ` B)" using orth by (force simp: pairwise_def orthogonal_clauses) have 3: "\x. x \ (\x. x /\<^sub>R norm x) ` B \ norm x = 1" by (metis (no_types, lifting) \0 \ B\ image_iff norm_sgn sgn_div_norm) have 4: "independent ((\x. x /\<^sub>R norm x) ` B)" by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) B" proof fix x y assume "x \ B" "y \ B" "x /\<^sub>R norm x = y /\<^sub>R norm y" moreover have "\i. i \ B \ norm (i /\<^sub>R norm i) = 1" using 3 by blast ultimately show "x = y" by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) qed then have 5: "card ((\x. x /\<^sub>R norm x) ` B) = dim S" by (metis \card B = dim S\ card_image) have 6: "span ((\x. x /\<^sub>R norm x) ` B) = S" by (metis "1" "4" "5" assms card_eq_dim independent_imp_finite span_subspace) show ?thesis by (rule that [OF 1 2 3 4 5 6]) qed proposition\<^marker>\tag unimportant\ orthogonal_to_subspace_exists_gen: fixes S :: "'a :: euclidean_space set" assumes "span S \ span T" obtains x where "x \ 0" "x \ span T" "\y. y \ span S \ orthogonal x y" proof - obtain B where "B \ span S" and orthB: "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = span S" by (metis dim_span orthonormal_basis_subspace subspace_span) with assms obtain u where spanBT: "span B \ span T" and "u \ span B" "u \ span T" by auto obtain C where orthBC: "pairwise orthogonal (B \ C)" and spanBC: "span (B \ C) = span (B \ {u})" by (blast intro: orthogonal_extension [OF orthB]) show thesis proof (cases "C \ insert 0 B") case True then have "C \ span B" using span_eq by (metis span_insert_0 subset_trans) moreover have "u \ span (B \ C)" using \span (B \ C) = span (B \ {u})\ span_superset by force ultimately show ?thesis using True \u \ span B\ by (metis Un_insert_left span_insert_0 sup.orderE) next case False then obtain x where "x \ C" "x \ 0" "x \ B" by blast then have "x \ span T" by (smt (verit, ccfv_SIG) Set.set_insert \u \ span T\ empty_subsetI insert_subset le_sup_iff spanBC spanBT span_mono span_span span_superset subset_trans) moreover have "orthogonal x y" if "y \ span B" for y using that proof (rule span_induct) show "subspace {a. orthogonal x a}" by (simp add: subspace_orthogonal_to_vector) show "\b. b \ B \ orthogonal x b" by (metis Un_iff \x \ C\ \x \ B\ orthBC pairwise_def) qed ultimately show ?thesis using \x \ 0\ that \span B = span S\ by auto qed qed corollary\<^marker>\tag unimportant\ orthogonal_to_subspace_exists: fixes S :: "'a :: euclidean_space set" assumes "dim S < DIM('a)" obtains x where "x \ 0" "\y. y \ span S \ orthogonal x y" proof - have "span S \ UNIV" by (metis assms dim_eq_full order_less_imp_not_less top.not_eq_extremum) with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis by (auto) qed corollary\<^marker>\tag unimportant\ orthogonal_to_vector_exists: fixes x :: "'a :: euclidean_space" assumes "2 \ DIM('a)" obtains y where "y \ 0" "orthogonal x y" proof - have "dim {x} < DIM('a)" using assms by auto then show thesis by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_base that) qed proposition\<^marker>\tag unimportant\ orthogonal_subspace_decomp_exists: fixes S :: "'a :: euclidean_space set" obtains y z where "y \ span S" and "\w. w \ span S \ orthogonal z w" and "x = y + z" proof - obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" using orthogonal_basis_subspace subspace_span by blast let ?a = "\b\T. (b \ x / (b \ b)) *\<^sub>R b" have orth: "orthogonal (x - ?a) w" if "w \ span S" for w by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) with that[of ?a "x-?a"] \T \ span S\ show ?thesis by (simp add: span_mul span_sum subsetD) qed lemma orthogonal_subspace_decomp_unique: fixes S :: "'a :: euclidean_space set" assumes "x + y = x' + y'" and ST: "x \ span S" "x' \ span S" "y \ span T" "y' \ span T" and orth: "\a b. \a \ S; b \ T\ \ orthogonal a b" shows "x = x' \ y = y'" proof - have "x + y - y' = x'" by (simp add: assms) moreover have "\a b. \a \ span S; b \ span T\ \ orthogonal a b" by (meson orth orthogonal_commute orthogonal_to_span) ultimately have "0 = x' - x" using assms by (metis add.commute add_diff_cancel_right' diff_right_commute orthogonal_self span_diff) with assms show ?thesis by auto qed lemma vector_in_orthogonal_spanningset: fixes a :: "'a::euclidean_space" obtains S where "a \ S" "pairwise orthogonal S" "span S = UNIV" by (metis UnI1 Un_UNIV_right insertI1 orthogonal_extension pairwise_singleton span_UNIV) lemma vector_in_orthogonal_basis: fixes a :: "'a::euclidean_space" assumes "a \ 0" obtains S where "a \ S" "0 \ S" "pairwise orthogonal S" "independent S" "finite S" "span S = UNIV" "card S = DIM('a)" proof - obtain S where S: "a \ S" "pairwise orthogonal S" "span S = UNIV" using vector_in_orthogonal_spanningset . show thesis proof show "pairwise orthogonal (S - {0})" using pairwise_mono S(2) by blast show "independent (S - {0})" by (simp add: \pairwise orthogonal (S - {0})\ pairwise_orthogonal_independent) show "finite (S - {0})" using \independent (S - {0})\ independent_imp_finite by blast show "card (S - {0}) = DIM('a)" using span_delete_0 [of S] S by (simp add: \independent (S - {0})\ indep_card_eq_dim_span) qed (use S \a \ 0\ in auto) qed lemma vector_in_orthonormal_basis: fixes a :: "'a::euclidean_space" assumes "norm a = 1" obtains S where "a \ S" "pairwise orthogonal S" "\x. x \ S \ norm x = 1" "independent S" "card S = DIM('a)" "span S = UNIV" proof - have "a \ 0" using assms by auto then obtain S where "a \ S" "0 \ S" "finite S" and S: "pairwise orthogonal S" "independent S" "span S = UNIV" "card S = DIM('a)" by (metis vector_in_orthogonal_basis) let ?S = "(\x. x /\<^sub>R norm x) ` S" show thesis proof show "a \ ?S" using \a \ S\ assms image_iff by fastforce next show "pairwise orthogonal ?S" using \pairwise orthogonal S\ by (auto simp: pairwise_def orthogonal_def) show "\x. x \ (\x. x /\<^sub>R norm x) ` S \ norm x = 1" using \0 \ S\ by (auto simp: field_split_simps) then show ind: "independent ?S" by (metis \pairwise orthogonal ((\x. x /\<^sub>R norm x) ` S)\ norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) S" unfolding inj_on_def by (metis (full_types) S(1) \0 \ S\ inverse_nonzero_iff_nonzero norm_eq_zero orthogonal_scaleR orthogonal_self pairwise_def) then show "card ?S = DIM('a)" by (simp add: card_image S) then show "span ?S = UNIV" by (metis ind dim_eq_card dim_eq_full) qed qed proposition dim_orthogonal_sum: fixes A :: "'a::euclidean_space set" assumes "\x y. \x \ A; y \ B\ \ x \ y = 0" shows "dim(A \ B) = dim A + dim B" proof - have 1: "\x y. \x \ span A; y \ B\ \ x \ y = 0" by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) have "\x y. \x \ span A; y \ span B\ \ x \ y = 0" using 1 by (simp add: span_induct [OF _ subspace_hyperplane]) then have 0: "\x y. \x \ span A; y \ span B\ \ x \ y = 0" by simp have "dim(A \ B) = dim (span (A \ B))" by (simp) also have "span (A \ B) = ((\(a, b). a + b) ` (span A \ span B))" by (auto simp add: span_Un image_def) also have "dim \ = dim {x + y |x y. x \ span A \ y \ span B}" by (auto intro!: arg_cong [where f=dim]) also have "\ = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" by (auto dest: 0) also have "\ = dim A + dim B" using dim_sums_Int by fastforce finally show ?thesis . qed lemma dim_subspace_orthogonal_to_vectors: fixes A :: "'a::euclidean_space set" assumes "subspace A" "subspace B" "A \ B" shows "dim {y \ B. \x \ A. orthogonal x y} + dim A = dim B" proof - have "dim (span ({y \ B. \x\A. orthogonal x y} \ A)) = dim (span B)" proof (rule arg_cong [where f=dim, OF subset_antisym]) show "span ({y \ B. \x\A. orthogonal x y} \ A) \ span B" by (simp add: \A \ B\ Collect_restrict span_mono) next have *: "x \ span ({y \ B. \x\A. orthogonal x y} \ A)" if "x \ B" for x proof - obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" using orthogonal_subspace_decomp_exists [of A x] that by auto moreover have "y \ span B" using \y \ span A\ assms(3) span_mono by blast ultimately have "z \ B \ (\x. x \ A \ orthogonal x z)" using assms by (metis orthogonal_commute span_add_eq span_eq_iff that) then have z: "z \ span {y \ B. \x\A. orthogonal x y}" by (simp add: span_base) then show ?thesis by (smt (verit, best) \x = y + z\ \y \ span A\ le_sup_iff span_add_eq span_subspace_induct span_superset subset_iff subspace_span) qed show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" by (rule span_minimal) (auto intro: * span_minimal) qed then show ?thesis by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) qed subsection\Linear functions are (uniformly) continuous on any set\ subsection\<^marker>\tag unimportant\ \Topological properties of linear functions\ lemma linear_lim_0: assumes "bounded_linear f" shows "(f \ 0) (at (0))" proof - interpret f: bounded_linear f by fact have "(f \ f 0) (at 0)" using tendsto_ident_at by (rule f.tendsto) then show ?thesis unfolding f.zero . qed lemma linear_continuous_at: "bounded_linear f \continuous (at a) f" by (simp add: bounded_linear.isUCont isUCont_isCont) lemma linear_continuous_within: "bounded_linear f \ continuous (at x within s) f" using continuous_at_imp_continuous_at_within linear_continuous_at by blast lemma linear_continuous_on: "bounded_linear f \ continuous_on s f" using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto lemma Lim_linear: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and h :: "'b \ 'c::real_normed_vector" assumes "(f \ l) F" "linear h" shows "((\x. h(f x)) \ h l) F" proof - obtain B where B: "B > 0" "\x. norm (h x) \ B * norm x" using linear_bounded_pos [OF \linear h\] by blast show ?thesis unfolding tendsto_iff by (simp add: assms bounded_linear.tendsto linear_linear tendstoD) qed lemma linear_continuous_compose: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" assumes "continuous F f" "linear g" shows "continuous F (\x. g(f x))" using assms unfolding continuous_def by (rule Lim_linear) lemma linear_continuous_on_compose: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" assumes "continuous_on S f" "linear g" shows "continuous_on S (\x. g(f x))" using assms by (simp add: continuous_on_eq_continuous_within linear_continuous_compose) text\Also bilinear functions, in composition form\ lemma bilinear_continuous_compose: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes "continuous F f" "continuous F g" "bilinear h" shows "continuous F (\x. h (f x) (g x))" using assms bilinear_conv_bounded_bilinear bounded_bilinear.continuous by blast lemma bilinear_continuous_on_compose: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" and f :: "'d::t2_space \ 'a" assumes "continuous_on S f" "continuous_on S g" "bilinear h" shows "continuous_on S (\x. h (f x) (g x))" using assms by (simp add: continuous_on_eq_continuous_within bilinear_continuous_compose) end diff --git a/src/HOL/Analysis/Urysohn.thy b/src/HOL/Analysis/Urysohn.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Urysohn.thy @@ -0,0 +1,2240 @@ +(* Title: HOL/Analysis/Arcwise_Connected.thy + Authors: LC Paulson, based on material from HOL Light +*) + +section \Urysohn lemma and its Consequences\ + +theory Urysohn +imports Abstract_Topological_Spaces Abstract_Metric_Spaces Infinite_Sum Arcwise_Connected +begin + +subsection \Urysohn lemma and Tietze's theorem\ + +lemma Urysohn_lemma: + fixes a b :: real + assumes "normal_space X" "closedin X S" "closedin X T" "disjnt S T" "a \ b" + obtains f where "continuous_map X (top_of_set {a..b}) f" "f ` S \ {a}" "f ` T \ {b}" +proof - + obtain U where "openin X U" "S \ U" "X closure_of U \ topspace X - T" + using assms unfolding normal_space_alt disjnt_def + by (metis Diff_mono Un_Diff_Int closedin_def subset_eq sup_bot_right) + have "\G :: real \ 'a set. G 0 = U \ G 1 = topspace X - T \ + (\x \ dyadics \ {0..1}. \y \ dyadics \ {0..1}. x < y \ openin X (G x) \ openin X (G y) \ X closure_of (G x) \ G y)" + proof (rule recursion_on_dyadic_fractions) + show "openin X U \ openin X (topspace X - T) \ X closure_of U \ topspace X - T" + using \X closure_of U \ topspace X - T\ \openin X U\ \closedin X T\ by blast + show "\z. (openin X x \ openin X z \ X closure_of x \ z) \ openin X z \ openin X y \ X closure_of z \ y" + if "openin X x \ openin X y \ X closure_of x \ y" for x y + by (meson that closedin_closure_of normal_space_alt \normal_space X\) + show "openin X x \ openin X z \ X closure_of x \ z" + if "openin X x \ openin X y \ X closure_of x \ y" and "openin X y \ openin X z \ X closure_of y \ z" for x y z + by (meson that closure_of_subset openin_subset subset_trans) + qed + then obtain G :: "real \ 'a set" + where G0: "G 0 = U" and G1: "G 1 = topspace X - T" + and G: "\x y. \x \ dyadics; y \ dyadics; 0 \ x; x < y; y \ 1\ + \ openin X (G x) \ openin X (G y) \ X closure_of (G x) \ G y" + by (smt (verit, del_insts) Int_iff atLeastAtMost_iff) + define f where "f \ \x. Inf(insert 1 {r. r \ dyadics \ {0..1} \ x \ G r})" + have f_ge: "f x \ 0" if "x \ topspace X" for x + unfolding f_def by (force intro: cInf_greatest) + moreover have f_le1: "f x \ 1" if "x \ topspace X" for x + proof - + have "bdd_below {r \ dyadics \ {0..1}. x \ G r}" + by (auto simp: bdd_below_def) + then show ?thesis + by (auto simp: f_def cInf_lower) + qed + ultimately have fim: "f ` topspace X \ {0..1}" + by (auto simp: f_def) + have 0: "0 \ dyadics \ {0..1::real}" and 1: "1 \ dyadics \ {0..1::real}" + by (force simp: dyadics_def)+ + then have opeG: "openin X (G r)" if "r \ dyadics \ {0..1}" for r + using G G0 \openin X U\ less_eq_real_def that by auto + have "x \ G 0" if "x \ S" for x + using G0 \S \ U\ that by blast + with 0 have fimS: "f ` S \ {0}" + unfolding f_def by (force intro!: cInf_eq_minimum) + have False if "r \ dyadics" "0 \ r" "r < 1" "x \ G r" "x \ T" for r x + using G [of r 1] 1 + by (smt (verit, best) DiffD2 G1 Int_iff closure_of_subset inf.orderE openin_subset that) + then have "r\1" if "r \ dyadics" "0 \ r" "r \ 1" "x \ G r" "x \ T" for r x + using linorder_not_le that by blast + then have fimT: "f ` T \ {1}" + unfolding f_def by (force intro!: cInf_eq_minimum) + have fle1: "f z \ 1" for z + by (force simp: f_def intro: cInf_lower) + have fle: "f z \ x" if "x \ dyadics \ {0..1}" "z \ G x" for z x + using that by (force simp: f_def intro: cInf_lower) + have *: "b \ f z" if "b \ 1" "\x. \x \ dyadics \ {0..1}; z \ G x\ \ b \ x" for z b + using that by (force simp: f_def intro: cInf_greatest) + have **: "r \ f x" if r: "r \ dyadics \ {0..1}" "x \ G r" for r x + proof (rule *) + show "r \ s" if "s \ dyadics \ {0..1}" and "x \ G s" for s :: real + using that r G [of s r] by (force simp add: dest: closure_of_subset openin_subset) + qed (use that in force) + + have "\U. openin X U \ x \ U \ (\y \ U. \f y - f x\ < \)" + if "x \ topspace X" and "0 < \" for x \ + proof - + have A: "\r. r \ dyadics \ {0..1} \ r < y \ \r - y\ < d" if "0 < y" "y \ 1" "0 < d" for y d::real + proof - + obtain n q r + where "of_nat q / 2^n < y" "y < of_nat r / 2^n" "\q / 2^n - r / 2^n \ < d" + by (smt (verit, del_insts) padic_rational_approximation_straddle_pos \0 < d\ \0 < y\) + then show ?thesis + unfolding dyadics_def + using divide_eq_0_iff that(2) by fastforce + qed + have B: "\r. r \ dyadics \ {0..1} \ y < r \ \r - y\ < d" if "0 \ y" "y < 1" "0 < d" for y d::real + proof - + obtain n q r + where "of_nat q / 2^n \ y" "y < of_nat r / 2^n" "\q / 2^n - r / 2^n \ < d" + using padic_rational_approximation_straddle_pos_le + by (smt (verit, del_insts) \0 < d\ \0 \ y\) + then show ?thesis + apply (clarsimp simp: dyadics_def) + using divide_eq_0_iff \y < 1\ + by (smt (verit) divide_nonneg_nonneg divide_self of_nat_0_le_iff of_nat_1 power_0 zero_le_power) + qed + show ?thesis + proof (cases "f x = 0") + case True + with B[of 0] obtain r where r: "r \ dyadics \ {0..1}" "0 < r" "\r\ < \/2" + by (smt (verit) \0 < \\ half_gt_zero) + show ?thesis + proof (intro exI conjI) + show "openin X (G r)" + using opeG r(1) by blast + show "x \ G r" + using True ** r by force + show "\y \ G r. \f y - f x\ < \" + using f_ge \openin X (G r)\ fle openin_subset r by (fastforce simp: True) + qed + next + case False + show ?thesis + proof (cases "f x = 1") + case True + with A[of 1] obtain r where r: "r \ dyadics \ {0..1}" "r < 1" "\r-1\ < \/2" + by (smt (verit) \0 < \\ half_gt_zero) + define G' where "G' \ topspace X - X closure_of G r" + show ?thesis + proof (intro exI conjI) + show "openin X G'" + unfolding G'_def by fastforce + obtain r' where "r' \ dyadics \ 0 \ r' \ r' \ 1 \ r < r' \ \r' - r\ < 1 - r" + using B r by force + moreover + have "1 - r \ dyadics" "0 \ r" + using 1 r dyadics_diff by force+ + ultimately have "x \ X closure_of G r" + using G True r fle by force + then show "x \ G'" + by (simp add: G'_def that) + show "\y \ G'. \f y - f x\ < \" + using ** f_le1 in_closure_of r by (fastforce simp add: True G'_def) + qed + next + case False + have "0 < f x" "f x < 1" + using fle1 f_ge that(1) \f x \ 0\ \f x \ 1\ by (metis order_le_less) + + obtain r where r: "r \ dyadics \ {0..1}" "r < f x" "\r - f x\ < \ / 2" + using A \0 < \\ \0 < f x\ \f x < 1\ by (smt (verit, best) half_gt_zero) + obtain r' where r': "r' \ dyadics \ {0..1}" "f x < r'" "\r' - f x\ < \ / 2" + using B \0 < \\ \0 < f x\ \f x < 1\ by (smt (verit, best) half_gt_zero) + have "r < 1" + using \f x < 1\ r(2) by force + show ?thesis + proof (intro conjI exI) + show "openin X (G r' - X closure_of G r)" + using closedin_closure_of opeG r' by blast + have "x \ X closure_of G r \ False" + using B [of r "f x - r"] r \r < 1\ G [of r] fle by force + then show "x \ G r' - X closure_of G r" + using ** r' by fastforce + show "\y\G r' - X closure_of G r. \f y - f x\ < \" + using r r' ** G closure_of_subset field_sum_of_halves fle openin_subset subset_eq + by (smt (verit) DiffE opeG) + qed + qed + qed + qed + then have contf: "continuous_map X (top_of_set {0..1}) f" + by (force simp add: Met_TC.continuous_map_to_metric dist_real_def continuous_map_in_subtopology fim simp flip: Met_TC.mtopology_is_euclideanreal) + define g where "g \ \x. a + (b - a) * f x" + show thesis + proof + have "continuous_map X euclideanreal g" + using contf \a \ b\ unfolding g_def by (auto simp: continuous_intros continuous_map_in_subtopology) + moreover have "g ` (topspace X) \ {a..b}" + using mult_left_le [of "f _" "b-a"] contf \a \ b\ + by (simp add: g_def add.commute continuous_map_in_subtopology image_subset_iff le_diff_eq) + ultimately show "continuous_map X (top_of_set {a..b}) g" + by (meson continuous_map_in_subtopology) + show "g ` S \ {a}" "g ` T \ {b}" + using fimS fimT by (auto simp: g_def) + qed +qed + +lemma Urysohn_lemma_alt: + fixes a b :: real + assumes "normal_space X" "closedin X S" "closedin X T" "disjnt S T" + obtains f where "continuous_map X euclideanreal f" "f ` S \ {a}" "f ` T \ {b}" + by (metis Urysohn_lemma assms continuous_map_in_subtopology disjnt_sym linear) + +lemma normal_space_iff_Urysohn_gen_alt: + assumes "a \ b" + shows "normal_space X \ + (\S T. closedin X S \ closedin X T \ disjnt S T + \ (\f. continuous_map X euclideanreal f \ f ` S \ {a} \ f ` T \ {b}))" + (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + by (metis Urysohn_lemma_alt) +next + assume R: ?rhs + show ?lhs + unfolding normal_space_def + proof clarify + fix S T + assume "closedin X S" and "closedin X T" and "disjnt S T" + with R obtain f where contf: "continuous_map X euclideanreal f" and "f ` S \ {a}" "f ` T \ {b}" + by meson + show "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" + proof (intro conjI exI) + show "openin X {x \ topspace X. f x \ ball a (\a - b\ / 2)}" + by (force intro!: openin_continuous_map_preimage [OF contf]) + show "openin X {x \ topspace X. f x \ ball b (\a - b\ / 2)}" + by (force intro!: openin_continuous_map_preimage [OF contf]) + show "S \ {x \ topspace X. f x \ ball a (\a - b\ / 2)}" + using \closedin X S\ closedin_subset \f ` S \ {a}\ assms by force + show "T \ {x \ topspace X. f x \ ball b (\a - b\ / 2)}" + using \closedin X T\ closedin_subset \f ` T \ {b}\ assms by force + have "\x. \x \ topspace X; dist a (f x) < \a-b\/2; dist b (f x) < \a-b\/2\ \ False" + by (smt (verit, best) dist_real_def dist_triangle_half_l) + then show "disjnt {x \ topspace X. f x \ ball a (\a-b\ / 2)} {x \ topspace X. f x \ ball b (\a-b\ / 2)}" + using disjnt_iff by fastforce + qed + qed +qed + +lemma normal_space_iff_Urysohn_gen: + fixes a b::real + shows + "a < b \ + normal_space X \ + (\S T. closedin X S \ closedin X T \ disjnt S T + \ (\f. continuous_map X (top_of_set {a..b}) f \ + f ` S \ {a} \ f ` T \ {b}))" + by (metis linear not_le Urysohn_lemma normal_space_iff_Urysohn_gen_alt continuous_map_in_subtopology) + +lemma normal_space_iff_Urysohn_alt: + "normal_space X \ + (\S T. closedin X S \ closedin X T \ disjnt S T + \ (\f. continuous_map X euclideanreal f \ + f ` S \ {0} \ f ` T \ {1}))" + by (rule normal_space_iff_Urysohn_gen_alt) auto + +lemma normal_space_iff_Urysohn: + "normal_space X \ + (\S T. closedin X S \ closedin X T \ disjnt S T + \ (\f::'a\real. continuous_map X (top_of_set {0..1}) f \ + f ` S \ {0} \ f ` T \ {1}))" + by (rule normal_space_iff_Urysohn_gen) auto + +lemma normal_space_perfect_map_image: + "\normal_space X; perfect_map X Y f\ \ normal_space Y" + unfolding perfect_map_def proper_map_def + using normal_space_continuous_closed_map_image by fastforce + +lemma Hausdorff_normal_space_closed_continuous_map_image: + "\normal_space X; closed_map X Y f; continuous_map X Y f; + f ` topspace X = topspace Y; t1_space Y\ + \ Hausdorff_space Y" + by (metis normal_space_continuous_closed_map_image normal_t1_imp_Hausdorff_space) + +lemma normal_Hausdorff_space_closed_continuous_map_image: + "\normal_space X; Hausdorff_space X; closed_map X Y f; + continuous_map X Y f; f ` topspace X = topspace Y\ + \ normal_space Y \ Hausdorff_space Y" + by (meson normal_space_continuous_closed_map_image normal_t1_eq_Hausdorff_space t1_space_closed_map_image) + +lemma Lindelof_cover: + assumes "regular_space X" and "Lindelof_space X" and "S \ {}" + and clo: "closedin X S" "closedin X T" "disjnt S T" + obtains h :: "nat \ 'a set" where + "\n. openin X (h n)" "\n. disjnt T (X closure_of (h n))" and "S \ \(range h)" +proof - + have "\U. openin X U \ x \ U \ disjnt T (X closure_of U)" + if "x \ S" for x + using \regular_space X\ unfolding regular_space + by (metis (full_types) Diff_iff \disjnt S T\ clo closure_of_eq disjnt_iff in_closure_of that) + then obtain h where oh: "\x. x \ S \ openin X (h x)" + and xh: "\x. x \ S \ x \ h x" + and dh: "\x. x \ S \ disjnt T (X closure_of h x)" + by metis + have "Lindelof_space(subtopology X S)" + by (simp add: Lindelof_space_closedin_subtopology \Lindelof_space X\ \closedin X S\) + then obtain \ where \: "countable \ \ \ \ h ` S \ S \ \\" + unfolding Lindelof_space_subtopology_subset [OF closedin_subset [OF \closedin X S\]] + by (smt (verit, del_insts) oh xh UN_I image_iff subsetI) + with \S \ {}\ have "\ \ {}" + by blast + show ?thesis + proof + show "openin X (from_nat_into \ n)" for n + by (metis \ from_nat_into image_iff \\ \ {}\ oh subsetD) + show "disjnt T (X closure_of (from_nat_into \) n)" for n + using dh from_nat_into [OF \\ \ {}\] + by (metis \ f_inv_into_f inv_into_into subset_eq) + show "S \ \(range (from_nat_into \))" + by (simp add: \ \\ \ {}\) + qed +qed + +lemma regular_Lindelof_imp_normal_space: + assumes "regular_space X" and "Lindelof_space X" + shows "normal_space X" + unfolding normal_space_def +proof clarify + fix S T + assume clo: "closedin X S" "closedin X T" and "disjnt S T" + show "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" + proof (cases "S={} \ T={}") + case True + with clo show ?thesis + by (meson closedin_def disjnt_empty1 disjnt_empty2 openin_empty openin_topspace subset_empty) + next + case False + obtain h :: "nat \ 'a set" where + opeh: "\n. openin X (h n)" and dish: "\n. disjnt T (X closure_of (h n))" + and Sh: "S \ \(range h)" + by (metis Lindelof_cover False \disjnt S T\ assms clo) + obtain k :: "nat \ 'a set" where + opek: "\n. openin X (k n)" and disk: "\n. disjnt S (X closure_of (k n))" + and Tk: "T \ \(range k)" + by (metis Lindelof_cover False \disjnt S T\ assms clo disjnt_sym) + define U where "U \ \i. h i - (\j \i. k i - (\j\i. X closure_of h j)" + show ?thesis + proof (intro exI conjI) + show "openin X U" "openin X V" + unfolding U_def V_def + by (force intro!: opek opeh closedin_Union closedin_closure_of)+ + show "S \ U" "T \ V" + using Sh Tk dish disk by (fastforce simp: U_def V_def disjnt_iff)+ + have "\x i j. \x \ k i; x \ h j; \j\i. x \ X closure_of h j\ + \ \i X closure_of k i" + by (metis in_closure_of linorder_not_less opek openin_subset subsetD) + then show "disjnt U V" + by (force simp add: U_def V_def disjnt_iff) + qed + qed +qed + +lemma Tietze_extension_closed_real_interval: + assumes "normal_space X" and "closedin X S" + and contf: "continuous_map (subtopology X S) euclideanreal f" + and fim: "f ` S \ {a..b}" and "a \ b" + obtains g + where "continuous_map X euclideanreal g" + "\x. x \ S \ g x = f x" "g ` topspace X \ {a..b}" +proof - + define c where "c \ max \a\ \b\ + 1" + have "0 < c" and c: "\x. x \ S \ \f x\ \ c" + using fim by (auto simp: c_def image_subset_iff) + define good where + "good \ \g n. continuous_map X euclideanreal g \ (\x \ S. \f x - g x\ \ c * (2/3)^n)" + have step: "\g. good g (Suc n) \ + (\x \ topspace X. \g x - h x\ \ c * (2/3)^n / 3)" + if h: "good h n" for n h + proof - + have pos: "0 < c * (2/3) ^ n" + by (simp add: \0 < c\) + have S_eq: "S = topspace(subtopology X S)" and "S \ topspace X" + using \closedin X S\ closedin_subset by auto + define d where "d \ c/3 * (2/3) ^ n" + define SA where "SA \ {x \ S. f x - h x \ {..-d}}" + define SB where "SB \ {x \ S. f x - h x \ {d..}}" + have contfh: "continuous_map (subtopology X S) euclideanreal (\x. f x - h x)" + using that + by (simp add: contf good_def continuous_map_diff continuous_map_from_subtopology) + then have "closedin (subtopology X S) SA" + unfolding SA_def + by (smt (verit, del_insts) closed_closedin continuous_map_closedin Collect_cong S_eq closed_real_atMost) + then have "closedin X SA" + using \closedin X S\ closedin_trans_full by blast + moreover have "closedin (subtopology X S) SB" + unfolding SB_def + using closedin_continuous_map_preimage_gen [OF contfh] + by (metis (full_types) S_eq closed_atLeast closed_closedin closedin_topspace) + then have "closedin X SB" + using \closedin X S\ closedin_trans_full by blast + moreover have "disjnt SA SB" + using pos by (auto simp: d_def disjnt_def SA_def SB_def) + moreover have "-d \ d" + using pos by (auto simp: d_def) + ultimately + obtain g where contg: "continuous_map X (top_of_set {- d..d}) g" + and ga: "g ` SA \ {- d}" and gb: "g ` SB \ {d}" + using Urysohn_lemma \normal_space X\ by metis + then have g_le_d: "\x. x \ topspace X \ \g x\ \ d" + by (simp add: abs_le_iff continuous_map_def minus_le_iff) + have g_eq_d: "\x. \x \ S; f x - h x \ -d\ \ g x = -d" + using ga by (auto simp: SA_def) + have g_eq_negd: "\x. \x \ S; f x - h x \ d\ \ g x = d" + using gb by (auto simp: SB_def) + show ?thesis + unfolding good_def + proof (intro conjI strip exI) + show "continuous_map X euclideanreal (\x. h x + g x)" + using contg continuous_map_add continuous_map_in_subtopology that + unfolding good_def by blast + show "\f x - (h x + g x)\ \ c * (2 / 3) ^ Suc n" if "x \ S" for x + proof - + have x: "x \ topspace X" + using \S \ topspace X\ that by auto + have "\f x - h x\ \ c * (2/3) ^ n" + using good_def h that by blast + with g_eq_d [OF that] g_eq_negd [OF that] g_le_d [OF x] + have "\f x - (h x + g x)\ \ d + d" + unfolding d_def by linarith + then show ?thesis + by (simp add: d_def) + qed + show "\h x + g x - h x\ \ c * (2 / 3) ^ n / 3" if "x \ topspace X" for x + using that d_def g_le_d by auto + qed + qed + then obtain nxtg where nxtg: "\h n. good h n \ + good (nxtg h n) (Suc n) \ (\x \ topspace X. \nxtg h n x - h x\ \ c * (2/3)^n / 3)" + by metis + define g where "g \ rec_nat (\x. 0) (\n r. nxtg r n)" + have [simp]: "g 0 x = 0" for x + by (auto simp: g_def) + have g_Suc: "g(Suc n) = nxtg (g n) n" for n + by (auto simp: g_def) + have good: "good (g n) n" for n + proof (induction n) + case 0 + with c show ?case + by (auto simp: good_def) + qed (simp add: g_Suc nxtg) + have *: "\n x. x \ topspace X \ \g(Suc n) x - g n x\ \ c * (2/3) ^ n / 3" + using nxtg g_Suc good by force + obtain h where conth: "continuous_map X euclideanreal h" + and h: "\\. 0 < \ \ \\<^sub>F n in sequentially. \x\topspace X. dist (g n x) (h x) < \" + proof (rule Met_TC.continuous_map_uniformly_Cauchy_limit) + show "\\<^sub>F n in sequentially. continuous_map X (Met_TC.mtopology) (g n)" + using good good_def by fastforce + show "\N. \m n x. N \ m \ N \ n \ x \ topspace X \ dist (g m x) (g n x) < \" + if "\ > 0" for \ + proof - + have "\\<^sub>F n in sequentially. \(2/3) ^ n\ < \/c" + proof (rule Archimedean_eventually_pow_inverse) + show "0 < \ / c" + by (simp add: \0 < c\ that) + qed auto + then obtain N where N: "\n. n \ N \ \(2/3) ^ n\ < \/c" + by (meson eventually_sequentially order_le_less_trans) + have "\g m x - g n x\ < \" + if "N \ m" "N \ n" and x: "x \ topspace X" "m \ n" for m n x + proof (cases "m < n") + case True + have 23: "(\k = m..m \ n\ + by (induction n) (auto simp: le_Suc_eq) + have "\g m x - g n x\ \ \\k = m.." + by (subst sum_Suc_diff' [OF \m \ n\]) linarith + also have "\ \ (\k = m..g (Suc k) x - g k x\)" + by (rule sum_abs) + also have "\ \ (\k = m.. = (c/3) * (\k = m.. = (c/3) * 3 * ((2/3) ^ m - (2/3) ^ n)" + by (simp add: sum_distrib_left 23) + also have "... < (c/3) * 3 * ((2/3) ^ m)" + using \0 < c\ by auto + also have "\ < \" + using N [OF \N \ m\] \0 < c\ by (simp add: field_simps) + finally show ?thesis . + qed (use \0 < \\ \m \ n\ in auto) + then show ?thesis + by (metis dist_commute_lessI dist_real_def nle_le) + qed + qed auto + define \ where "\ \ \x. max a (min (h x) b)" + show thesis + proof + show "continuous_map X euclidean \" + unfolding \_def using conth by (intro continuous_intros) auto + show "\ x = f x" if "x \ S" for x + proof - + have x: "x \ topspace X" + using \closedin X S\ closedin_subset that by blast + have "h x = f x" + proof (rule Met_TC.limitin_metric_unique) + show "limitin Met_TC.mtopology (\n. g n x) (h x) sequentially" + using h x by (force simp: tendsto_iff eventually_sequentially) + show "limitin Met_TC.mtopology (\n. g n x) (f x) sequentially" + proof (clarsimp simp: tendsto_iff) + fix \::real + assume "\ > 0" + then have "\\<^sub>F n in sequentially. \(2/3) ^ n\ < \/c" + by (intro Archimedean_eventually_pow_inverse) (auto simp: \c > 0\) + then show "\\<^sub>F n in sequentially. dist (g n x) (f x) < \" + apply eventually_elim + by (smt (verit) good x good_def \c > 0\ dist_real_def mult.commute pos_less_divide_eq that) + qed + qed auto + then show ?thesis + using that fim by (auto simp: \_def) + qed + then show "\ ` topspace X \ {a..b}" + using fim \a \ b\ by (auto simp: \_def) + qed +qed + + +lemma Tietze_extension_realinterval: + assumes XS: "normal_space X" "closedin X S" and T: "is_interval T" "T \ {}" + and contf: "continuous_map (subtopology X S) euclideanreal f" + and "f ` S \ T" + obtains g where "continuous_map X euclideanreal g" "g ` topspace X \ T" "\x. x \ S \ g x = f x" +proof - + define \ where + "\ \ \T::real set. \f. continuous_map (subtopology X S) euclidean f \ f`S \ T + \ (\g. continuous_map X euclidean g \ g ` topspace X \ T \ (\x \ S. g x = f x))" + have "\ T" + if *: "\T. \bounded T; is_interval T; T \ {}\ \ \ T" + and "is_interval T" "T \ {}" for T + unfolding \_def + proof (intro strip) + fix f + assume contf: "continuous_map (subtopology X S) euclideanreal f" + and "f ` S \ T" + have \T: "\ ((\x. x / (1 + \x\)) ` T)" + proof (rule *) + show "bounded ((\x. x / (1 + \x\)) ` T)" + using shrink_range [of T] by (force intro: boundedI [where B=1]) + show "is_interval ((\x. x / (1 + \x\)) ` T)" + using connected_shrink that(2) is_interval_connected_1 by blast + show "(\x. x / (1 + \x\)) ` T \ {}" + using \T \ {}\ by auto + qed + moreover have "continuous_map (subtopology X S) euclidean ((\x. x / (1 + \x\)) \ f)" + by (metis contf continuous_map_compose continuous_map_into_fulltopology continuous_map_real_shrink) + moreover have "((\x. x / (1 + \x\)) \ f) ` S \ (\x. x / (1 + \x\)) ` T" + using \f ` S \ T\ by auto + ultimately obtain g + where contg: "continuous_map X euclidean g" + and gim: "g ` topspace X \ (\x. x / (1 + \x\)) ` T" + and geq: "\x. x \ S \ g x = ((\x. x / (1 + \x\)) \ f) x" + using \T unfolding \_def by force + show "\g. continuous_map X euclideanreal g \ g ` topspace X \ T \ (\x\S. g x = f x)" + proof (intro conjI exI) + have "continuous_map X (top_of_set {-1<..<1}) g" + using contg continuous_map_in_subtopology gim shrink_range by blast + then show "continuous_map X euclideanreal ((\x. x / (1 - \x\)) \ g)" + by (rule continuous_map_compose) (auto simp: continuous_on_real_grow) + show "((\x. x / (1 - \x\)) \ g) ` topspace X \ T" + using gim real_grow_shrink by fastforce + show "\x\S. ((\x. x / (1 - \x\)) \ g) x = f x" + using geq real_grow_shrink by force + qed + qed + moreover have "\ T" + if "bounded T" "is_interval T" "T \ {}" for T + unfolding \_def + proof (intro strip) + fix f + assume contf: "continuous_map (subtopology X S) euclideanreal f" + and "f ` S \ T" + obtain a b where ab: "closure T = {a..b}" + by (meson \bounded T\ \is_interval T\ compact_closure connected_compact_interval_1 + connected_imp_connected_closure is_interval_connected) + with \T \ {}\ have "a \ b" by auto + have "f ` S \ {a..b}" + using \f ` S \ T\ ab closure_subset by auto + then obtain g where contg: "continuous_map X euclideanreal g" + and gf: "\x. x \ S \ g x = f x" and gim: "g ` topspace X \ {a..b}" + using Tietze_extension_closed_real_interval [OF XS contf _ \a \ b\] by metis + define W where "W \ {x \ topspace X. g x \ closure T - T}" + have "{a..b} - {a, b} \ T" + using that + by (metis ab atLeastAtMost_diff_ends convex_interior_closure interior_atLeastAtMost_real + interior_subset is_interval_convex) + with finite_imp_compact have "compact (closure T - T)" + by (metis Diff_eq_empty_iff Diff_insert2 ab finite.emptyI finite_Diff_insert) + then have "closedin X W" + unfolding W_def using closedin_continuous_map_preimage [OF contg] compact_imp_closed by force + moreover have "disjnt W S" + unfolding W_def disjnt_iff using \f ` S \ T\ gf by blast + ultimately obtain h :: "'a \ real" + where conth: "continuous_map X (top_of_set {0..1}) h" + and him: "h ` W \ {0}" "h ` S \ {1}" + by (metis XS normal_space_iff_Urysohn) + then have him01: "h ` topspace X \ {0..1}" + by (meson continuous_map_in_subtopology) + obtain z where "z \ T" + using \T \ {}\ by blast + define g' where "g' \ \x. z + h x * (g x - z)" + show "\g. continuous_map X euclidean g \ g ` topspace X \ T \ (\x\S. g x = f x)" + proof (intro exI conjI) + show "continuous_map X euclideanreal g'" + unfolding g'_def using contg conth continuous_map_in_subtopology + by (intro continuous_intros) auto + show "g' ` topspace X \ T" + unfolding g'_def + proof clarify + fix x + assume "x \ topspace X" + show "z + h x * (g x - z) \ T" + proof (cases "g x \ T") + case True + define w where "w \ z + h x * (g x - z)" + have "\h x\ * \g x - z\ \ \g x - z\" "\1 - h x\ * \g x - z\ \ \g x - z\" + using him01 \x \ topspace X\ by (force simp: intro: mult_left_le_one_le)+ + then consider "z \ w \ w \ g x" | "g x \ w \ w \ z" + unfolding w_def by (smt (verit) left_diff_distrib mult_cancel_right2 mult_minus_right zero_less_mult_iff) + then show ?thesis + using \is_interval T\ unfolding w_def is_interval_1 by (metis True \z \ T\) + next + case False + then have "g x \ closure T" + using \x \ topspace X\ ab gim by blast + then have "h x = 0" + using him False \x \ topspace X\ by (auto simp: W_def image_subset_iff) + then show ?thesis + by (simp add: \z \ T\) + qed + qed + show "\x\S. g' x = f x" + using gf him by (auto simp: W_def g'_def) + qed + qed + ultimately show thesis + using assms that unfolding \_def by best +qed + +lemma normal_space_iff_Tietze: + "normal_space X \ + (\f S. closedin X S \ + continuous_map (subtopology X S) euclidean f + \ (\g:: 'a \ real. continuous_map X euclidean g \ (\x \ S. g x = f x)))" + (is "?lhs \ ?rhs") +proof + assume ?lhs + then show ?rhs + by (metis Tietze_extension_realinterval empty_not_UNIV is_interval_univ subset_UNIV) +next + assume R: ?rhs + show ?lhs + unfolding normal_space_iff_Urysohn_alt + proof clarify + fix S T + assume "closedin X S" + and "closedin X T" + and "disjnt S T" + then have cloST: "closedin X (S \ T)" + by (simp add: closedin_Un) + moreover + have "continuous_map (subtopology X (S \ T)) euclideanreal (\x. if x \ S then 0 else 1)" + unfolding continuous_map_closedin + proof (intro conjI strip) + fix C :: "real set" + define D where "D \ {x \ topspace X. if x \ S then 0 \ C else x \ T \ 1 \ C}" + have "D \ {{}, S, T, S \ T}" + unfolding D_def + using closedin_subset [OF \closedin X S\] closedin_subset [OF \closedin X T\] \disjnt S T\ + by (auto simp: disjnt_iff) + then have "closedin X D" + using \closedin X S\ \closedin X T\ closedin_empty by blast + with closedin_subset_topspace + show "closedin (subtopology X (S \ T)) {x \ topspace (subtopology X (S \ T)). (if x \ S then 0::real else 1) \ C}" + apply (simp add: D_def) + by (smt (verit, best) Collect_cong Collect_mono_iff Un_def closedin_subset_topspace) + qed auto + ultimately obtain g :: "'a \ real" where + contg: "continuous_map X euclidean g" and gf: "\x \ S \ T. g x = (if x \ S then 0 else 1)" + using R by blast + then show "\f. continuous_map X euclideanreal f \ f ` S \ {0} \ f ` T \ {1}" + by (smt (verit) Un_iff \disjnt S T\ disjnt_iff image_subset_iff insert_iff) + qed +qed + +subsection \random metric space stuff\ + + +lemma metrizable_imp_k_space: + assumes "metrizable_space X" + shows "k_space X" +proof - + obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d" + using assms unfolding metrizable_space_def by metis + then interpret Metric_space M d + by blast + show ?thesis + unfolding k_space Xeq + proof clarsimp + fix S + assume "S \ M" and S: "\K. compactin mtopology K \ closedin (subtopology mtopology K) (K \ S)" + have "l \ S" + if \: "range \ \ S" and l: "limitin mtopology \ l sequentially" for \ l + proof - + define K where "K \ insert l (range \)" + interpret Submetric M d K + proof + show "K \ M" + unfolding K_def using l limitin_mspace \S \ M\ \ by blast + qed + have "compactin mtopology K" + unfolding K_def using \S \ M\ \ + by (force intro: compactin_sequence_with_limit [OF l]) + then have *: "closedin sub.mtopology (K \ S)" + by (simp add: S mtopology_submetric) + have "\ n \ K \ S" for n + by (simp add: K_def range_subsetD \) + moreover have "limitin sub.mtopology \ l sequentially" + using l + unfolding sub.limit_metric_sequentially limit_metric_sequentially + by (force simp: K_def) + ultimately have "l \ K \ S" + by (meson * \ image_subsetI sub.metric_closedin_iff_sequentially_closed) + then show ?thesis + by simp + qed + then show "closedin mtopology S" + unfolding metric_closedin_iff_sequentially_closed + using \S \ M\ by blast + qed +qed + +lemma (in Metric_space) k_space_mtopology: "k_space mtopology" + by (simp add: metrizable_imp_k_space metrizable_space_mtopology) + +lemma k_space_euclideanreal: "k_space (euclidean :: 'a::metric_space topology)" + using metrizable_imp_k_space metrizable_space_euclidean by auto + + +subsection\Hereditarily normal spaces\ + +lemma hereditarily_B: + assumes "\S T. separatedin X S T + \ \U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" + shows "hereditarily normal_space X" + unfolding hereditarily_def +proof (intro strip) + fix W + assume "W \ topspace X" + show "normal_space (subtopology X W)" + unfolding normal_space_def + proof clarify + fix S T + assume clo: "closedin (subtopology X W) S" "closedin (subtopology X W) T" + and "disjnt S T" + then have "separatedin (subtopology X W) S T" + by (simp add: separatedin_closed_sets) + then obtain U V where "openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" + using assms [of S T] by (meson separatedin_subtopology) + then show "\U V. openin (subtopology X W) U \ openin (subtopology X W) V \ S \ U \ T \ V \ disjnt U V" + apply (simp add: openin_subtopology_alt) + by (meson clo closedin_imp_subset disjnt_subset1 disjnt_subset2 inf_le2) + qed +qed + +lemma hereditarily_C: + assumes "separatedin X S T" and norm: "\U. openin X U \ normal_space (subtopology X U)" + shows "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" +proof - + define ST where "ST \ X closure_of S \ X closure_of T" + have subX: "S \ topspace X" "T \ topspace X" + by (meson \separatedin X S T\ separation_closedin_Un_gen)+ + have sub: "S \ topspace X - ST" "T \ topspace X - ST" + unfolding ST_def + by (metis Diff_mono Diff_triv \separatedin X S T\ Int_lower1 Int_lower2 separatedin_def)+ + have "normal_space (subtopology X (topspace X - ST))" + by (simp add: ST_def norm closedin_Int openin_diff) + moreover have " disjnt (subtopology X (topspace X - ST) closure_of S) + (subtopology X (topspace X - ST) closure_of T)" + using Int_absorb1 ST_def sub by (fastforce simp: disjnt_iff closure_of_subtopology) + ultimately show ?thesis + using sub subX + apply (simp add: normal_space_closures) + by (metis ST_def closedin_Int closedin_closure_of closedin_def openin_trans_full) +qed + +lemma hereditarily_normal_space: + "hereditarily normal_space X \ (\U. openin X U \ normal_space(subtopology X U))" + by (metis hereditarily_B hereditarily_C hereditarily) + +lemma hereditarily_normal_separation: + "hereditarily normal_space X \ + (\S T. separatedin X S T + \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" + by (metis hereditarily_B hereditarily_C hereditarily) + + +lemma metrizable_imp_hereditarily_normal_space: + "metrizable_space X \ hereditarily normal_space X" + by (simp add: hereditarily metrizable_imp_normal_space metrizable_space_subtopology) + +lemma metrizable_space_separation: + "\metrizable_space X; separatedin X S T\ + \ \U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" + by (metis hereditarily hereditarily_C metrizable_imp_hereditarily_normal_space) + +lemma hereditarily_normal_separation_pairwise: + "hereditarily normal_space X \ + (\\. finite \ \ (\S \ \. S \ topspace X) \ pairwise (separatedin X) \ + \ (\\. (\S \ \. openin X (\ S) \ S \ \ S) \ + pairwise (\S T. disjnt (\ S) (\ T)) \))" + (is "?lhs \ ?rhs") +proof + assume L: ?lhs + show ?rhs + proof clarify + fix \ + assume "finite \" and \: "\S\\. S \ topspace X" + and pw\: "pairwise (separatedin X) \" + have "\V W. openin X V \ openin X W \ S \ V \ + (\T. T \ \ \ T \ S \ T \ W) \ disjnt V W" + if "S \ \" for S + proof - + have "separatedin X S (\(\ - {S}))" + by (metis \ \finite \\ pw\ finite_Diff pairwise_alt separatedin_Union(1) that) + with L show ?thesis + unfolding hereditarily_normal_separation + by (smt (verit) Diff_iff UnionI empty_iff insert_iff subset_iff) + qed + then obtain \ \ + where *: "\S. S \ \ \ S \ \ S \ (\T. T \ \ \ T \ S \ T \ \ S)" + and ope: "\S. S \ \ \ openin X (\ S) \ openin X (\ S)" + and dis: "\S. S \ \ \ disjnt (\ S) (\ S)" + by metis + define \ where "\ \ \S. \ S \ (\T \ \ - {S}. \ T)" + show "\\. (\S\\. openin X (\ S) \ S \ \ S) \ pairwise (\S T. disjnt (\ S) (\ T)) \" + proof (intro exI conjI strip) + show "openin X (\ S)" if "S \ \" for S + unfolding \_def + by (smt (verit) ope that DiffD1 \finite \\ finite_Diff finite_imageI imageE openin_Int_Inter) + show "S \ \ S" if "S \ \" for S + unfolding \_def using "*" that by auto + show "pairwise (\S T. disjnt (\ S) (\ T)) \" + using dis by (fastforce simp: disjnt_iff pairwise_alt \_def) + qed + qed +next + assume R: ?rhs + show ?lhs + unfolding hereditarily_normal_separation + proof (intro strip) + fix S T + assume "separatedin X S T" + show "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" + proof (cases "T=S") + case True + then show ?thesis + using \separatedin X S T\ by force + next + case False + have "pairwise (separatedin X) {S, T}" + by (simp add: \separatedin X S T\ pairwise_insert separatedin_sym) + moreover have "\S\{S, T}. S \ topspace X" + by (metis \separatedin X S T\ insertE separatedin_def singletonD) + ultimately show ?thesis + using R by (smt (verit) False finite.emptyI finite.insertI insertCI pairwiseD) + qed + qed +qed + +lemma hereditarily_normal_space_perfect_map_image: + "\hereditarily normal_space X; perfect_map X Y f\ \ hereditarily normal_space Y" + unfolding perfect_map_def proper_map_def + by (meson hereditarily_normal_space_continuous_closed_map_image) + +lemma regular_second_countable_imp_hereditarily_normal_space: + assumes "regular_space X \ second_countable X" + shows "hereditarily normal_space X" + unfolding hereditarily + proof (intro regular_Lindelof_imp_normal_space strip) + show "regular_space (subtopology X S)" for S + by (simp add: assms regular_space_subtopology) + show "Lindelof_space (subtopology X S)" for S + using assms by (simp add: second_countable_imp_Lindelof_space second_countable_subtopology) +qed + + +subsection\Completely regular spaces\ + +definition completely_regular_space where + "completely_regular_space X \ + \S x. closedin X S \ x \ topspace X - S + \ (\f::'a\real. continuous_map X (top_of_set {0..1}) f \ + f x = 0 \ (f ` S \ {1}))" + +lemma homeomorphic_completely_regular_space_aux: + assumes X: "completely_regular_space X" and hom: "X homeomorphic_space Y" + shows "completely_regular_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)" + using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce + show ?thesis + unfolding completely_regular_space_def + proof clarify + fix S x + assume A: "closedin Y S" and x: "x \ topspace Y" and "x \ S" + then have "closedin X (g`S)" + using hmg homeomorphic_map_closedness_eq by blast + moreover have "g x \ g`S" + by (meson A x \x \ S\ closedin_subset hmg homeomorphic_imp_injective_map inj_on_image_mem_iff) + ultimately obtain \ where \: "continuous_map X (top_of_set {0..1::real}) \ \ \ (g x) = 0 \ \ ` g`S \ {1}" + by (metis DiffI X completely_regular_space_def hmg homeomorphic_imp_surjective_map image_eqI x) + then have "continuous_map Y (top_of_set {0..1::real}) (\ \ g)" + by (meson continuous_map_compose hmg homeomorphic_imp_continuous_map) + then show "\\. continuous_map Y (top_of_set {0..1::real}) \ \ \ x = 0 \ \ ` S \ {1}" + by (metis \ comp_apply image_comp) + qed +qed + +lemma homeomorphic_completely_regular_space: + assumes "X homeomorphic_space Y" + shows "completely_regular_space X \ completely_regular_space Y" + by (meson assms homeomorphic_completely_regular_space_aux homeomorphic_space_sym) + +lemma completely_regular_space_alt: + "completely_regular_space X \ + (\S x. closedin X S \ x \ topspace X - S + \ (\f. continuous_map X euclideanreal f \ f x = 0 \ f ` S \ {1}))" +proof - + have "\f. continuous_map X (top_of_set {0..1::real}) f \ f x = 0 \ f ` S \ {1}" + if "closedin X S" "x \ topspace X - S" and f: "continuous_map X euclideanreal f \ f x = 0 \ f ` S \ {1}" + for S x f + proof (intro exI conjI) + show "continuous_map X (top_of_set {0..1}) (\x. max 0 (min (f x) 1))" + using that + by (auto simp: continuous_map_in_subtopology intro!: continuous_map_real_max continuous_map_real_min) + qed (use that in auto) + with continuous_map_in_subtopology show ?thesis + unfolding completely_regular_space_def by metis +qed + +text \As above, but with @{term openin}\ +lemma completely_regular_space_alt': + "completely_regular_space X \ + (\S x. openin X S \ x \ S + \ (\f. continuous_map X euclideanreal f \ f x = 0 \ f ` (topspace X - S) \ {1}))" + apply (simp add: completely_regular_space_alt all_closedin) + by (meson openin_subset subsetD) + +lemma completely_regular_space_gen_alt: + fixes a b::real + assumes "a \ b" + shows "completely_regular_space X \ + (\S x. closedin X S \ x \ topspace X - S + \ (\f. continuous_map X euclidean f \ f x = a \ (f ` S \ {b})))" +proof - + have "\f. continuous_map X euclideanreal f \ f x = 0 \ f ` S \ {1}" + if "closedin X S" "x \ topspace X - S" + and f: "continuous_map X euclidean f \ f x = a \ f ` S \ {b}" + for S x f + proof (intro exI conjI) + show "continuous_map X euclideanreal ((\x. inverse(b - a) * (x - a)) \ f)" + using that by (intro continuous_intros) auto + qed (use that assms in auto) + moreover + have "\f. continuous_map X euclidean f \ f x = a \ f ` S \ {b}" + if "closedin X S" "x \ topspace X - S" + and f: "continuous_map X euclideanreal f \ f x = 0 \ f ` S \ {1}" + for S x f + proof (intro exI conjI) + show "continuous_map X euclideanreal ((\x. a + (b - a) * x) \ f)" + using that by (intro continuous_intros) auto + qed (use that in auto) + ultimately show ?thesis + unfolding completely_regular_space_alt by meson +qed + +text \As above, but with @{term openin}\ +lemma completely_regular_space_gen_alt': + fixes a b::real + assumes "a \ b" + shows "completely_regular_space X \ + (\S x. openin X S \ x \ S + \ (\f. continuous_map X euclidean f \ f x = a \ (f ` (topspace X - S) \ {b})))" + apply (simp add: completely_regular_space_gen_alt[OF assms] all_closedin) + by (meson openin_subset subsetD) + +lemma completely_regular_space_gen: + fixes a b::real + assumes "a < b" + shows "completely_regular_space X \ + (\S x. closedin X S \ x \ topspace X - S + \ (\f. continuous_map X (top_of_set {a..b}) f \ f x = a \ f ` S \ {b}))" +proof - + have "\f. continuous_map X (top_of_set {a..b}) f \ f x = a \ f ` S \ {b}" + if "closedin X S" "x \ topspace X - S" + and f: "continuous_map X euclidean f \ f x = a \ f ` S \ {b}" + for S x f + proof (intro exI conjI) + show "continuous_map X (top_of_set {a..b}) (\x. max a (min (f x) b))" + using that assms + by (auto simp: continuous_map_in_subtopology intro!: continuous_map_real_max continuous_map_real_min) + qed (use that assms in auto) + with continuous_map_in_subtopology assms show ?thesis + using completely_regular_space_gen_alt [of a b] + by (smt (verit) atLeastAtMost_singleton atLeastatMost_empty singletonI) +qed + +lemma normal_imp_completely_regular_space_A: + assumes "normal_space X" "t1_space X" + shows "completely_regular_space X" + unfolding completely_regular_space_alt +proof clarify + fix x S + assume A: "closedin X S" "x \ S" + assume "x \ topspace X" + then have "closedin X {x}" + by (simp add: \t1_space X\ closedin_t1_singleton) + with A \normal_space X\ have "\f. continuous_map X euclideanreal f \ f ` {x} \ {0} \ f ` S \ {1}" + using assms unfolding normal_space_iff_Urysohn_alt disjnt_iff by blast + then show "\f. continuous_map X euclideanreal f \ f x = 0 \ f ` S \ {1}" + by auto +qed + +lemma normal_imp_completely_regular_space_B: + assumes "normal_space X" "regular_space X" + shows "completely_regular_space X" + unfolding completely_regular_space_alt +proof clarify + fix x S + assume "closedin X S" "x \ S" "x \ topspace X" + then obtain U C where "openin X U" "closedin X C" "x \ U" "U \ C" "C \ topspace X - S" + using assms + unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of closedin_def by (metis Diff_iff) + then obtain f where "continuous_map X euclideanreal f \ f ` C \ {0} \ f ` S \ {1}" + using assms unfolding normal_space_iff_Urysohn_alt + by (metis Diff_iff \closedin X S\ disjnt_iff subsetD) + then show "\f. continuous_map X euclideanreal f \ f x = 0 \ f ` S \ {1}" + by (meson \U \ C\ \x \ U\ image_subset_iff singletonD subsetD) +qed + +lemma normal_imp_completely_regular_space_gen: + "\normal_space X; t1_space X \ Hausdorff_space X \ regular_space X\ \ completely_regular_space X" + using normal_imp_completely_regular_space_A normal_imp_completely_regular_space_B t1_or_Hausdorff_space by blast + +lemma normal_imp_completely_regular_space: + "\normal_space X; Hausdorff_space X \ regular_space X\ \ completely_regular_space X" + by (simp add: normal_imp_completely_regular_space_gen) + +lemma (in Metric_space) completely_regular_space_mtopology: + "completely_regular_space mtopology" + by (simp add: normal_imp_completely_regular_space normal_space_mtopology regular_space_mtopology) + +lemma metrizable_imp_completely_regular_space: + "metrizable_space X \ completely_regular_space X" + by (simp add: metrizable_imp_normal_space metrizable_imp_regular_space normal_imp_completely_regular_space) + +lemma completely_regular_space_discrete_topology: + "completely_regular_space(discrete_topology U)" + by (simp add: normal_imp_completely_regular_space normal_space_discrete_topology) + +lemma completely_regular_space_subtopology: + assumes "completely_regular_space X" + shows "completely_regular_space (subtopology X S)" + unfolding completely_regular_space_def +proof clarify + fix A x + assume "closedin (subtopology X S) A" and x: "x \ topspace (subtopology X S)" and "x \ A" + then obtain T where "closedin X T" "A = S \ T" "x \ topspace X" "x \ S" + by (force simp: closedin_subtopology_alt image_iff) + then show "\f. continuous_map (subtopology X S) (top_of_set {0::real..1}) f \ f x = 0 \ f ` A \ {1}" + using assms \x \ A\ + apply (simp add: completely_regular_space_def continuous_map_from_subtopology) + using continuous_map_from_subtopology by fastforce +qed + +lemma completely_regular_space_retraction_map_image: + " \retraction_map X Y r; completely_regular_space X\ \ completely_regular_space Y" + using completely_regular_space_subtopology hereditary_imp_retractive_property homeomorphic_completely_regular_space by blast + +lemma completely_regular_imp_regular_space: + assumes "completely_regular_space X" + shows "regular_space X" +proof - + have *: "\U V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V" + if contf: "continuous_map X euclideanreal f" and a: "a \ topspace X - C" and "closedin X C" + and fim: "f ` topspace X \ {0..1}" and f0: "f a = 0" and f1: "f ` C \ {1}" + for C a f + proof (intro exI conjI) + show "openin X {x \ topspace X. f x \ {..<1 / 2}}" "openin X {x \ topspace X. f x \ {1 / 2<..}}" + using openin_continuous_map_preimage [OF contf] + by (meson open_lessThan open_greaterThan open_openin)+ + show "a \ {x \ topspace X. f x \ {..<1 / 2}}" + using a f0 by auto + show "C \ {x \ topspace X. f x \ {1 / 2<..}}" + using \closedin X C\ f1 closedin_subset by auto + qed (auto simp: disjnt_iff) + show ?thesis + using assms + unfolding completely_regular_space_def regular_space_def continuous_map_in_subtopology + by (meson "*") +qed + + +lemma locally_compact_regular_imp_completely_regular_space: + assumes "locally_compact_space X" "Hausdorff_space X \ regular_space X" + shows "completely_regular_space X" + unfolding completely_regular_space_def +proof clarify + fix S x + assume "closedin X S" and "x \ topspace X" and "x \ S" + have "regular_space X" + using assms locally_compact_Hausdorff_imp_regular_space by blast + then have nbase: "neighbourhood_base_of (\C. compactin X C \ closedin X C) X" + using assms(1) locally_compact_regular_space_neighbourhood_base by blast + then obtain U M where "openin X U" "compactin X M" "closedin X M" "x \ U" "U \ M" "M \ topspace X - S" + unfolding neighbourhood_base_of by (metis (no_types, lifting) Diff_iff \closedin X S\ \x \ topspace X\ \x \ S\ closedin_def) + then have "M \ topspace X" + by blast + obtain V K where "openin X V" "closedin X K" "x \ V" "V \ K" "K \ U" + by (metis (no_types, lifting) \openin X U\ \x \ U\ neighbourhood_base_of nbase) + have "compact_space (subtopology X M)" + by (simp add: \compactin X M\ compact_space_subtopology) + then have "normal_space (subtopology X M)" + by (simp add: \regular_space X\ compact_Hausdorff_or_regular_imp_normal_space regular_space_subtopology) + moreover have "closedin (subtopology X M) K" + using \K \ U\ \U \ M\ \closedin X K\ closedin_subset_topspace by fastforce + moreover have "closedin (subtopology X M) (M - U)" + by (simp add: \closedin X M\ \openin X U\ closedin_diff closedin_subset_topspace) + moreover have "disjnt K (M - U)" + by (meson DiffD2 \K \ U\ disjnt_iff subsetD) + ultimately obtain f::"'a\real" where contf: "continuous_map (subtopology X M) (top_of_set {0..1}) f" + and f0: "f ` K \ {0}" and f1: "f ` (M - U) \ {1}" + using Urysohn_lemma [of "subtopology X M" K "M-U" 0 1] by auto + then obtain g::"'a\real" where contg: "continuous_map (subtopology X M) euclidean g" and gim: "g ` M \ {0..1}" + and g0: "\x. x \ K \ g x = 0" and g1: "\x. \x \ M; x \ U\ \ g x = 1" + using \M \ topspace X\ by (force simp add: continuous_map_in_subtopology image_subset_iff) + show "\f::'a\real. continuous_map X (top_of_set {0..1}) f \ f x = 0 \ f ` S \ {1}" + proof (intro exI conjI) + show "continuous_map X (top_of_set {0..1}) (\x. if x \ M then g x else 1)" + unfolding continuous_map_closedin + proof (intro strip conjI) + fix C + assume C: "closedin (top_of_set {0::real..1}) C" + have eq: "{x \ topspace X. (if x \ M then g x else 1) \ C} = {x \ M. g x \ C} \ (if 1 \ C then topspace X - U else {})" + using \U \ M\ \M \ topspace X\ g1 by auto + show "closedin X {x \ topspace X. (if x \ M then g x else 1) \ C}" + unfolding eq + proof (intro closedin_Un) + have "closedin euclidean C" + using C closed_closedin closedin_closed_trans by blast + then have "closedin (subtopology X M) {x \ M. g x \ C}" + using closedin_continuous_map_preimage_gen [OF contg] \M \ topspace X\ by auto + then show "closedin X {x \ M. g x \ C}" + using \closedin X M\ closedin_trans_full by blast + qed (use \openin X U\ in force) + qed (use gim in force) + show "(if x \ M then g x else 1) = 0" + using \U \ M\ \V \ K\ g0 \x \ U\ \x \ V\ by auto + show "(\x. if x \ M then g x else 1) ` S \ {1}" + using \M \ topspace X - S\ by auto + qed +qed + +lemma completely_regular_eq_regular_space: + "locally_compact_space X + \ (completely_regular_space X \ regular_space X)" + using completely_regular_imp_regular_space locally_compact_regular_imp_completely_regular_space + by blast + +lemma completely_regular_space_prod_topology: + "completely_regular_space (prod_topology X Y) \ + topspace (prod_topology X Y) = {} \ + completely_regular_space X \ completely_regular_space Y" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (rule topological_property_of_prod_component) + (auto simp: completely_regular_space_subtopology homeomorphic_completely_regular_space) +next + assume R: ?rhs + show ?lhs + proof (cases "topspace(prod_topology X Y) = {}") + case False + then have X: "completely_regular_space X" and Y: "completely_regular_space Y" + using R by blast+ + show ?thesis + unfolding completely_regular_space_alt' + proof clarify + fix W x y + assume "openin (prod_topology X Y) W" and "(x, y) \ W" + then obtain U V where "openin X U" "openin Y V" "x \ U" "y \ V" "U\V \ W" + by (force simp: openin_prod_topology_alt) + then have "x \ topspace X" "y \ topspace Y" + using openin_subset by fastforce+ + obtain f where contf: "continuous_map X euclideanreal f" and "f x = 0" + and f1: "\x. x \ topspace X \ x \ U \ f x = 1" + using X \openin X U\ \x \ U\ unfolding completely_regular_space_alt' + by (smt (verit, best) Diff_iff image_subset_iff singletonD) + obtain g where contg: "continuous_map Y euclideanreal g" and "g y = 0" + and g1: "\y. y \ topspace Y \ y \ V \ g y = 1" + using Y \openin Y V\ \y \ V\ unfolding completely_regular_space_alt' + by (smt (verit, best) Diff_iff image_subset_iff singletonD) + define h where "h \ \(x,y). 1 - (1 - f x) * (1 - g y)" + show "\h. continuous_map (prod_topology X Y) euclideanreal h \ h (x,y) = 0 \ h ` (topspace (prod_topology X Y) - W) \ {1}" + proof (intro exI conjI) + have "continuous_map (prod_topology X Y) euclideanreal (f \ fst)" + using contf continuous_map_of_fst by blast + moreover + have "continuous_map (prod_topology X Y) euclideanreal (g \ snd)" + using contg continuous_map_of_snd by blast + ultimately + show "continuous_map (prod_topology X Y) euclideanreal h" + unfolding o_def h_def case_prod_unfold + by (intro continuous_intros) auto + show "h (x, y) = 0" + by (simp add: h_def \f x = 0\ \g y = 0\) + show "h ` (topspace (prod_topology X Y) - W) \ {1}" + using \U \ V \ W\ f1 g1 by (force simp: h_def) + qed + qed + qed (force simp: completely_regular_space_def) +qed + + +lemma completely_regular_space_product_topology: + "completely_regular_space (product_topology X I) \ + (\\<^sub>E i\I. topspace(X i)) = {} \ (\i \ I. completely_regular_space (X i))" + (is "?lhs \ ?rhs") +proof + assume ?lhs then show ?rhs + by (rule topological_property_of_product_component) + (auto simp: completely_regular_space_subtopology homeomorphic_completely_regular_space) +next + assume R: ?rhs + show ?lhs + proof (cases "(\\<^sub>E i\I. topspace(X i)) = {}") + case False + show ?thesis + unfolding completely_regular_space_alt' + proof clarify + fix W x + assume W: "openin (product_topology X I) W" and "x \ W" + then obtain U where finU: "finite {i \ I. U i \ topspace (X i)}" + and ope: "\i. i\I \ openin (X i) (U i)" + and x: "x \ Pi\<^sub>E I U" and "Pi\<^sub>E I U \ W" + by (auto simp: openin_product_topology_alt) + have "\i \ I. openin (X i) (U i) \ x i \ U i + \ (\f. continuous_map (X i) euclideanreal f \ + f (x i) = 0 \ (\x \ topspace(X i). x \ U i \ f x = 1))" + using R unfolding completely_regular_space_alt' + by (smt (verit) DiffI False image_subset_iff singletonD) + with ope x have "\i. \f. i \ I \ continuous_map (X i) euclideanreal f \ + f (x i) = 0 \ (\x \ topspace (X i). x \ U i \ f x = 1)" + by auto + then obtain f where f: "\i. i \ I \ continuous_map (X i) euclideanreal (f i) \ + f i (x i) = 0 \ (\x \ topspace (X i). x \ U i \ f i x = 1)" + by metis + define h where "h \ \z. 1 - prod (\i. 1 - f i (z i)) {i\I. U i \ topspace(X i)}" + show "\h. continuous_map (product_topology X I) euclideanreal h \ h x = 0 \ + h ` (topspace (product_topology X I) - W) \ {1}" + proof (intro conjI exI) + have "continuous_map (product_topology X I) euclidean (f i \ (\x. x i))" if "i\I" for i + using f that + by (blast intro: continuous_intros continuous_map_product_projection) + then show "continuous_map (product_topology X I) euclideanreal h" + unfolding h_def o_def by (intro continuous_intros) (auto simp: finU) + show "h x = 0" + by (simp add: h_def f) + show "h ` (topspace (product_topology X I) - W) \ {1}" + proof - + have "\i. i \ I \ U i \ topspace (X i) \ f i (x' i) = 1" + if "x' \ (\\<^sub>E i\I. topspace (X i))" "x' \ W" for x' + using that \Pi\<^sub>E I U \ W\ by (smt (verit, best) PiE_iff f in_mono) + then show ?thesis + by (auto simp: f h_def finU) + qed + qed + qed + qed (force simp: completely_regular_space_def) +qed + + +lemma (in Metric_space) t1_space_mtopology: + "t1_space mtopology" + using Hausdorff_space_mtopology t1_or_Hausdorff_space by blast + + +subsection\More generally, the k-ification functor\ + +definition kification_open + where "kification_open \ + \X S. S \ topspace X \ (\K. compactin X K \ openin (subtopology X K) (K \ S))" + +definition kification + where "kification X \ topology (kification_open X)" + +lemma istopology_kification_open: "istopology (kification_open X)" + unfolding istopology_def +proof (intro conjI strip) + show "kification_open X (S \ T)" + if "kification_open X S" and "kification_open X T" for S T + using that unfolding kification_open_def + by (smt (verit, best) inf.idem inf_commute inf_left_commute le_infI2 openin_Int) + show "kification_open X (\ \)" if "\K\\. kification_open X K" for \ + using that unfolding kification_open_def Int_Union by blast +qed + +lemma openin_kification: + "openin (kification X) U \ + U \ topspace X \ + (\K. compactin X K \ openin (subtopology X K) (K \ U))" + by (metis topology_inverse' kification_def istopology_kification_open kification_open_def) + +lemma openin_kification_finer: + "openin X S \ openin (kification X) S" + by (simp add: openin_kification openin_subset openin_subtopology_Int2) + +lemma topspace_kification [simp]: + "topspace(kification X) = topspace X" + by (meson openin_kification openin_kification_finer openin_topspace subset_antisym) + +lemma closedin_kification: + "closedin (kification X) U \ + U \ topspace X \ + (\K. compactin X K \ closedin (subtopology X K) (K \ U))" +proof (cases "U \ topspace X") + case True + then show ?thesis + by (simp add: closedin_def Diff_Int_distrib inf_commute le_infI2 openin_kification) +qed (simp add: closedin_def) + +lemma closedin_kification_finer: "closedin X S \ closedin (kification X) S" + by (simp add: closedin_def openin_kification_finer) + +lemma kification_eq_self: "kification X = X \ k_space X" + unfolding fun_eq_iff openin_kification k_space_alt openin_inject [symmetric] + by (metis openin_closedin_eq) + +lemma compactin_kification [simp]: + "compactin (kification X) K \ compactin X K" (is "?lhs \ ?rhs") +proof + assume ?lhs then show ?rhs + by (simp add: compactin_contractive openin_kification_finer) +next + assume R: ?rhs + show ?lhs + unfolding compactin_def + proof (intro conjI strip) + show "K \ topspace (kification X)" + by (simp add: R compactin_subset_topspace) + fix \ + assume \: "Ball \ (openin (kification X)) \ K \ \ \" + then have *: "\U. U \ \ \ U \ topspace X \ openin (subtopology X K) (K \ U)" + by (simp add: R openin_kification) + have "K \ topspace X" "compact_space (subtopology X K)" + using R compactin_subspace by force+ + then have "\V. finite V \ V \ (\U. K \ U) ` \ \ \ V = topspace (subtopology X K)" + unfolding compact_space + by (smt (verit, del_insts) Int_Union \ * image_iff inf.order_iff inf_commute topspace_subtopology) + then show "\\. finite \ \ \ \ \ \ K \ \ \" + by (metis Int_Union \K \ topspace X\ finite_subset_image inf.orderI topspace_subtopology_subset) + qed +qed + +lemma compact_space_kification [simp]: + "compact_space(kification X) \ compact_space X" + by (simp add: compact_space_def) + +lemma kification_kification [simp]: + "kification(kification X) = kification X" + unfolding openin_inject [symmetric] +proof + fix U + show "openin (kification (kification X)) U = openin (kification X) U" + proof + show "openin (kification (kification X)) U \ openin (kification X) U" + by (metis compactin_kification inf_commute openin_kification openin_subtopology topspace_kification) + qed (simp add: openin_kification_finer) +qed + +lemma k_space_kification [iff]: "k_space(kification X)" + using kification_eq_self by fastforce + +lemma continuous_map_into_kification: + assumes "k_space X" + shows "continuous_map X (kification Y) f \ continuous_map X Y f" (is "?lhs \ ?rhs") +proof + assume ?lhs then show ?rhs + by (simp add: continuous_map_def openin_kification_finer) +next + assume R: ?rhs + have "openin X {x \ topspace X. f x \ V}" if V: "openin (kification Y) V" for V + proof - + have "openin (subtopology X K) (K \ {x \ topspace X. f x \ V})" + if "compactin X K" for K + proof - + have "compactin Y (f ` K)" + using R image_compactin that by blast + then have "openin (subtopology Y (f ` K)) (f ` K \ V)" + by (meson V openin_kification) + then obtain U where U: "openin Y U" "f`K \ V = U \ f`K" + by (meson openin_subtopology) + show ?thesis + unfolding openin_subtopology + proof (intro conjI exI) + show "openin X {x \ topspace X. f x \ U}" + using R U openin_continuous_map_preimage_gen by (simp add: continuous_map_def) + qed (use U in auto) + qed + then show ?thesis + by (metis (full_types) Collect_subset assms k_space_open) + qed + with R show ?lhs + by (simp add: continuous_map_def) +qed + +lemma continuous_map_from_kification: + "continuous_map X Y f \ continuous_map (kification X) Y f" + by (simp add: continuous_map_openin_preimage_eq openin_kification_finer) + +lemma continuous_map_kification: + "continuous_map X Y f \ continuous_map (kification X) (kification Y) f" + by (simp add: continuous_map_from_kification continuous_map_into_kification) + +lemma subtopology_kification_compact: + assumes "compactin X K" + shows "subtopology (kification X) K = subtopology X K" + unfolding openin_inject [symmetric] +proof + fix U + show "openin (subtopology (kification X) K) U = openin (subtopology X K) U" + by (metis assms inf_commute openin_kification openin_subset openin_subtopology) +qed + + +lemma subtopology_kification_finer: + assumes "openin (subtopology (kification X) S) U" + shows "openin (kification (subtopology X S)) U" + using assms + by (fastforce simp: openin_subtopology_alt image_iff openin_kification subtopology_subtopology compactin_subtopology) + +lemma proper_map_from_kification: + assumes "k_space Y" + shows "proper_map (kification X) Y f \ proper_map X Y f" (is "?lhs \ ?rhs") +proof + assume ?lhs then show ?rhs + by (simp add: closed_map_def closedin_kification_finer proper_map_alt) +next + assume R: ?rhs + have "compactin Y K \ compactin X {x \ topspace X. f x \ K}" for K + using R proper_map_alt by auto + with R show ?lhs + by (simp add: assms proper_map_into_k_space_eq subtopology_kification_compact) +qed + +lemma perfect_map_from_kification: + "\k_space Y; perfect_map X Y f\ \ perfect_map(kification X) Y f" + by (simp add: continuous_map_from_kification perfect_map_def proper_map_from_kification) + +lemma k_space_perfect_map_image_eq: + assumes "Hausdorff_space X" "perfect_map X Y f" + shows "k_space X \ k_space Y" +proof + show "k_space X \ k_space Y" + using k_space_perfect_map_image assms by blast + assume "k_space Y" + have "homeomorphic_map (kification X) X id" + unfolding homeomorphic_eq_injective_perfect_map + proof (intro conjI perfect_map_from_composition_right [where f = id]) + show "perfect_map (kification X) Y (f \ id)" + by (simp add: \k_space Y\ assms(2) perfect_map_from_kification) + show "continuous_map (kification X) X id" + by (simp add: continuous_map_from_kification) +qed (use assms perfect_map_def in auto) + then show "k_space X" + using homeomorphic_k_space homeomorphic_space by blast +qed + + +subsection\One-point compactifications and the Alexandroff extension construction\ + +lemma one_point_compactification_dense: + "\compact_space X; \ compactin X (topspace X - {a})\ \ X closure_of (topspace X - {a}) = topspace X" + unfolding closure_of_complement + by (metis Diff_empty closedin_compact_space interior_of_eq_empty openin_closedin_eq subset_singletonD) + +lemma one_point_compactification_interior: + "\compact_space X; \ compactin X (topspace X - {a})\ \ X interior_of {a} = {}" + by (simp add: interior_of_eq_empty_complement one_point_compactification_dense) + +lemma kc_space_one_point_compactification_gen: + assumes "compact_space X" + shows "kc_space X \ + openin X (topspace X - {a}) \ (\K. compactin X K \ a\K \ closedin X K) \ + k_space (subtopology X (topspace X - {a})) \ kc_space (subtopology X (topspace X - {a}))" + (is "?lhs \ ?rhs") +proof + assume L: ?lhs show ?rhs + proof (intro conjI strip) + show "openin X (topspace X - {a})" + using L kc_imp_t1_space t1_space_openin_delete_alt by auto + then show "k_space (subtopology X (topspace X - {a}))" + by (simp add: L assms k_space_open_subtopology_aux) + show "closedin X k" if "compactin X k \ a \ k" for k :: "'a set" + using L kc_space_def that by blast + show "kc_space (subtopology X (topspace X - {a}))" + by (simp add: L kc_space_subtopology) + qed +next + assume R: ?rhs + show ?lhs + unfolding kc_space_def + proof (intro strip) + fix S + assume "compactin X S" + then have "S \topspace X" + by (simp add: compactin_subset_topspace) + show "closedin X S" + proof (cases "a \ S") + case True + then have "topspace X - S = topspace X - {a} - (S - {a})" + by auto + moreover have "openin X (topspace X - {a} - (S - {a}))" + proof (rule openin_trans_full) + show "openin (subtopology X (topspace X - {a})) (topspace X - {a} - (S - {a}))" + proof + show "openin (subtopology X (topspace X - {a})) (topspace X - {a})" + using R openin_open_subtopology by blast + have "closedin (subtopology X ((topspace X - {a}) \ K)) (K \ (S - {a}))" + if "compactin X K" "K \ topspace X - {a}" for K + proof (intro closedin_subset_topspace) + show "closedin X (K \ (S - {a}))" + using that + by (metis IntD1 Int_insert_right_if0 R True \compactin X S\ closed_Int_compactin insert_Diff subset_Diff_insert) + qed (use that in auto) + moreover have "k_space (subtopology X (topspace X - {a}))" + using R by blast + moreover have "S-{a} \ topspace X \ S-{a} \ topspace X - {a}" + using \S \ topspace X\ by auto + ultimately show "closedin (subtopology X (topspace X - {a})) (S - {a})" + using \S \ topspace X\ True + by (simp add: k_space_def compactin_subtopology subtopology_subtopology) + qed + show "openin X (topspace X - {a})" + by (simp add: R) + qed + ultimately show ?thesis + by (simp add: \S \ topspace X\ closedin_def) + next + case False + then show ?thesis + by (simp add: R \compactin X S\) + qed + qed +qed + + +inductive Alexandroff_open for X where + base: "openin X U \ Alexandroff_open X (Some ` U)" +| ext: "\compactin X C; closedin X C\ \ Alexandroff_open X (insert None (Some ` (topspace X - C)))" + +hide_fact (open) base ext + +lemma Alexandroff_open_iff: "Alexandroff_open X S \ + (\U. (S = Some ` U \ openin X U) \ (S = insert None (Some ` (topspace X - U)) \ compactin X U \ closedin X U))" + by (meson Alexandroff_open.cases Alexandroff_open.ext Alexandroff_open.base) + +lemma Alexandroff_open_Un_aux: + assumes U: "openin X U" and "Alexandroff_open X T" + shows "Alexandroff_open X (Some ` U \ T)" + using \Alexandroff_open X T\ +proof (induction rule: Alexandroff_open.induct) + case (base V) + then show ?case + by (metis Alexandroff_open.base U image_Un openin_Un) +next + case (ext C) + have "U \ topspace X" + by (simp add: U openin_subset) + then have eq: "Some ` U \ insert None (Some ` (topspace X - C)) = insert None (Some ` (topspace X - (C \ (topspace X - U))))" + by force + have "closedin X (C \ (topspace X - U))" + using U ext.hyps(2) by blast + moreover + have "compactin X (C \ (topspace X - U))" + using U compact_Int_closedin ext.hyps(1) by blast + ultimately show ?case + unfolding eq using Alexandroff_open.ext by blast +qed + +lemma Alexandroff_open_Un: + assumes "Alexandroff_open X S" and "Alexandroff_open X T" + shows "Alexandroff_open X (S \ T)" + using assms +proof (induction rule: Alexandroff_open.induct) + case (base U) + then show ?case + by (simp add: Alexandroff_open_Un_aux) +next + case (ext C) + then show ?case + by (smt (verit, best) Alexandroff_open_Un_aux Alexandroff_open_iff Un_commute Un_insert_left closedin_def insert_absorb2) +qed + +lemma Alexandroff_open_Int_aux: + assumes U: "openin X U" and "Alexandroff_open X T" + shows "Alexandroff_open X (Some ` U \ T)" + using \Alexandroff_open X T\ +proof (induction rule: Alexandroff_open.induct) + case (base V) + then show ?case + by (metis Alexandroff_open.base U image_Int inj_Some openin_Int) +next + case (ext C) + have eq: "Some ` U \ insert None (Some ` (topspace X - C)) = Some ` (topspace X - (C \ (topspace X - U)))" + by force + have "openin X (topspace X - (C \ (topspace X - U)))" + using U ext.hyps(2) by blast + then show ?case + unfolding eq using Alexandroff_open.base by blast +qed + +lemma istopology_Alexandroff_open: "istopology (Alexandroff_open X)" + unfolding istopology_def +proof (intro conjI strip) + fix S T + assume "Alexandroff_open X S" and "Alexandroff_open X T" + then show "Alexandroff_open X (S \ T)" + proof (induction rule: Alexandroff_open.induct) + case (base U) + then show ?case + using Alexandroff_open_Int_aux by blast + next + case EC: (ext C) + show ?case + using \Alexandroff_open X T\ + proof (induction rule: Alexandroff_open.induct) + case (base V) + then show ?case + by (metis Alexandroff_open.ext Alexandroff_open_Int_aux EC.hyps inf_commute) + next + case (ext D) + have eq: "insert None (Some ` (topspace X - C)) \ insert None (Some ` (topspace X - D)) + = insert None (Some ` (topspace X - (C \ D)))" + by auto + show ?case + unfolding eq + by (simp add: Alexandroff_open.ext EC.hyps closedin_Un compactin_Un ext.hyps) + qed + qed +next + fix \ + assume \
: "\K\\. Alexandroff_open X K" + show "Alexandroff_open X (\\)" + proof (cases "None \ \\") + case True + have "\K\\. \U. (openin X U \ K = Some ` U) \ (K = insert None (Some ` (topspace X - U)) \ compactin X U \ closedin X U)" + by (metis \
Alexandroff_open_iff) + then obtain U where U: + "\K. K \ \ \ openin X (U K) \ K = Some ` (U K) + \ (K = insert None (Some ` (topspace X - U K)) \ compactin X (U K) \ closedin X (U K))" + by metis + define \N where "\N \ {K \ \. None \ K}" + define A where "A \ \K\\-\N. U K" + define B where "B \ \K\\N. U K" + have U1: "\K. K \ \-\N \ openin X (U K) \ K = Some ` (U K)" + using U \N_def by auto + have U2: "\K. K \ \N \ K = insert None (Some ` (topspace X - U K)) \ compactin X (U K) \ closedin X (U K)" + using U \N_def by auto + have eqA: "\(\-\N) = Some ` A" + proof + show "\ (\ - \N) \ Some ` A" + by (metis A_def Sup_le_iff U1 UN_upper subset_image_iff) + show "Some ` A \ \ (\ - \N)" + using A_def U1 by blast + qed + have eqB: "\\N = insert None (Some ` (topspace X - B))" + using U2 True + by (auto simp: B_def image_iff \N_def) + have "\\ = \\N \ \(\-\N)" + by (auto simp: \N_def) + then have eq: "\\ = (Some ` A) \ (insert None (Some ` (topspace X - B)))" + by (simp add: eqA eqB Un_commute) + show ?thesis + unfolding eq + proof (intro Alexandroff_open_Un Alexandroff_open.intros) + show "openin X A" + using A_def U1 by blast + show "closedin X B" + unfolding B_def using U2 True \N_def by auto + show "compactin X B" + by (metis B_def U2 eqB Inf_lower Union_iff \closedin X B\ closed_compactin imageI insertI1) + qed + next + case False + then have "\K\\. \U. openin X U \ K = Some ` U" + by (metis Alexandroff_open.simps UnionI \
insertCI) + then obtain U where U: "\K\\. openin X (U K) \ K = Some ` (U K)" + by metis + then have eq: "\\ = Some ` (\ K\\. U K)" + using image_iff by fastforce + show ?thesis + unfolding eq by (simp add: U Alexandroff_open.base openin_clauses(3)) + qed +qed + + +definition Alexandroff_compactification where + "Alexandroff_compactification X \ topology (Alexandroff_open X)" + +lemma openin_Alexandroff_compactification: + "openin(Alexandroff_compactification X) V \ + (\U. openin X U \ V = Some ` U) \ + (\C. compactin X C \ closedin X C \ V = insert None (Some ` (topspace X - C)))" + by (auto simp: Alexandroff_compactification_def istopology_Alexandroff_open Alexandroff_open.simps) + + +lemma topspace_Alexandroff_compactification [simp]: + "topspace(Alexandroff_compactification X) = insert None (Some ` topspace X)" + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (force simp add: topspace_def openin_Alexandroff_compactification) + have "None \ topspace (Alexandroff_compactification X)" + by (meson closedin_empty compactin_empty insert_subset openin_Alexandroff_compactification openin_subset) + moreover have "Some x \ topspace (Alexandroff_compactification X)" + if "x \ topspace X" for x + by (meson that imageI openin_Alexandroff_compactification openin_subset openin_topspace subsetD) + ultimately show "?rhs \ ?lhs" + by (auto simp: image_subset_iff) +qed + +lemma closedin_Alexandroff_compactification: + "closedin (Alexandroff_compactification X) C \ + (\K. compactin X K \ closedin X K \ C = Some ` K) \ + (\U. openin X U \ C = topspace(Alexandroff_compactification X) - Some ` U)" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + apply (clarsimp simp: closedin_def openin_Alexandroff_compactification) + by (smt (verit) Diff_Diff_Int None_notin_image_Some image_set_diff inf.absorb_iff2 inj_Some insert_Diff_if subset_insert) + show "?rhs \ ?lhs" + using openin_subset + by (fastforce simp: closedin_def openin_Alexandroff_compactification) +qed + +lemma openin_Alexandroff_compactification_image_Some [simp]: + "openin(Alexandroff_compactification X) (Some ` U) \ openin X U" + by (auto simp: openin_Alexandroff_compactification inj_image_eq_iff) + +lemma closedin_Alexandroff_compactification_image_Some [simp]: + "closedin (Alexandroff_compactification X) (Some ` K) \ compactin X K \ closedin X K" + by (auto simp: closedin_Alexandroff_compactification inj_image_eq_iff) + +lemma open_map_Some: "open_map X (Alexandroff_compactification X) Some" + using open_map_def openin_Alexandroff_compactification by blast + +lemma continuous_map_Some: "continuous_map X (Alexandroff_compactification X) Some" + unfolding continuous_map_def +proof (intro conjI strip) + fix U + assume "openin (Alexandroff_compactification X) U" + then consider V where "openin X V" "U = Some ` V" + | C where "compactin X C" "closedin X C" "U = insert None (Some ` (topspace X - C))" + by (auto simp: openin_Alexandroff_compactification) + then show "openin X {x \ topspace X. Some x \ U}" + proof cases + case 1 + then show ?thesis + using openin_subopen openin_subset by fastforce + next + case 2 + then show ?thesis + by (simp add: closedin_def image_iff set_diff_eq) + qed +qed auto + + +lemma embedding_map_Some: "embedding_map X (Alexandroff_compactification X) Some" + by (simp add: continuous_map_Some injective_open_imp_embedding_map open_map_Some) + +lemma compact_space_Alexandroff_compactification [simp]: + "compact_space(Alexandroff_compactification X)" +proof (clarsimp simp: compact_space_alt image_subset_iff) + fix \ U + assume ope [rule_format]: "\U\\. openin (Alexandroff_compactification X) U" + and cover: "\x\topspace X. \X\\. Some x \ X" + and "U \ \" "None \ U" + then have Usub: "U \ insert None (Some ` topspace X)" + by (metis openin_subset topspace_Alexandroff_compactification) + with ope [OF \U \ \\] \None \ U\ + obtain C where C: "compactin X C \ closedin X C \ + insert None (Some ` topspace X) - U = Some ` C" + by (auto simp: openin_closedin closedin_Alexandroff_compactification) + then have D: "compactin (Alexandroff_compactification X) (insert None (Some ` topspace X) - U)" + by (metis continuous_map_Some image_compactin) + consider V where "openin X V" "U = Some ` V" + | C where "compactin X C" "closedin X C" "U = insert None (Some ` (topspace X - C))" + using ope [OF \U \ \\] by (auto simp: openin_Alexandroff_compactification) + then show "\\. finite \ \ \ \ \ \ (\X\\. None \ X) \ (\x\topspace X. \X\\. Some x \ X)" + proof cases + case 1 + then show ?thesis + using \None \ U\ by blast + next + case 2 + obtain \ where "finite \" "\ \ \" "insert None (Some ` topspace X) - U \ \\" + by (smt (verit, del_insts) C D Union_iff compactinD compactin_subset_topspace cover image_subset_iff ope subsetD) + with \U \ \\ show ?thesis + by (rule_tac x="insert U \" in exI) auto + qed +qed + +lemma topspace_Alexandroff_compactification_delete: + "topspace(Alexandroff_compactification X) - {None} = Some ` topspace X" + by simp + +lemma Alexandroff_compactification_dense: + assumes "\ compact_space X" + shows "(Alexandroff_compactification X) closure_of (Some ` topspace X) = + topspace(Alexandroff_compactification X)" + unfolding topspace_Alexandroff_compactification_delete [symmetric] +proof (intro one_point_compactification_dense) + show "\ compactin (Alexandroff_compactification X) (topspace (Alexandroff_compactification X) - {None})" + using assms compact_space_proper_map_preimage compact_space_subtopology embedding_map_Some embedding_map_def homeomorphic_imp_proper_map by fastforce +qed auto + + +lemma t0_space_one_point_compactification: + assumes "compact_space X \ openin X (topspace X - {a})" + shows "t0_space X \ t0_space (subtopology X (topspace X - {a}))" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + using t0_space_subtopology by blast + show "?rhs \ ?lhs" + using assms + unfolding t0_space_def by (bestsimp simp flip: Int_Diff dest: openin_trans_full) +qed + +lemma t0_space_Alexandroff_compactification [simp]: + "t0_space (Alexandroff_compactification X) \ t0_space X" + using t0_space_one_point_compactification [of "Alexandroff_compactification X" None] + using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t0_space by fastforce + +lemma t1_space_one_point_compactification: + assumes Xa: "openin X (topspace X - {a})" + and \
: "\K. \compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\ \ closedin X K" + shows "t1_space X \ t1_space (subtopology X (topspace X - {a}))" (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + using t1_space_subtopology by blast + assume R: ?rhs + show ?lhs + unfolding t1_space_closedin_singleton + proof (intro strip) + fix x + assume "x \ topspace X" + show "closedin X {x}" + proof (cases "x=a") + case True + then show ?thesis + using \x \ topspace X\ Xa closedin_def by blast + next + case False + show ?thesis + by (simp add: "\
" False R \x \ topspace X\ closedin_t1_singleton) + qed + qed +qed + +lemma closedin_Alexandroff_I: + assumes "compactin (Alexandroff_compactification X) K" "K \ Some ` topspace X" + "closedin (Alexandroff_compactification X) T" "K = T \ Some ` topspace X" + shows "closedin (Alexandroff_compactification X) K" +proof - + obtain S where S: "S \ topspace X" "K = Some ` S" + by (meson \K \ Some ` topspace X\ subset_imageE) + with assms have "compactin X S" + by (metis compactin_subtopology embedding_map_Some embedding_map_def homeomorphic_map_compactness) + moreover have "closedin X S" + using assms S + by (metis closedin_subtopology embedding_map_Some embedding_map_def homeomorphic_map_closedness) + ultimately show ?thesis + by (simp add: S) +qed + + +lemma t1_space_Alexandroff_compactification [simp]: + "t1_space(Alexandroff_compactification X) \ t1_space X" +proof - + have "openin (Alexandroff_compactification X) (topspace (Alexandroff_compactification X) - {None})" + by auto + then show ?thesis + using t1_space_one_point_compactification [of "Alexandroff_compactification X" None] + using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t1_space + by (fastforce simp add: compactin_subtopology closedin_Alexandroff_I closedin_subtopology) +qed + + +lemma kc_space_one_point_compactification: + assumes "compact_space X" + and ope: "openin X (topspace X - {a})" + and \
: "\K. \compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\ + \ closedin X K" + shows "kc_space X \ + k_space (subtopology X (topspace X - {a})) \ kc_space (subtopology X (topspace X - {a}))" +proof - + have "closedin X K" + if "kc_space (subtopology X (topspace X - {a}))" and "compactin X K" "a \ K" for K + using that unfolding kc_space_def + by (metis "\
" Diff_empty compactin_subspace compactin_subtopology subset_Diff_insert) + then show ?thesis + by (metis \compact_space X\ kc_space_one_point_compactification_gen ope) +qed + +lemma kc_space_Alexandroff_compactification: + "kc_space(Alexandroff_compactification X) \ (k_space X \ kc_space X)" (is "kc_space ?Y = _") +proof - + have "kc_space (Alexandroff_compactification X) \ + (k_space (subtopology ?Y (topspace ?Y - {None})) \ kc_space (subtopology ?Y (topspace ?Y - {None})))" + by (rule kc_space_one_point_compactification) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I) + also have "... \ k_space X \ kc_space X" + using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_k_space homeomorphic_kc_space by simp blast + finally show ?thesis . +qed + + +lemma regular_space_one_point_compactification: + assumes "compact_space X" and ope: "openin X (topspace X - {a})" + and \
: "\K. \compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\ \ closedin X K" + shows "regular_space X \ + regular_space (subtopology X (topspace X - {a})) \ locally_compact_space (subtopology X (topspace X - {a}))" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + using assms(1) compact_imp_locally_compact_space locally_compact_space_open_subset ope regular_space_subtopology by blast + assume R: ?rhs + let ?Xa = "subtopology X (topspace X - {a})" + show ?lhs + unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of imp_conjL + proof (intro strip) + fix W x + assume "openin X W" and "x \ W" + show "\U V. openin X U \ closedin X V \ x \ U \ U \ V \ V \ W" + proof (cases "x=a") + case True + have "compactin ?Xa (topspace X - W)" + using \openin X W\ assms(1) closedin_compact_space + by (metis Diff_mono True \x \ W\ compactin_subtopology empty_subsetI insert_subset openin_closedin_eq order_refl) + then obtain V K where V: "openin ?Xa V" and K: "compactin ?Xa K" "closedin ?Xa K" and "topspace X - W \ V" "V \ K" + by (metis locally_compact_space_compact_closed_compact R) + show ?thesis + proof (intro exI conjI) + show "openin X (topspace X - K)" + using "\
" K by blast + show "closedin X (topspace X - V)" + using V ope openin_trans_full by blast + show "x \ topspace X - K" + proof (rule) + show "x \ topspace X" + using \openin X W\ \x \ W\ openin_subset by blast + show "x \ K" + using K True closedin_imp_subset by blast + qed + show "topspace X - K \ topspace X - V" + by (simp add: Diff_mono \V \ K\) + show "topspace X - V \ W" + using \topspace X - W \ V\ by auto + qed + next + case False + have "openin ?Xa ((topspace X - {a}) \ W)" + using \openin X W\ openin_subtopology_Int2 by blast + moreover have "x \ (topspace X - {a}) \ W" + using \openin X W\ \x \ W\ openin_subset False by blast + ultimately obtain U V where "openin ?Xa U" "compactin ?Xa V" "closedin ?Xa V" + "x \ U" "U \ V" "V \ (topspace X - {a}) \ W" + using R locally_compact_regular_space_neighbourhood_base neighbourhood_base_of + by (metis (no_types, lifting)) + then show ?thesis + by (meson "\
" le_infE ope openin_trans_full) + qed + qed +qed + + +lemma regular_space_Alexandroff_compactification: + "regular_space(Alexandroff_compactification X) \ regular_space X \ locally_compact_space X" + (is "regular_space ?Y = ?rhs") +proof - + have "regular_space ?Y \ + regular_space (subtopology ?Y (topspace ?Y - {None})) \ locally_compact_space (subtopology ?Y (topspace ?Y - {None}))" + by (rule regular_space_one_point_compactification) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I) + also have "... \ regular_space X \ locally_compact_space X" + using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_locally_compact_space homeomorphic_regular_space + by fastforce + finally show ?thesis . +qed + + +lemma Hausdorff_space_one_point_compactification: + assumes "compact_space X" and "openin X (topspace X - {a})" + and "\K. \compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\ \ closedin X K" + shows "Hausdorff_space X \ + Hausdorff_space (subtopology X (topspace X - {a})) \ locally_compact_space (subtopology X (topspace X - {a}))" + (is "?lhs \ ?rhs") +proof + show ?rhs if ?lhs + proof - + have "locally_compact_space (subtopology X (topspace X - {a}))" + using assms that compact_imp_locally_compact_space locally_compact_space_open_subset + by blast + with that show ?rhs + by (simp add: Hausdorff_space_subtopology) + qed +next + show "?rhs \ ?lhs" + by (metis assms locally_compact_Hausdorff_or_regular regular_space_one_point_compactification + regular_t1_eq_Hausdorff_space t1_space_one_point_compactification) +qed + +lemma Hausdorff_space_Alexandroff_compactification: + "Hausdorff_space(Alexandroff_compactification X) \ Hausdorff_space X \ locally_compact_space X" + by (meson compact_Hausdorff_imp_regular_space compact_space_Alexandroff_compactification + locally_compact_Hausdorff_or_regular regular_space_Alexandroff_compactification + regular_t1_eq_Hausdorff_space t1_space_Alexandroff_compactification) + +lemma completely_regular_space_Alexandroff_compactification: + "completely_regular_space(Alexandroff_compactification X) \ + completely_regular_space X \ locally_compact_space X" + by (metis regular_space_Alexandroff_compactification completely_regular_eq_regular_space + compact_imp_locally_compact_space compact_space_Alexandroff_compactification) + +lemma Hausdorff_space_one_point_compactification_asymmetric_prod: + assumes "compact_space X" + shows "Hausdorff_space X \ + kc_space (prod_topology X (subtopology X (topspace X - {a}))) \ + k_space (prod_topology X (subtopology X (topspace X - {a})))" (is "?lhs \ ?rhs") +proof (cases "a \ topspace X") + case True + show ?thesis + proof + show ?rhs if ?lhs + proof + show "kc_space (prod_topology X (subtopology X (topspace X - {a})))" + using Hausdorff_imp_kc_space kc_space_prod_topology_right kc_space_subtopology that by blast + show "k_space (prod_topology X (subtopology X (topspace X - {a})))" + by (meson Hausdorff_imp_kc_space assms compact_imp_locally_compact_space k_space_prod_topology_left + kc_space_one_point_compactification_gen that) + qed + next + assume R: ?rhs + define D where "D \ (\x. (x,x)) ` (topspace X - {a})" + show ?lhs + proof (cases "topspace X = {a}") + case True + then show ?thesis + by (simp add: Hausdorff_space_def) + next + case False + then have "kc_space X" + using kc_space_retraction_map_image [of "prod_topology X (subtopology X (topspace X - {a}))" X fst] + by (metis Diff_subset R True insert_Diff retraction_map_fst topspace_subtopology_subset) + have "closedin (subtopology (prod_topology X (subtopology X (topspace X - {a}))) K) (K \ D)" + if "compactin (prod_topology X (subtopology X (topspace X - {a}))) K" for K + proof (intro closedin_subtopology_Int_subset[where V=K] closedin_subset_topspace) + show "fst ` K \ snd ` K \ D \ fst ` K \ snd ` K" "K \ fst ` K \ snd ` K" + by force+ + have eq: "(fst ` K \ snd ` K \ D) = ((\x. (x,x)) ` (fst ` K \ snd ` K))" + using compactin_subset_topspace that by (force simp: D_def image_iff) + have "compactin (prod_topology X (subtopology X (topspace X - {a}))) (fst ` K \ snd ` K \ D)" + unfolding eq + proof (rule image_compactin [of "subtopology X (topspace X - {a})"]) + have "compactin X (fst ` K)" "compactin X (snd ` K)" + by (meson compactin_subtopology continuous_map_fst continuous_map_snd image_compactin that)+ + moreover have "fst ` K \ snd ` K \ topspace X - {a}" + using compactin_subset_topspace that by force + ultimately + show "compactin (subtopology X (topspace X - {a})) (fst ` K \ snd ` K)" + unfolding compactin_subtopology + by (meson \kc_space X\ closed_Int_compactin kc_space_def) + show "continuous_map (subtopology X (topspace X - {a})) (prod_topology X (subtopology X (topspace X - {a}))) (\x. (x,x))" + by (simp add: continuous_map_paired) + qed + then show "closedin (prod_topology X (subtopology X (topspace X - {a}))) (fst ` K \ snd ` K \ D)" + using R compactin_imp_closedin_gen by blast + qed + moreover have "D \ topspace X \ (topspace X \ (topspace X - {a}))" + by (auto simp: D_def) + ultimately have *: "closedin (prod_topology X (subtopology X (topspace X - {a}))) D" + using R by (auto simp: k_space) + have "x=y" + if "x \ topspace X" "y \ topspace X" + and \
: "\T. \(x,y) \ T; openin (prod_topology X X) T\ \ \z \ topspace X. (z,z) \ T" for x y + proof (cases "x=a \ y=a") + case False + then consider "x\a" | "y\a" + by blast + then show ?thesis + proof cases + case 1 + have "\z \ topspace X - {a}. (z,z) \ T" + if "(y,x) \ T" "openin (prod_topology X (subtopology X (topspace X - {a}))) T" for T + proof - + have "(x,y) \ {z \ topspace (prod_topology X X). (snd z,fst z) \ T \ topspace X \ (topspace X - {a})}" + by (simp add: "1" \x \ topspace X\ \y \ topspace X\ that) + moreover have "openin (prod_topology X X) {z \ topspace (prod_topology X X). (snd z,fst z) \ T \ topspace X \ (topspace X - {a})}" + proof (rule openin_continuous_map_preimage) + show "continuous_map (prod_topology X X) (prod_topology X X) (\x. (snd x, fst x))" + by (simp add: continuous_map_fst continuous_map_pairedI continuous_map_snd) + have "openin (prod_topology X X) (topspace X \ (topspace X - {a}))" + using \kc_space X\ assms kc_space_one_point_compactification_gen openin_prod_Times_iff by fastforce + moreover have "openin (prod_topology X X) T" + using kc_space_one_point_compactification_gen [OF \compact_space X\] \kc_space X\ that + by (metis openin_prod_Times_iff openin_topspace openin_trans_full prod_topology_subtopology(2)) + ultimately show "openin (prod_topology X X) (T \ topspace X \ (topspace X - {a}))" + by blast + qed + ultimately show ?thesis + by (smt (verit) \
Int_iff fst_conv mem_Collect_eq mem_Sigma_iff snd_conv) + qed + then have "(y,x) \ prod_topology X (subtopology X (topspace X - {a})) closure_of D" + by (simp add: "1" D_def in_closure_of that) + then show ?thesis + using that "*" D_def closure_of_closedin by fastforce + next + case 2 + have "\z \ topspace X - {a}. (z,z) \ T" + if "(x,y) \ T" "openin (prod_topology X (subtopology X (topspace X - {a}))) T" for T + proof - + have "openin (prod_topology X X) (topspace X \ (topspace X - {a}))" + using \kc_space X\ assms kc_space_one_point_compactification_gen openin_prod_Times_iff by fastforce + moreover have XXT: "openin (prod_topology X X) T" + using kc_space_one_point_compactification_gen [OF \compact_space X\] \kc_space X\ that + by (metis openin_prod_Times_iff openin_topspace openin_trans_full prod_topology_subtopology(2)) + ultimately have "openin (prod_topology X X) (T \ topspace X \ (topspace X - {a}))" + by blast + then show ?thesis + by (smt (verit) "\
" Diff_subset XXT mem_Sigma_iff openin_subset subsetD that topspace_prod_topology topspace_subtopology_subset) + qed + then have "(x,y) \ prod_topology X (subtopology X (topspace X - {a})) closure_of D" + by (simp add: "2" D_def in_closure_of that) + then show ?thesis + using that "*" D_def closure_of_closedin by fastforce + qed + qed auto + then show ?thesis + unfolding Hausdorff_space_closedin_diagonal closure_of_subset_eq [symmetric] + by (force simp add: closure_of_def) + qed + qed +next + case False + then show ?thesis + by (simp add: assms compact_imp_k_space compact_space_prod_topology kc_space_compact_prod_topology) +qed + + +lemma Hausdorff_space_Alexandroff_compactification_asymmetric_prod: + "Hausdorff_space(Alexandroff_compactification X) \ + kc_space(prod_topology (Alexandroff_compactification X) X) \ + k_space(prod_topology (Alexandroff_compactification X) X)" + (is "Hausdorff_space ?Y = ?rhs") +proof - + have *: "subtopology (Alexandroff_compactification X) + (topspace (Alexandroff_compactification X) - + {None}) homeomorphic_space X" + using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_space_sym by fastforce + have "Hausdorff_space (Alexandroff_compactification X) \ + (kc_space (prod_topology ?Y (subtopology ?Y (topspace ?Y - {None}))) \ + k_space (prod_topology ?Y (subtopology ?Y (topspace ?Y - {None}))))" + by (rule Hausdorff_space_one_point_compactification_asymmetric_prod) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I) + also have "... \ ?rhs" + using homeomorphic_k_space homeomorphic_kc_space homeomorphic_space_prod_topology + homeomorphic_space_refl * by blast + finally show ?thesis . +qed + +lemma kc_space_as_compactification_unique: + assumes "kc_space X" "compact_space X" + shows "openin X U \ + (if a \ U then U \ topspace X \ compactin X (topspace X - U) + else openin (subtopology X (topspace X - {a})) U)" +proof (cases "a \ U") + case True + then show ?thesis + by (meson assms closedin_compact_space compactin_imp_closedin_gen openin_closedin_eq) +next + case False + then show ?thesis + by (metis Diff_empty kc_space_one_point_compactification_gen openin_open_subtopology openin_subset subset_Diff_insert assms) +qed + +lemma kc_space_as_compactification_unique_explicit: + assumes "kc_space X" "compact_space X" + shows "openin X U \ + (if a \ U then U \ topspace X \ + compactin (subtopology X (topspace X - {a})) (topspace X - U) \ + closedin (subtopology X (topspace X - {a})) (topspace X - U) + else openin (subtopology X (topspace X - {a})) U)" + apply (simp add: kc_space_subtopology compactin_imp_closedin_gen assms compactin_subtopology cong: conj_cong) + by (metis Diff_mono assms bot.extremum insert_subset kc_space_as_compactification_unique subset_refl) + +lemma Alexandroff_compactification_unique: + assumes "kc_space X" "compact_space X" and a: "a \ topspace X" + shows "Alexandroff_compactification (subtopology X (topspace X - {a})) homeomorphic_space X" + (is "?Y homeomorphic_space X") +proof - + have [simp]: "topspace X \ (topspace X - {a}) = topspace X - {a}" + by auto + have [simp]: "insert None (Some ` A) = insert None (Some ` B) \ A = B" + "insert None (Some ` A) \ Some ` B" for A B + by auto + have "quotient_map X ?Y (\x. if x = a then None else Some x)" + unfolding quotient_map_def + proof (intro conjI strip) + show "(\x. if x = a then None else Some x) ` topspace X = topspace ?Y" + using \a \ topspace X\ by force + show "openin X {x \ topspace X. (if x = a then None else Some x) \ U} = openin ?Y U" (is "?L = ?R") + if "U \ topspace ?Y" for U + proof (cases "None \ U") + case True + then obtain T where T[simp]: "U = insert None (Some ` T)" + by (metis Int_insert_right UNIV_I UNIV_option_conv inf.orderE inf_le2 subsetI subset_imageE) + have Tsub: "T \ topspace X - {a}" + using in_these_eq that by auto + then have "{x \ topspace X. (if x = a then None else Some x) \ U} = insert a T" + by (auto simp: a image_iff cong: conj_cong) + then have "?L \ openin X (insert a T)" + by metis + also have "\ \ ?R" + using Tsub assms + apply (simp add: openin_Alexandroff_compactification kc_space_as_compactification_unique_explicit [where a=a]) + by (smt (verit, best) Diff_insert2 Diff_subset closedin_imp_subset double_diff) + finally show ?thesis . + next + case False + then obtain T where [simp]: "U = Some ` T" + by (metis Int_insert_right UNIV_I UNIV_option_conv inf.orderE inf_le2 subsetI subset_imageE) + have **: "\V. openin X V \ openin X (V - {a})" + by (simp add: assms compactin_imp_closedin_gen openin_diff) + have Tsub: "T \ topspace X - {a}" + using in_these_eq that by auto + then have "{x \ topspace X. (if x = a then None else Some x) \ U} = T" + by (auto simp: image_iff cong: conj_cong) + then show ?thesis + by (simp add: "**" Tsub openin_open_subtopology) + qed + qed + moreover have "inj_on (\x. if x = a then None else Some x) (topspace X)" + by (auto simp: inj_on_def) + ultimately show ?thesis + using homeomorphic_space_sym homeomorphic_space homeomorphic_map_def by blast +qed + +end diff --git a/src/HOL/NthRoot.thy b/src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy +++ b/src/HOL/NthRoot.thy @@ -1,824 +1,824 @@ (* Title: HOL/NthRoot.thy Author: Jacques D. Fleuriot, 1998 Author: Lawrence C Paulson, 2004 *) section \Nth Roots of Real Numbers\ theory NthRoot imports Deriv begin subsection \Existence of Nth Root\ text \Existence follows from the Intermediate Value Theorem\ lemma realpow_pos_nth: fixes a :: real assumes n: "0 < n" and a: "0 < a" shows "\r>0. r ^ n = a" proof - have "\r\0. r \ (max 1 a) \ r ^ n = a" proof (rule IVT) show "0 ^ n \ a" using n a by (simp add: power_0_left) show "0 \ max 1 a" by simp from n have n1: "1 \ n" by simp have "a \ max 1 a ^ 1" by simp also have "max 1 a ^ 1 \ max 1 a ^ n" using n1 by (rule power_increasing) simp finally show "a \ max 1 a ^ n" . show "\r. 0 \ r \ r \ max 1 a \ isCont (\x. x ^ n) r" by simp qed then obtain r where r: "0 \ r \ r ^ n = a" by fast with n a have "r \ 0" by (auto simp add: power_0_left) with r have "0 < r \ r ^ n = a" by simp then show ?thesis .. qed (* Used by Integration/RealRandVar.thy in AFP *) lemma realpow_pos_nth2: "(0::real) < a \ \r>0. r ^ Suc n = a" by (blast intro: realpow_pos_nth) text \Uniqueness of nth positive root.\ lemma realpow_pos_nth_unique: "0 < n \ 0 < a \ \!r. 0 < r \ r ^ n = a" for a :: real by (auto intro!: realpow_pos_nth simp: power_eq_iff_eq_base) subsection \Nth Root\ text \ We define roots of negative reals such that \root n (- x) = - root n x\. This allows us to omit side conditions from many theorems. \ lemma inj_sgn_power: assumes "0 < n" shows "inj (\y. sgn y * \y\^n :: real)" (is "inj ?f") proof (rule injI) have x: "(0 < a \ b < 0) \ (a < 0 \ 0 < b) \ a \ b" for a b :: real by auto fix x y assume "?f x = ?f y" with power_eq_iff_eq_base[of n "\x\" "\y\"] \0 < n\ show "x = y" by (cases rule: linorder_cases[of 0 x, case_product linorder_cases[of 0 y]]) (simp_all add: x) qed lemma sgn_power_injE: "sgn a * \a\ ^ n = x \ x = sgn b * \b\ ^ n \ 0 < n \ a = b" for a b :: real using inj_sgn_power[THEN injD, of n a b] by simp definition root :: "nat \ real \ real" where "root n x = (if n = 0 then 0 else the_inv (\y. sgn y * \y\^n) x)" lemma root_0 [simp]: "root 0 x = 0" by (simp add: root_def) lemma root_sgn_power: "0 < n \ root n (sgn y * \y\^n) = y" using the_inv_f_f[OF inj_sgn_power] by (simp add: root_def) lemma sgn_power_root: assumes "0 < n" shows "sgn (root n x) * \(root n x)\^n = x" (is "?f (root n x) = x") proof (cases "x = 0") case True with assms root_sgn_power[of n 0] show ?thesis by simp next case False with realpow_pos_nth[OF \0 < n\, of "\x\"] obtain r where "0 < r" "r ^ n = \x\" by auto with \x \ 0\ have S: "x \ range ?f" by (intro image_eqI[of _ _ "sgn x * r"]) (auto simp: abs_mult sgn_mult power_mult_distrib abs_sgn_eq mult_sgn_abs) from \0 < n\ f_the_inv_into_f[OF inj_sgn_power[OF \0 < n\] this] show ?thesis by (simp add: root_def) qed lemma split_root: "P (root n x) \ (n = 0 \ P 0) \ (0 < n \ (\y. sgn y * \y\^n = x \ P y))" proof (cases "n = 0") case True then show ?thesis by simp next case False then show ?thesis by simp (metis root_sgn_power sgn_power_root) qed lemma real_root_zero [simp]: "root n 0 = 0" by (simp split: split_root add: sgn_zero_iff) lemma real_root_minus: "root n (- x) = - root n x" by (clarsimp split: split_root elim!: sgn_power_injE simp: sgn_minus) lemma real_root_less_mono: "0 < n \ x < y \ root n x < root n y" proof (clarsimp split: split_root) have *: "0 < b \ a < 0 \ \ a > b" for a b :: real by auto fix a b :: real assume "0 < n" "sgn a * \a\ ^ n < sgn b * \b\ ^ n" then show "a < b" using power_less_imp_less_base[of a n b] power_less_imp_less_base[of "- b" n "- a"] by (simp add: sgn_real_def * [of "a ^ n" "- ((- b) ^ n)"] split: if_split_asm) qed lemma real_root_gt_zero: "0 < n \ 0 < x \ 0 < root n x" using real_root_less_mono[of n 0 x] by simp lemma real_root_ge_zero: "0 \ x \ 0 \ root n x" using real_root_gt_zero[of n x] by (cases "n = 0") (auto simp add: le_less) lemma real_root_pow_pos: "0 < n \ 0 < x \ root n x ^ n = x" (* TODO: rename *) using sgn_power_root[of n x] real_root_gt_zero[of n x] by simp lemma real_root_pow_pos2 [simp]: "0 < n \ 0 \ x \ root n x ^ n = x" (* TODO: rename *) by (auto simp add: order_le_less real_root_pow_pos) lemma sgn_root: "0 < n \ sgn (root n x) = sgn x" by (auto split: split_root simp: sgn_real_def) lemma odd_real_root_pow: "odd n \ root n x ^ n = x" using sgn_power_root[of n x] by (simp add: odd_pos sgn_real_def split: if_split_asm) lemma real_root_power_cancel: "0 < n \ 0 \ x \ root n (x ^ n) = x" using root_sgn_power[of n x] by (auto simp add: le_less power_0_left) lemma odd_real_root_power_cancel: "odd n \ root n (x ^ n) = x" using root_sgn_power[of n x] by (simp add: odd_pos sgn_real_def power_0_left split: if_split_asm) lemma real_root_pos_unique: "0 < n \ 0 \ y \ y ^ n = x \ root n x = y" using root_sgn_power[of n y] by (auto simp add: le_less power_0_left) lemma odd_real_root_unique: "odd n \ y ^ n = x \ root n x = y" by (erule subst, rule odd_real_root_power_cancel) lemma real_root_one [simp]: "0 < n \ root n 1 = 1" by (simp add: real_root_pos_unique) text \Root function is strictly monotonic, hence injective.\ lemma real_root_le_mono: "0 < n \ x \ y \ root n x \ root n y" by (auto simp add: order_le_less real_root_less_mono) lemma real_root_less_iff [simp]: "0 < n \ root n x < root n y \ x < y" by (cases "x < y") (simp_all add: real_root_less_mono linorder_not_less real_root_le_mono) lemma real_root_le_iff [simp]: "0 < n \ root n x \ root n y \ x \ y" by (cases "x \ y") (simp_all add: real_root_le_mono linorder_not_le real_root_less_mono) lemma real_root_eq_iff [simp]: "0 < n \ root n x = root n y \ x = y" by (simp add: order_eq_iff) lemmas real_root_gt_0_iff [simp] = real_root_less_iff [where x=0, simplified] lemmas real_root_lt_0_iff [simp] = real_root_less_iff [where y=0, simplified] lemmas real_root_ge_0_iff [simp] = real_root_le_iff [where x=0, simplified] lemmas real_root_le_0_iff [simp] = real_root_le_iff [where y=0, simplified] lemmas real_root_eq_0_iff [simp] = real_root_eq_iff [where y=0, simplified] lemma real_root_gt_1_iff [simp]: "0 < n \ 1 < root n y \ 1 < y" using real_root_less_iff [where x=1] by simp lemma real_root_lt_1_iff [simp]: "0 < n \ root n x < 1 \ x < 1" using real_root_less_iff [where y=1] by simp lemma real_root_ge_1_iff [simp]: "0 < n \ 1 \ root n y \ 1 \ y" using real_root_le_iff [where x=1] by simp lemma real_root_le_1_iff [simp]: "0 < n \ root n x \ 1 \ x \ 1" using real_root_le_iff [where y=1] by simp lemma real_root_eq_1_iff [simp]: "0 < n \ root n x = 1 \ x = 1" using real_root_eq_iff [where y=1] by simp text \Roots of multiplication and division.\ lemma real_root_mult: "root n (x * y) = root n x * root n y" by (auto split: split_root elim!: sgn_power_injE simp: sgn_mult abs_mult power_mult_distrib) lemma real_root_inverse: "root n (inverse x) = inverse (root n x)" by (auto split: split_root elim!: sgn_power_injE simp: power_inverse) lemma real_root_divide: "root n (x / y) = root n x / root n y" by (simp add: divide_inverse real_root_mult real_root_inverse) lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" by (simp add: abs_if real_root_minus) lemma root_abs_power: "n > 0 \ abs (root n (y ^n)) = abs y" using root_sgn_power [of n] by (metis abs_ge_zero power_abs real_root_abs real_root_power_cancel) lemma real_root_power: "0 < n \ root n (x ^ k) = root n x ^ k" by (induct k) (simp_all add: real_root_mult) text \Roots of roots.\ lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x" by (simp add: odd_real_root_unique) lemma real_root_mult_exp: "root (m * n) x = root m (root n x)" by (auto split: split_root elim!: sgn_power_injE simp: sgn_zero_iff sgn_mult power_mult[symmetric] abs_mult power_mult_distrib abs_sgn_eq) lemma real_root_commute: "root m (root n x) = root n (root m x)" by (simp add: real_root_mult_exp [symmetric] mult.commute) text \Monotonicity in first argument.\ lemma real_root_strict_decreasing: assumes "0 < n" "n < N" "1 < x" shows "root N x < root n x" proof - from assms have "root n (root N x) ^ n < root N (root n x) ^ N" by (simp add: real_root_commute power_strict_increasing del: real_root_pow_pos2) with assms show ?thesis by simp qed lemma real_root_strict_increasing: assumes "0 < n" "n < N" "0 < x" "x < 1" shows "root n x < root N x" proof - from assms have "root N (root n x) ^ N < root n (root N x) ^ n" by (simp add: real_root_commute power_strict_decreasing del: real_root_pow_pos2) with assms show ?thesis by simp qed lemma real_root_decreasing: "0 < n \ n \ N \ 1 \ x \ root N x \ root n x" by (auto simp add: order_le_less real_root_strict_decreasing) lemma real_root_increasing: "0 < n \ n \ N \ 0 \ x \ x \ 1 \ root n x \ root N x" by (auto simp add: order_le_less real_root_strict_increasing) text \Continuity and derivatives.\ lemma isCont_real_root: "isCont (root n) x" proof (cases "n > 0") case True let ?f = "\y::real. sgn y * \y\^n" have "continuous_on ({0..} \ {.. 0}) (\x. if 0 < x then x ^ n else - ((-x) ^ n) :: real)" using True by (intro continuous_on_If continuous_intros) auto then have "continuous_on UNIV ?f" by (rule continuous_on_cong[THEN iffD1, rotated 2]) (auto simp: not_less le_less True) then have [simp]: "isCont ?f x" for x by (simp add: continuous_on_eq_continuous_at) have "isCont (root n) (?f (root n x))" by (rule isCont_inverse_function [where f="?f" and d=1]) (auto simp: root_sgn_power True) then show ?thesis by (simp add: sgn_power_root True) next case False then show ?thesis by (simp add: root_def[abs_def]) qed lemma tendsto_real_root [tendsto_intros]: "(f \ x) F \ ((\x. root n (f x)) \ root n x) F" using isCont_tendsto_compose[OF isCont_real_root, of f x F] . lemma continuous_real_root [continuous_intros]: "continuous F f \ continuous F (\x. root n (f x))" unfolding continuous_def by (rule tendsto_real_root) lemma continuous_on_real_root [continuous_intros]: "continuous_on s f \ continuous_on s (\x. root n (f x))" unfolding continuous_on_def by (auto intro: tendsto_real_root) lemma DERIV_real_root: assumes n: "0 < n" and x: "0 < x" shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" proof (rule DERIV_inverse_function) show "0 < x" using x . show "x < x + 1" by simp show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" by (rule DERIV_pow) show "real n * root n x ^ (n - Suc 0) \ 0" using n x by simp show "isCont (root n) x" by (rule isCont_real_root) qed (use n in auto) lemma DERIV_odd_real_root: assumes n: "odd n" and x: "x \ 0" shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" proof (rule DERIV_inverse_function) show "x - 1 < x" "x < x + 1" by auto show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" by (rule DERIV_pow) show "real n * root n x ^ (n - Suc 0) \ 0" using odd_pos [OF n] x by simp show "isCont (root n) x" by (rule isCont_real_root) qed (use n odd_real_root_pow in auto) lemma DERIV_even_real_root: assumes n: "0 < n" and "even n" and x: "x < 0" shows "DERIV (root n) x :> inverse (- real n * root n x ^ (n - Suc 0))" proof (rule DERIV_inverse_function) show "x - 1 < x" by simp show "x < 0" using x . show "- (root n y ^ n) = y" if "x - 1 < y" and "y < 0" for y proof - have "root n (-y) ^ n = -y" using that \0 < n\ by simp with real_root_minus and \even n\ show "- (root n y ^ n) = y" by simp qed show "DERIV (\x. - (x ^ n)) (root n x) :> - real n * root n x ^ (n - Suc 0)" by (auto intro!: derivative_eq_intros) show "- real n * root n x ^ (n - Suc 0) \ 0" using n x by simp show "isCont (root n) x" by (rule isCont_real_root) qed lemma DERIV_real_root_generic: assumes "0 < n" and "x \ 0" and "even n \ 0 < x \ D = inverse (real n * root n x ^ (n - Suc 0))" and "even n \ x < 0 \ D = - inverse (real n * root n x ^ (n - Suc 0))" and "odd n \ D = inverse (real n * root n x ^ (n - Suc 0))" shows "DERIV (root n) x :> D" using assms by (cases "even n", cases "0 < x") (auto intro: DERIV_real_root[THEN DERIV_cong] DERIV_odd_real_root[THEN DERIV_cong] DERIV_even_real_root[THEN DERIV_cong]) lemma power_tendsto_0_iff [simp]: fixes f :: "'a \ real" assumes "n > 0" shows "((\x. f x ^ n) \ 0) F \ (f \ 0) F" proof - have "((\x. \root n (f x ^ n)\) \ 0) F \ (f \ 0) F" by (auto simp: assms root_abs_power tendsto_rabs_zero_iff) then have "((\x. f x ^ n) \ 0) F \ (f \ 0) F" by (metis tendsto_real_root abs_0 real_root_zero tendsto_rabs) with assms show ?thesis by (auto simp: tendsto_null_power) qed subsection \Square Root\ definition sqrt :: "real \ real" where "sqrt = root 2" lemma pos2: "0 < (2::nat)" by simp lemma real_sqrt_unique: "y\<^sup>2 = x \ 0 \ y \ sqrt x = y" unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) lemma real_sqrt_abs [simp]: "sqrt (x\<^sup>2) = \x\" by (metis power2_abs abs_ge_zero real_sqrt_unique) lemma real_sqrt_pow2 [simp]: "0 \ x \ (sqrt x)\<^sup>2 = x" unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) lemma real_sqrt_pow2_iff [simp]: "(sqrt x)\<^sup>2 = x \ 0 \ x" by (metis real_sqrt_pow2 zero_le_power2) lemma real_sqrt_zero [simp]: "sqrt 0 = 0" unfolding sqrt_def by (rule real_root_zero) lemma real_sqrt_one [simp]: "sqrt 1 = 1" unfolding sqrt_def by (rule real_root_one [OF pos2]) lemma real_sqrt_four [simp]: "sqrt 4 = 2" using real_sqrt_abs[of 2] by simp lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" unfolding sqrt_def by (rule real_root_minus) lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" unfolding sqrt_def by (rule real_root_mult) lemma real_sqrt_mult_self[simp]: "sqrt a * sqrt a = \a\" using real_sqrt_abs[of a] unfolding power2_eq_square real_sqrt_mult . lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" unfolding sqrt_def by (rule real_root_inverse) lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" unfolding sqrt_def by (rule real_root_divide) lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k" unfolding sqrt_def by (rule real_root_power [OF pos2]) lemma real_sqrt_gt_zero: "0 < x \ 0 < sqrt x" unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) lemma real_sqrt_ge_zero: "0 \ x \ 0 \ sqrt x" unfolding sqrt_def by (rule real_root_ge_zero) lemma real_sqrt_less_mono: "x < y \ sqrt x < sqrt y" unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) lemma real_sqrt_le_mono: "x \ y \ sqrt x \ sqrt y" unfolding sqrt_def by (rule real_root_le_mono [OF pos2]) lemma real_sqrt_less_iff [simp]: "sqrt x < sqrt y \ x < y" unfolding sqrt_def by (rule real_root_less_iff [OF pos2]) lemma real_sqrt_le_iff [simp]: "sqrt x \ sqrt y \ x \ y" unfolding sqrt_def by (rule real_root_le_iff [OF pos2]) lemma real_sqrt_eq_iff [simp]: "sqrt x = sqrt y \ x = y" unfolding sqrt_def by (rule real_root_eq_iff [OF pos2]) -lemma real_less_lsqrt: "0 \ x \ 0 \ y \ x < y\<^sup>2 \ sqrt x < y" +lemma real_less_lsqrt: "0 \ y \ x < y\<^sup>2 \ sqrt x < y" using real_sqrt_less_iff[of x "y\<^sup>2"] by simp -lemma real_le_lsqrt: "0 \ x \ 0 \ y \ x \ y\<^sup>2 \ sqrt x \ y" +lemma real_le_lsqrt: "0 \ y \ x \ y\<^sup>2 \ sqrt x \ y" using real_sqrt_le_iff[of x "y\<^sup>2"] by simp lemma real_le_rsqrt: "x\<^sup>2 \ y \ x \ sqrt y" using real_sqrt_le_mono[of "x\<^sup>2" y] by simp lemma real_less_rsqrt: "x\<^sup>2 < y \ x < sqrt y" using real_sqrt_less_mono[of "x\<^sup>2" y] by simp lemma real_sqrt_power_even: assumes "even n" "x \ 0" shows "sqrt x ^ n = x ^ (n div 2)" proof - from assms obtain k where "n = 2*k" by (auto elim!: evenE) with assms show ?thesis by (simp add: power_mult) qed lemma sqrt_le_D: "sqrt x \ y \ x \ y\<^sup>2" by (meson not_le real_less_rsqrt) lemma sqrt_ge_absD: "\x\ \ sqrt y \ x\<^sup>2 \ y" using real_sqrt_le_iff[of "x\<^sup>2"] by simp lemma sqrt_even_pow2: assumes n: "even n" shows "sqrt (2 ^ n) = 2 ^ (n div 2)" proof - from n obtain m where m: "n = 2 * m" .. from m have "sqrt (2 ^ n) = sqrt ((2 ^ m)\<^sup>2)" by (simp only: power_mult[symmetric] mult.commute) then show ?thesis using m by simp qed lemmas real_sqrt_gt_0_iff [simp] = real_sqrt_less_iff [where x=0, unfolded real_sqrt_zero] lemmas real_sqrt_lt_0_iff [simp] = real_sqrt_less_iff [where y=0, unfolded real_sqrt_zero] lemmas real_sqrt_ge_0_iff [simp] = real_sqrt_le_iff [where x=0, unfolded real_sqrt_zero] lemmas real_sqrt_le_0_iff [simp] = real_sqrt_le_iff [where y=0, unfolded real_sqrt_zero] lemmas real_sqrt_eq_0_iff [simp] = real_sqrt_eq_iff [where y=0, unfolded real_sqrt_zero] lemmas real_sqrt_gt_1_iff [simp] = real_sqrt_less_iff [where x=1, unfolded real_sqrt_one] lemmas real_sqrt_lt_1_iff [simp] = real_sqrt_less_iff [where y=1, unfolded real_sqrt_one] lemmas real_sqrt_ge_1_iff [simp] = real_sqrt_le_iff [where x=1, unfolded real_sqrt_one] lemmas real_sqrt_le_1_iff [simp] = real_sqrt_le_iff [where y=1, unfolded real_sqrt_one] lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, unfolded real_sqrt_one] lemma sqrt_add_le_add_sqrt: assumes "0 \ x" "0 \ y" shows "sqrt (x + y) \ sqrt x + sqrt y" by (rule power2_le_imp_le) (simp_all add: power2_sum assms) lemma isCont_real_sqrt: "isCont sqrt x" unfolding sqrt_def by (rule isCont_real_root) lemma tendsto_real_sqrt [tendsto_intros]: "(f \ x) F \ ((\x. sqrt (f x)) \ sqrt x) F" unfolding sqrt_def by (rule tendsto_real_root) lemma continuous_real_sqrt [continuous_intros]: "continuous F f \ continuous F (\x. sqrt (f x))" unfolding sqrt_def by (rule continuous_real_root) lemma continuous_on_real_sqrt [continuous_intros]: "continuous_on s f \ continuous_on s (\x. sqrt (f x))" unfolding sqrt_def by (rule continuous_on_real_root) lemma DERIV_real_sqrt_generic: assumes "x \ 0" and "x > 0 \ D = inverse (sqrt x) / 2" and "x < 0 \ D = - inverse (sqrt x) / 2" shows "DERIV sqrt x :> D" using assms unfolding sqrt_def by (auto intro!: DERIV_real_root_generic) lemma DERIV_real_sqrt: "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" using DERIV_real_sqrt_generic by simp declare DERIV_real_sqrt_generic[THEN DERIV_chain2, derivative_intros] DERIV_real_root_generic[THEN DERIV_chain2, derivative_intros] lemmas has_derivative_real_sqrt[derivative_intros] = DERIV_real_sqrt[THEN DERIV_compose_FDERIV] lemma not_real_square_gt_zero [simp]: "\ 0 < x * x \ x = 0" for x :: real apply auto using linorder_less_linear [where x = x and y = 0] apply (simp add: zero_less_mult_iff) done lemma real_sqrt_abs2 [simp]: "sqrt (x * x) = \x\" apply (subst power2_eq_square [symmetric]) apply (rule real_sqrt_abs) done lemma real_inv_sqrt_pow2: "0 < x \ (inverse (sqrt x))\<^sup>2 = inverse x" by (simp add: power_inverse) lemma real_sqrt_eq_zero_cancel: "0 \ x \ sqrt x = 0 \ x = 0" by simp lemma real_sqrt_ge_one: "1 \ x \ 1 \ sqrt x" by simp lemma sqrt_divide_self_eq: assumes nneg: "0 \ x" shows "sqrt x / x = inverse (sqrt x)" proof (cases "x = 0") case True then show ?thesis by simp next case False then have pos: "0 < x" using nneg by arith show ?thesis proof (rule right_inverse_eq [THEN iffD1, symmetric]) show "sqrt x / x \ 0" by (simp add: divide_inverse nneg False) show "inverse (sqrt x) / (sqrt x / x) = 1" by (simp add: divide_inverse mult.assoc [symmetric] power2_eq_square [symmetric] real_inv_sqrt_pow2 pos False) qed qed lemma real_div_sqrt: "0 \ x \ x / sqrt x = sqrt x" by (cases "x = 0") (simp_all add: sqrt_divide_self_eq [of x] field_simps) lemma real_divide_square_eq [simp]: "(r * a) / (r * r) = a / r" for a r :: real by (cases "r = 0") (simp_all add: divide_inverse ac_simps) lemma lemma_real_divide_sqrt_less: "0 < u \ u / sqrt 2 < u" by (simp add: divide_less_eq) lemma four_x_squared: "4 * x\<^sup>2 = (2 * x)\<^sup>2" for x :: real by (simp add: power2_eq_square) lemma sqrt_at_top: "LIM x at_top. sqrt x :: real :> at_top" by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g="power2"]) (auto intro: eventually_gt_at_top) subsection \Square Root of Sum of Squares\ lemma sum_squares_bound: "2 * x * y \ x\<^sup>2 + y\<^sup>2" for x y :: "'a::linordered_field" proof - have "(x - y)\<^sup>2 = x * x - 2 * x * y + y * y" by algebra then have "0 \ x\<^sup>2 - 2 * x * y + y\<^sup>2" by (metis sum_power2_ge_zero zero_le_double_add_iff_zero_le_single_add power2_eq_square) then show ?thesis by arith qed lemma arith_geo_mean: fixes u :: "'a::linordered_field" assumes "u\<^sup>2 = x * y" "x \ 0" "y \ 0" shows "u \ (x + y)/2" apply (rule power2_le_imp_le) using sum_squares_bound assms apply (auto simp: zero_le_mult_iff) apply (auto simp: algebra_simps power2_eq_square) done lemma arith_geo_mean_sqrt: fixes x :: real assumes "x \ 0" "y \ 0" shows "sqrt (x * y) \ (x + y)/2" apply (rule arith_geo_mean) using assms apply (auto simp: zero_le_mult_iff) done lemma real_sqrt_sum_squares_mult_ge_zero [simp]: "0 \ sqrt ((x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2))" by (metis real_sqrt_ge_0_iff split_mult_pos_le sum_power2_ge_zero) lemma real_sqrt_sum_squares_mult_squared_eq [simp]: "(sqrt ((x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2)))\<^sup>2 = (x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2)" by (simp add: zero_le_mult_iff) lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\<^sup>2 + y\<^sup>2) = x \ y = 0" by (drule arg_cong [where f = "\x. x\<^sup>2"]) simp lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\<^sup>2 + y\<^sup>2) = y \ x = 0" by (drule arg_cong [where f = "\x. x\<^sup>2"]) simp lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt (x\<^sup>2 + y\<^sup>2)" by (rule power2_le_imp_le) simp_all lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt (x\<^sup>2 + y\<^sup>2)" by (rule power2_le_imp_le) simp_all lemma real_sqrt_ge_abs1 [simp]: "\x\ \ sqrt (x\<^sup>2 + y\<^sup>2)" by (rule power2_le_imp_le) simp_all lemma real_sqrt_ge_abs2 [simp]: "\y\ \ sqrt (x\<^sup>2 + y\<^sup>2)" by (rule power2_le_imp_le) simp_all lemma le_real_sqrt_sumsq [simp]: "x \ sqrt (x * x + y * y)" by (simp add: power2_eq_square [symmetric]) lemma sqrt_sum_squares_le_sum: "\0 \ x; 0 \ y\ \ sqrt (x\<^sup>2 + y\<^sup>2) \ x + y" by (rule power2_le_imp_le) (simp_all add: power2_sum) lemma L2_set_mult_ineq_lemma: fixes a b c d :: real shows "2 * (a * c) * (b * d) \ a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2" proof - have "0 \ (a * d - b * c)\<^sup>2" by simp also have "\ = a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2 - 2 * (a * d) * (b * c)" by (simp only: power2_diff power_mult_distrib) also have "\ = a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2 - 2 * (a * c) * (b * d)" by simp finally show "2 * (a * c) * (b * d) \ a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2" by simp qed lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<^sup>2 + y\<^sup>2) \ \x\ + \y\" by (rule power2_le_imp_le) (simp_all add: power2_sum) lemma real_sqrt_sum_squares_triangle_ineq: "sqrt ((a + c)\<^sup>2 + (b + d)\<^sup>2) \ sqrt (a\<^sup>2 + b\<^sup>2) + sqrt (c\<^sup>2 + d\<^sup>2)" proof - have "(a * c + b * d) \ (sqrt (a\<^sup>2 + b\<^sup>2) * sqrt (c\<^sup>2 + d\<^sup>2))" by (rule power2_le_imp_le) (simp_all add: power2_sum power_mult_distrib ring_distribs L2_set_mult_ineq_lemma add.commute) then have "(a + c)\<^sup>2 + (b + d)\<^sup>2 \ (sqrt (a\<^sup>2 + b\<^sup>2) + sqrt (c\<^sup>2 + d\<^sup>2))\<^sup>2" by (simp add: power2_sum) then show ?thesis by (auto intro: power2_le_imp_le) qed lemma real_sqrt_sum_squares_less: "\x\ < u / sqrt 2 \ \y\ < u / sqrt 2 \ sqrt (x\<^sup>2 + y\<^sup>2) < u" apply (rule power2_less_imp_less) apply simp apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) apply (simp add: power_divide) apply (drule order_le_less_trans [OF abs_ge_zero]) apply (simp add: zero_less_divide_iff) done lemma sqrt2_less_2: "sqrt 2 < (2::real)" by (metis not_less not_less_iff_gr_or_eq numeral_less_iff real_sqrt_four real_sqrt_le_iff semiring_norm(75) semiring_norm(78) semiring_norm(85)) lemma sqrt_sum_squares_half_less: "x < u/2 \ y < u/2 \ 0 \ x \ 0 \ y \ sqrt (x\<^sup>2 + y\<^sup>2) < u" apply (rule real_sqrt_sum_squares_less) apply (auto simp add: abs_if field_simps) apply (rule le_less_trans [where y = "x*2"]) using less_eq_real_def sqrt2_less_2 apply force apply assumption apply (rule le_less_trans [where y = "y*2"]) using less_eq_real_def sqrt2_less_2 mult_le_cancel_left apply auto done lemma LIMSEQ_root: "(\n. root n n) \ 1" proof - define x where "x n = root n n - 1" for n have "x \ sqrt 0" proof (rule tendsto_sandwich[OF _ _ tendsto_const]) show "(\x. sqrt (2 / x)) \ sqrt 0" by (intro tendsto_intros tendsto_divide_0[OF tendsto_const] filterlim_mono[OF filterlim_real_sequentially]) (simp_all add: at_infinity_eq_at_top_bot) have "x n \ sqrt (2 / real n)" if "2 < n" for n :: nat proof - have "1 + (real (n - 1) * n) / 2 * (x n)\<^sup>2 = 1 + of_nat (n choose 2) * (x n)\<^sup>2" by (auto simp add: choose_two field_char_0_class.of_nat_div mod_eq_0_iff_dvd) also have "\ \ (\k\{0, 2}. of_nat (n choose k) * x n^k)" by (simp add: x_def) also have "\ \ (\k\n. of_nat (n choose k) * x n^k)" using \2 < n\ by (intro sum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) also have "\ = (x n + 1) ^ n" by (simp add: binomial_ring) also have "\ = n" using \2 < n\ by (simp add: x_def) finally have "real (n - 1) * (real n / 2 * (x n)\<^sup>2) \ real (n - 1) * 1" by simp then have "(x n)\<^sup>2 \ 2 / real n" using \2 < n\ unfolding mult_le_cancel_left by (simp add: field_simps) from real_sqrt_le_mono[OF this] show ?thesis by simp qed then show "eventually (\n. x n \ sqrt (2 / real n)) sequentially" by (auto intro!: exI[of _ 3] simp: eventually_sequentially) show "eventually (\n. sqrt 0 \ x n) sequentially" by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def) qed from tendsto_add[OF this tendsto_const[of 1]] show ?thesis by (simp add: x_def) qed lemma LIMSEQ_root_const: assumes "0 < c" shows "(\n. root n c) \ 1" proof - have ge_1: "(\n. root n c) \ 1" if "1 \ c" for c :: real proof - define x where "x n = root n c - 1" for n have "x \ 0" proof (rule tendsto_sandwich[OF _ _ tendsto_const]) show "(\n. c / n) \ 0" by (intro tendsto_divide_0[OF tendsto_const] filterlim_mono[OF filterlim_real_sequentially]) (simp_all add: at_infinity_eq_at_top_bot) have "x n \ c / n" if "1 < n" for n :: nat proof - have "1 + x n * n = 1 + of_nat (n choose 1) * x n^1" by (simp add: choose_one) also have "\ \ (\k\{0, 1}. of_nat (n choose k) * x n^k)" by (simp add: x_def) also have "\ \ (\k\n. of_nat (n choose k) * x n^k)" using \1 < n\ \1 \ c\ by (intro sum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) also have "\ = (x n + 1) ^ n" by (simp add: binomial_ring) also have "\ = c" using \1 < n\ \1 \ c\ by (simp add: x_def) finally show ?thesis using \1 \ c\ \1 < n\ by (simp add: field_simps) qed then show "eventually (\n. x n \ c / n) sequentially" by (auto intro!: exI[of _ 3] simp: eventually_sequentially) show "eventually (\n. 0 \ x n) sequentially" using \1 \ c\ by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def) qed from tendsto_add[OF this tendsto_const[of 1]] show ?thesis by (simp add: x_def) qed show ?thesis proof (cases "1 \ c") case True with ge_1 show ?thesis by blast next case False with \0 < c\ have "1 \ 1 / c" by simp then have "(\n. 1 / root n (1 / c)) \ 1 / 1" by (intro tendsto_divide tendsto_const ge_1 \1 \ 1 / c\ one_neq_zero) then show ?thesis by (rule filterlim_cong[THEN iffD1, rotated 3]) (auto intro!: exI[of _ 1] simp: eventually_sequentially real_root_divide) qed qed text "Legacy theorem names:" lemmas real_root_pos2 = real_root_power_cancel lemmas real_root_pos_pos = real_root_gt_zero [THEN order_less_imp_le] lemmas real_root_pos_pos_le = real_root_ge_zero lemmas real_sqrt_eq_zero_cancel_iff = real_sqrt_eq_0_iff end