diff --git a/src/HOL/Analysis/Abstract_Topological_Spaces.thy b/src/HOL/Analysis/Abstract_Topological_Spaces.thy --- a/src/HOL/Analysis/Abstract_Topological_Spaces.thy +++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy @@ -1,3468 +1,4437 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Various Forms of Topological Spaces\ theory Abstract_Topological_Spaces imports Lindelof_Spaces Locally Sum_Topology FSigma begin subsection\Connected topological spaces\ lemma connected_space_eq_frontier_eq_empty: "connected_space X \ (\S. S \ topspace X \ X frontier_of S = {} \ S = {} \ S = topspace X)" by (meson clopenin_eq_frontier_of connected_space_clopen_in) lemma connected_space_frontier_eq_empty: "connected_space X \ S \ topspace X \ (X frontier_of S = {} \ S = {} \ S = topspace X)" by (meson connected_space_eq_frontier_eq_empty frontier_of_empty frontier_of_topspace) lemma connectedin_eq_subset_separated_union: "connectedin X C \ C \ topspace X \ (\S T. separatedin X S T \ C \ S \ T \ C \ S \ C \ T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using connectedin_subset_topspace connectedin_subset_separated_union by blast next assume ?rhs then show ?lhs by (metis closure_of_subset connectedin_separation dual_order.eq_iff inf.orderE separatedin_def sup.boundedE) qed lemma connectedin_clopen_cases: "\connectedin X C; closedin X T; openin X T\ \ C \ T \ disjnt C T" by (metis Diff_eq_empty_iff Int_empty_right clopenin_eq_frontier_of connectedin_Int_frontier_of disjnt_def) 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 connected_space_retraction_map_image: "\retraction_map X X' r; connected_space X\ \ connected_space X'" using connected_space_quotient_map_image retraction_imp_quotient_map by blast lemma connectedin_imp_perfect_gen: assumes X: "t1_space X" and S: "connectedin X S" and nontriv: "\a. S = {a}" shows "S \ X derived_set_of S" unfolding derived_set_of_def proof (intro subsetI CollectI conjI strip) show XS: "x \ topspace X" if "x \ S" for x using that S connectedin by fastforce show "\y. y \ x \ y \ S \ y \ T" if "x \ S" and "x \ T \ openin X T" for x T proof - have opeXx: "openin X (topspace X - {x})" by (meson X openin_topspace t1_space_openin_delete_alt) moreover have "S \ T \ (topspace X - {x})" using XS that(2) by auto moreover have "(topspace X - {x}) \ S \ {}" by (metis Diff_triv S connectedin double_diff empty_subsetI inf_commute insert_subsetI nontriv that(1)) ultimately show ?thesis using that connectedinD [OF S, of T "topspace X - {x}"] by blast qed qed lemma connectedin_imp_perfect: "\Hausdorff_space X; connectedin X S; \a. S = {a}\ \ S \ X derived_set_of S" by (simp add: Hausdorff_imp_t1_space connectedin_imp_perfect_gen) subsection\The notion of "separated between" (complement of "connected between)"\ definition separated_between where "separated_between X S T \ \U V. openin X U \ openin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V" lemma separated_between_alt: "separated_between X S T \ (\U V. closedin X U \ closedin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V)" unfolding separated_between_def by (metis separatedin_open_sets separation_closedin_Un_gen subtopology_topspace separatedin_closed_sets separation_openin_Un_gen) lemma separated_between: "separated_between X S T \ (\U. closedin X U \ openin X U \ S \ U \ T \ topspace X - U)" unfolding separated_between_def closedin_def disjnt_def by (smt (verit, del_insts) Diff_cancel Diff_disjoint Diff_partition Un_Diff Un_Diff_Int openin_subset) lemma separated_between_mono: "\separated_between X S T; S' \ S; T' \ T\ \ separated_between X S' T'" by (meson order.trans separated_between) lemma separated_between_refl: "separated_between X S S \ S = {}" unfolding separated_between_def by (metis Un_empty_right disjnt_def disjnt_empty2 disjnt_subset2 disjnt_sym le_iff_inf openin_empty openin_topspace) lemma separated_between_sym: "separated_between X S T \ separated_between X T S" by (metis disjnt_sym separated_between_alt sup_commute) lemma separated_between_imp_subset: "separated_between X S T \ S \ topspace X \ T \ topspace X" by (metis le_supI1 le_supI2 separated_between_def) lemma separated_between_empty: "(separated_between X {} S \ S \ topspace X) \ (separated_between X S {} \ S \ topspace X)" by (metis Diff_empty bot.extremum closedin_empty openin_empty separated_between separated_between_imp_subset separated_between_sym) lemma separated_between_Un: "separated_between X S (T \ U) \ separated_between X S T \ separated_between X S U" by (auto simp: separated_between) lemma separated_between_Un': "separated_between X (S \ T) U \ separated_between X S U \ separated_between X T U" by (simp add: separated_between_Un separated_between_sym) lemma separated_between_imp_disjoint: "separated_between X S T \ disjnt S T" by (meson disjnt_iff separated_between_def subsetD) lemma separated_between_imp_separatedin: "separated_between X S T \ separatedin X S T" by (meson separated_between_def separatedin_mono separatedin_open_sets) lemma separated_between_full: assumes "S \ T = topspace X" shows "separated_between X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" proof - have "separated_between X S T \ separatedin X S T" by (simp add: separated_between_imp_separatedin) then show ?thesis unfolding separated_between_def by (metis assms separation_closedin_Un_gen separation_openin_Un_gen subset_refl subtopology_topspace) qed lemma separated_between_eq_separatedin: "S \ T = topspace X \ (separated_between X S T \ separatedin X S T)" by (simp add: separated_between_full separatedin_full) lemma separated_between_pointwise_left: assumes "compactin X S" shows "separated_between X S T \ (S = {} \ T \ topspace X) \ (\x \ S. separated_between X {x} T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using separated_between_imp_subset separated_between_mono by fastforce next assume R: ?rhs then have "T \ topspace X" by (meson equals0I separated_between_imp_subset) show ?lhs proof - obtain U where U: "\x \ S. openin X (U x)" "\x \ S. \V. openin X V \ U x \ V = topspace X \ disjnt (U x) V \ {x} \ U x \ T \ V" using R unfolding separated_between_def by metis then have "S \ \(U ` S)" by blast then obtain K where "finite K" "K \ S" and K: "S \ (\i\K. U i)" using assms U unfolding compactin_def by (smt (verit) finite_subset_image imageE) show ?thesis unfolding separated_between proof (intro conjI exI) have "\x. x \ K \ closedin X (U x)" by (smt (verit) \K \ S\ Diff_cancel U(2) Un_Diff Un_Diff_Int disjnt_def openin_closedin_eq subsetD) then show "closedin X (\ (U ` K))" by (metis (mono_tags, lifting) \finite K\ closedin_Union finite_imageI image_iff) show "openin X (\ (U ` K))" using U(1) \K \ S\ by blast show "S \ \ (U ` K)" by (simp add: K) have "\x i. \x \ T; i \ K; x \ U i\ \ False" by (meson U(2) \K \ S\ disjnt_iff subsetD) then show "T \ topspace X - \ (U ` K)" using \T \ topspace X\ by auto qed qed qed lemma separated_between_pointwise_right: "compactin X T \ separated_between X S T \ (T = {} \ S \ topspace X) \ (\y \ T. separated_between X S {y})" by (meson separated_between_pointwise_left separated_between_sym) lemma separated_between_closure_of: "S \ topspace X \ separated_between X (X closure_of S) T \ separated_between X S T" by (meson closure_of_minimal_eq separated_between_alt) lemma separated_between_closure_of': "T \ topspace X \ separated_between X S (X closure_of T) \ separated_between X S T" by (meson separated_between_closure_of separated_between_sym) lemma separated_between_closure_of_eq: "separated_between X S T \ S \ topspace X \ separated_between X (X closure_of S) T" by (metis separated_between_closure_of separated_between_imp_subset) lemma separated_between_closure_of_eq': "separated_between X S T \ T \ topspace X \ separated_between X S (X closure_of T)" by (metis separated_between_closure_of' separated_between_imp_subset) lemma separated_between_frontier_of_eq': "separated_between X S T \ T \ topspace X \ disjnt S T \ separated_between X S (X frontier_of T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis interior_of_union_frontier_of separated_between_Un separated_between_closure_of_eq' separated_between_imp_disjoint) next assume R: ?rhs then obtain U where U: "closedin X U" "openin X U" "S \ U" "X closure_of T - X interior_of T \ topspace X - U" by (metis frontier_of_def separated_between) show ?lhs proof (rule separated_between_mono [of _ S "X closure_of T"]) have "separated_between X S T" unfolding separated_between proof (intro conjI exI) show "S \ U - T" "T \ topspace X - (U - T)" using R U(3) by (force simp: disjnt_iff)+ have "T \ X closure_of T" by (simp add: R closure_of_subset) then have *: "U - T = U - X interior_of T" using U(4) interior_of_subset by fastforce then show "closedin X (U - T)" by (simp add: U(1) closedin_diff) have "U \ X frontier_of T = {}" using U(4) frontier_of_def by fastforce then show "openin X (U - T)" by (metis * Diff_Un U(2) Un_Diff_Int Un_Int_eq(1) closedin_closure_of interior_of_union_frontier_of openin_diff sup_bot_right) qed then show "separated_between X S (X closure_of T)" by (simp add: R separated_between_closure_of') qed (auto simp: R closure_of_subset) qed lemma separated_between_frontier_of_eq: "separated_between X S T \ S \ topspace X \ disjnt S T \ separated_between X (X frontier_of S) T" by (metis disjnt_sym separated_between_frontier_of_eq' separated_between_sym) lemma separated_between_frontier_of: "\S \ topspace X; disjnt S T\ \ (separated_between X (X frontier_of S) T \ separated_between X S T)" using separated_between_frontier_of_eq by blast lemma separated_between_frontier_of': "\T \ topspace X; disjnt S T\ \ (separated_between X S (X frontier_of T) \ separated_between X S T)" using separated_between_frontier_of_eq' by auto lemma connected_space_separated_between: "connected_space X \ (\S T. separated_between X S T \ S = {} \ T = {})" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Diff_cancel connected_space_clopen_in separated_between subset_empty) next assume ?rhs then show ?lhs by (meson connected_space_eq_not_separated separated_between_eq_separatedin) qed lemma connected_space_imp_separated_between_trivial: "connected_space X \ (separated_between X S T \ S = {} \ T \ topspace X \ S \ topspace X \ T = {})" by (metis connected_space_separated_between separated_between_empty) subsection\Connected components\ lemma connected_component_of_subtopology_eq: "connected_component_of (subtopology X U) a = connected_component_of X a \ connected_component_of_set X a \ U" by (force simp: connected_component_of_set connectedin_subtopology connected_component_of_def fun_eq_iff subset_iff) lemma connected_components_of_subtopology: assumes "C \ connected_components_of X" "C \ U" shows "C \ connected_components_of (subtopology X U)" proof - obtain a where a: "connected_component_of_set X a \ U" and "a \ topspace X" and Ceq: "C = connected_component_of_set X a" using assms by (force simp: connected_components_of_def) then have "a \ U" by (simp add: connected_component_of_refl in_mono) then have "connected_component_of_set X a = connected_component_of_set (subtopology X U) a" by (metis a connected_component_of_subtopology_eq) then show ?thesis by (simp add: Ceq \a \ U\ \a \ topspace X\ connected_component_in_connected_components_of) qed thm connected_space_iff_components_eq lemma open_in_finite_connected_components: assumes "finite(connected_components_of X)" "C \ connected_components_of X" shows "openin X C" proof - have "closedin X (topspace X - C)" by (metis DiffD1 assms closedin_Union closedin_connected_components_of complement_connected_components_of_Union finite_Diff) then show ?thesis by (simp add: assms connected_components_of_subset openin_closedin) qed thm connected_component_of_eq_overlap lemma connected_components_of_disjoint: assumes "C \ connected_components_of X" "C' \ connected_components_of X" shows "(disjnt C C' \ (C \ C'))" proof - have "C \ {}" using nonempty_connected_components_of assms by blast with assms show ?thesis by (metis disjnt_self_iff_empty pairwiseD pairwise_disjoint_connected_components_of) qed lemma connected_components_of_overlap: "\C \ connected_components_of X; C' \ connected_components_of X\ \ C \ C' \ {} \ C = C'" by (meson connected_components_of_disjoint disjnt_def) lemma pairwise_separated_connected_components_of: "pairwise (separatedin X) (connected_components_of X)" by (simp add: closedin_connected_components_of connected_components_of_disjoint pairwiseI separatedin_closed_sets) lemma finite_connected_components_of_finite: "finite(topspace X) \ finite(connected_components_of X)" by (simp add: Union_connected_components_of finite_UnionD) lemma connected_component_of_unique: "\x \ C; connectedin X C; \C'. x \ C' \ connectedin X C' \ C' \ C\ \ connected_component_of_set X x = C" by (meson connected_component_of_maximal connectedin_connected_component_of subsetD subset_antisym) lemma closedin_connected_component_of_subtopology: "\C \ connected_components_of (subtopology X s); X closure_of C \ s\ \ closedin X C" by (metis closedin_Int_closure_of closedin_connected_components_of closure_of_eq inf.absorb_iff2) lemma connected_component_of_discrete_topology: "connected_component_of_set (discrete_topology U) x = (if x \ U then {x} else {})" by (simp add: locally_path_connected_space_discrete_topology flip: path_component_eq_connected_component_of) lemma connected_components_of_discrete_topology: "connected_components_of (discrete_topology U) = (\x. {x}) ` U" by (simp add: connected_component_of_discrete_topology connected_components_of_def) lemma connected_component_of_continuous_image: "\continuous_map X Y f; connected_component_of X x y\ \ connected_component_of Y (f x) (f y)" by (meson connected_component_of_def connectedin_continuous_map_image image_eqI) lemma homeomorphic_map_connected_component_of: assumes "homeomorphic_map X Y f" and x: "x \ topspace X" shows "connected_component_of_set Y (f x) = f ` (connected_component_of_set X x)" proof - obtain g where g: "continuous_map X Y f" "continuous_map Y X g " "\x. x \ topspace X \ g (f x) = x" "\y. y \ topspace Y \ f (g y) = y" using assms(1) homeomorphic_map_maps homeomorphic_maps_def by fastforce show ?thesis using connected_component_in_topspace [of Y] x g connected_component_of_continuous_image [of X Y f] connected_component_of_continuous_image [of Y X g] by force qed lemma homeomorphic_map_connected_components_of: assumes "homeomorphic_map X Y f" shows "connected_components_of Y = (image f) ` (connected_components_of X)" proof - have "topspace Y = f ` topspace X" by (metis assms homeomorphic_imp_surjective_map) with homeomorphic_map_connected_component_of [OF assms] show ?thesis by (auto simp: connected_components_of_def image_iff) qed lemma connected_component_of_pair: "connected_component_of_set (prod_topology X Y) (x,y) = connected_component_of_set X x \ connected_component_of_set Y y" proof (cases "x \ topspace X \ y \ topspace Y") case True show ?thesis proof (rule connected_component_of_unique) show "(x, y) \ connected_component_of_set X x \ connected_component_of_set Y y" using True by (simp add: connected_component_of_refl) show "connectedin (prod_topology X Y) (connected_component_of_set X x \ connected_component_of_set Y y)" by (metis connectedin_Times connectedin_connected_component_of) show "C \ connected_component_of_set X x \ connected_component_of_set Y y" if "(x, y) \ C \ connectedin (prod_topology X Y) C" for C using that unfolding connected_component_of_def apply clarsimp by (metis (no_types) connectedin_continuous_map_image continuous_map_fst continuous_map_snd fst_conv imageI snd_conv) qed next case False then show ?thesis by (metis Sigma_empty1 Sigma_empty2 connected_component_of_eq_empty mem_Sigma_iff topspace_prod_topology) qed lemma connected_components_of_prod_topology: "connected_components_of (prod_topology X Y) = {C \ D |C D. C \ connected_components_of X \ D \ connected_components_of Y}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" apply (clarsimp simp: connected_components_of_def) by (metis (no_types) connected_component_of_pair imageI) next show "?rhs \ ?lhs" using connected_component_of_pair by (fastforce simp: connected_components_of_def) qed lemma connected_component_of_product_topology: "connected_component_of_set (product_topology X I) x = (if x \ extensional I then PiE I (\i. connected_component_of_set (X i) (x i)) else {})" (is "?lhs = If _ ?R _") proof (cases "x \ topspace(product_topology X I)") case True have "?lhs = (\\<^sub>E i\I. connected_component_of_set (X i) (x i))" if xX: "\i. i\I \ x i \ topspace (X i)" and ext: "x \ extensional I" proof (rule connected_component_of_unique) show "x \ ?R" by (simp add: PiE_iff connected_component_of_refl local.ext xX) show "connectedin (product_topology X I) ?R" by (simp add: connectedin_PiE connectedin_connected_component_of) show "C \ ?R" if "x \ C \ connectedin (product_topology X I) C" for C proof - have "C \ extensional I" using PiE_def connectedin_subset_topspace that by fastforce have "\y. y \ C \ y \ (\ i\I. connected_component_of_set (X i) (x i))" apply (simp add: connected_component_of_def Pi_def) by (metis connectedin_continuous_map_image continuous_map_product_projection imageI that) then show ?thesis using PiE_def \C \ extensional I\ by fastforce qed qed with True show ?thesis by (simp add: PiE_iff) next case False then show ?thesis apply (simp add: PiE_iff) by (smt (verit) Collect_empty_eq False PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty) qed lemma connected_components_of_product_topology: "connected_components_of (product_topology X I) = {PiE I B |B. \i \ I. B i \ connected_components_of(X i)}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (auto simp: connected_components_of_def connected_component_of_product_topology PiE_iff) show "?rhs \ ?lhs" proof fix F assume "F \ ?rhs" then obtain B where Feq: "F = Pi\<^sub>E I B" and "\i\I. \x\topspace (X i). B i = connected_component_of_set (X i) x" by (force simp: connected_components_of_def connected_component_of_product_topology image_iff) then obtain f where f: "\i. i \ I \ f i \ topspace (X i) \ B i = connected_component_of_set (X i) (f i)" by metis then have "(\i\I. f i) \ ((\\<^sub>E i\I. topspace (X i)) \ extensional I)" by simp with f show "F \ ?lhs" unfolding Feq connected_components_of_def connected_component_of_product_topology image_iff by (smt (verit, del_insts) PiE_cong restrict_PiE_iff restrict_apply' restrict_extensional topspace_product_topology) qed qed subsection \Monotone maps (in the general topological sense)\ definition monotone_map where "monotone_map X Y f == f ` (topspace X) \ topspace Y \ (\y \ topspace Y. connectedin X {x \ topspace X. f x = y})" lemma monotone_map: "monotone_map X Y f \ f ` (topspace X) \ topspace Y \ (\y. connectedin X {x \ topspace X. f x = y})" apply (simp add: monotone_map_def) by (metis (mono_tags, lifting) connectedin_empty [of X] Collect_empty_eq image_subset_iff) lemma monotone_map_in_subtopology: "monotone_map X (subtopology Y S) f \ monotone_map X Y f \ f ` (topspace X) \ S" by (smt (verit, del_insts) le_inf_iff monotone_map topspace_subtopology) lemma monotone_map_from_subtopology: assumes "monotone_map X Y f" "\x y. x \ topspace X \ y \ topspace X \ x \ S \ f x = f y \ y \ S" shows "monotone_map (subtopology X S) Y f" using assms unfolding monotone_map_def connectedin_subtopology by (smt (verit, del_insts) Collect_cong Collect_empty_eq IntE IntI connectedin_empty image_subset_iff mem_Collect_eq subsetI topspace_subtopology) lemma monotone_map_restriction: "monotone_map X Y f \ {x \ topspace X. f x \ v} = u \ monotone_map (subtopology X u) (subtopology Y v) f" by (smt (verit, best) IntI Int_Collect image_subset_iff mem_Collect_eq monotone_map monotone_map_from_subtopology topspace_subtopology) lemma injective_imp_monotone_map: assumes "f ` topspace X \ topspace Y" "inj_on f (topspace X)" shows "monotone_map X Y f" unfolding monotone_map_def proof (intro conjI assms strip) fix y assume "y \ topspace Y" then have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using assms(2) unfolding inj_on_def by blast then show "connectedin X {x \ topspace X. f x = y}" by (metis (no_types, lifting) connectedin_empty connectedin_sing) qed lemma embedding_imp_monotone_map: "embedding_map X Y f \ monotone_map X Y f" by (metis (no_types) embedding_map_def homeomorphic_eq_everything_map inf.absorb_iff2 injective_imp_monotone_map topspace_subtopology) lemma section_imp_monotone_map: "section_map X Y f \ monotone_map X Y f" by (simp add: embedding_imp_monotone_map section_imp_embedding_map) lemma homeomorphic_imp_monotone_map: "homeomorphic_map X Y f \ monotone_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_monotone_map) proposition connected_space_monotone_quotient_map_preimage: assumes f: "monotone_map X Y f" "quotient_map X Y f" and "connected_space Y" shows "connected_space X" proof (rule ccontr) assume "\ connected_space X" then obtain U V where "openin X U" "openin X V" "U \ V = {}" "U \ {}" "V \ {}" and topUV: "topspace X \ U \ V" by (auto simp: connected_space_def) then have UVsub: "U \ topspace X" "V \ topspace X" by (auto simp: openin_subset) have "\ connected_space Y" unfolding connected_space_def not_not proof (intro exI conjI) show "topspace Y \ f`U \ f`V" by (metis f(2) image_Un quotient_imp_surjective_map subset_Un_eq topUV) show "f`U \ {}" by (simp add: \U \ {}\) show "(f`V) \ {}" by (simp add: \V \ {}\) have *: "y \ f ` V" if "y \ f ` U" for y proof - have \
: "connectedin X {x \ topspace X. f x = y}" using f(1) monotone_map by fastforce show ?thesis using connectedinD [OF \
\openin X U\ \openin X V\] UVsub topUV \U \ V = {}\ that by (force simp: disjoint_iff) qed then show "f`U \ f`V = {}" by blast show "openin Y (f`U)" using f \openin X U\ topUV * unfolding quotient_map_saturated_open by force show "openin Y (f`V)" using f \openin X V\ topUV * unfolding quotient_map_saturated_open by force qed then show False by (simp add: assms) qed lemma connectedin_monotone_quotient_map_preimage: assumes "monotone_map X Y f" "quotient_map X Y f" "connectedin Y C" "openin Y C \ closedin Y C" shows "connectedin X {x \ topspace X. f x \ C}" proof - have "connected_space (subtopology X {x \ topspace X. f x \ C})" proof - have "connected_space (subtopology Y C)" using \connectedin Y C\ connectedin_def by blast moreover have "quotient_map (subtopology X {a \ topspace X. f a \ C}) (subtopology Y C) f" by (simp add: assms quotient_map_restriction) ultimately show ?thesis using \monotone_map X Y f\ connected_space_monotone_quotient_map_preimage monotone_map_restriction by blast qed then show ?thesis by (simp add: connectedin_def) qed lemma monotone_open_map: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding connectedin_def proof (intro strip conjI) fix C assume C: "C \ topspace Y \ connected_space (subtopology Y C)" show "connected_space (subtopology X {x \ topspace X. f x \ C})" proof (rule connected_space_monotone_quotient_map_preimage) show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" by (simp add: L monotone_map_restriction) show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" proof (rule continuous_open_imp_quotient_map) show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce qed (use open_map_restriction assms in fastforce)+ qed (simp add: C) qed auto next assume ?rhs then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" by (smt (verit) Collect_cong singletonD singletonI) then show ?lhs by (simp add: fim monotone_map_def) qed lemma monotone_closed_map: assumes "continuous_map X Y f" "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding connectedin_def proof (intro strip conjI) fix C assume C: "C \ topspace Y \ connected_space (subtopology Y C)" show "connected_space (subtopology X {x \ topspace X. f x \ C})" proof (rule connected_space_monotone_quotient_map_preimage) show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" by (simp add: L monotone_map_restriction) show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" proof (rule continuous_closed_imp_quotient_map) show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce qed (use closed_map_restriction assms in fastforce)+ qed (simp add: C) qed auto next assume ?rhs then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" by (smt (verit) Collect_cong singletonD singletonI) then show ?lhs by (simp add: fim monotone_map_def) qed subsection\Other countability properties\ definition second_countable where "second_countable X \ \\. countable \ \ (\V \ \. openin X V) \ (\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" definition first_countable where "first_countable X \ \x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ (\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" definition separable_space where "separable_space X \ \C. countable C \ C \ topspace X \ X closure_of C = topspace X" lemma second_countable: "second_countable X \ (\\. countable \ \ openin X = arbitrary union_of (\x. x \ \))" by (smt (verit) openin_topology_base_unique second_countable_def) lemma second_countable_subtopology: assumes "second_countable X" shows "second_countable (subtopology X S)" proof - obtain \ where \: "countable \" "\V. V \ \ \ openin X V" "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) show ?thesis unfolding second_countable_def proof (intro exI conjI) show "\V\((\)S) ` \. openin (subtopology X S) V" using openin_subtopology_Int2 \ by blast show "\U x. openin (subtopology X S) U \ x \ U \ (\V\((\)S) ` \. x \ V \ V \ U)" using \ unfolding openin_subtopology by (smt (verit, del_insts) IntI image_iff inf_commute inf_le1 subset_iff) qed (use \ in auto) qed lemma second_countable_discrete_topology: "second_countable(discrete_topology U) \ countable U" (is "?lhs=?rhs") proof assume L: ?lhs then obtain \ where \: "countable \" "\V. V \ \ \ V \ U" "\W x. W \ U \ x \ W \ (\V \ \. x \ V \ V \ W)" by (auto simp: second_countable_def) then have "{x} \ \" if "x \ U" for x by (metis empty_subsetI insertCI insert_subset subset_antisym that) then show ?rhs by (smt (verit) countable_subset image_subsetI \countable \\ countable_image_inj_on [OF _ inj_singleton]) next assume ?rhs then show ?lhs unfolding second_countable_def by (rule_tac x="(\x. {x}) ` U" in exI) auto qed lemma second_countable_open_map_image: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "second_countable X" shows "second_countable Y" proof - have openXYf: "\U. openin X U \ openin Y (f ` U)" using assms by (auto simp: open_map_def) obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) show ?thesis unfolding second_countable_def proof (intro exI conjI strip) fix V y assume V: "openin Y V \ y \ V" then obtain x where "x \ topspace X" and x: "f x = y" by (metis fim image_iff openin_subset subsetD) then obtain W where "W\\" "x \ W" "W \ {x \ topspace X. f x \ V}" using * [of "{x \ topspace X. f x \ V}" x] V assms openin_continuous_map_preimage by force then show "\W \ (image f) ` \. y \ W \ W \ V" using x by auto qed (use \ openXYf in auto) qed lemma homeomorphic_space_second_countability: "X homeomorphic_space Y \ (second_countable X \ second_countable Y)" by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym second_countable_open_map_image) lemma second_countable_retraction_map_image: "\retraction_map X Y r; second_countable X\ \ second_countable Y" using hereditary_imp_retractive_property homeomorphic_space_second_countability second_countable_subtopology by blast lemma second_countable_imp_first_countable: "second_countable X \ first_countable X" by (metis first_countable_def second_countable_def) lemma first_countable_subtopology: assumes "first_countable X" shows "first_countable (subtopology X S)" unfolding first_countable_def proof fix x assume "x \ topspace (subtopology X S)" then obtain \ where "countable \" and \: "\V. V \ \ \ openin X V" "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms first_countable_def by force show "\\. countable \ \ (\V\\. openin (subtopology X S) V) \ (\U. openin (subtopology X S) U \ x \ U \ (\V\\. x \ V \ V \ U))" proof (intro exI conjI strip) show "countable (((\)S) ` \)" using \countable \\ by blast show "openin (subtopology X S) V" if "V \ ((\)S) ` \" for V using \ openin_subtopology_Int2 that by fastforce show "\V\((\)S) ` \. x \ V \ V \ U" if "openin (subtopology X S) U \ x \ U" for U using that \(2) by (clarsimp simp: openin_subtopology) (meson le_infI2) qed qed lemma first_countable_discrete_topology: "first_countable (discrete_topology U)" unfolding first_countable_def topspace_discrete_topology openin_discrete_topology proof fix x assume "x \ U" show "\\. countable \ \ (\V\\. V \ U) \ (\Ua. Ua \ U \ x \ Ua \ (\V\\. x \ V \ V \ Ua))" using \x \ U\ by (rule_tac x="{{x}}" in exI) auto qed lemma first_countable_open_map_image: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "first_countable X" shows "first_countable Y" unfolding first_countable_def proof fix y assume "y \ topspace Y" have openXYf: "\U. openin X U \ openin Y (f ` U)" using assms by (auto simp: open_map_def) then obtain x where x: "x \ topspace X" "f x = y" by (metis \y \ topspace Y\ fim imageE) obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms x first_countable_def by force show "\\. countable \ \ (\V\\. openin Y V) \ (\U. openin Y U \ y \ U \ (\V\\. y \ V \ V \ U))" proof (intro exI conjI strip) fix V assume "openin Y V \ y \ V" then have "\W\\. x \ W \ W \ {x \ topspace X. f x \ V}" using * [of "{x \ topspace X. f x \ V}"] assms openin_continuous_map_preimage x by fastforce then show "\V' \ (image f) ` \. y \ V' \ V' \ V" using image_mono x by auto qed (use \ openXYf in force)+ qed lemma homeomorphic_space_first_countability: "X homeomorphic_space Y \ first_countable X \ first_countable Y" by (meson first_countable_open_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) lemma first_countable_retraction_map_image: "\retraction_map X Y r; first_countable X\ \ first_countable Y" using first_countable_subtopology hereditary_imp_retractive_property homeomorphic_space_first_countability by blast lemma separable_space_open_subset: assumes "separable_space X" "openin X S" shows "separable_space (subtopology X S)" proof - obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" by (meson assms separable_space_def) then have "\x T. \x \ topspace X; x \ T; openin (subtopology X S) T\ \ \y. y \ S \ y \ C \ y \ T" by (smt (verit) \openin X S\ in_closure_of openin_open_subtopology subsetD) with C \openin X S\ show ?thesis unfolding separable_space_def by (rule_tac x="S \ C" in exI) (force simp: in_closure_of) qed lemma separable_space_continuous_map_image: assumes "separable_space X" "continuous_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "separable_space Y" proof - have cont: "\S. f ` (X closure_of S) \ Y closure_of f ` S" by (simp add: assms continuous_map_image_closure_subset) obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" by (meson assms separable_space_def) then show ?thesis unfolding separable_space_def by (metis cont fim closure_of_subset_topspace countable_image image_mono subset_antisym) qed lemma separable_space_quotient_map_image: "\quotient_map X Y q; separable_space X\ \ separable_space Y" by (meson quotient_imp_continuous_map quotient_imp_surjective_map separable_space_continuous_map_image) lemma separable_space_retraction_map_image: "\retraction_map X Y r; separable_space X\ \ separable_space Y" using retraction_imp_quotient_map separable_space_quotient_map_image by blast lemma homeomorphic_separable_space: "X homeomorphic_space Y \ (separable_space X \ separable_space Y)" by (meson homeomorphic_eq_everything_map homeomorphic_maps_map homeomorphic_space_def separable_space_continuous_map_image) lemma separable_space_discrete_topology: "separable_space(discrete_topology U) \ countable U" by (metis countable_Int2 discrete_topology_closure_of dual_order.refl inf.orderE separable_space_def topspace_discrete_topology) lemma second_countable_imp_separable_space: assumes "second_countable X" shows "separable_space X" proof - obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) obtain c where c: "\V. \V \ \; V \ {}\ \ c V \ V" by (metis all_not_in_conv) then have **: "\x. x \ topspace X \ x \ X closure_of c ` (\ - {{}})" using * by (force simp: closure_of_def) show ?thesis unfolding separable_space_def proof (intro exI conjI) show "countable (c ` (\-{{}}))" using \(1) by blast show "(c ` (\-{{}})) \ topspace X" using \(2) c openin_subset by fastforce show "X closure_of (c ` (\-{{}})) = topspace X" by (meson ** closure_of_subset_topspace subsetI subset_antisym) qed qed lemma second_countable_imp_Lindelof_space: assumes "second_countable X" shows "Lindelof_space X" unfolding Lindelof_space_def proof clarify fix \ assume "\U \ \. openin X U" and UU: "\\ = topspace X" obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) define \' where "\' = {B \ \. \U. U \ \ \ B \ U}" have \': "countable \'" "\\' = \\" using \ using "*" \\U\\. openin X U\ by (fastforce simp: \'_def)+ have "\b. \U. b \ \' \ U \ \ \ b \ U" by (simp add: \'_def) then obtain G where G: "\b. b \ \' \ G b \ \ \ b \ G b" by metis with \' UU show "\\. countable \ \ \ \ \ \ \\ = topspace X" by (rule_tac x="G ` \'" in exI) fastforce qed subsection \Neigbourhood bases EXTRAS\ (* Neigbourhood bases (useful for "local" properties of various kind). *) lemma openin_topology_neighbourhood_base_unique: "openin X = arbitrary union_of P \ (\u. P u \ openin X u) \ neighbourhood_base_of P X" by (smt (verit, best) open_neighbourhood_base_of openin_topology_base_unique) lemma neighbourhood_base_at_topology_base: " openin X = arbitrary union_of b \ (neighbourhood_base_at x P X \ (\w. b w \ x \ w \ (\u v. openin X u \ P v \ x \ u \ u \ v \ v \ w)))" apply (simp add: neighbourhood_base_at_def) by (smt (verit, del_insts) openin_topology_base_unique subset_trans) lemma neighbourhood_base_of_unlocalized: assumes "\S t. P S \ openin X t \ (t \ {}) \ t \ S \ P t" shows "neighbourhood_base_of P X \ (\x \ topspace X. \u v. openin X u \ P v \ x \ u \ u \ v \ v \ topspace X)" apply (simp add: neighbourhood_base_of_def) by (smt (verit, ccfv_SIG) assms empty_iff neighbourhood_base_at_unlocalized) lemma neighbourhood_base_at_discrete_topology: "neighbourhood_base_at x P (discrete_topology u) \ x \ u \ P {x}" apply (simp add: neighbourhood_base_at_def) by (smt (verit) empty_iff empty_subsetI insert_subset singletonI subsetD subset_singletonD) lemma neighbourhood_base_of_discrete_topology: "neighbourhood_base_of P (discrete_topology u) \ (\x \ u. P {x})" apply (simp add: neighbourhood_base_of_def) using neighbourhood_base_at_discrete_topology[of _ P u] by (metis empty_subsetI insert_subset neighbourhood_base_at_def openin_discrete_topology singletonI) lemma second_countable_neighbourhood_base_alt: "second_countable X \ (\\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_of (\A. A\\) X)" by (metis (full_types) openin_topology_neighbourhood_base_unique second_countable) lemma first_countable_neighbourhood_base_alt: "first_countable X \ (\x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_at x (\V. V \ \) X)" unfolding first_countable_def apply (intro ball_cong refl ex_cong conj_cong) by (metis (mono_tags, lifting) open_neighbourhood_base_at) lemma second_countable_neighbourhood_base: "second_countable X \ (\\. countable \ \ neighbourhood_base_of (\V. V \ \) X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using second_countable_neighbourhood_base_alt by blast next assume ?rhs then obtain \ where "countable \" and \: "\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. V \ \ \ x \ U \ U \ V \ V \ W))" by (metis neighbourhood_base_of) then show ?lhs unfolding second_countable_neighbourhood_base_alt neighbourhood_base_of apply (rule_tac x="(\u. X interior_of u) ` \" in exI) by (smt (verit, best) interior_of_eq interior_of_mono countable_image image_iff openin_interior_of) qed lemma first_countable_neighbourhood_base: "first_countable X \ (\x \ topspace X. \\. countable \ \ neighbourhood_base_at x (\V. V \ \) X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis first_countable_neighbourhood_base_alt) next assume R: ?rhs show ?lhs unfolding first_countable_neighbourhood_base_alt proof fix x assume "x \ topspace X" with R obtain \ where "countable \" and \: "neighbourhood_base_at x (\V. V \ \) X" by blast then show "\\. countable \ \ Ball \ (openin X) \ neighbourhood_base_at x (\V. V \ \) X" unfolding neighbourhood_base_at_def apply (rule_tac x="(\u. X interior_of u) ` \" in exI) by (smt (verit, best) countable_image image_iff interior_of_eq interior_of_mono openin_interior_of) qed qed subsection\$T_0$ spaces and the Kolmogorov quotient\ definition t0_space where "t0_space X \ \x \ topspace X. \y \ topspace X. x \ y \ (\U. openin X U \ (x \ U \ y \ U))" lemma t0_space_expansive: "\topspace Y = topspace X; \U. openin X U \ openin Y U\ \ t0_space X \ t0_space Y" by (metis t0_space_def) lemma t1_imp_t0_space: "t1_space X \ t0_space X" by (metis t0_space_def t1_space_def) lemma t1_eq_symmetric_t0_space_alt: "t1_space X \ t0_space X \ (\x \ topspace X. \y \ topspace X. x \ X closure_of {y} \ y \ X closure_of {x})" apply (simp add: t0_space_def t1_space_def closure_of_def) by (smt (verit, best) openin_topspace) lemma t1_eq_symmetric_t0_space: "t1_space X \ t0_space X \ (\x y. x \ X closure_of {y} \ y \ X closure_of {x})" by (auto simp: t1_eq_symmetric_t0_space_alt in_closure_of) lemma Hausdorff_imp_t0_space: "Hausdorff_space X \ t0_space X" by (simp add: Hausdorff_imp_t1_space t1_imp_t0_space) lemma t0_space: "t0_space X \ (\x \ topspace X. \y \ topspace X. x \ y \ (\C. closedin X C \ (x \ C \ y \ C)))" unfolding t0_space_def by (metis Diff_iff closedin_def openin_closedin_eq) lemma homeomorphic_t0_space: assumes "X homeomorphic_space Y" shows "t0_space X \ t0_space Y" proof - obtain f where f: "homeomorphic_map X Y f" and F: "inj_on f (topspace X)" and "topspace Y = f ` topspace X" by (metis assms homeomorphic_imp_injective_map homeomorphic_imp_surjective_map homeomorphic_space) with inj_on_image_mem_iff [OF F] show ?thesis apply (simp add: t0_space_def homeomorphic_eq_everything_map continuous_map_def open_map_def inj_on_def) by (smt (verit) mem_Collect_eq openin_subset) qed lemma t0_space_closure_of_sing: "t0_space X \ (\x \ topspace X. \y \ topspace X. X closure_of {x} = X closure_of {y} \ x = y)" by (simp add: t0_space_def closure_of_def set_eq_iff) (smt (verit)) lemma t0_space_discrete_topology: "t0_space (discrete_topology S)" by (simp add: Hausdorff_imp_t0_space) lemma t0_space_subtopology: "t0_space X \ t0_space (subtopology X U)" by (simp add: t0_space_def openin_subtopology) (metis Int_iff) lemma t0_space_retraction_map_image: "\retraction_map X Y r; t0_space X\ \ t0_space Y" using hereditary_imp_retractive_property homeomorphic_t0_space t0_space_subtopology by blast lemma XY: "{x}\{y} = {(x,y)}" by simp lemma t0_space_prod_topologyI: "\t0_space X; t0_space Y\ \ t0_space (prod_topology X Y)" by (simp add: t0_space_closure_of_sing closure_of_Times closure_of_eq_empty_gen times_eq_iff flip: XY insert_Times_insert) lemma t0_space_prod_topology_iff: "t0_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ t0_space X \ t0_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Sigma_empty1 Sigma_empty2 retraction_map_fst retraction_map_snd t0_space_retraction_map_image topspace_prod_topology) qed (metis empty_iff t0_space_def t0_space_prod_topologyI) proposition t0_space_product_topology: "t0_space (product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. t0_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (meson retraction_map_product_projection t0_space_retraction_map_image) next assume R: ?rhs show ?lhs proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis by (simp add: t0_space_def) next case False show ?thesis unfolding t0_space proof (intro strip) fix x y assume x: "x \ topspace (product_topology X I)" and y: "y \ topspace (product_topology X I)" and "x \ y" then obtain i where "i \ I" "x i \ y i" by (metis PiE_ext topspace_product_topology) then have "t0_space (X i)" using False R by blast then obtain U where "closedin (X i) U" "(x i \ U \ y i \ U)" by (metis t0_space PiE_mem \i \ I\ \x i \ y i\ topspace_product_topology x y) with \i \ I\ x y show "\U. closedin (product_topology X I) U \ (x \ U) = (y \ U)" by (rule_tac x="PiE I (\j. if j = i then U else topspace(X j))" in exI) (simp add: closedin_product_topology PiE_iff) qed qed qed subsection \Kolmogorov quotients\ definition Kolmogorov_quotient where "Kolmogorov_quotient X \ \x. @y. \U. openin X U \ (y \ U \ x \ U)" lemma Kolmogorov_quotient_in_open: "openin X U \ (Kolmogorov_quotient X x \ U \ x \ U)" by (smt (verit, ccfv_SIG) Kolmogorov_quotient_def someI_ex) lemma Kolmogorov_quotient_in_topspace: "Kolmogorov_quotient X x \ topspace X \ x \ topspace X" by (simp add: Kolmogorov_quotient_in_open) lemma Kolmogorov_quotient_in_closed: "closedin X C \ (Kolmogorov_quotient X x \ C \ x \ C)" unfolding closedin_def by (meson DiffD2 DiffI Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace in_mono) lemma continuous_map_Kolmogorov_quotient: "continuous_map X X (Kolmogorov_quotient X)" using Kolmogorov_quotient_in_open openin_subopen openin_subset by (fastforce simp: continuous_map_def Kolmogorov_quotient_in_topspace) lemma open_map_Kolmogorov_quotient_explicit: "openin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" using Kolmogorov_quotient_in_open openin_subset by fastforce lemma open_map_Kolmogorov_quotient_gen: "open_map (subtopology X S) (subtopology X (image (Kolmogorov_quotient X) S)) (Kolmogorov_quotient X)" proof (clarsimp simp: open_map_def openin_subtopology_alt image_iff) fix U assume "openin X U" then have "Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ U" using Kolmogorov_quotient_in_open [of X U] by auto then show "\V. openin X V \ Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ V" using \openin X U\ by blast qed lemma open_map_Kolmogorov_quotient: "open_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis open_map_Kolmogorov_quotient_gen subtopology_topspace) lemma closed_map_Kolmogorov_quotient_explicit: "closedin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" using closedin_subset by (fastforce simp: Kolmogorov_quotient_in_closed) lemma closed_map_Kolmogorov_quotient_gen: "closed_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" using Kolmogorov_quotient_in_closed by (force simp: closed_map_def closedin_subtopology_alt image_iff) lemma closed_map_Kolmogorov_quotient: "closed_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis closed_map_Kolmogorov_quotient_gen subtopology_topspace) lemma quotient_map_Kolmogorov_quotient_gen: "quotient_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" proof (intro continuous_open_imp_quotient_map) show "continuous_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" by (simp add: continuous_map_Kolmogorov_quotient continuous_map_from_subtopology continuous_map_in_subtopology image_mono) show "open_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" using open_map_Kolmogorov_quotient_gen by blast show "Kolmogorov_quotient X ` topspace (subtopology X S) = topspace (subtopology X (Kolmogorov_quotient X ` S))" by (force simp: Kolmogorov_quotient_in_open) qed lemma quotient_map_Kolmogorov_quotient: "quotient_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis quotient_map_Kolmogorov_quotient_gen subtopology_topspace) lemma Kolmogorov_quotient_eq: "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ (\U. openin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Kolmogorov_quotient_in_open) next assume ?rhs then show ?lhs by (simp add: Kolmogorov_quotient_def) qed lemma Kolmogorov_quotient_eq_alt: "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ (\U. closedin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Kolmogorov_quotient_in_closed) next assume ?rhs then show ?lhs by (smt (verit) Diff_iff Kolmogorov_quotient_eq closedin_topspace in_mono openin_closedin_eq) qed lemma Kolmogorov_quotient_continuous_map: assumes "continuous_map X Y f" "t0_space Y" and x: "x \ topspace X" shows "f (Kolmogorov_quotient X x) = f x" using assms unfolding continuous_map_def t0_space_def by (smt (verit, ccfv_SIG) Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace x mem_Collect_eq) lemma t0_space_Kolmogorov_quotient: "t0_space (subtopology X (Kolmogorov_quotient X ` topspace X))" apply (clarsimp simp: t0_space_def ) by (smt (verit, best) Kolmogorov_quotient_eq imageE image_eqI open_map_Kolmogorov_quotient open_map_def) lemma Kolmogorov_quotient_id: "t0_space X \ x \ topspace X \ Kolmogorov_quotient X x = x" by (metis Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace t0_space_def) lemma Kolmogorov_quotient_idemp: "Kolmogorov_quotient X (Kolmogorov_quotient X x) = Kolmogorov_quotient X x" by (simp add: Kolmogorov_quotient_eq Kolmogorov_quotient_in_open) lemma retraction_maps_Kolmogorov_quotient: "retraction_maps X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X) id" unfolding retraction_maps_def continuous_map_in_subtopology using Kolmogorov_quotient_idemp continuous_map_Kolmogorov_quotient by force lemma retraction_map_Kolmogorov_quotient: "retraction_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" using retraction_map_def retraction_maps_Kolmogorov_quotient by blast lemma retract_of_space_Kolmogorov_quotient_image: "Kolmogorov_quotient X ` topspace X retract_of_space X" proof - have "continuous_map X X (Kolmogorov_quotient X)" by (simp add: continuous_map_Kolmogorov_quotient) then have "Kolmogorov_quotient X ` topspace X \ topspace X" by (simp add: continuous_map_image_subset_topspace) then show ?thesis by (meson retract_of_space_retraction_maps retraction_maps_Kolmogorov_quotient) qed lemma Kolmogorov_quotient_lift_exists: assumes "S \ topspace X" "t0_space Y" and f: "continuous_map (subtopology X S) Y f" obtains g where "continuous_map (subtopology X (image (Kolmogorov_quotient X) S)) Y g" "\x. x \ S \ g(Kolmogorov_quotient X x) = f x" proof - have "\x y. \x \ S; y \ S; Kolmogorov_quotient X x = Kolmogorov_quotient X y\ \ f x = f y" using assms apply (simp add: Kolmogorov_quotient_eq t0_space_def continuous_map_def Int_absorb1 openin_subtopology) by (smt (verit, del_insts) Int_iff mem_Collect_eq) then obtain g where g: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) Y g" "g ` (topspace X \ Kolmogorov_quotient X ` S) = f ` S" "\x. x \ S \ g (Kolmogorov_quotient X x) = f x" using quotient_map_lift_exists [OF quotient_map_Kolmogorov_quotient_gen [of X S] f] by (metis assms(1) topspace_subtopology topspace_subtopology_subset) show ?thesis proof qed (use g in auto) qed subsection\Closed diagonals and graphs\ lemma Hausdorff_space_closedin_diagonal: "Hausdorff_space X \ closedin (prod_topology X X) ((\x. (x,x)) ` topspace X)" proof - have \
: "((\x. (x, x)) ` topspace X) \ topspace X \ topspace X" by auto show ?thesis apply (simp add: closedin_def openin_prod_topology_alt Hausdorff_space_def disjnt_iff \
) apply (intro all_cong1 imp_cong ex_cong1 conj_cong refl) by (force dest!: openin_subset)+ qed lemma closed_map_diag_eq: "closed_map X (prod_topology X X) (\x. (x,x)) \ Hausdorff_space X" proof - have "section_map X (prod_topology X X) (\x. (x, x))" unfolding section_map_def retraction_maps_def by (smt (verit) continuous_map_fst continuous_map_of_fst continuous_map_on_empty continuous_map_pairwise fst_conv fst_diag_fst snd_diag_fst) then have "embedding_map X (prod_topology X X) (\x. (x, x))" by (rule section_imp_embedding_map) then show ?thesis using Hausdorff_space_closedin_diagonal embedding_imp_closed_map_eq by blast qed lemma closedin_continuous_maps_eq: assumes "Hausdorff_space Y" and f: "continuous_map X Y f" and g: "continuous_map X Y g" shows "closedin X {x \ topspace X. f x = g x}" proof - have \
:"{x \ topspace X. f x = g x} = {x \ topspace X. (f x,g x) \ ((\y.(y,y)) ` topspace Y)}" using f continuous_map_image_subset_topspace by fastforce show ?thesis unfolding \
proof (intro closedin_continuous_map_preimage) show "continuous_map X (prod_topology Y Y) (\x. (f x, g x))" by (simp add: continuous_map_pairedI f g) show "closedin (prod_topology Y Y) ((\y. (y, y)) ` topspace Y)" using Hausdorff_space_closedin_diagonal assms by blast qed qed lemma retract_of_space_imp_closedin: assumes "Hausdorff_space X" and S: "S retract_of_space X" shows "closedin X S" proof - obtain r where r: "continuous_map X (subtopology X S) r" "\x\S. r x = x" using assms by (meson retract_of_space_def) then have \
: "S = {x \ topspace X. r x = x}" using S retract_of_space_imp_subset by (force simp: continuous_map_def) show ?thesis unfolding \
using r continuous_map_into_fulltopology assms by (force intro: closedin_continuous_maps_eq) qed lemma homeomorphic_maps_graph: "homeomorphic_maps X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` (topspace X))) (\x. (x, f x)) fst \ continuous_map X Y f" (is "?lhs=?rhs") proof assume ?lhs then have h: "homeomorphic_map X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` topspace X)) (\x. (x, f x))" by (auto simp: homeomorphic_maps_map) have "f = snd \ (\x. (x, f x))" by force then show ?rhs by (metis (no_types, lifting) h continuous_map_in_subtopology continuous_map_snd_of homeomorphic_eq_everything_map) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by (smt (verit, ccfv_threshold) continuous_map_eq continuous_map_subtopology_fst embedding_map_def embedding_map_graph homeomorphic_eq_everything_map image_cong image_iff prod.collapse prod.inject) qed subsection \ KC spaces, those where all compact sets are closed.\ definition kc_space where "kc_space X \ \S. compactin X S \ closedin X S" lemma kc_space_euclidean: "kc_space (euclidean :: 'a::metric_space topology)" by (simp add: compact_imp_closed kc_space_def) lemma kc_space_expansive: "\kc_space X; topspace Y = topspace X; \U. openin X U \ openin Y U\ \ kc_space Y" by (meson compactin_contractive kc_space_def topology_finer_closedin) lemma compactin_imp_closedin_gen: "\kc_space X; compactin X S\ \ closedin X S" using kc_space_def by blast lemma Hausdorff_imp_kc_space: "Hausdorff_space X \ kc_space X" by (simp add: compactin_imp_closedin kc_space_def) lemma kc_imp_t1_space: "kc_space X \ t1_space X" by (simp add: finite_imp_compactin kc_space_def t1_space_closedin_finite) lemma kc_space_subtopology: "kc_space X \ kc_space(subtopology X S)" by (metis closedin_Int_closure_of closure_of_eq compactin_subtopology inf.absorb2 kc_space_def) lemma kc_space_discrete_topology: "kc_space(discrete_topology U)" using Hausdorff_space_discrete_topology compactin_imp_closedin kc_space_def by blast lemma kc_space_continuous_injective_map_preimage: assumes "kc_space Y" "continuous_map X Y f" and injf: "inj_on f (topspace X)" shows "kc_space X" unfolding kc_space_def proof (intro strip) fix S assume S: "compactin X S" have "S = {x \ topspace X. f x \ f ` S}" using S compactin_subset_topspace inj_onD [OF injf] by blast with assms S show "closedin X S" by (metis (no_types, lifting) Collect_cong closedin_continuous_map_preimage compactin_imp_closedin_gen image_compactin) qed lemma kc_space_retraction_map_image: assumes "retraction_map X Y r" "kc_space X" shows "kc_space Y" proof - obtain s where s: "continuous_map X Y r" "continuous_map Y X s" "\x. x \ topspace Y \ r (s x) = x" using assms by (force simp: retraction_map_def retraction_maps_def) then have inj: "inj_on s (topspace Y)" by (metis inj_on_def) show ?thesis unfolding kc_space_def proof (intro strip) fix S assume S: "compactin Y S" have "S = {x \ topspace Y. s x \ s ` S}" using S compactin_subset_topspace inj_onD [OF inj] by blast with assms S show "closedin Y S" by (meson compactin_imp_closedin_gen inj kc_space_continuous_injective_map_preimage s(2)) qed qed lemma homeomorphic_kc_space: "X homeomorphic_space Y \ kc_space X \ kc_space Y" by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym kc_space_continuous_injective_map_preimage) lemma compact_kc_eq_maximal_compact_space: assumes "compact_space X" shows "kc_space X \ (\Y. topspace Y = topspace X \ (\S. openin X S \ openin Y S) \ compact_space Y \ Y = X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis closedin_compact_space compactin_contractive kc_space_def topology_eq topology_finer_closedin) next assume R: ?rhs show ?lhs unfolding kc_space_def proof (intro strip) fix S assume S: "compactin X S" define Y where "Y \ topology (arbitrary union_of (finite intersection_of (\A. A = topspace X - S \ openin X A) relative_to (topspace X)))" have "topspace Y = topspace X" by (auto simp: Y_def) have "openin X T \ openin Y T" for T by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase openin_subset relative_to_subset) have "compact_space Y" proof (rule Alexander_subbase_alt) show "\\'. finite \' \ \' \ \ \ topspace X \ \ \'" if \: "\ \ insert (topspace X - S) (Collect (openin X))" and sub: "topspace X \ \\" for \ proof - consider "\ \ Collect (openin X)" | \ where "\ = insert (topspace X - S) \" "\ \ Collect (openin X)" using \ by (metis insert_Diff subset_insert_iff) then show ?thesis proof cases case 1 then show ?thesis by (metis assms compact_space_alt mem_Collect_eq subsetD that(2)) next case 2 then have "S \ \\" using S sub compactin_subset_topspace by blast with 2 obtain \ where "finite \ \ \ \ \ \ S \ \\" using S unfolding compactin_def by (metis Ball_Collect) with 2 show ?thesis by (rule_tac x="insert (topspace X - S) \" in exI) blast qed qed qed (auto simp: Y_def) have "Y = X" using R \\S. openin X S \ openin Y S\ \compact_space Y\ \topspace Y = topspace X\ by blast moreover have "openin Y (topspace X - S)" by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase relative_to_subset) ultimately show "closedin X S" unfolding closedin_def using S compactin_subset_topspace by blast qed qed lemma continuous_imp_closed_map_gen: "\compact_space X; kc_space Y; continuous_map X Y f\ \ closed_map X Y f" by (meson closed_map_def closedin_compact_space compactin_imp_closedin_gen image_compactin) lemma kc_space_compact_subtopologies: "kc_space X \ (\K. compactin X K \ kc_space(subtopology X K))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (auto simp: kc_space_def closedin_closed_subtopology compactin_subtopology) next assume R: ?rhs show ?lhs unfolding kc_space_def proof (intro strip) fix K assume K: "compactin X K" then have "K \ topspace X" by (simp add: compactin_subset_topspace) moreover have "X closure_of K \ K" proof fix x assume x: "x \ X closure_of K" have "kc_space (subtopology X K)" by (simp add: R \compactin X K\) have "compactin X (insert x K)" by (metis K x compactin_Un compactin_sing in_closure_of insert_is_Un) then show "x \ K" by (metis R x K Int_insert_left_if1 closedin_Int_closure_of compact_imp_compactin_subtopology insertCI kc_space_def subset_insertI) qed ultimately show "closedin X K" using closure_of_subset_eq by blast qed qed lemma kc_space_compact_prod_topology: assumes "compact_space X" shows "kc_space(prod_topology X X) \ Hausdorff_space X" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding closed_map_diag_eq [symmetric] proof (intro continuous_imp_closed_map_gen) show "continuous_map X (prod_topology X X) (\x. (x, x))" by (intro continuous_intros) qed (use L assms in auto) next assume ?rhs then show ?lhs by (simp add: Hausdorff_imp_kc_space Hausdorff_space_prod_topology) qed lemma kc_space_prod_topology: "kc_space(prod_topology X X) \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis compactin_subspace kc_space_compact_prod_topology kc_space_subtopology subtopology_Times) next assume R: ?rhs have "kc_space (subtopology (prod_topology X X) L)" if "compactin (prod_topology X X) L" for L proof - define K where "K \ fst ` L \ snd ` L" have "L \ K \ K" by (force simp: K_def) have "compactin X K" by (metis K_def compactin_Un continuous_map_fst continuous_map_snd image_compactin that) then have "Hausdorff_space (subtopology X K)" by (simp add: R) then have "kc_space (prod_topology (subtopology X K) (subtopology X K))" by (simp add: \compactin X K\ compact_space_subtopology kc_space_compact_prod_topology) then have "kc_space (subtopology (prod_topology (subtopology X K) (subtopology X K)) L)" using kc_space_subtopology by blast then show ?thesis using \L \ K \ K\ subtopology_Times subtopology_subtopology by (metis (no_types, lifting) Sigma_cong inf.absorb_iff2) qed then show ?lhs using kc_space_compact_subtopologies by blast qed lemma kc_space_prod_topology_alt: "kc_space(prod_topology X X) \ kc_space X \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" using Hausdorff_imp_kc_space kc_space_compact_subtopologies kc_space_prod_topology by blast proposition kc_space_prod_topology_left: assumes X: "kc_space X" and Y: "Hausdorff_space Y" shows "kc_space (prod_topology X Y)" unfolding kc_space_def proof (intro strip) fix K assume K: "compactin (prod_topology X Y) K" then have "K \ topspace X \ topspace Y" using compactin_subset_topspace topspace_prod_topology by blast moreover have "\T. openin (prod_topology X Y) T \ (a,b) \ T \ T \ (topspace X \ topspace Y) - K" if ab: "(a,b) \ (topspace X \ topspace Y) - K" for a b proof - have "compactin Y {b}" using that by force moreover have "compactin Y {y \ topspace Y. (a,y) \ K}" proof - have "compactin (prod_topology X Y) (K \ {a} \ topspace Y)" using that compact_Int_closedin [OF K] by (simp add: X closedin_prod_Times_iff compactin_imp_closedin_gen) moreover have "subtopology (prod_topology X Y) (K \ {a} \ topspace Y) homeomorphic_space subtopology Y {y \ topspace Y. (a, y) \ K}" unfolding homeomorphic_space_def homeomorphic_maps_def using that apply (rule_tac x="snd" in exI) apply (rule_tac x="Pair a" in exI) by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology continuous_map_subtopology_snd continuous_map_paired) ultimately show ?thesis by (simp add: compactin_subspace homeomorphic_compact_space) qed moreover have "disjnt {b} {y \ topspace Y. (a,y) \ K}" using ab by force ultimately obtain V U where VU: "openin Y V" "openin Y U" "{b} \ V" "{y \ topspace Y. (a,y) \ K} \ U" "disjnt V U" using Hausdorff_space_compact_separation [OF Y] by blast define V' where "V' \ topspace Y - U" have W: "closedin Y V'" "{y \ topspace Y. (a,y) \ K} \ topspace Y - V'" "disjnt V (topspace Y - V')" using VU by (auto simp: V'_def disjnt_iff) with VU obtain "V \ topspace Y" "V' \ topspace Y" by (meson closedin_subset openin_closedin_eq) then obtain "b \ V" "disjnt {y \ topspace Y. (a,y) \ K} V'" "V \ V'" using VU unfolding disjnt_iff V'_def by force define C where "C \ image fst (K \ {z \ topspace(prod_topology X Y). snd z \ V'})" have "closedin (prod_topology X Y) {z \ topspace (prod_topology X Y). snd z \ V'}" using closedin_continuous_map_preimage \closedin Y V'\ continuous_map_snd by blast then have "compactin X C" unfolding C_def by (meson K compact_Int_closedin continuous_map_fst image_compactin) then have "closedin X C" using assms by (auto simp: kc_space_def) show ?thesis proof (intro exI conjI) show "openin (prod_topology X Y) ((topspace X - C) \ V)" by (simp add: VU \closedin X C\ openin_diff openin_prod_Times_iff) have "a \ C" using VU by (auto simp: C_def V'_def) then show "(a, b) \ (topspace X - C) \ V" using \a \ C\ \b \ V\ that by blast show "(topspace X - C) \ V \ topspace X \ topspace Y - K" using \V \ V'\ \V \ topspace Y\ apply (simp add: C_def ) by (smt (verit, ccfv_threshold) DiffE DiffI IntI SigmaE SigmaI image_eqI mem_Collect_eq prod.sel(1) snd_conv subset_iff) qed qed ultimately show "closedin (prod_topology X Y) K" by (metis surj_pair closedin_def openin_subopen topspace_prod_topology) qed lemma kc_space_prod_topology_right: "\Hausdorff_space X; kc_space Y\ \ kc_space (prod_topology X Y)" using kc_space_prod_topology_left homeomorphic_kc_space homeomorphic_space_prod_topology_swap by blast subsection \Regular spaces\ text \Regular spaces are *not* a priori assumed to be Hausdorff or $T_1$\ definition regular_space where "regular_space X \ \C a. closedin X C \ a \ topspace X - C \ (\U V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V)" lemma homeomorphic_regular_space_aux: assumes hom: "X homeomorphic_space Y" and X: "regular_space X" shows "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 regular_space_def proof clarify fix C a assume Y: "closedin Y C" "a \ topspace Y" and "a \ C" then obtain "closedin X (g ` C)" "g a \ topspace X" "g a \ g ` C" using \closedin Y C\ hmg homeomorphic_map_closedness_eq by (smt (verit, ccfv_SIG) fg homeomorphic_imp_surjective_map image_iff in_mono) then obtain S T where ST: "openin X S" "openin X T" "g a \ S" "g`C \ T" "disjnt S T" using X unfolding regular_space_def by (metis DiffI) then have "openin Y (f`S)" "openin Y (f`T)" by (meson hmf homeomorphic_map_openness_eq)+ moreover have "a \ f`S \ C \ f`T" using ST by (smt (verit, best) Y closedin_subset fg image_eqI subset_iff) moreover have "disjnt (f`S) (f`T)" using ST by (smt (verit, ccfv_SIG) disjnt_iff fg image_iff openin_subset subsetD) ultimately show "\U V. openin Y U \ openin Y V \ a \ U \ C \ V \ disjnt U V" by metis qed qed lemma homeomorphic_regular_space: "X homeomorphic_space Y \ (regular_space X \ regular_space Y)" by (meson homeomorphic_regular_space_aux homeomorphic_space_sym) lemma regular_space: "regular_space X \ (\C a. closedin X C \ a \ topspace X - C \ (\U. openin X U \ a \ U \ disjnt C (X closure_of U)))" unfolding regular_space_def proof (intro all_cong1 imp_cong refl ex_cong1) fix C a U assume C: "closedin X C \ a \ topspace X - C" show "(\V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V) \ (openin X U \ a \ U \ disjnt C (X closure_of U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (smt (verit, best) disjnt_iff in_closure_of subsetD) next assume R: ?rhs then have "disjnt U (topspace X - X closure_of U)" by (meson DiffD2 closure_of_subset disjnt_iff openin_subset subsetD) moreover have "C \ topspace X - X closure_of U" by (meson C DiffI R closedin_subset disjnt_iff subset_eq) ultimately show ?lhs using R by (rule_tac x="topspace X - X closure_of U" in exI) auto qed qed lemma neighbourhood_base_of_closedin: "neighbourhood_base_of (closedin X) X \ regular_space X" (is "?lhs=?rhs") proof - have "?lhs \ (\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ W)))" by (simp add: neighbourhood_base_of) also have "\ \ (\W x. closedin X W \ x \ topspace X - W \ (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W)))" by (smt (verit) Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) also have "\ \ ?rhs" proof - have \
: "(\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W) \ (\V. openin X V \ x \ U \ W \ V \ disjnt U V)" (is "?lhs=?rhs") if "openin X U" "closedin X W" "x \ topspace X" "x \ W" for W U x proof assume ?lhs with \closedin X W\ show ?rhs unfolding closedin_def by (smt (verit) Diff_mono disjnt_Diff1 double_diff subset_eq) next assume ?rhs with \openin X U\ show ?lhs unfolding openin_closedin_eq disjnt_def by (smt (verit) Diff_Diff_Int Diff_disjoint Diff_eq_empty_iff Int_Diff inf.orderE) qed show ?thesis unfolding regular_space_def by (intro all_cong1 ex_cong1 imp_cong refl) (metis \
DiffE) qed finally show ?thesis . qed lemma regular_space_discrete_topology: "regular_space (discrete_topology S)" using neighbourhood_base_of_closedin neighbourhood_base_of_discrete_topology by fastforce lemma regular_space_subtopology: "regular_space X \ regular_space (subtopology X S)" unfolding regular_space_def openin_subtopology_alt closedin_subtopology_alt disjnt_iff by clarsimp (smt (verit, best) inf.orderE inf_le1 le_inf_iff) lemma regular_space_retraction_map_image: "\retraction_map X Y r; regular_space X\ \ regular_space Y" using hereditary_imp_retractive_property homeomorphic_regular_space regular_space_subtopology by blast lemma regular_t0_imp_Hausdorff_space: "\regular_space X; t0_space X\ \ Hausdorff_space X" apply (clarsimp simp: regular_space_def t0_space Hausdorff_space_def) by (metis disjnt_sym subsetD) lemma regular_t0_eq_Hausdorff_space: "regular_space X \ (t0_space X \ Hausdorff_space X)" using Hausdorff_imp_t0_space regular_t0_imp_Hausdorff_space by blast lemma regular_t1_imp_Hausdorff_space: "\regular_space X; t1_space X\ \ Hausdorff_space X" by (simp add: regular_t0_imp_Hausdorff_space t1_imp_t0_space) lemma regular_t1_eq_Hausdorff_space: "regular_space X \ t1_space X \ Hausdorff_space X" using regular_t0_imp_Hausdorff_space t1_imp_t0_space t1_or_Hausdorff_space by blast lemma compact_Hausdorff_imp_regular_space: assumes "compact_space X" "Hausdorff_space X" shows "regular_space X" unfolding regular_space_def proof clarify fix S a assume "closedin X S" and "a \ topspace X" and "a \ S" then show "\U V. openin X U \ openin X V \ a \ U \ S \ V \ disjnt U V" using assms unfolding Hausdorff_space_compact_sets by (metis closedin_compact_space compactin_sing disjnt_empty1 insert_subset disjnt_insert1) qed lemma regular_space_topspace_empty: "topspace X = {} \ regular_space X" by (simp add: Hausdorff_space_topspace_empty compact_Hausdorff_imp_regular_space compact_space_topspace_empty) lemma neighbourhood_base_of_closed_Hausdorff_space: "regular_space X \ Hausdorff_space X \ neighbourhood_base_of (\C. closedin X C \ Hausdorff_space(subtopology X C)) X" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (simp add: Hausdorff_space_subtopology neighbourhood_base_of_closedin) next assume ?rhs then show ?lhs by (metis (mono_tags, lifting) Hausdorff_space_closed_neighbourhood neighbourhood_base_of neighbourhood_base_of_closedin openin_topspace) qed lemma locally_compact_imp_kc_eq_Hausdorff_space: "neighbourhood_base_of (compactin X) X \ kc_space X \ Hausdorff_space X" by (metis Hausdorff_imp_kc_space kc_imp_t1_space kc_space_def neighbourhood_base_of_closedin neighbourhood_base_of_mono regular_t1_imp_Hausdorff_space) lemma regular_space_compact_closed_separation: assumes X: "regular_space X" and S: "compactin X S" and T: "closedin X T" and "disjnt S T" shows "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" proof (cases "S={}") case True then show ?thesis by (meson T closedin_def disjnt_empty1 empty_subsetI openin_empty openin_topspace) next case False then have "\U V. x \ S \ openin X U \ openin X V \ x \ U \ T \ V \ disjnt U V" for x using assms unfolding regular_space_def by (smt (verit) Diff_iff compactin_subset_topspace disjnt_iff subsetD) then obtain U V where UV: "\x. x \ S \ openin X (U x) \ openin X (V x) \ x \ (U x) \ T \ (V x) \ disjnt (U x) (V x)" by metis then obtain \ where "finite \" "\ \ U ` S" "S \ \ \" using S unfolding compactin_def by (smt (verit) UN_iff image_iff subsetI) then obtain K where "finite K" "K \ S" and K: "S \ \(U ` K)" by (metis finite_subset_image) show ?thesis proof (intro exI conjI) show "openin X (\(U ` K))" using \K \ S\ UV by blast show "openin X (\(V ` K))" using False K UV \K \ S\ \finite K\ by blast show "S \ \(U ` K)" by (simp add: K) show "T \ \(V ` K)" using UV \K \ S\ by blast show "disjnt (\(U ` K)) (\(V ` K))" by (smt (verit) Inter_iff UN_E UV \K \ S\ disjnt_iff image_eqI subset_iff) qed qed lemma regular_space_compact_closed_sets: "regular_space X \ (\S T. compactin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using regular_space_compact_closed_separation by fastforce next assume R: ?rhs show ?lhs unfolding regular_space proof clarify fix S x assume "closedin X S" and "x \ topspace X" and "x \ S" then obtain U V where "openin X U \ openin X V \ {x} \ U \ S \ V \ disjnt U V" by (metis R compactin_sing disjnt_empty1 disjnt_insert1) then show "\U. openin X U \ x \ U \ disjnt S (X closure_of U)" by (smt (verit, best) disjnt_iff in_closure_of insert_subset subsetD) qed qed lemma regular_space_prod_topology: "regular_space (prod_topology X Y) \ topspace X = {} \ topspace Y = {} \ regular_space X \ regular_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis regular_space_retraction_map_image retraction_map_fst retraction_map_snd) next assume R: ?rhs show ?lhs proof (cases "topspace X = {} \ topspace Y = {}") case True then show ?thesis by (simp add: regular_space_topspace_empty) next case False then have "regular_space X" "regular_space Y" using R by auto show ?thesis unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of proof clarify fix W x y assume W: "openin (prod_topology X Y) W" and "(x,y) \ W" then obtain U V where U: "openin X U" "x \ U" and V: "openin Y V" "y \ V" and "U \ V \ W" by (metis openin_prod_topology_alt) obtain D1 C1 where 1: "openin X D1" "closedin X C1" "x \ D1" "D1 \ C1" "C1 \ U" by (metis \regular_space X\ U neighbourhood_base_of neighbourhood_base_of_closedin) obtain D2 C2 where 2: "openin Y D2" "closedin Y C2" "y \ D2" "D2 \ C2" "C2 \ V" by (metis \regular_space Y\ V neighbourhood_base_of neighbourhood_base_of_closedin) show "\U V. openin (prod_topology X Y) U \ closedin (prod_topology X Y) V \ (x,y) \ U \ U \ V \ V \ W" proof (intro conjI exI) show "openin (prod_topology X Y) (D1 \ D2)" by (simp add: 1 2 openin_prod_Times_iff) show "closedin (prod_topology X Y) (C1 \ C2)" by (simp add: 1 2 closedin_prod_Times_iff) qed (use 1 2 \U \ V \ W\ in auto) qed qed qed lemma regular_space_product_topology: "regular_space (product_topology X I) \ topspace (product_topology X I) = {} \ (\i \ I. regular_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (meson regular_space_retraction_map_image retraction_map_product_projection) next assume R: ?rhs show ?lhs proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis by (simp add: regular_space_topspace_empty) next case False then obtain x where x: "x \ topspace (product_topology X I)" by blast define \ where "\ \ {Pi\<^sub>E I U |U. finite {i \ I. U i \ topspace (X i)} \ (\i\I. openin (X i) (U i))}" have oo: "openin (product_topology X I) = arbitrary union_of (\W. W \ \)" by (simp add: \_def openin_product_topology product_topology_base_alt) have "\U V. openin (product_topology X I) U \ closedin (product_topology X I) V \ x \ U \ U \ V \ V \ Pi\<^sub>E I W" if fin: "finite {i \ I. W i \ topspace (X i)}" and opeW: "\k. k \ I \ openin (X k) (W k)" and x: "x \ PiE I W" for W x proof - have "\i. i \ I \ \U V. openin (X i) U \ closedin (X i) V \ x i \ U \ U \ V \ V \ W i" by (metis False PiE_iff R neighbourhood_base_of neighbourhood_base_of_closedin opeW x) then obtain U C where UC: "\i. i \ I \ openin (X i) (U i) \ closedin (X i) (C i) \ x i \ U i \ U i \ C i \ C i \ W i" by metis define PI where "PI \ \V. PiE I (\i. if W i = topspace(X i) then topspace(X i) else V i)" show ?thesis proof (intro conjI exI) have "\i\I. W i \ topspace (X i) \ openin (X i) (U i)" using UC by force with fin show "openin (product_topology X I) (PI U)" by (simp add: Collect_mono_iff PI_def openin_PiE_gen rev_finite_subset) show "closedin (product_topology X I) (PI C)" by (simp add: UC closedin_product_topology PI_def) show "x \ PI U" using UC x by (fastforce simp: PI_def) show "PI U \ PI C" by (smt (verit) UC Orderings.order_eq_iff PiE_mono PI_def) show "PI C \ Pi\<^sub>E I W" by (simp add: UC PI_def subset_PiE) qed qed then have "neighbourhood_base_of (closedin (product_topology X I)) (product_topology X I)" unfolding neighbourhood_base_of_topology_base [OF oo] by (force simp: \_def) then show ?thesis by (simp add: neighbourhood_base_of_closedin) qed qed lemma closed_map_paired_gen_aux: assumes "regular_space Y" and f: "closed_map Z X f" and g: "closed_map Z Y g" and clo: "\y. y \ topspace X \ closedin Z {x \ topspace Z. f x = y}" and contg: "continuous_map Z Y g" shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" unfolding closed_map_def proof (intro strip) fix C assume "closedin Z C" then have "C \ topspace Z" by (simp add: closedin_subset) have "f ` topspace Z \ topspace X" "g ` topspace Z \ topspace Y" by (simp_all add: assms closed_map_imp_subset_topspace) show "closedin (prod_topology X Y) ((\x. (f x, g x)) ` C)" unfolding closedin_def topspace_prod_topology proof (intro conjI) have "closedin Y (g ` C)" using \closedin Z C\ assms(3) closed_map_def by blast with assms show "(\x. (f x, g x)) ` C \ topspace X \ topspace Y" using \C \ topspace Z\ \f ` topspace Z \ topspace X\ continuous_map_closedin subsetD by fastforce have *: "\T. openin (prod_topology X Y) T \ (y1,y2) \ T \ T \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" if "(y1,y2) \ (\x. (f x, g x)) ` C" and y1: "y1 \ topspace X" and y2: "y2 \ topspace Y" for y1 y2 proof - define A where "A \ topspace Z - (C \ {x \ topspace Z. f x = y1})" have A: "openin Z A" "{x \ topspace Z. g x = y2} \ A" using that \closedin Z C\ clo that(2) by (auto simp: A_def) obtain V0 where "openin Y V0 \ y2 \ V0" and UA: "{x \ topspace Z. g x \ V0} \ A" using g A y2 unfolding closed_map_fibre_neighbourhood by blast then obtain V V' where VV: "openin Y V \ closedin Y V' \ y2 \ V \ V \ V'" and "V' \ V0" by (metis (no_types, lifting) \regular_space Y\ neighbourhood_base_of neighbourhood_base_of_closedin) with UA have subA: "{x \ topspace Z. g x \ V'} \ A" by blast show ?thesis proof - define B where "B \ topspace Z - (C \ {x \ topspace Z. g x \ V'})" have "openin Z B" using VV \closedin Z C\ contg by (fastforce simp: B_def continuous_map_closedin) have "{x \ topspace Z. f x = y1} \ B" using A_def subA by (auto simp: A_def B_def) then obtain U where "openin X U" "y1 \ U" and subB: "{x \ topspace Z. f x \ U} \ B" using \openin Z B\ y1 f unfolding closed_map_fibre_neighbourhood by meson show ?thesis proof (intro conjI exI) show "openin (prod_topology X Y) (U \ V)" by (metis VV \openin X U\ openin_prod_Times_iff) show "(y1, y2) \ U \ V" by (simp add: VV \y1 \ U\) show "U \ V \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" using VV \C \ topspace Z\ \openin X U\ subB by (force simp: image_iff B_def subset_iff dest: openin_subset) qed qed qed then show "openin (prod_topology X Y) (topspace X \ topspace Y - (\x. (f x, g x)) ` C)" by (smt (verit, ccfv_threshold) Diff_iff SigmaE openin_subopen) qed qed lemma closed_map_paired_gen: assumes f: "closed_map Z X f" and g: "closed_map Z Y g" and D: "(regular_space X \ continuous_map Z X f \ (\z \ topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y \ topspace X. closedin Z {x \ topspace Z. f x = y}))" (is "?RSX \ ?RSY") shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" using D proof assume RSX: ?RSX have eq: "(\x. (f x, g x)) = (\(x,y). (y,x)) \ (\x. (g x, f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map Z (prod_topology Y X) (\x. (g x, f x))" using RSX closed_map_paired_gen_aux f g by fastforce show "closed_map (prod_topology Y X) (prod_topology X Y) (\(x, y). (y, x))" using homeomorphic_imp_closed_map homeomorphic_map_swap by blast qed qed (blast intro: assms closed_map_paired_gen_aux) lemma closed_map_paired: assumes "closed_map Z X f" and contf: "continuous_map Z X f" "closed_map Z Y g" and contg: "continuous_map Z Y g" and D: "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" proof (rule closed_map_paired_gen) show "regular_space X \ continuous_map Z X f \ (\z\topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y\topspace X. closedin Z {x \ topspace Z. f x = y})" using D contf contg by (smt (verit, del_insts) Collect_cong closedin_continuous_map_preimage t1_space_closedin_singleton singleton_iff) qed (use assms in auto) lemma closed_map_pairwise: assumes "closed_map Z X (fst \ f)" "continuous_map Z X (fst \ f)" "closed_map Z Y (snd \ f)" "continuous_map Z Y (snd \ f)" "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" shows "closed_map Z (prod_topology X Y) f" proof - have "closed_map Z (prod_topology X Y) (\a. ((fst \ f) a, (snd \ f) a))" using assms closed_map_paired by blast then show ?thesis by auto qed lemma tube_lemma_right: assumes W: "openin (prod_topology X Y) W" and C: "compactin Y C" and x: "x \ topspace X" and subW: "{x} \ C \ W" shows "\U V. openin X U \ openin Y V \ x \ U \ C \ V \ U \ V \ W" proof (cases "C = {}") case True with x show ?thesis by auto next case False have "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ W" if "y \ C" for y using W openin_prod_topology_alt subW subsetD that by fastforce then obtain U V where UV: "\y. y \ C \ openin X (U y) \ openin Y (V y) \ x \ U y \ y \ V y \ U y \ V y \ W" by metis then obtain D where D: "finite D" "D \ C" "C \ \ (V ` D)" using compactinD [OF C, of "V`C"] by (smt (verit) UN_I finite_subset_image imageE subsetI) show ?thesis proof (intro exI conjI) show "openin X (\ (U ` D))" "openin Y (\ (V ` D))" using D False UV by blast+ show "x \ \ (U ` D)" "C \ \ (V ` D)" "\ (U ` D) \ \ (V ` D) \ W" using D UV by force+ qed qed lemma closed_map_fst: assumes "compact_space Y" shows "closed_map (prod_topology X Y) X fst" proof - have *: "{x \ topspace X \ topspace Y. fst x \ U} = U \ topspace Y" if "U \ topspace X" for U using that by force have **: "\U y. \openin (prod_topology X Y) U; y \ topspace X; {x \ topspace X \ topspace Y. fst x = y} \ U\ \ \V. openin X V \ y \ V \ V \ topspace Y \ U" using tube_lemma_right[of X Y _ "topspace Y"] assms compact_space_def by force show ?thesis unfolding closed_map_fibre_neighbourhood by (force simp: * openin_subset cong: conj_cong intro: **) qed lemma closed_map_snd: assumes "compact_space X" shows "closed_map (prod_topology X Y) Y snd" proof - have "snd = fst o prod.swap" by force moreover have "closed_map (prod_topology X Y) Y (fst o prod.swap)" proof (rule closed_map_compose) show "closed_map (prod_topology X Y) (prod_topology Y X) prod.swap" by (metis (no_types, lifting) homeomorphic_imp_closed_map homeomorphic_map_eq homeomorphic_map_swap prod.swap_def split_beta) show "closed_map (prod_topology Y X) Y fst" by (simp add: closed_map_fst assms) qed ultimately show ?thesis by metis qed lemma closed_map_paired_closed_map_right: "\closed_map X Y f; regular_space X; \y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}\ \ closed_map X (prod_topology X Y) (\x. (x, f x))" by (rule closed_map_paired_gen [OF closed_map_id, unfolded id_def]) auto lemma closed_map_paired_closed_map_left: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "closed_map X (prod_topology Y X) (\x. (f x, x))" proof - have eq: "(\x. (f x, x)) = (\(x,y). (y,x)) \ (\x. (x, f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map X (prod_topology X Y) (\x. (x, f x))" by (simp add: assms closed_map_paired_closed_map_right) show "closed_map (prod_topology X Y) (prod_topology Y X) (\(x, y). (y, x))" using homeomorphic_imp_closed_map homeomorphic_map_swap by blast qed qed lemma closed_map_imp_closed_graph: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "closedin (prod_topology X Y) ((\x. (x, f x)) ` topspace X)" using assms closed_map_def closed_map_paired_closed_map_right by blast lemma proper_map_paired_closed_map_right: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "proper_map X (prod_topology X Y) (\x. (x, f x))" by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_right) lemma proper_map_paired_closed_map_left: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "proper_map X (prod_topology Y X) (\x. (f x, x))" by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_left) proposition regular_space_continuous_proper_map_image: assumes "regular_space X" and contf: "continuous_map X Y f" and pmapf: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "regular_space Y" unfolding regular_space_def proof clarify fix C y assume "closedin Y C" and "y \ topspace Y" and "y \ C" have "closed_map X Y f" "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" using pmapf proper_map_def by force+ moreover have "closedin X {z \ topspace X. f z \ C}" using \closedin Y C\ contf continuous_map_closedin by fastforce moreover have "disjnt {z \ topspace X. f z = y} {z \ topspace X. f z \ C}" using \y \ C\ disjnt_iff by blast ultimately obtain U V where UV: "openin X U" "openin X V" "{z \ topspace X. f z = y} \ U" "{z \ topspace X. f z \ C} \ V" and dUV: "disjnt U V" using \y \ topspace Y\ \regular_space X\ unfolding regular_space_compact_closed_sets by meson have *: "\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)" using \closed_map X Y f\ unfolding closed_map_preimage_neighbourhood by blast obtain V1 where V1: "openin Y V1 \ y \ V1" and sub1: "{x \ topspace X. f x \ V1} \ U" using * [of U "{y}"] UV \y \ topspace Y\ by auto moreover obtain V2 where "openin Y V2 \ C \ V2" and sub2: "{x \ topspace X. f x \ V2} \ V" by (smt (verit, ccfv_SIG) * UV \closedin Y C\ closedin_subset mem_Collect_eq subset_iff) moreover have "disjnt V1 V2" proof - have "\x. \\x. x \ U \ x \ V; x \ V1; x \ V2\ \ False" by (smt (verit) V1 fim image_iff mem_Collect_eq openin_subset sub1 sub2 subsetD) with dUV show ?thesis by (auto simp: disjnt_iff) qed ultimately show "\U V. openin Y U \ openin Y V \ y \ U \ C \ V \ disjnt U V" by meson qed lemma regular_space_perfect_map_image: "\regular_space X; perfect_map X Y f\ \ regular_space Y" by (meson perfect_map_def regular_space_continuous_proper_map_image) proposition regular_space_perfect_map_image_eq: assumes "Hausdorff_space X" and perf: "perfect_map X Y f" shows "regular_space X \ regular_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using perf regular_space_perfect_map_image by blast next assume R: ?rhs have "continuous_map X Y f" "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" using perf by (auto simp: perfect_map_def) then have "closed_map X Y f" and preYf: "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" by (simp_all add: proper_map_def) show ?lhs unfolding regular_space_def proof clarify fix C x assume "closedin X C" and "x \ topspace X" and "x \ C" obtain U1 U2 where "openin X U1" "openin X U2" "{x} \ U1" and "disjnt U1 U2" and subV: "C \ {z \ topspace X. f z = f x} \ U2" proof (rule Hausdorff_space_compact_separation [of X "{x}" "C \ {z \ topspace X. f z = f x}", OF \Hausdorff_space X\]) show "compactin X {x}" by (simp add: \x \ topspace X\) show "compactin X (C \ {z \ topspace X. f z = f x})" using \closedin X C\ fim \x \ topspace X\ closed_Int_compactin preYf by fastforce show "disjnt {x} (C \ {z \ topspace X. f z = f x})" using \x \ C\ by force qed have "closedin Y (f ` (C - U2))" using \closed_map X Y f\ \closedin X C\ \openin X U2\ closed_map_def by blast moreover have "f x \ topspace Y - f ` (C - U2)" using \closedin X C\ \continuous_map X Y f\ \x \ topspace X\ closedin_subset continuous_map_def subV by fastforce ultimately obtain V1 V2 where VV: "openin Y V1" "openin Y V2" "f x \ V1" and subV2: "f ` (C - U2) \ V2" and "disjnt V1 V2" by (meson R regular_space_def) show "\U U'. openin X U \ openin X U' \ x \ U \ C \ U' \ disjnt U U'" proof (intro exI conjI) show "openin X (U1 \ {x \ topspace X. f x \ V1})" using VV(1) \continuous_map X Y f\ \openin X U1\ continuous_map by fastforce show "openin X (U2 \ {x \ topspace X. f x \ V2})" using VV(2) \continuous_map X Y f\ \openin X U2\ continuous_map by fastforce show "x \ U1 \ {x \ topspace X. f x \ V1}" using VV(3) \x \ topspace X\ \{x} \ U1\ by auto show "C \ U2 \ {x \ topspace X. f x \ V2}" using \closedin X C\ closedin_subset subV2 by auto show "disjnt (U1 \ {x \ topspace X. f x \ V1}) (U2 \ {x \ topspace X. f x \ V2})" using \disjnt U1 U2\ \disjnt V1 V2\ by (auto simp: disjnt_iff) qed qed qed subsection\Locally compact spaces\ definition locally_compact_space where "locally_compact_space X \ \x \ topspace X. \U K. openin X U \ compactin X K \ x \ U \ U \ K" lemma homeomorphic_locally_compact_spaceD: assumes X: "locally_compact_space X" and "X homeomorphic_space Y" shows "locally_compact_space Y" proof - obtain f where hmf: "homeomorphic_map X Y f" using assms homeomorphic_space by blast then have eq: "topspace Y = f ` (topspace X)" by (simp add: homeomorphic_imp_surjective_map) have "\V K. openin Y V \ compactin Y K \ f x \ V \ V \ K" if "x \ topspace X" "openin X U" "compactin X K" "x \ U" "U \ K" for x U K using that by (meson hmf homeomorphic_map_compactness_eq homeomorphic_map_openness_eq image_mono image_eqI) with X show ?thesis by (smt (verit) eq image_iff locally_compact_space_def) qed lemma homeomorphic_locally_compact_space: assumes "X homeomorphic_space Y" shows "locally_compact_space X \ locally_compact_space Y" by (meson assms homeomorphic_locally_compact_spaceD homeomorphic_space_sym) lemma locally_compact_space_retraction_map_image: assumes "retraction_map X Y r" and X: "locally_compact_space X" shows "locally_compact_space Y" proof - obtain s where s: "retraction_maps X Y r s" using assms retraction_map_def by blast obtain T where "T retract_of_space X" and Teq: "T = s ` topspace Y" using retraction_maps_section_image1 s by blast then obtain r where r: "continuous_map X (subtopology X T) r" "\x\T. r x = x" by (meson retract_of_space_def) have "locally_compact_space (subtopology X T)" unfolding locally_compact_space_def openin_subtopology_alt proof clarsimp fix x assume "x \ topspace X" "x \ T" obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" by (meson X \x \ topspace X\ locally_compact_space_def) then have "compactin (subtopology X T) (r ` K) \ T \ U \ r ` K" by (smt (verit) IntD1 image_compactin image_iff inf_le2 r subset_iff) then show "\U. openin X U \ (\K. compactin (subtopology X T) K \ x \ U \ T \ U \ K)" using UK by auto qed with Teq show ?thesis using homeomorphic_locally_compact_space retraction_maps_section_image2 s by blast qed lemma compact_imp_locally_compact_space: "compact_space X \ locally_compact_space X" using compact_space_def locally_compact_space_def by blast lemma neighbourhood_base_imp_locally_compact_space: "neighbourhood_base_of (compactin X) X \ locally_compact_space X" by (metis locally_compact_space_def neighbourhood_base_of openin_topspace) lemma locally_compact_imp_neighbourhood_base: assumes loc: "locally_compact_space X" and reg: "regular_space X" shows "neighbourhood_base_of (compactin X) X" unfolding neighbourhood_base_of proof clarify fix W x assume "openin X W" and "x \ W" then obtain U K where "openin X U" "compactin X K" "x \ U" "U \ K" by (metis loc locally_compact_space_def openin_subset subsetD) moreover have "openin X (U \ W) \ x \ U \ W" using \openin X W\ \x \ W\ \openin X U\ \x \ U\ by blast then have "\u' v. openin X u' \ closedin X v \ x \ u' \ u' \ v \ v \ U \ v \ W" using reg by (metis le_infE neighbourhood_base_of neighbourhood_base_of_closedin) then show "\U V. openin X U \ compactin X V \ x \ U \ U \ V \ V \ W" by (meson \U \ K\ \compactin X K\ closed_compactin subset_trans) qed lemma Hausdorff_regular: "\Hausdorff_space X; neighbourhood_base_of (compactin X) X\ \ regular_space X" by (metis compactin_imp_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono) lemma locally_compact_Hausdorff_imp_regular_space: assumes loc: "locally_compact_space X" and "Hausdorff_space X" shows "regular_space X" unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of proof clarify fix W x assume "openin X W" and "x \ W" then have "x \ topspace X" using openin_subset by blast then obtain U K where "openin X U" "compactin X K" and UK: "x \ U" "U \ K" by (meson loc locally_compact_space_def) with \Hausdorff_space X\ have "regular_space (subtopology X K)" using Hausdorff_space_subtopology compact_Hausdorff_imp_regular_space compact_space_subtopology by blast then have "\U' V'. openin (subtopology X K) U' \ closedin (subtopology X K) V' \ x \ U' \ U' \ V' \ V' \ K \ W" unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of by (meson IntI \U \ K\ \openin X W\ \x \ U\ \x \ W\ openin_subtopology_Int2 subsetD) then obtain V C where "openin X V" "closedin X C" and VC: "x \ K \ V" "K \ V \ K \ C" "K \ C \ K \ W" by (metis Int_commute closedin_subtopology openin_subtopology) show "\U V. openin X U \ closedin X V \ x \ U \ U \ V \ V \ W" proof (intro conjI exI) show "openin X (U \ V)" using \openin X U\ \openin X V\ by blast show "closedin X (K \ C)" using \closedin X C\ \compactin X K\ compactin_imp_closedin \Hausdorff_space X\ by blast qed (use UK VC in auto) qed lemma locally_compact_space_neighbourhood_base: "Hausdorff_space X \ regular_space X \ locally_compact_space X \ neighbourhood_base_of (compactin X) X" by (metis locally_compact_imp_neighbourhood_base locally_compact_Hausdorff_imp_regular_space neighbourhood_base_imp_locally_compact_space) lemma locally_compact_Hausdorff_or_regular: "locally_compact_space X \ (Hausdorff_space X \ regular_space X) \ locally_compact_space X \ regular_space X" using locally_compact_Hausdorff_imp_regular_space by blast lemma locally_compact_space_compact_closedin: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\x \ topspace X. \U K. openin X U \ compactin X K \ closedin X K \ x \ U \ U \ K)" using locally_compact_Hausdorff_or_regular unfolding locally_compact_space_def by (metis assms closed_compactin inf.absorb_iff2 le_infE neighbourhood_base_of neighbourhood_base_of_closedin) lemma locally_compact_space_compact_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\x \ topspace X. \U. openin X U \ compactin X (X closure_of U) \ x \ U)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis assms closed_compactin closedin_closure_of closure_of_eq closure_of_mono locally_compact_space_compact_closedin) next assume ?rhs then show ?lhs by (meson closure_of_subset locally_compact_space_def openin_subset) qed lemma locally_compact_space_neighbourhood_base_closedin: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast with L have "neighbourhood_base_of (compactin X) X" by (simp add: locally_compact_imp_neighbourhood_base) with \regular_space X\ show ?rhs by (smt (verit, ccfv_threshold) closed_compactin neighbourhood_base_of subset_trans neighbourhood_base_of_closedin) next assume ?rhs then show ?lhs using neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_mono by blast qed lemma locally_compact_space_neighbourhood_base_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\T. compactin X (X closure_of T)) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast with L have "neighbourhood_base_of (\A. compactin X A \ closedin X A) X" using locally_compact_space_neighbourhood_base_closedin by blast then show ?rhs by (simp add: closure_of_closedin neighbourhood_base_of_mono) next assume ?rhs then show ?lhs unfolding locally_compact_space_def neighbourhood_base_of by (meson closure_of_subset openin_topspace subset_trans) qed lemma locally_compact_space_neighbourhood_base_open_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\U. openin X U \ compactin X (X closure_of U)) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast then have "neighbourhood_base_of (\T. compactin X (X closure_of T)) X" using L locally_compact_space_neighbourhood_base_closure_of by auto with L show ?rhs unfolding neighbourhood_base_of by (meson closed_compactin closure_of_closure_of closure_of_eq closure_of_mono subset_trans) next assume ?rhs then show ?lhs unfolding locally_compact_space_def neighbourhood_base_of by (meson closure_of_subset openin_topspace subset_trans) qed lemma locally_compact_space_compact_closed_compact: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\K. compactin X K \ (\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L))" (is "?lhs=?rhs") proof assume L: ?lhs then obtain U L where UL: "\x \ topspace X. openin X (U x) \ compactin X (L x) \ closedin X (L x) \ x \ U x \ U x \ L x" unfolding locally_compact_space_compact_closedin [OF assms] by metis show ?rhs proof clarify fix K assume "compactin X K" then have "K \ topspace X" by (simp add: compactin_subset_topspace) then have *: "(\U\U ` K. openin X U) \ K \ \ (U ` K)" using UL by blast with \compactin X K\ obtain KF where KF: "finite KF" "KF \ K" "K \ \(U ` KF)" by (metis compactinD finite_subset_image) show "\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L" proof (intro conjI exI) show "openin X (\ (U ` KF))" using "*" \KF \ K\ by fastforce show "compactin X (\ (L ` KF))" by (smt (verit) UL \K \ topspace X\ KF compactin_Union finite_imageI imageE subset_iff) show "closedin X (\ (L ` KF))" by (smt (verit) UL \K \ topspace X\ KF closedin_Union finite_imageI imageE subsetD) qed (use UL \K \ topspace X\ KF in auto) qed next assume ?rhs then show ?lhs by (metis compactin_sing insert_subset locally_compact_space_def) qed lemma locally_compact_regular_space_neighbourhood_base: "locally_compact_space X \ regular_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" using locally_compact_space_neighbourhood_base_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono by blast lemma locally_compact_kc_space: "neighbourhood_base_of (compactin X) X \ kc_space X \ locally_compact_space X \ Hausdorff_space X" using Hausdorff_imp_kc_space locally_compact_imp_kc_eq_Hausdorff_space locally_compact_space_neighbourhood_base by blast lemma locally_compact_kc_space_alt: "neighbourhood_base_of (compactin X) X \ kc_space X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" using Hausdorff_regular locally_compact_kc_space by blast lemma locally_compact_kc_imp_regular_space: "\neighbourhood_base_of (compactin X) X; kc_space X\ \ regular_space X" using Hausdorff_regular locally_compact_imp_kc_eq_Hausdorff_space by blast lemma kc_locally_compact_space: "kc_space X \ neighbourhood_base_of (compactin X) X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" using Hausdorff_regular locally_compact_kc_space by blast lemma locally_compact_space_closed_subset: assumes loc: "locally_compact_space X" and "closedin X S" shows "locally_compact_space (subtopology X S)" proof (clarsimp simp: locally_compact_space_def) fix x assume x: "x \ topspace X" "x \ S" then obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" by (meson loc locally_compact_space_def) show "\U. openin (subtopology X S) U \ (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" proof (intro conjI exI) show "openin (subtopology X S) (S \ U)" by (simp add: UK openin_subtopology_Int2) show "compactin (subtopology X S) (S \ K)" by (simp add: UK assms(2) closed_Int_compactin compactin_subtopology) qed (use UK x in auto) qed lemma locally_compact_space_open_subset: assumes reg: "regular_space X" and loc: "locally_compact_space X" and "openin X S" shows "locally_compact_space (subtopology X S)" proof (clarsimp simp: locally_compact_space_def) fix x assume x: "x \ topspace X" "x \ S" then obtain U K where UK: "openin X U" "compactin X K" "x \ U" "U \ K" by (meson loc locally_compact_space_def) have "openin X (U \ S)" by (simp add: UK \openin X S\ openin_Int) with UK reg x obtain V C where VC: "openin X V" "closedin X C" "x \ V" "V \ C" "C \ U" "C \ S" by (metis IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin) show "\U. openin (subtopology X S) U \ (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" proof (intro conjI exI) show "openin (subtopology X S) V" using VC by (meson \openin X S\ openin_open_subtopology order_trans) show "compactin (subtopology X S) (C \ K)" using UK VC closed_Int_compactin compactin_subtopology by fastforce qed (use UK VC x in auto) qed lemma locally_compact_space_discrete_topology: "locally_compact_space (discrete_topology U)" by (simp add: neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_discrete_topology) lemma locally_compact_space_continuous_open_map_image: "\continuous_map X X' f; open_map X X' f; f ` topspace X = topspace X'; locally_compact_space X\ \ locally_compact_space X'" unfolding locally_compact_space_def open_map_def by (smt (verit, ccfv_SIG) image_compactin image_iff image_mono) lemma locally_compact_subspace_openin_closure_of: assumes "Hausdorff_space X" and S: "S \ topspace X" and loc: "locally_compact_space (subtopology X S)" shows "openin (subtopology X (X closure_of S)) S" unfolding openin_subopen [where S=S] proof clarify fix a assume "a \ S" then obtain T K where *: "openin X T" "compactin X K" "K \ S" "a \ S" "a \ T" "S \ T \ K" using loc unfolding locally_compact_space_def by (metis IntE S compactin_subtopology inf_commute openin_subtopology topspace_subtopology_subset) have "T \ X closure_of S \ X closure_of (T \ S)" by (simp add: "*"(1) openin_Int_closure_of_subset) also have "... \ S" using * \Hausdorff_space X\ by (metis closure_of_minimal compactin_imp_closedin order.trans inf_commute) finally have "T \ X closure_of S \ T \ S" by simp then have "openin (subtopology X (X closure_of S)) (T \ S)" unfolding openin_subtopology using \openin X T\ S closure_of_subset by fastforce with * show "\T. openin (subtopology X (X closure_of S)) T \ a \ T \ T \ S" by blast qed lemma locally_compact_subspace_closed_Int_openin: "\Hausdorff_space X \ S \ topspace X \ locally_compact_space(subtopology X S)\ \ \C U. closedin X C \ openin X U \ C \ U = S" by (metis closedin_closure_of inf_commute locally_compact_subspace_openin_closure_of openin_subtopology) lemma locally_compact_subspace_open_in_closure_of_eq: assumes "Hausdorff_space X" and loc: "locally_compact_space X" shows "openin (subtopology X (X closure_of S)) S \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") proof assume L: ?lhs then obtain "S \ topspace X" "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space openin_subset by fastforce then have "locally_compact_space (subtopology (subtopology X (X closure_of S)) S)" by (simp add: L loc locally_compact_space_closed_subset locally_compact_space_open_subset regular_space_subtopology) then show ?rhs by (metis L inf.orderE inf_commute le_inf_iff openin_subset subtopology_subtopology topspace_subtopology) next assume ?rhs then show ?lhs using assms locally_compact_subspace_openin_closure_of by blast qed lemma locally_compact_subspace_closed_Int_openin_eq: assumes "Hausdorff_space X" and loc: "locally_compact_space X" shows "(\C U. closedin X C \ openin X U \ C \ U = S) \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") proof assume L: ?lhs then obtain C U where "closedin X C" "openin X U" and Seq: "S = C \ U" by blast then have "C \ topspace X" by (simp add: closedin_subset) have "locally_compact_space (subtopology (subtopology X C) (topspace (subtopology X C) \ U))" proof (rule locally_compact_space_open_subset) show "regular_space (subtopology X C)" by (simp add: \Hausdorff_space X\ loc locally_compact_Hausdorff_imp_regular_space regular_space_subtopology) show "locally_compact_space (subtopology X C)" by (simp add: \closedin X C\ loc locally_compact_space_closed_subset) show "openin (subtopology X C) (topspace (subtopology X C) \ U)" by (simp add: \openin X U\ Int_left_commute inf_commute openin_Int openin_subtopology_Int2) qed then show ?rhs by (metis Seq \C \ topspace X\ inf.coboundedI1 subtopology_subtopology subtopology_topspace) next assume ?rhs then show ?lhs using assms locally_compact_subspace_closed_Int_openin by blast qed lemma dense_locally_compact_openin_Hausdorff_space: "\Hausdorff_space X; S \ topspace X; X closure_of S = topspace X; locally_compact_space (subtopology X S)\ \ openin X S" by (metis locally_compact_subspace_openin_closure_of subtopology_topspace) lemma locally_compact_space_prod_topology: "locally_compact_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ locally_compact_space X \ locally_compact_space Y" (is "?lhs=?rhs") proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis unfolding locally_compact_space_def by blast next case False then obtain w z where wz: "w \ topspace X" "z \ topspace Y" by auto show ?thesis proof assume L: ?lhs then show ?rhs by (metis wz empty_iff locally_compact_space_retraction_map_image retraction_map_fst retraction_map_snd) next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix x y assume "x \ topspace X" and "y \ topspace Y" obtain U C where "openin X U" "compactin X C" "x \ U" "U \ C" by (meson False R \x \ topspace X\ locally_compact_space_def) obtain V D where "openin Y V" "compactin Y D" "y \ V" "V \ D" by (meson False R \y \ topspace Y\ locally_compact_space_def) show "\U. openin (prod_topology X Y) U \ (\K. compactin (prod_topology X Y) K \ (x, y) \ U \ U \ K)" proof (intro exI conjI) show "openin (prod_topology X Y) (U \ V)" by (simp add: \openin X U\ \openin Y V\ openin_prod_Times_iff) show "compactin (prod_topology X Y) (C \ D)" by (simp add: \compactin X C\ \compactin Y D\ compactin_Times) show "(x, y) \ U \ V" by (simp add: \x \ U\ \y \ V\) show "U \ V \ C \ D" by (simp add: Sigma_mono \U \ C\ \V \ D\) qed qed qed qed lemma locally_compact_space_product_topology: "locally_compact_space(product_topology X I) \ topspace(product_topology X I) = {} \ finite {i \ I. \ compact_space(X i)} \ (\i \ I. locally_compact_space(X i))" (is "?lhs=?rhs") proof (cases "topspace (product_topology X I) = {}") case True then show ?thesis by (simp add: locally_compact_space_def) next case False show ?thesis proof assume L: ?lhs obtain z where z: "z \ topspace (product_topology X I)" using False by auto with L z obtain U C where "openin (product_topology X I) U" "compactin (product_topology X I) C" "z \ U" "U \ C" by (meson locally_compact_space_def) then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" and "\i \ I. openin (X i) (V i)" and "z \ PiE I V" "PiE I V \ U" by (auto simp: openin_product_topology_alt) have "compact_space (X i)" if "i \ I" "V i = topspace (X i)" for i proof - have "compactin (X i) ((\x. x i) ` C)" using \compactin (product_topology X I) C\ image_compactin by (metis continuous_map_product_projection \i \ I\) moreover have "V i \ (\x. x i) ` C" proof - have "V i \ (\x. x i) ` Pi\<^sub>E I V" by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl \i \ I\) also have "\ \ (\x. x i) ` C" using \U \ C\ \Pi\<^sub>E I V \ U\ by blast finally show ?thesis . qed ultimately show ?thesis by (metis closed_compactin closedin_topspace compact_space_def that(2)) qed with finV have "finite {i \ I. \ compact_space (X i)}" by (metis (mono_tags, lifting) mem_Collect_eq finite_subset subsetI) moreover have "locally_compact_space (X i)" if "i \ I" for i by (meson False L locally_compact_space_retraction_map_image retraction_map_product_projection that) ultimately show ?rhs by metis next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix z assume z: "z \ (\\<^sub>E i\I. topspace (X i))" have "\U C. openin (X i) U \ compactin (X i) C \ z i \ U \ U \ C \ (compact_space(X i) \ U = topspace(X i) \ C = topspace(X i))" if "i \ I" for i using that R z unfolding locally_compact_space_def compact_space_def by (metis (no_types, lifting) False PiE_mem openin_topspace set_eq_subset) then obtain U C where UC: "\i. i \ I \ openin (X i) (U i) \ compactin (X i) (C i) \ z i \ U i \ U i \ C i \ (compact_space(X i) \ U i = topspace(X i) \ C i = topspace(X i))" by metis show "\U. openin (product_topology X I) U \ (\K. compactin (product_topology X I) K \ z \ U \ U \ K)" proof (intro exI conjI) show "openin (product_topology X I) (Pi\<^sub>E I U)" by (smt (verit) Collect_cong False R UC compactin_subspace openin_PiE_gen subset_antisym subtopology_topspace) show "compactin (product_topology X I) (Pi\<^sub>E I C)" by (simp add: UC compactin_PiE) qed (use UC z in blast)+ qed qed qed lemma locally_compact_space_sum_topology: "locally_compact_space (sum_topology X I) \ (\i \ I. locally_compact_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis closed_map_component_injection embedding_map_imp_homeomorphic_space embedding_map_component_injection embedding_imp_closed_map_eq homeomorphic_locally_compact_space locally_compact_space_closed_subset) next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix i y assume "i \ I" and y: "y \ topspace (X i)" then obtain U K where UK: "openin (X i) U" "compactin (X i) K" "y \ U" "U \ K" using R by (fastforce simp: locally_compact_space_def) then show "\U. openin (sum_topology X I) U \ (\K. compactin (sum_topology X I) K \ (i, y) \ U \ U \ K)" by (metis \i \ I\ continuous_map_component_injection image_compactin image_mono imageI open_map_component_injection open_map_def) qed qed proposition quotient_map_prod_right: assumes loc: "locally_compact_space Z" and reg: "Hausdorff_space Z \ regular_space Z" and f: "quotient_map X Y f" shows "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x,y). (x,f y))" proof - define h where "h \ (\(x::'a,y). (x,f y))" have "continuous_map (prod_topology Z X) Y (f o snd)" by (simp add: continuous_map_of_snd f quotient_imp_continuous_map) then have cmh: "continuous_map (prod_topology Z X) (prod_topology Z Y) h" by (simp add: h_def continuous_map_paired split_def continuous_map_fst o_def) have fim: "f ` topspace X = topspace Y" by (simp add: f quotient_imp_surjective_map) moreover have "openin (prod_topology Z X) {u \ topspace Z \ topspace X. h u \ W} \ openin (prod_topology Z Y) W" (is "?lhs=?rhs") if W: "W \ topspace Z \ topspace Y" for W proof define S where "S \ {u \ topspace Z \ topspace X. h u \ W}" assume ?lhs then have L: "openin (prod_topology Z X) S" using S_def by blast have "\T. openin (prod_topology Z Y) T \ (x0, z0) \ T \ T \ W" if \
: "(x0,z0) \ W" for x0 z0 proof - have x0: "x0 \ topspace Z" using W that by blast obtain y0 where y0: "y0 \ topspace X" "f y0 = z0" by (metis W fim imageE insert_absorb insert_subset mem_Sigma_iff \
) then have "(x0, y0) \ S" by (simp add: S_def h_def that x0) have "continuous_map Z (prod_topology Z X) (\x. (x, y0))" by (simp add: continuous_map_paired y0) with openin_continuous_map_preimage [OF _ L] have ope_ZS: "openin Z {x \ topspace Z. (x,y0) \ S}" by blast obtain U U' where "openin Z U" "compactin Z U'" "closedin Z U'" "x0 \ U" "U \ U'" "U' \ {x \ topspace Z. (x,y0) \ S}" using loc ope_ZS x0 \(x0, y0) \ S\ by (force simp: locally_compact_space_neighbourhood_base_closedin [OF reg] neighbourhood_base_of) then have D: "U' \ {y0} \ S" by (auto simp: ) define V where "V \ {z \ topspace Y. U' \ {y \ topspace X. f y = z} \ S}" have "z0 \ V" using D y0 Int_Collect fim by (fastforce simp: h_def V_def S_def) have "openin X {x \ topspace X. f x \ V} \ openin Y V" using f unfolding V_def quotient_map_def subset_iff by (smt (verit, del_insts) Collect_cong mem_Collect_eq) moreover have "openin X {x \ topspace X. f x \ V}" proof - let ?Z = "subtopology Z U'" have *: "{x \ topspace X. f x \ V} = topspace X - snd ` (U' \ topspace X - S)" by (force simp: V_def S_def h_def simp flip: fim) have "compact_space ?Z" using \compactin Z U'\ compactin_subspace by auto moreover have "closedin (prod_topology ?Z X) (U' \ topspace X - S)" by (simp add: L \closedin Z U'\ closedin_closed_subtopology closedin_diff closedin_prod_Times_iff prod_topology_subtopology(1)) ultimately show ?thesis using "*" closed_map_snd closed_map_def by fastforce qed ultimately have "openin Y V" by metis show ?thesis proof (intro conjI exI) show "openin (prod_topology Z Y) (U \ V)" by (simp add: openin_prod_Times_iff \openin Z U\ \openin Y V\) show "(x0, z0) \ U \ V" by (simp add: \x0 \ U\ \z0 \ V\) show "U \ V \ W" using \U \ U'\ by (force simp: V_def S_def h_def simp flip: fim) qed qed with openin_subopen show ?rhs by force next assume ?rhs then show ?lhs using openin_continuous_map_preimage cmh by fastforce qed ultimately show ?thesis by (fastforce simp: image_iff quotient_map_def h_def) qed lemma quotient_map_prod_left: assumes loc: "locally_compact_space Z" and reg: "Hausdorff_space Z \ regular_space Z" and f: "quotient_map X Y f" shows "quotient_map (prod_topology X Z) (prod_topology Y Z) (\(x,y). (f x,y))" proof - have "(\(x,y). (f x,y)) = prod.swap \ (\(x,y). (x,f y)) \ prod.swap" by force then show ?thesis apply (rule ssubst) proof (intro quotient_map_compose) show "quotient_map (prod_topology X Z) (prod_topology Z X) prod.swap" "quotient_map (prod_topology Z Y) (prod_topology Y Z) prod.swap" using homeomorphic_map_def homeomorphic_map_swap quotient_map_eq by fastforce+ show "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x, y). (x, f y))" by (simp add: f loc quotient_map_prod_right reg) qed qed lemma locally_compact_space_perfect_map_preimage: assumes "locally_compact_space X'" and f: "perfect_map X X' f" shows "locally_compact_space X" unfolding locally_compact_space_def proof (intro strip) fix x assume x: "x \ topspace X" then obtain U K where "openin X' U" "compactin X' K" "f x \ U" "U \ K" using assms unfolding locally_compact_space_def perfect_map_def by (metis (no_types, lifting) continuous_map_closedin) show "\U K. openin X U \ compactin X K \ x \ U \ U \ K" proof (intro exI conjI) have "continuous_map X X' f" using f perfect_map_def by blast then show "openin X {x \ topspace X. f x \ U}" by (simp add: \openin X' U\ continuous_map) show "compactin X {x \ topspace X. f x \ K}" using \compactin X' K\ f perfect_imp_proper_map proper_map_alt by blast qed (use x \f x \ U\ \U \ K\ in auto) qed subsection\Special characterizations of classes of functions into and out of R\ lemma monotone_map_into_euclideanreal_alt: assumes "continuous_map X euclideanreal f" shows "(\k. is_interval k \ connectedin X {x \ topspace X. f x \ k}) \ connected_space X \ monotone_map X euclideanreal f" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof show "connected_space X" using L connected_space_subconnected by blast have "connectedin X {x \ topspace X. f x \ {y}}" for y by (metis L is_interval_1 nle_le singletonD) then show "monotone_map X euclideanreal f" by (simp add: monotone_map) qed next assume R: ?rhs then have *: False if "a < b" "closedin X U" "closedin X V" "U \ {}" "V \ {}" "disjnt U V" and UV: "{x \ topspace X. f x \ {a..b}} = U \ V" and dis: "disjnt U {x \ topspace X. f x = b}" "disjnt V {x \ topspace X. f x = a}" for a b U V proof - define E1 where "E1 \ U \ {x \ topspace X. f x \ {c. c \ a}}" define E2 where "E2 \ V \ {x \ topspace X. f x \ {c. b \ c}}" have "closedin X {x \ topspace X. f x \ a}" "closedin X {x \ topspace X. b \ f x}" using assms continuous_map_upper_lower_semicontinuous_le by blast+ then have "closedin X E1" "closedin X E2" unfolding E1_def E2_def using that by auto moreover have "E1 \ E2 = {}" unfolding E1_def E2_def using \a \disjnt U V\ dis UV by (simp add: disjnt_def set_eq_iff) (smt (verit)) have "topspace X \ E1 \ E2" unfolding E1_def E2_def using UV by fastforce have "E1 = {} \ E2 = {}" using R connected_space_closedin using \E1 \ E2 = {}\ \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ by blast then show False using E1_def E2_def \U \ {}\ \V \ {}\ by fastforce qed show ?lhs proof (intro strip) fix K :: "real set" assume "is_interval K" have False if "a \ K" "b \ K" and clo: "closedin X U" "closedin X V" and UV: "{x. x \ topspace X \ f x \ K} \ U \ V" "U \ V \ {x. x \ topspace X \ f x \ K} = {}" and nondis: "\ disjnt U {x. x \ topspace X \ f x = a}" "\ disjnt V {x. x \ topspace X \ f x = b}" for a b U V proof - have "\y. connectedin X {x. x \ topspace X \ f x = y}" using R monotone_map by fastforce then have **: False if "p \ U \ q \ V \ f p = f q \ f q \ K" for p q unfolding connectedin_closedin using \a \ K\ \b \ K\ UV clo that by (smt (verit, ccfv_threshold) closedin_subset disjoint_iff mem_Collect_eq subset_iff) consider "a < b" | "a = b" | "b < a" by linarith then show ?thesis proof cases case 1 define W where "W \ {x \ topspace X. f x \ {a..b}}" have "closedin X W" unfolding W_def by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) show ?thesis proof (rule * [OF 1 , of "U \ W" "V \ W"]) show "closedin X (U \ W)" "closedin X (V \ W)" using \closedin X W\ clo by auto show "U \ W \ {}" "V \ W \ {}" using nondis 1 by (auto simp: disjnt_iff W_def) show "disjnt (U \ W) (V \ W)" using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) have "\x. \x \ topspace X; a \ f x; f x \ b\ \ x \ U \ x \ V" using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff by blast then show "{x \ topspace X. f x \ {a..b}} = U \ W \ V \ W" by (auto simp: W_def) show "disjnt (U \ W) {x \ topspace X. f x = b}" "disjnt (V \ W) {x \ topspace X. f x = a}" using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ qed next case 2 then show ?thesis using ** nondis \b \ K\ by (force simp add: disjnt_iff) next case 3 define W where "W \ {x \ topspace X. f x \ {b..a}}" have "closedin X W" unfolding W_def by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) show ?thesis proof (rule * [OF 3, of "V \ W" "U \ W"]) show "closedin X (U \ W)" "closedin X (V \ W)" using \closedin X W\ clo by auto show "U \ W \ {}" "V \ W \ {}" using nondis 3 by (auto simp: disjnt_iff W_def) show "disjnt (V \ W) (U \ W)" using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) have "\x. \x \ topspace X; b \ f x; f x \ a\ \ x \ U \ x \ V" using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff by blast then show "{x \ topspace X. f x \ {b..a}} = V \ W \ U \ W" by (auto simp: W_def) show "disjnt (V \ W) {x \ topspace X. f x = a}" "disjnt (U \ W) {x \ topspace X. f x = b}" using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ qed qed qed then show "connectedin X {x \ topspace X. f x \ K}" unfolding connectedin_closedin disjnt_iff by blast qed qed lemma monotone_map_into_euclideanreal: "\connected_space X; continuous_map X euclideanreal f\ \ monotone_map X euclideanreal f \ (\k. is_interval k \ connectedin X {x \ topspace X. f x \ k})" by (simp add: monotone_map_into_euclideanreal_alt) lemma monotone_map_euclideanreal_alt: "(\I::real set. is_interval I \ is_interval {x::real. x \ S \ f x \ I}) \ is_interval S \ (mono_on S f \ antimono_on S f)" (is "?lhs=?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof show "is_interval S" using L is_interval_1 by auto have False if "a \ S" "b \ S" "c \ S" "a f c < f b \ f a > f b \ f c > f b" for a b c using d proof assume "f a < f b \ f c < f b" then show False using L [of "{y. y < f b}"] unfolding is_interval_1 by (smt (verit, best) mem_Collect_eq that) next assume "f b < f a \ f b < f c" then show False using L [of "{y. y > f b}"] unfolding is_interval_1 by (smt (verit, best) mem_Collect_eq that) qed then show "mono_on S f \ monotone_on S (\) (\) f" unfolding monotone_on_def by (smt (verit)) qed next assume ?rhs then show ?lhs unfolding is_interval_1 monotone_on_def by simp meson qed lemma monotone_map_euclideanreal: fixes S :: "real set" shows "\is_interval S; continuous_on S f\ \ monotone_map (top_of_set S) euclideanreal f \ (mono_on S f \ monotone_on S (\) (\) f)" using monotone_map_euclideanreal_alt by (simp add: monotone_map_into_euclideanreal connectedin_subtopology is_interval_connected_1) lemma injective_eq_monotone_map: fixes f :: "real \ real" assumes "is_interval S" "continuous_on S f" shows "inj_on f S \ strict_mono_on S f \ strict_antimono_on S f" by (metis assms injective_imp_monotone_map monotone_map_euclideanreal strict_antimono_iff_antimono strict_mono_iff_mono top_greatest topspace_euclidean topspace_euclidean_subtopology) subsection\Normal spaces including Urysohn's lemma and the Tietze extension theorem\ definition normal_space where "normal_space X \ \S T. closedin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V)" lemma normal_space_retraction_map_image: assumes r: "retraction_map X Y r" and X: "normal_space X" shows "normal_space Y" unfolding normal_space_def proof clarify fix S T assume "closedin Y S" and "closedin Y T" and "disjnt S T" obtain r' where r': "retraction_maps X Y r r'" using r retraction_map_def by blast have "closedin X {x \ topspace X. r x \ S}" "closedin X {x \ topspace X. r x \ T}" using closedin_continuous_map_preimage \closedin Y S\ \closedin Y T\ r' by (auto simp: retraction_maps_def) moreover have "disjnt {x \ topspace X. r x \ S} {x \ topspace X. r x \ T}" using \disjnt S T\ by (auto simp: disjnt_def) ultimately obtain U V where UV: "openin X U \ openin X V \ {x \ topspace X. r x \ S} \ U \ {x \ topspace X. r x \ T} \ V" "disjnt U V" by (meson X normal_space_def) show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" proof (intro exI conjI) show "openin Y {x \ topspace Y. r' x \ U}" "openin Y {x \ topspace Y. r' x \ V}" using openin_continuous_map_preimage UV r' by (auto simp: retraction_maps_def) show "S \ {x \ topspace Y. r' x \ U}" "T \ {x \ topspace Y. r' x \ V}" using openin_continuous_map_preimage UV r' \closedin Y S\ \closedin Y T\ by (auto simp add: closedin_def continuous_map_closedin retraction_maps_def subset_iff) show "disjnt {x \ topspace Y. r' x \ U} {x \ topspace Y. r' x \ V}" using \disjnt U V\ by (auto simp: disjnt_def) qed qed lemma homeomorphic_normal_space: "X homeomorphic_space Y \ normal_space X \ normal_space Y" unfolding homeomorphic_space_def by (meson homeomorphic_imp_retraction_maps homeomorphic_maps_sym normal_space_retraction_map_image retraction_map_def) lemma normal_space: "normal_space X \ (\S T. closedin X S \ closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U)))" proof - have "(\V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V) \ openin X U \ S \ U \ disjnt T (X closure_of U)" (is "?lhs=?rhs") if "closedin X S" "closedin X T" "disjnt S T" for S T U proof show "?lhs \ ?rhs" by (smt (verit, best) disjnt_iff in_closure_of subsetD) assume R: ?rhs then have "(U \ S) \ (topspace X - X closure_of U) = {}" by (metis Diff_eq_empty_iff Int_Diff Int_Un_eq(4) closure_of_subset inf.orderE openin_subset) moreover have "T \ topspace X - X closure_of U" by (meson DiffI R closedin_subset disjnt_iff subsetD subsetI that(2)) ultimately show ?lhs by (metis R closedin_closure_of closedin_def disjnt_def sup.orderE) qed then show ?thesis unfolding normal_space_def by meson qed lemma normal_space_alt: "normal_space X \ (\S U. closedin X S \ openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U))" proof - have "\V. openin X V \ S \ V \ X closure_of V \ U" if "\T. closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U))" "closedin X S" "openin X U" "S \ U" for S U using that by (smt (verit) Diff_eq_empty_iff Int_Diff closure_of_subset_topspace disjnt_def inf.orderE inf_commute openin_closedin_eq) moreover have "\U. openin X U \ S \ U \ disjnt T (X closure_of U)" if "\U. openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U)" and "closedin X S" "closedin X T" "disjnt S T" for S T using that by (smt (verit) Diff_Diff_Int Diff_eq_empty_iff Int_Diff closedin_def disjnt_def inf.absorb_iff2 inf.orderE) ultimately show ?thesis by (fastforce simp: normal_space) qed lemma normal_space_closures: "normal_space X \ (\S T. S \ topspace X \ T \ topspace X \ disjnt (X closure_of S) (X closure_of T) \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (meson closedin_closure_of closure_of_subset normal_space_def order.trans) show "?rhs \ ?lhs" by (metis closedin_subset closure_of_eq normal_space_def) qed lemma normal_space_disjoint_closures: "normal_space X \ (\S T. closedin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt (X closure_of U) (X closure_of V)))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (metis closedin_closure_of normal_space) show "?rhs \ ?lhs" by (smt (verit) closure_of_subset disjnt_iff normal_space openin_subset subset_eq) qed lemma normal_space_dual: "normal_space X \ (\U V. openin X U \ openin X V \ U \ V = topspace X \ (\S T. closedin X S \ closedin X T \ S \ U \ T \ V \ S \ T = topspace X))" (is "_ = ?rhs") proof - have "normal_space X \ (\U V. closedin X U \ closedin X V \ disjnt U V \ (\S T. \ (openin X S \ openin X T \ \ (U \ S \ V \ T \ disjnt S T))))" unfolding normal_space_def by meson also have "... \ (\U V. openin X U \ openin X V \ disjnt (topspace X - U) (topspace X - V) \ (\S T. \ (openin X S \ openin X T \ \ (topspace X - U \ S \ topspace X - V \ T \ disjnt S T))))" by (auto simp: all_closedin) also have "... \ ?rhs" proof - have *: "disjnt (topspace X - U) (topspace X - V) \ U \ V = topspace X" if "U \ topspace X" "V \ topspace X" for U V using that by (auto simp: disjnt_iff) show ?thesis using ex_closedin * apply (simp add: ex_closedin * [OF openin_subset openin_subset] cong: conj_cong) apply (intro all_cong1 ex_cong1 imp_cong refl) by (smt (verit, best) "*" Diff_Diff_Int Diff_subset Diff_subset_conv inf.orderE inf_commute openin_subset sup_commute) qed finally show ?thesis . qed lemma normal_t1_imp_Hausdorff_space: assumes "normal_space X" "t1_space X" shows "Hausdorff_space X" unfolding Hausdorff_space_def proof clarify fix x y assume xy: "x \ topspace X" "y \ topspace X" "x \ y" then have "disjnt {x} {y}" by (auto simp: disjnt_iff) then show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" using assms xy closedin_t1_singleton normal_space_def by (metis singletonI subsetD) qed lemma normal_t1_eq_Hausdorff_space: "normal_space X \ t1_space X \ Hausdorff_space X" using normal_t1_imp_Hausdorff_space t1_or_Hausdorff_space by blast lemma normal_t1_imp_regular_space: "\normal_space X; t1_space X\ \ regular_space X" by (metis compactin_imp_closedin normal_space_def normal_t1_eq_Hausdorff_space regular_space_compact_closed_sets) lemma compact_Hausdorff_or_regular_imp_normal_space: "\compact_space X; Hausdorff_space X \ regular_space X\ \ normal_space X" by (metis Hausdorff_space_compact_sets closedin_compact_space normal_space_def regular_space_compact_closed_sets) lemma normal_space_discrete_topology: "normal_space(discrete_topology U)" by (metis discrete_topology_closure_of inf_le2 normal_space_alt) lemma normal_space_fsigmas: "normal_space X \ (\S T. fsigma_in X S \ fsigma_in X T \ separatedin X S T \ (\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof clarify fix S T assume "fsigma_in X S" then obtain C where C: "\n. closedin X (C n)" "\n. C n \ C (Suc n)" "\ (range C) = S" by (meson fsigma_in_ascending) assume "fsigma_in X T" then obtain D where D: "\n. closedin X (D n)" "\n. D n \ D (Suc n)" "\ (range D) = T" by (meson fsigma_in_ascending) assume "separatedin X S T" have "\n. disjnt (D n) (X closure_of S)" by (metis D(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) then have "\n. \V V'. openin X V \ openin X V' \ D n \ V \ X closure_of S \ V' \ disjnt V V'" by (metis D(1) L closedin_closure_of normal_space_def) then obtain V V' where V: "\n. openin X (V n)" and "\n. openin X (V' n)" "\n. disjnt (V n) (V' n)" and DV: "\n. D n \ V n" and subV': "\n. X closure_of S \ V' n" by metis then have VV: "V' n \ X closure_of V n = {}" for n using openin_Int_closure_of_eq_empty [of X "V' n" "V n"] by (simp add: Int_commute disjnt_def) have "\n. disjnt (C n) (X closure_of T)" by (metis C(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) then have "\n. \U U'. openin X U \ openin X U' \ C n \ U \ X closure_of T \ U' \ disjnt U U'" by (metis C(1) L closedin_closure_of normal_space_def) then obtain U U' where U: "\n. openin X (U n)" and "\n. openin X (U' n)" "\n. disjnt (U n) (U' n)" and CU: "\n. C n \ U n" and subU': "\n. X closure_of T \ U' n" by metis then have UU: "U' n \ X closure_of U n = {}" for n using openin_Int_closure_of_eq_empty [of X "U' n" "U n"] by (simp add: Int_commute disjnt_def) show "\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B" proof (intro conjI exI) have "\S n. closedin X (\m\n. X closure_of V m)" by (force intro: closedin_Union) then show "openin X (\n. U n - (\m\n. X closure_of V m))" using U by blast have "\S n. closedin X (\m\n. X closure_of U m)" by (force intro: closedin_Union) then show "openin X (\n. V n - (\m\n. X closure_of U m))" using V by blast have "S \ topspace X" by (simp add: \fsigma_in X S\ fsigma_in_subset) then show "S \ (\n. U n - (\m\n. X closure_of V m))" apply (clarsimp simp: Ball_def) by (metis VV C(3) CU IntI UN_E closure_of_subset empty_iff subV' subsetD) have "T \ topspace X" by (simp add: \fsigma_in X T\ fsigma_in_subset) then show "T \ (\n. V n - (\m\n. X closure_of U m))" apply (clarsimp simp: Ball_def) by (metis UU D(3) DV IntI UN_E closure_of_subset empty_iff subU' subsetD) have "\x m n. \x \ U n; x \ V m; \k\m. x \ X closure_of U k\ \ \k\n. x \ X closure_of V k" by (meson U V closure_of_subset nat_le_linear openin_subset subsetD) then show "disjnt (\n. U n - (\m\n. X closure_of V m)) (\n. V n - (\m\n. X closure_of U m))" by (force simp: disjnt_iff) qed qed next show "?rhs \ ?lhs" by (simp add: closed_imp_fsigma_in normal_space_def separatedin_closed_sets) qed lemma normal_space_fsigma_subtopology: assumes "normal_space X" "fsigma_in X S" shows "normal_space (subtopology X S)" unfolding normal_space_fsigmas proof clarify fix T U assume "fsigma_in (subtopology X S) T" and "fsigma_in (subtopology X S) U" and TU: "separatedin (subtopology X S) T U" then obtain A B where "openin X A \ openin X B \ T \ A \ U \ B \ disjnt A B" by (metis assms fsigma_in_fsigma_subtopology normal_space_fsigmas separatedin_subtopology) then show "\A B. openin (subtopology X S) A \ openin (subtopology X S) B \ T \ A \ U \ B \ disjnt A B" using TU by (force simp add: separatedin_subtopology openin_subtopology_alt disjnt_iff) qed lemma normal_space_closed_subtopology: assumes "normal_space X" "closedin X S" shows "normal_space (subtopology X S)" by (simp add: assms closed_imp_fsigma_in normal_space_fsigma_subtopology) lemma normal_space_continuous_closed_map_image: assumes "normal_space X" and contf: "continuous_map X Y f" and clof: "closed_map X Y f" and fim: "f ` topspace X = topspace Y" shows "normal_space Y" unfolding normal_space_def proof clarify fix S T assume "closedin Y S" and "closedin Y T" and "disjnt S T" have "closedin X {x \ topspace X. f x \ S}" "closedin X {x \ topspace X. f x \ T}" using \closedin Y S\ \closedin Y T\ closedin_continuous_map_preimage contf by auto moreover have "disjnt {x \ topspace X. f x \ S} {x \ topspace X. f x \ T}" using \disjnt S T\ by (auto simp: disjnt_iff) ultimately obtain U V where "closedin X U" "closedin X V" and subXU: "{x \ topspace X. f x \ S} \ topspace X - U" and subXV: "{x \ topspace X. f x \ T} \ topspace X - V" and dis: "disjnt (topspace X - U) (topspace X -V)" using \normal_space X\ by (force simp add: normal_space_def ex_openin) have "closedin Y (f ` U)" "closedin Y (f ` V)" using \closedin X U\ \closedin X V\ clof closed_map_def by blast+ moreover have "S \ topspace Y - f ` U" using \closedin Y S\ \closedin X U\ subXU by (force dest: closedin_subset) moreover have "T \ topspace Y - f ` V" using \closedin Y T\ \closedin X V\ subXV by (force dest: closedin_subset) moreover have "disjnt (topspace Y - f ` U) (topspace Y - f ` V)" using fim dis by (force simp add: disjnt_iff) ultimately show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" by (force simp add: ex_openin) qed subsection \Hereditary topological properties\ definition hereditarily where "hereditarily P X \ \S. S \ topspace X \ P(subtopology X S)" lemma hereditarily: "hereditarily P X \ (\S. P(subtopology X S))" by (metis Int_lower1 hereditarily_def subtopology_restrict) lemma hereditarily_mono: "\hereditarily P X; \x. P x \ Q x\ \ hereditarily Q X" by (simp add: hereditarily) lemma hereditarily_inc: "hereditarily P X \ P X" by (metis hereditarily subtopology_topspace) lemma hereditarily_subtopology: "hereditarily P X \ hereditarily P (subtopology X S)" by (simp add: hereditarily subtopology_subtopology) lemma hereditarily_normal_space_continuous_closed_map_image: assumes X: "hereditarily normal_space X" and contf: "continuous_map X Y f" and clof: "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "hereditarily normal_space Y" unfolding hereditarily_def proof (intro strip) fix T assume "T \ topspace Y" then have nx: "normal_space (subtopology X {x \ topspace X. f x \ T})" by (meson X hereditarily) moreover have "continuous_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" by (simp add: contf continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) moreover have "closed_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" by (simp add: clof closed_map_restriction) ultimately show "normal_space (subtopology Y T)" using fim normal_space_continuous_closed_map_image by fastforce qed lemma homeomorphic_hereditarily_normal_space: "X homeomorphic_space Y \ (hereditarily normal_space X \ hereditarily normal_space Y)" by (meson hereditarily_normal_space_continuous_closed_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) lemma hereditarily_normal_space_retraction_map_image: "\retraction_map X Y r; hereditarily normal_space X\ \ hereditarily normal_space Y" by (smt (verit) hereditarily_subtopology hereditary_imp_retractive_property homeomorphic_hereditarily_normal_space) subsection\Limits in a topological space\ lemma limitin_const_iff: assumes "t1_space X" "\ trivial_limit F" shows "limitin X (\k. a) l F \ l \ topspace X \ a = l" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using assms unfolding limitin_def t1_space_def by (metis eventually_const openin_topspace) next assume ?rhs then show ?lhs using assms by (auto simp: limitin_def t1_space_def) qed lemma compactin_sequence_with_limit: assumes lim: "limitin X \ l sequentially" and "S \ range \" and SX: "S \ topspace X" shows "compactin X (insert l S)" unfolding compactin_def proof (intro conjI strip) show "insert l S \ topspace X" by (meson SX insert_subset lim limitin_topspace) fix \ assume \
: "Ball \ (openin X) \ insert l S \ \ \" have "\V. finite V \ V \ \ \ (\t \ V. l \ t) \ S \ \ V" if *: "\x \ S. \T \ \. x \ T" and "T \ \" "l \ T" for T proof - obtain V where V: "\x. x \ S \ V x \ \ \ x \ V x" using * by metis obtain N where N: "\n. N \ n \ \ n \ T" by (meson "\
" \T \ \\ \l \ T\ lim limitin_sequentially) show ?thesis proof (intro conjI exI) have "x \ T" if "x \ S" and "\A. (\x \ S. (\n\N. x \ \ n) \ A \ V x) \ x \ A" for x by (metis (no_types) N V that assms(2) imageE nle_le subsetD) then show "S \ \ (insert T (V ` (S \ \ ` {0..N})))" by force qed (use V that in auto) qed then show "\\. finite \ \ \ \ \ \ insert l S \ \ \" by (smt (verit, best) Union_iff \
insert_subset subsetD) qed lemma limitin_Hausdorff_unique: assumes "limitin X f l1 F" "limitin X f l2 F" "\ trivial_limit F" "Hausdorff_space X" shows "l1 = l2" proof (rule ccontr) assume "l1 \ l2" with assms obtain U V where "openin X U" "openin X V" "l1 \ U" "l2 \ V" "disjnt U V" by (metis Hausdorff_space_def limitin_topspace) then have "eventually (\x. f x \ U) F" "eventually (\x. f x \ V) F" using assms by (fastforce simp: limitin_def)+ then have "\x. f x \ U \ f x \ V" using assms eventually_elim2 filter_eq_iff by fastforce with assms \disjnt U V\ show False by (meson disjnt_iff) qed lemma limitin_kc_unique: assumes "kc_space X" and lim1: "limitin X f l1 sequentially" and lim2: "limitin X f l2 sequentially" shows "l1 = l2" proof (rule ccontr) assume "l1 \ l2" define A where "A \ insert l1 (range f - {l2})" have "l1 \ topspace X" using lim1 limitin_def by fastforce moreover have "compactin X (insert l1 (topspace X \ (range f - {l2})))" by (meson Diff_subset compactin_sequence_with_limit inf_le1 inf_le2 lim1 subset_trans) ultimately have "compactin X (topspace X \ A)" by (simp add: A_def) then have OXA: "openin X (topspace X - A)" by (metis Diff_Diff_Int Diff_subset \kc_space X\ kc_space_def openin_closedin_eq) have "l2 \ topspace X - A" using \l1 \ l2\ A_def lim2 limitin_topspace by fastforce then have "\\<^sub>F x in sequentially. f x = l2" using limitinD [OF lim2 OXA] by (auto simp: A_def eventually_conj_iff) then show False using limitin_transform_eventually [OF _ lim1] limitin_const_iff [OF kc_imp_t1_space trivial_limit_sequentially] using \l1 \ l2\ \kc_space X\ by fastforce qed lemma limitin_closedin: assumes lim: "limitin X f l F" and "closedin X S" and ev: "eventually (\x. f x \ S) F" "\ trivial_limit F" shows "l \ S" proof (rule ccontr) assume "l \ S" have "\\<^sub>F x in F. f x \ topspace X - S" by (metis Diff_iff \l \ S\ \closedin X S\ closedin_def lim limitin_def) with ev eventually_elim2 trivial_limit_def show False by force qed +subsection\Quasi-components\ + +definition quasi_component_of :: "'a topology \ 'a \ 'a \ bool" + where + "quasi_component_of X x y \ + x \ topspace X \ y \ topspace X \ + (\T. closedin X T \ openin X T \ (x \ T \ y \ T))" + +abbreviation "quasi_component_of_set S x \ Collect (quasi_component_of S x)" + +definition quasi_components_of :: "'a topology \ ('a set) set" + where + "quasi_components_of X = quasi_component_of_set X ` topspace X" + +lemma quasi_component_in_topspace: + "quasi_component_of X x y \ x \ topspace X \ y \ topspace X" + by (simp add: quasi_component_of_def) + +lemma quasi_component_of_refl [simp]: + "quasi_component_of X x x \ x \ topspace X" + by (simp add: quasi_component_of_def) + +lemma quasi_component_of_sym: + "quasi_component_of X x y \ quasi_component_of X y x" + by (meson quasi_component_of_def) + +lemma quasi_component_of_trans: + "\quasi_component_of X x y; quasi_component_of X y z\ \ quasi_component_of X x z" + by (simp add: quasi_component_of_def) + +lemma quasi_component_of_subset_topspace: + "quasi_component_of_set X x \ topspace X" + using quasi_component_of_def by fastforce + +lemma quasi_component_of_eq_empty: + "quasi_component_of_set X x = {} \ (x \ topspace X)" + using quasi_component_of_def by fastforce + +lemma quasi_component_of: + "quasi_component_of X x y \ + x \ topspace X \ y \ topspace X \ (\T. x \ T \ closedin X T \ openin X T \ y \ T)" + unfolding quasi_component_of_def by blast + +lemma quasi_component_of_alt: + "quasi_component_of X x y \ + x \ topspace X \ y \ topspace X \ + \ (\U V. openin X U \ openin X V \ U \ V = topspace X \ disjnt U V \ x \ U \ y \ V)" + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + unfolding quasi_component_of_def + by (metis disjnt_iff separatedin_full separatedin_open_sets) + show "?rhs \ ?lhs" + unfolding quasi_component_of_def + by (metis Diff_disjoint Diff_iff Un_Diff_cancel closedin_def disjnt_def inf_commute sup.orderE sup_commute) +qed + +lemma quasi_component_of_separated: + "quasi_component_of X x y \ + x \ topspace X \ y \ topspace X \ + \ (\U V. separatedin X U V \ U \ V = topspace X \ x \ U \ y \ V)" + by (meson quasi_component_of_alt separatedin_full separatedin_open_sets) + +lemma quasi_component_of_subtopology: + "quasi_component_of (subtopology X s) x y \ quasi_component_of X x y" + unfolding quasi_component_of_def + by (simp add: closedin_subtopology) (metis Int_iff inf_commute openin_subtopology_Int2) + +lemma quasi_component_of_mono: + "quasi_component_of (subtopology X S) x y \ S \ T + \ quasi_component_of (subtopology X T) x y" + by (metis inf.absorb_iff2 quasi_component_of_subtopology subtopology_subtopology) + +lemma quasi_component_of_equiv: + "quasi_component_of X x y \ + x \ topspace X \ y \ topspace X \ quasi_component_of X x = quasi_component_of X y" + using quasi_component_of_def by fastforce + +lemma quasi_component_of_disjoint [simp]: + "disjnt (quasi_component_of_set X x) (quasi_component_of_set X y) \ \ (quasi_component_of X x y)" + by (metis disjnt_iff quasi_component_of_equiv mem_Collect_eq) + +lemma quasi_component_of_eq: + "quasi_component_of X x = quasi_component_of X y \ + (x \ topspace X \ y \ topspace X) + \ x \ topspace X \ y \ topspace X \ quasi_component_of X x y" + by (metis Collect_empty_eq_bot quasi_component_of_eq_empty quasi_component_of_equiv) + +lemma topspace_imp_quasi_components_of: + assumes "x \ topspace X" + obtains C where "C \ quasi_components_of X" "x \ C" + by (metis assms imageI mem_Collect_eq quasi_component_of_refl quasi_components_of_def) + +lemma Union_quasi_components_of: "\ (quasi_components_of X) = topspace X" + by (auto simp: quasi_components_of_def quasi_component_of_def) + +lemma pairwise_disjoint_quasi_components_of: + "pairwise disjnt (quasi_components_of X)" + by (auto simp: quasi_components_of_def quasi_component_of_def disjoint_def) + +lemma complement_quasi_components_of_Union: + assumes "C \ quasi_components_of X" + shows "topspace X - C = \ (quasi_components_of X - {C})" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + using Union_quasi_components_of by fastforce + show "?rhs \ ?lhs" + using assms + using quasi_component_of_equiv by (fastforce simp add: quasi_components_of_def image_iff subset_iff) +qed + +lemma nonempty_quasi_components_of: + "C \ quasi_components_of X \ C \ {}" + by (metis imageE quasi_component_of_eq_empty quasi_components_of_def) + +lemma quasi_components_of_subset: + "C \ quasi_components_of X \ C \ topspace X" + using Union_quasi_components_of by force + +lemma quasi_component_in_quasi_components_of: + "quasi_component_of_set X a \ quasi_components_of X \ a \ topspace X" + by (metis (no_types, lifting) image_iff quasi_component_of_eq_empty quasi_components_of_def) + +lemma quasi_components_of_eq_empty [simp]: + "quasi_components_of X = {} \ topspace X = {}" + by (simp add: quasi_components_of_def) + +lemma quasi_components_of_empty_space: + "topspace X = {} \ quasi_components_of X = {}" + by simp + +lemma quasi_component_of_set: + "quasi_component_of_set X x = + (if x \ topspace X + then \ {t. closedin X t \ openin X t \ x \ t} + else {})" + by (auto simp: quasi_component_of) + +lemma closedin_quasi_component_of: "closedin X (quasi_component_of_set X x)" + by (auto simp: quasi_component_of_set) + +lemma closedin_quasi_components_of: + "C \ quasi_components_of X \ closedin X C" + by (auto simp: quasi_components_of_def closedin_quasi_component_of) + +lemma openin_finite_quasi_components: + "\finite(quasi_components_of X); C \ quasi_components_of X\ \ openin X C" + apply (simp add:openin_closedin_eq quasi_components_of_subset complement_quasi_components_of_Union) + by (meson DiffD1 closedin_Union closedin_quasi_components_of finite_Diff) + +lemma quasi_component_of_eq_overlap: + "quasi_component_of X x = quasi_component_of X y \ + (x \ topspace X \ y \ topspace X) \ + \ (quasi_component_of_set X x \ quasi_component_of_set X y = {})" + using quasi_component_of_equiv by fastforce + +lemma quasi_component_of_nonoverlap: + "quasi_component_of_set X x \ quasi_component_of_set X y = {} \ + (x \ topspace X) \ (y \ topspace X) \ + \ (quasi_component_of X x = quasi_component_of X y)" + by (metis inf.idem quasi_component_of_eq_empty quasi_component_of_eq_overlap) + +lemma quasi_component_of_overlap: + "\ (quasi_component_of_set X x \ quasi_component_of_set X y = {}) \ + x \ topspace X \ y \ topspace X \ quasi_component_of X x = quasi_component_of X y" + by (meson quasi_component_of_nonoverlap) + +lemma quasi_components_of_disjoint: + "\C \ quasi_components_of X; D \ quasi_components_of X\ \ disjnt C D \ C \ D" + by (metis disjnt_self_iff_empty nonempty_quasi_components_of pairwiseD pairwise_disjoint_quasi_components_of) + +lemma quasi_components_of_overlap: + "\C \ quasi_components_of X; D \ quasi_components_of X\ \ \ (C \ D = {}) \ C = D" + by (metis disjnt_def quasi_components_of_disjoint) + +lemma pairwise_separated_quasi_components_of: + "pairwise (separatedin X) (quasi_components_of X)" + by (metis closedin_quasi_components_of pairwise_def pairwise_disjoint_quasi_components_of separatedin_closed_sets) + +lemma finite_quasi_components_of_finite: + "finite(topspace X) \ finite(quasi_components_of X)" + by (simp add: Union_quasi_components_of finite_UnionD) + +lemma connected_imp_quasi_component_of: + assumes "connected_component_of X x y" + shows "quasi_component_of X x y" +proof - + have "x \ topspace X" "y \ topspace X" + by (meson assms connected_component_of_equiv)+ + with assms show ?thesis + apply (clarsimp simp add: quasi_component_of connected_component_of_def) + by (meson connectedin_clopen_cases disjnt_iff subsetD) +qed + +lemma connected_component_subset_quasi_component_of: + "connected_component_of_set X x \ quasi_component_of_set X x" + using connected_imp_quasi_component_of by force + +lemma quasi_component_as_connected_component_Union: + "quasi_component_of_set X x = + \ (connected_component_of_set X ` quasi_component_of_set X x)" + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + using connected_component_of_refl quasi_component_of by fastforce + show "?rhs \ ?lhs" + apply (rule SUP_least) + by (simp add: connected_component_subset_quasi_component_of quasi_component_of_equiv) +qed + +lemma quasi_components_as_connected_components_Union: + assumes "C \ quasi_components_of X" + obtains \ where "\ \ connected_components_of X" "\\ = C" +proof - + obtain x where "x \ topspace X" and Ceq: "C = quasi_component_of_set X x" + by (metis assms imageE quasi_components_of_def) + define \ where "\ \ connected_component_of_set X ` quasi_component_of_set X x" + show thesis + proof + show "\ \ connected_components_of X" + by (simp add: \_def connected_components_of_def image_mono quasi_component_of_subset_topspace) + show "\\ = C" + by (metis \_def Ceq quasi_component_as_connected_component_Union) + qed +qed + +lemma path_imp_quasi_component_of: + "path_component_of X x y \ quasi_component_of X x y" + by (simp add: connected_imp_quasi_component_of path_imp_connected_component_of) + +lemma path_component_subset_quasi_component_of: + "path_component_of_set X x \ quasi_component_of_set X x" + by (simp add: Collect_mono path_imp_quasi_component_of) + +lemma connected_space_iff_quasi_component: + "connected_space X \ (\x \ topspace X. \y \ topspace X. quasi_component_of X x y)" + unfolding connected_space_clopen_in closedin_def quasi_component_of + by blast + +lemma connected_space_imp_quasi_component_of: + " \connected_space X; a \ topspace X; b \ topspace X\ \ quasi_component_of X a b" + by (simp add: connected_space_iff_quasi_component) + +lemma connected_space_quasi_component_set: + "connected_space X \ (\x \ topspace X. quasi_component_of_set X x = topspace X)" + by (metis Ball_Collect connected_space_iff_quasi_component quasi_component_of_subset_topspace subset_antisym) + +lemma connected_space_iff_quasi_components_eq: + "connected_space X \ + (\C \ quasi_components_of X. \D \ quasi_components_of X. C = D)" + apply (simp add: quasi_components_of_def) + by (metis connected_space_iff_quasi_component mem_Collect_eq quasi_component_of_equiv) + +lemma quasi_components_of_subset_sing: + "quasi_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" +proof (cases "quasi_components_of X = {}") + case True + then show ?thesis + by (simp add: connected_space_topspace_empty subset_singleton_iff) +next + case False + then show ?thesis + apply (simp add: connected_space_iff_quasi_components_eq subset_iff Ball_def) + by (metis quasi_components_of_subset subsetI subset_antisym subset_empty topspace_imp_quasi_components_of) +qed + +lemma connected_space_iff_quasi_components_subset_sing: + "connected_space X \ (\a. quasi_components_of X \ {a})" + by (simp add: quasi_components_of_subset_sing) + +lemma quasi_components_of_eq_singleton: + "quasi_components_of X = {S} \ + connected_space X \ \ (topspace X = {}) \ S = topspace X" + by (metis ccpo_Sup_singleton insert_not_empty quasi_components_of_subset_sing subset_singleton_iff) + +lemma quasi_components_of_connected_space: + "connected_space X + \ quasi_components_of X = (if topspace X = {} then {} else {topspace X})" + by (simp add: quasi_components_of_eq_singleton) + +lemma separated_between_singletons: + "separated_between X {x} {y} \ + x \ topspace X \ y \ topspace X \ \ (quasi_component_of X x y)" +proof (cases "x \ topspace X \ y \ topspace X") + case True + then show ?thesis + by (auto simp add: separated_between_def quasi_component_of_alt) +qed (use separated_between_imp_subset in blast) + +lemma quasi_component_nonseparated: + "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ \ (separated_between X {x} {y})" + by (metis quasi_component_of_equiv separated_between_singletons) + +lemma separated_between_quasi_component_pointwise_left: + assumes "C \ quasi_components_of X" + shows "separated_between X C S \ (\x \ C. separated_between X {x} S)" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + using assms quasi_components_of_disjoint separated_between_mono by fastforce +next + assume ?rhs + then obtain y where "separated_between X {y} S" and "y \ C" + by metis + with assms show ?lhs + by (force simp add: separated_between quasi_components_of_def quasi_component_of_def) +qed + +lemma separated_between_quasi_component_pointwise_right: + "C \ quasi_components_of X \ separated_between X S C \ (\x \ C. separated_between X S {x})" + by (simp add: separated_between_quasi_component_pointwise_left separated_between_sym) + +lemma separated_between_quasi_component_point: + assumes "C \ quasi_components_of X" + shows "separated_between X C {x} \ x \ topspace X - C" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (meson DiffI disjnt_insert2 insert_subset separated_between_imp_disjoint separated_between_imp_subset) +next + assume ?rhs + with assms show ?lhs + unfolding quasi_components_of_def image_iff Diff_iff separated_between_quasi_component_pointwise_left [OF assms] + by (metis mem_Collect_eq quasi_component_of_refl separated_between_singletons) +qed + +lemma separated_between_point_quasi_component: + "C \ quasi_components_of X \ separated_between X {x} C \ x \ topspace X - C" + by (simp add: separated_between_quasi_component_point separated_between_sym) + +lemma separated_between_quasi_component_compact: + "\C \ quasi_components_of X; compactin X K\ \ (separated_between X C K \ disjnt C K)" + unfolding disjnt_iff + using compactin_subset_topspace quasi_components_of_subset separated_between_pointwise_right separated_between_quasi_component_point by fastforce + +lemma separated_between_compact_quasi_component: + "\compactin X K; C \ quasi_components_of X\ \ separated_between X K C \ disjnt K C" + using disjnt_sym separated_between_quasi_component_compact separated_between_sym by blast + +lemma separated_between_quasi_components: + assumes C: "C \ quasi_components_of X" and D: "D \ quasi_components_of X" + shows "separated_between X C D \ disjnt C D" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: separated_between_imp_disjoint) +next + assume ?rhs + obtain x y where x: "C = quasi_component_of_set X x" and "x \ C" + and y: "D = quasi_component_of_set X y" and "y \ D" + using assms by (auto simp: quasi_components_of_def) + then have "separated_between X {x} {y}" + using \disjnt C D\ separated_between_singletons by fastforce + with \x \ C\ \y \ D\ show ?lhs + by (auto simp: assms separated_between_quasi_component_pointwise_left separated_between_quasi_component_pointwise_right) +qed + +lemma quasi_eq_connected_component_of_eq: + "quasi_component_of X x = connected_component_of X x \ + connectedin X (quasi_component_of_set X x)" (is "?lhs = ?rhs") +proof (cases "x \ topspace X") + case True + show ?thesis + proof + show "?lhs \ ?rhs" + by (simp add: connectedin_connected_component_of) + next + assume ?rhs + then have "\y. quasi_component_of X x y = connected_component_of X x y" + by (metis connected_component_of_def connected_imp_quasi_component_of mem_Collect_eq quasi_component_of_equiv) + then show ?lhs + by force + qed +next + case False + then show ?thesis + by (metis Collect_empty_eq_bot connected_component_of_eq_empty connectedin_empty quasi_component_of_eq_empty) +qed + +lemma connected_quasi_component_of: + assumes "C \ quasi_components_of X" + shows "C \ connected_components_of X \ connectedin X C" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + using assms + by (simp add: connectedin_connected_components_of) +next + assume ?rhs + with assms show ?lhs + unfolding quasi_components_of_def connected_components_of_def image_iff + by (metis quasi_eq_connected_component_of_eq) +qed + +lemma quasi_component_of_clopen_cases: + "\C \ quasi_components_of X; closedin X T; openin X T\ \ C \ T \ disjnt C T" + by (smt (verit) disjnt_iff image_iff mem_Collect_eq quasi_component_of_def quasi_components_of_def subset_iff) + +lemma quasi_components_of_set: + assumes "C \ quasi_components_of X" + shows "\ {T. closedin X T \ openin X T \ C \ T} = C" (is "?lhs = ?rhs") +proof + have "x \ C" if "x \ \ {T. closedin X T \ openin X T \ C \ T}" for x + proof (rule ccontr) + assume "x \ C" + have "x \ topspace X" + using assms quasi_components_of_subset that by force + then have "separated_between X C {x}" + by (simp add: \x \ C\ assms separated_between_quasi_component_point) + with that show False + by (auto simp: separated_between) + qed + then show "?lhs \ ?rhs" + by auto +qed blast + +lemma open_quasi_eq_connected_components_of: + assumes "openin X C" + shows "C \ quasi_components_of X \ C \ connected_components_of X" (is "?lhs = ?rhs") +proof (cases "closedin X C") + case True + show ?thesis + proof + assume L: ?lhs + have "T = {} \ T = topspace X \ C" + if "openin (subtopology X C) T" "closedin (subtopology X C) T" for T + proof - + have "C \ T \ disjnt C T" + by (meson L True assms closedin_trans_full openin_trans_full quasi_component_of_clopen_cases that) + with that show ?thesis + by (metis Int_absorb2 True closedin_imp_subset closure_of_subset_eq disjnt_def inf_absorb2) + qed + with L assms show "?rhs" + by (simp add: connected_quasi_component_of connected_space_clopen_in connectedin_def openin_subset) + next + assume ?rhs + then obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" + by (metis connected_components_of_def imageE) + have "C = quasi_component_of_set X x" + using True assms connected_component_of_refl connected_imp_quasi_component_of quasi_component_of_def x by fastforce + then show ?lhs + using \x \ topspace X\ quasi_components_of_def by fastforce + qed +next + case False + then show ?thesis + using closedin_connected_components_of closedin_quasi_components_of by blast +qed + +lemma quasi_component_of_continuous_image: + assumes f: "continuous_map X Y f" and qc: "quasi_component_of X x y" + shows "quasi_component_of Y (f x) (f y)" + unfolding quasi_component_of_def +proof (intro strip conjI) + show "f x \ topspace Y" "f y \ topspace Y" + using assms by (simp_all add: continuous_map_def quasi_component_of_def) + fix T + assume "closedin Y T \ openin Y T" + with assms show "(f x \ T) = (f y \ T)" + by (smt (verit) continuous_map_closedin continuous_map_def mem_Collect_eq quasi_component_of_def) +qed + +lemma quasi_component_of_discrete_topology: + "quasi_component_of_set (discrete_topology U) x = (if x \ U then {x} else {})" +proof - + have "quasi_component_of_set (discrete_topology U) y = {y}" if "y \ U" for y + using that + apply (simp add: set_eq_iff quasi_component_of_def) + by (metis Set.set_insert insertE subset_insertI) + then show ?thesis + by (simp add: quasi_component_of) +qed + +lemma quasi_components_of_discrete_topology: + "quasi_components_of (discrete_topology U) = (\x. {x}) ` U" + by (auto simp add: quasi_components_of_def quasi_component_of_discrete_topology) + +lemma homeomorphic_map_quasi_component_of: + assumes hmf: "homeomorphic_map X Y f" and "x \ topspace X" + shows "quasi_component_of_set Y (f x) = f ` (quasi_component_of_set X x)" +proof - + obtain g where hmg: "homeomorphic_map Y X g" + and contf: "continuous_map X Y f" and contg: "continuous_map Y X g" + and fg: "(\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" + by (smt (verit, best) hmf homeomorphic_map_maps homeomorphic_maps_def) + show ?thesis + proof + show "quasi_component_of_set Y (f x) \ f ` quasi_component_of_set X x" + using quasi_component_of_continuous_image [OF contg] + \x \ topspace X\ fg image_iff quasi_component_of_subset_topspace by fastforce + show "f ` quasi_component_of_set X x \ quasi_component_of_set Y (f x)" + using quasi_component_of_continuous_image [OF contf] by blast + qed +qed + + +lemma homeomorphic_map_quasi_components_of: + assumes "homeomorphic_map X Y f" + shows "quasi_components_of Y = image (image f) (quasi_components_of X)" + using assms +proof - + have "\x\topspace X. quasi_component_of_set Y y = f ` quasi_component_of_set X x" + if "y \ topspace Y" for y + by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of image_iff) + moreover have "\x\topspace Y. f ` quasi_component_of_set X u = quasi_component_of_set Y x" + if "u \ topspace X" for u + by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of imageI) + ultimately show ?thesis + by (auto simp: quasi_components_of_def image_iff) +qed + +lemma openin_quasi_component_of_locally_connected_space: + assumes "locally_connected_space X" + shows "openin X (quasi_component_of_set X x)" +proof - + have *: "openin X (connected_component_of_set X x)" + by (simp add: assms openin_connected_component_of_locally_connected_space) + moreover have "connected_component_of_set X x = quasi_component_of_set X x" + using * closedin_connected_component_of connected_component_of_refl connected_imp_quasi_component_of + quasi_component_of_def by fastforce + ultimately show ?thesis + by simp +qed + +lemma openin_quasi_components_of_locally_connected_space: + "locally_connected_space X \ c \ quasi_components_of X + \ openin X c" + by (smt (verit, best) image_iff openin_quasi_component_of_locally_connected_space quasi_components_of_def) + +lemma quasi_eq_connected_components_of_alt: + "quasi_components_of X = connected_components_of X \ (\C \ quasi_components_of X. connectedin X C)" + (is "?lhs = ?rhs") +proof + assume R: ?rhs + moreover have "connected_components_of X \ quasi_components_of X" + using R unfolding quasi_components_of_def connected_components_of_def + by (force simp flip: quasi_eq_connected_component_of_eq) + ultimately show ?lhs + using connected_quasi_component_of by blast +qed (use connected_quasi_component_of in blast) + +lemma connected_subset_quasi_components_of_pointwise: + "connected_components_of X \ quasi_components_of X \ + (\x \ topspace X. quasi_component_of X x = connected_component_of X x)" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + have "connectedin X (quasi_component_of_set X x)" if "x \ topspace X" for x + proof - + have "\y\topspace X. connected_component_of_set X x = quasi_component_of_set X y" + using L that by (force simp: quasi_components_of_def connected_components_of_def image_subset_iff) + then show ?thesis + by (metis connected_component_of_equiv connectedin_connected_component_of mem_Collect_eq quasi_component_of_eq) + qed + then show ?rhs + by (simp add: quasi_eq_connected_component_of_eq) +qed (simp add: connected_components_of_def quasi_components_of_def) + +lemma quasi_subset_connected_components_of_pointwise: + "quasi_components_of X \ connected_components_of X \ + (\x \ topspace X. quasi_component_of X x = connected_component_of X x)" + by (simp add: connected_quasi_component_of image_subset_iff quasi_components_of_def quasi_eq_connected_component_of_eq) + +lemma quasi_eq_connected_components_of_pointwise: + "quasi_components_of X = connected_components_of X \ + (\x \ topspace X. quasi_component_of X x = connected_component_of X x)" + using connected_subset_quasi_components_of_pointwise quasi_subset_connected_components_of_pointwise by fastforce + +lemma quasi_eq_connected_components_of_pointwise_alt: + "quasi_components_of X = connected_components_of X \ + (\x. quasi_component_of X x = connected_component_of X x)" + unfolding quasi_eq_connected_components_of_pointwise + by (metis connectedin_empty quasi_component_of_eq_empty quasi_eq_connected_component_of_eq) + +lemma quasi_eq_connected_components_of_inclusion: + "quasi_components_of X = connected_components_of X \ + connected_components_of X \ quasi_components_of X \ + quasi_components_of X \ connected_components_of X" + by (simp add: connected_subset_quasi_components_of_pointwise dual_order.eq_iff quasi_subset_connected_components_of_pointwise) + + +lemma quasi_eq_connected_components_of: + "finite(connected_components_of X) \ + finite(quasi_components_of X) \ + locally_connected_space X \ + compact_space X \ (Hausdorff_space X \ regular_space X \ normal_space X) + \ quasi_components_of X = connected_components_of X" +proof (elim disjE) + show "quasi_components_of X = connected_components_of X" + if "finite (connected_components_of X)" + unfolding quasi_eq_connected_components_of_inclusion + using that open_in_finite_connected_components open_quasi_eq_connected_components_of by blast + show "quasi_components_of X = connected_components_of X" + if "finite (quasi_components_of X)" + unfolding quasi_eq_connected_components_of_inclusion + using that open_quasi_eq_connected_components_of openin_finite_quasi_components by blast + show "quasi_components_of X = connected_components_of X" + if "locally_connected_space X" + unfolding quasi_eq_connected_components_of_inclusion + using that open_quasi_eq_connected_components_of openin_quasi_components_of_locally_connected_space by auto + show "quasi_components_of X = connected_components_of X" + if "compact_space X \ (Hausdorff_space X \ regular_space X \ normal_space X)" + proof - + show ?thesis + unfolding quasi_eq_connected_components_of_alt + proof (intro strip) + fix C + assume C: "C \ quasi_components_of X" + then have cloC: "closedin X C" + by (simp add: closedin_quasi_components_of) + have "normal_space X" + using that compact_Hausdorff_or_regular_imp_normal_space by blast + show "connectedin X C" + proof (clarsimp simp add: connectedin_def connected_space_closedin_eq closedin_closed_subtopology cloC closedin_subset [OF cloC]) + fix S T + assume "S \ C" and "closedin X S" and "S \ T = {}" and SUT: "S \ T = topspace X \ C" + and T: "T \ C" "T \ {}" and "closedin X T" + with \normal_space X\ obtain U V where UV: "openin X U" "openin X V" "S \ U" "T \ V" "disjnt U V" + by (meson disjnt_def normal_space_def) + moreover have "compactin X (topspace X - (U \ V))" + using UV that by (intro closedin_compact_space closedin_diff openin_Un) auto + ultimately have "separated_between X C (topspace X - (U \ V)) \ disjnt C (topspace X - (U \ V))" + by (simp add: \C \ quasi_components_of X\ separated_between_quasi_component_compact) + moreover have "disjnt C (topspace X - (U \ V))" + using UV SUT disjnt_def by fastforce + ultimately have "separated_between X C (topspace X - (U \ V))" + by simp + then obtain A B where "openin X A" "openin X B" "A \ B = topspace X" "disjnt A B" "C \ A" + and subB: "topspace X - (U \ V) \ B" + by (meson separated_between_def) + have "B \ U = topspace X - (A \ V)" + proof + show "B \ U \ topspace X - A \ V" + using \openin X U\ \disjnt U V\ \disjnt A B\ \openin X B\ disjnt_iff openin_closedin_eq by fastforce + show "topspace X - A \ V \ B \ U" + using \A \ B = topspace X\ subB by fastforce + qed + then have "closedin X (B \ U)" + using \openin X V\ \openin X A\ by auto + then have "C \ B \ U \ disjnt C (B \ U)" + using quasi_component_of_clopen_cases [OF C] \openin X U\ \openin X B\ by blast + with UV show "S = {}" + by (metis UnE \C \ A\ \S \ C\ T \disjnt A B\ all_not_in_conv disjnt_Un2 disjnt_iff subset_eq) + qed + qed + qed +qed + + +lemma quasi_eq_connected_component_of: + "finite(connected_components_of X) \ + finite(quasi_components_of X) \ + locally_connected_space X \ + compact_space X \ (Hausdorff_space X \ regular_space X \ normal_space X) + \ quasi_component_of X x = connected_component_of X x" + by (metis quasi_eq_connected_components_of quasi_eq_connected_components_of_pointwise_alt) + + +subsection\Additional quasicomponent and continuum properties like Boundary Bumping\ + +lemma cut_wire_fence_theorem_gen: + assumes "compact_space X" and X: "Hausdorff_space X \ regular_space X \ normal_space X" + and S: "compactin X S" and T: "closedin X T" + and dis: "\C. connectedin X C \ disjnt C S \ disjnt C T" + shows "separated_between X S T" + proof - + have "x \ topspace X" if "x \ S" and "T = {}" for x + using that S compactin_subset_topspace by auto + moreover have "separated_between X {x} {y}" if "x \ S" and "y \ T" for x y + proof (cases "x \ topspace X \ y \ topspace X") + case True + then have "\ connected_component_of X x y" + by (meson dis connected_component_of_def disjnt_iff that) + with True X \compact_space X\ show ?thesis + by (metis quasi_component_nonseparated quasi_eq_connected_component_of) + next + case False + then show ?thesis + using S T compactin_subset_topspace closedin_subset that by blast + qed + ultimately show ?thesis + using assms + by (simp add: separated_between_pointwise_left separated_between_pointwise_right + closedin_compact_space closedin_subset) +qed + +lemma cut_wire_fence_theorem: + "\compact_space X; Hausdorff_space X; closedin X S; closedin X T; + \C. connectedin X C \ disjnt C S \ disjnt C T\ + \ separated_between X S T" + by (simp add: closedin_compact_space cut_wire_fence_theorem_gen) + +lemma separated_between_from_closed_subtopology: + assumes XC: "separated_between (subtopology X C) S (X frontier_of C)" + and ST: "separated_between (subtopology X C) S T" + shows "separated_between X S T" +proof - + obtain U where clo: "closedin (subtopology X C) U" and ope: "openin (subtopology X C) U" + and "S \ U" and sub: "X frontier_of C \ T \ topspace (subtopology X C) - U" + by (meson assms separated_between separated_between_Un) + then have "X frontier_of C \ T \ topspace X \ C - U" + by auto + have "closedin X (topspace X \ C)" + by (metis XC frontier_of_restrict frontier_of_subset_eq inf_le1 separated_between_imp_subset topspace_subtopology) + then have "closedin X U" + by (metis clo closedin_closed_subtopology subtopology_restrict) + moreover have "openin (subtopology X C) U \ openin X U \ U \ C" + using disjnt_iff sub by (force intro!: openin_subset_topspace_eq) + with ope have "openin X U" + by blast + moreover have "T \ topspace X - U" + using ope openin_closedin_eq sub by auto + ultimately show ?thesis + using \S \ U\ separated_between by blast +qed + +lemma separated_between_from_closed_subtopology_frontier: + "separated_between (subtopology X T) S (X frontier_of T) + \ separated_between X S (X frontier_of T)" + using separated_between_from_closed_subtopology by blast + +lemma separated_between_from_frontier_of_closed_subtopology: + assumes "separated_between (subtopology X T) S (X frontier_of T)" + shows "separated_between X S (topspace X - T)" +proof - + have "disjnt S (topspace X - T)" + using assms disjnt_iff separated_between_imp_subset by fastforce + then show ?thesis + by (metis Diff_subset assms frontier_of_complement separated_between_from_closed_subtopology separated_between_frontier_of_eq') +qed + +lemma separated_between_compact_connected_component: + assumes "locally_compact_space X" "Hausdorff_space X" + and C: "C \ connected_components_of X" + and "compactin X C" "closedin X T" "disjnt C T" + shows "separated_between X C T" +proof - + have Csub: "C \ topspace X" + by (simp add: assms(4) compactin_subset_topspace) + have "Hausdorff_space (subtopology X (topspace X - T))" + using Hausdorff_space_subtopology assms(2) by blast + moreover have "compactin (subtopology X (topspace X - T)) C" + using assms Csub by (metis Diff_Int_distrib Diff_empty compact_imp_compactin_subtopology disjnt_def le_iff_inf) + moreover have "locally_compact_space (subtopology X (topspace X - T))" + by (meson assms closedin_def locally_compact_Hausdorff_imp_regular_space locally_compact_space_open_subset) + ultimately + obtain N L where "openin X N" "compactin X L" "closedin X L" "C \ N" "N \ L" + and Lsub: "L \ topspace X - T" + using \Hausdorff_space X\ \closedin X T\ + apply (simp add: locally_compact_space_compact_closed_compact compactin_subtopology) + by (meson closedin_def compactin_imp_closedin openin_trans_full) + then have disC: "disjnt C (topspace X - L)" + by (meson DiffD2 disjnt_iff subset_iff) + have "separated_between (subtopology X L) C (X frontier_of L)" + proof (rule cut_wire_fence_theorem) + show "compact_space (subtopology X L)" + by (simp add: \compactin X L\ compact_space_subtopology) + show "Hausdorff_space (subtopology X L)" + by (simp add: Hausdorff_space_subtopology \Hausdorff_space X\) + show "closedin (subtopology X L) C" + by (meson \C \ N\ \N \ L\ \Hausdorff_space X\ \compactin X C\ closedin_subset_topspace compactin_imp_closedin subset_trans) + show "closedin (subtopology X L) (X frontier_of L)" + by (simp add: \closedin X L\ closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin) + show "disjnt D C \ disjnt D (X frontier_of L)" + if "connectedin (subtopology X L) D" for D + proof (rule ccontr) + assume "\ (disjnt D C \ disjnt D (X frontier_of L))" + moreover have "connectedin X D" + using connectedin_subtopology that by blast + ultimately show False + using that connected_components_of_maximal [of C X D] C + apply (simp add: disjnt_iff) + by (metis Diff_eq_empty_iff \C \ N\ \N \ L\ \openin X N\ disjoint_iff frontier_of_openin_straddle_Int(2) subsetD) + qed + qed + then have "separated_between X (X frontier_of C) (topspace X - L)" + using separated_between_from_frontier_of_closed_subtopology separated_between_frontier_of_eq by blast + with \closedin X T\ + separated_between_frontier_of [OF Csub disC] + show ?thesis + unfolding separated_between by (smt (verit) Diff_iff Lsub closedin_subset subset_iff) +qed + +lemma wilder_locally_compact_component_thm: + assumes "locally_compact_space X" "Hausdorff_space X" + and "C \ connected_components_of X" "compactin X C" "openin X W" "C \ W" + obtains U V where "openin X U" "openin X V" "disjnt U V" "U \ V = topspace X" "C \ U" "U \ W" +proof - + have "closedin X (topspace X - W)" + using \openin X W\ by blast + moreover have "disjnt C (topspace X - W)" + using \C \ W\ disjnt_def by fastforce + ultimately have "separated_between X C (topspace X - W)" + using separated_between_compact_connected_component assms by blast + then show thesis + by (smt (verit, del_insts) DiffI disjnt_iff openin_subset separated_between_def subset_iff that) +qed + +lemma compact_quasi_eq_connected_components_of: + assumes "locally_compact_space X" "Hausdorff_space X" "compactin X C" + shows "C \ quasi_components_of X \ C \ connected_components_of X" +proof - + have "compactin X (connected_component_of_set X x)" + if "x \ topspace X" "compactin X (quasi_component_of_set X x)" for x + proof (rule closed_compactin) + show "compactin X (quasi_component_of_set X x)" + by (simp add: that) + show "connected_component_of_set X x \ quasi_component_of_set X x" + by (simp add: connected_component_subset_quasi_component_of) + show "closedin X (connected_component_of_set X x)" + by (simp add: closedin_connected_component_of) + qed + moreover have "connected_component_of X x = quasi_component_of X x" + if \
: "x \ topspace X" "compactin X (connected_component_of_set X x)" for x + proof - + have "\y. connected_component_of X x y \ quasi_component_of X x y" + by (simp add: connected_imp_quasi_component_of) + moreover have False if non: "\ connected_component_of X x y" and quasi: "quasi_component_of X x y" for y + proof - + have "y \ topspace X" + by (meson quasi_component_of_equiv that) + then have "closedin X {y}" + by (simp add: \Hausdorff_space X\ compactin_imp_closedin) + moreover have "disjnt (connected_component_of_set X x) {y}" + by (simp add: non) + moreover have "\ separated_between X (connected_component_of_set X x) {y}" + using \
quasi separated_between_pointwise_left + by (fastforce simp: quasi_component_nonseparated connected_component_of_refl) + ultimately show False + using assms by (metis \
connected_component_in_connected_components_of separated_between_compact_connected_component) + qed + ultimately show ?thesis + by blast + qed + ultimately show ?thesis + using \compactin X C\ unfolding connected_components_of_def image_iff quasi_components_of_def by metis +qed + + +lemma boundary_bumping_theorem_closed_gen: + assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" "closedin X S" + "S \ topspace X" and C: "compactin X C" "C \ connected_components_of (subtopology X S)" + shows "C \ X frontier_of S \ {}" +proof + assume \
: "C \ X frontier_of S = {}" + consider "C \ {}" "X frontier_of S \ topspace X" | "C \ topspace X" "S = {}" + using C by (metis frontier_of_subset_topspace nonempty_connected_components_of) + then show False + proof cases + case 1 + have "separated_between (subtopology X S) C (X frontier_of S)" + proof (rule separated_between_compact_connected_component) + show "compactin (subtopology X S) C" + using C compact_imp_compactin_subtopology connected_components_of_subset by fastforce + show "closedin (subtopology X S) (X frontier_of S)" + by (simp add: \closedin X S\ closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin) + show "disjnt C (X frontier_of S)" + using \
by (simp add: disjnt_def) + qed (use assms Hausdorff_space_subtopology locally_compact_space_closed_subset in auto) + then have "separated_between X C (X frontier_of S)" + using separated_between_from_closed_subtopology by auto + then have "X frontier_of S = {}" + using \C \ {}\ \connected_space X\ connected_space_separated_between by blast + moreover have "C \ S" + using C connected_components_of_subset by fastforce + ultimately show False + using 1 assms by (metis closedin_subset connected_space_eq_frontier_eq_empty subset_empty) + next + case 2 + then show False + using C connected_components_of_eq_empty by fastforce + qed +qed + +lemma boundary_bumping_theorem_closed: + assumes "connected_space X" "compact_space X" "Hausdorff_space X" "closedin X S" + "S \ topspace X" "C \ connected_components_of(subtopology X S)" + shows "C \ X frontier_of S \ {}" + by (meson assms boundary_bumping_theorem_closed_gen closedin_compact_space closedin_connected_components_of + closedin_trans_full compact_imp_locally_compact_space) + + +lemma intermediate_continuum_exists: + assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" + and C: "compactin X C" "connectedin X C" "C \ {}" "C \ topspace X" + and U: "openin X U" "C \ U" + obtains D where "compactin X D" "connectedin X D" "C \ D" "D \ U" +proof - + have "C \ topspace X" + by (simp add: C compactin_subset_topspace) + with C obtain a where a: "a \ topspace X" "a \ C" + by blast + moreover have "compactin (subtopology X (U - {a})) C" + by (simp add: C U a compact_imp_compactin_subtopology subset_Diff_insert) + moreover have "Hausdorff_space (subtopology X (U - {a}))" + using Hausdorff_space_subtopology assms(3) by blast + moreover + have "locally_compact_space (subtopology X (U - {a}))" + by (rule locally_compact_space_open_subset) + (auto simp: locally_compact_Hausdorff_imp_regular_space open_in_Hausdorff_delete assms) + ultimately obtain V K where V: "openin X V" "a \ V" "V \ U" and K: "compactin X K" "a \ K" "K \ U" + and cloK: "closedin (subtopology X (U - {a})) K" and "C \ V" "V \ K" + using locally_compact_space_compact_closed_compact [of "subtopology X (U - {a})"] assms + by (smt (verit, del_insts) Diff_empty compactin_subtopology open_in_Hausdorff_delete openin_open_subtopology subset_Diff_insert) + then obtain D where D: "D \ connected_components_of (subtopology X K)" and "C \ D" + using C by (metis bot.extremum_unique connectedin_subtopology order.trans exists_connected_component_of_superset subtopology_topspace) + show thesis + proof + have cloD: "closedin (subtopology X K) D" + by (simp add: D closedin_connected_components_of) + then have XKD: "compactin (subtopology X K) D" + by (simp add: K closedin_compact_space compact_space_subtopology) + then show "compactin X D" + using compactin_subtopology_imp_compact by blast + show "connectedin X D" + using D connectedin_connected_components_of connectedin_subtopology by blast + have "K \ topspace X" + using K a by blast + moreover have "V \ X interior_of K" + by (simp add: \openin X V\ \V \ K\ interior_of_maximal) + ultimately have "C \ D" + using boundary_bumping_theorem_closed_gen [of X K C] D \C \ V\ + by (auto simp add: assms K compactin_imp_closedin frontier_of_def) + then show "C \ D" + using \C \ D\ by blast + have "D \ U" + using K(3) \closedin (subtopology X K) D\ closedin_imp_subset by blast + moreover have "D \ U" + using K XKD \C \ D\ assms + by (metis \K \ topspace X\ cloD closedin_imp_subset compactin_imp_closedin connected_space_clopen_in + inf_bot_left inf_le2 subset_antisym) + ultimately + show "D \ U" by blast + qed +qed + +lemma boundary_bumping_theorem_gen: + assumes X: "connected_space X" "locally_compact_space X" "Hausdorff_space X" + and "S \ topspace X" and C: "C \ connected_components_of(subtopology X S)" + and compC: "compactin X (X closure_of C)" + shows "X frontier_of C \ X frontier_of S \ {}" +proof - + have Csub: "C \ topspace X" "C \ S" and "connectedin X C" + using C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology + by fastforce+ + have "C \ {}" + using C nonempty_connected_components_of by blast + obtain "X interior_of C \ X interior_of S" "X closure_of C \ X closure_of S" + by (simp add: Csub closure_of_mono interior_of_mono) + moreover have False if "X closure_of C \ X interior_of S" + proof - + have "X closure_of C = C" + by (meson C closedin_connected_component_of_subtopology closure_of_eq interior_of_subset order_trans that) + with that have "C \ X interior_of S" + by simp + then obtain D where "compactin X D" and "connectedin X D" and "C \ D" and "D \ X interior_of S" + using intermediate_continuum_exists assms \X closure_of C = C\ compC Csub + by (metis \C \ {}\ \connectedin X C\ openin_interior_of psubsetE) + then have "D \ C" + by (metis C \C \ {}\ connected_components_of_maximal connectedin_subtopology disjnt_def inf.orderE interior_of_subset order_trans psubsetE) + then show False + using \C \ D\ by blast + qed + ultimately show ?thesis + by (smt (verit, ccfv_SIG) DiffI disjoint_iff_not_equal frontier_of_def subset_eq) +qed + +lemma boundary_bumping_theorem: + "\connected_space X; compact_space X; Hausdorff_space X; S \ topspace X; + C \ connected_components_of(subtopology X S)\ + \ X frontier_of C \ X frontier_of S \ {}" + by (simp add: boundary_bumping_theorem_gen closedin_compact_space compact_imp_locally_compact_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,4994 +1,5024 @@ (* 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 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_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 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 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) 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) end