merged
authorwenzelm
Sun, 07 May 2023 22:51:23 +0200
changeset 77985 e30a56be9c7b
parent 77943 ffdad62bc235 (diff)
parent 77984 c1b8fdd6588a (current diff)
child 77991 bdb5de00379a
child 77992 1de3db73618e
merged
--- a/src/HOL/Analysis/Abstract_Limits.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Abstract_Limits.thy	Sun May 07 22:51:23 2023 +0200
@@ -41,12 +41,38 @@
   finally show ?thesis by (simp add: True)
 qed auto
 
+lemma nontrivial_limit_atin:
+   "atin X a \<noteq> bot \<longleftrightarrow> a \<in> X derived_set_of topspace X"
+proof 
+  assume L: "atin X a \<noteq> bot"
+  then have "a \<in> topspace X"
+    by (meson atin_degenerate)
+  moreover have "\<not> openin X {a}"
+    using L by (auto simp: eventually_atin trivial_limit_eq)
+  ultimately
+  show "a \<in> X derived_set_of topspace X"
+    by (auto simp: derived_set_of_topspace)
+next
+  assume a: "a \<in> X derived_set_of topspace X"
+  show "atin X a \<noteq> bot"
+  proof
+    assume "atin X a = bot"
+    then have "eventually (\<lambda>_. False) (atin X a)"
+      by simp
+    then show False
+      by (smt (verit, best) a eventually_atin in_derived_set_of insertE insert_Diff)
+  qed
+qed
+
 
 subsection\<open>Limits in a topological space\<close>
 
 definition limitin :: "'a topology \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" where
   "limitin X f l F \<equiv> l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> eventually (\<lambda>x. f x \<in> U) F)"
 
+lemma limitinD: "\<lbrakk>limitin X f l F; openin X U; l \<in> U\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. f x \<in> U) F"
+  by (simp add: limitin_def)
+
 lemma limitin_canonical_iff [simp]: "limitin euclidean f l F \<longleftrightarrow> (f \<longlongrightarrow> l) F"
   by (auto simp: limitin_def tendsto_def)
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy	Sun May 07 22:51:23 2023 +0200
@@ -0,0 +1,3531 @@
+(*  Author:     L C Paulson, University of Cambridge [ported from HOL Light] *)
+
+section \<open>Various Forms of Topological Spaces\<close>
+
+theory Abstract_Topological_Spaces
+  imports Lindelof_Spaces Locally Sum_Topology FSigma
+begin
+
+
+subsection\<open>Connected topological spaces\<close>
+
+lemma connected_space_eq_frontier_eq_empty:
+   "connected_space X \<longleftrightarrow> (\<forall>S. S \<subseteq> topspace X \<and> X frontier_of S = {} \<longrightarrow> S = {} \<or> S = topspace X)"
+  by (meson clopenin_eq_frontier_of connected_space_clopen_in)
+
+lemma connected_space_frontier_eq_empty:
+   "connected_space X \<and> S \<subseteq> topspace X
+        \<Longrightarrow> (X frontier_of S = {} \<longleftrightarrow> S = {} \<or> 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 \<longleftrightarrow>
+        C \<subseteq> topspace X \<and> (\<forall>S T. separatedin X S T \<and> C \<subseteq> S \<union> T \<longrightarrow> C \<subseteq> S \<or> C \<subseteq> 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:
+   "\<lbrakk>connectedin X C; closedin X T; openin X T\<rbrakk> \<Longrightarrow> C \<subseteq> T \<or> 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:
+   "\<lbrakk>quotient_map X X' q; connected_space X\<rbrakk> \<Longrightarrow> 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:
+   "\<lbrakk>retraction_map X X' r; connected_space X\<rbrakk> \<Longrightarrow> 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: "\<nexists>a. S = {a}"
+  shows "S \<subseteq> X derived_set_of S"
+unfolding derived_set_of_def
+proof (intro subsetI CollectI conjI strip)
+  show XS: "x \<in> topspace X" if "x \<in> S" for x
+    using that S connectedin by fastforce 
+  show "\<exists>y. y \<noteq> x \<and> y \<in> S \<and> y \<in> T"
+    if "x \<in> S" and "x \<in> T \<and> 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 \<subseteq> T \<union> (topspace X - {x})"
+      using XS that(2) by auto
+    moreover have "(topspace X - {x}) \<inter> S \<noteq> {}"
+      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:
+  "\<lbrakk>Hausdorff_space X; connectedin X S; \<nexists>a. S = {a}\<rbrakk> \<Longrightarrow> S \<subseteq> X derived_set_of S"
+  by (simp add: Hausdorff_imp_t1_space connectedin_imp_perfect_gen)
+
+
+proposition connected_space_prod_topology:
+   "connected_space(prod_topology X Y) \<longleftrightarrow>
+    topspace(prod_topology X Y) = {} \<or> connected_space X \<and> connected_space Y" (is "?lhs=?rhs")
+proof (cases "topspace(prod_topology X Y) = {}")
+  case True
+  then show ?thesis
+    using connected_space_topspace_empty by blast
+next
+  case False
+  then have nonempty: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
+    by force+
+  show ?thesis 
+  proof
+    assume ?lhs
+    then show ?rhs
+      by (meson connected_space_quotient_map_image nonempty quotient_map_fst quotient_map_snd)
+  next
+    assume ?rhs
+    then have conX: "connected_space X" and conY: "connected_space Y"
+      using False by blast+
+    have False
+      if "openin (prod_topology X Y) U" and "openin (prod_topology X Y) V"
+        and UV: "topspace X \<times> topspace Y \<subseteq> U \<union> V" "U \<inter> V = {}" 
+        and "U \<noteq> {}" and "V \<noteq> {}"
+      for U V
+    proof -
+      have Usub: "U \<subseteq> topspace X \<times> topspace Y" and Vsub: "V \<subseteq> topspace X \<times> topspace Y"
+        using that by (metis openin_subset topspace_prod_topology)+
+      obtain a b where ab: "(a,b) \<in> U" and a: "a \<in> topspace X" and b: "b \<in> topspace Y"
+        using \<open>U \<noteq> {}\<close> Usub by auto
+      have "\<not> topspace X \<times> topspace Y \<subseteq> U"
+        using Usub Vsub \<open>U \<inter> V = {}\<close> \<open>V \<noteq> {}\<close> by auto
+      then obtain x y where x: "x \<in> topspace X" and y: "y \<in> topspace Y" and "(x,y) \<notin> U"
+        by blast
+      have oX: "openin X {x \<in> topspace X. (x,y) \<in> U}" "openin X {x \<in> topspace X. (x,y) \<in> V}"
+       and oY: "openin Y {y \<in> topspace Y. (a,y) \<in> U}" "openin Y {y \<in> topspace Y. (a,y) \<in> V}"
+        by (force intro: openin_continuous_map_preimage [where Y = "prod_topology X Y"] 
+            simp: that continuous_map_pairwise o_def x y a)+
+      have 1: "topspace Y \<subseteq> {y \<in> topspace Y. (a,y) \<in> U} \<union> {y \<in> topspace Y. (a,y) \<in> V}"
+        using a that(3) by auto
+      have 2: "{y \<in> topspace Y. (a,y) \<in> U} \<inter> {y \<in> topspace Y. (a,y) \<in> V} = {}"
+        using that(4) by auto
+      have 3: "{y \<in> topspace Y. (a,y) \<in> U} \<noteq> {}"
+        using ab b by auto
+      have 4: "{y \<in> topspace Y. (a,y) \<in> V} \<noteq> {}"
+      proof -
+        show ?thesis
+          using connected_spaceD [OF conX oX] UV \<open>(x,y) \<notin> U\<close> a x y
+                disjoint_iff_not_equal by blast
+      qed
+      show ?thesis
+        using connected_spaceD [OF conY oY 1 2 3 4] by auto
+    qed
+    then show ?lhs
+      unfolding connected_space_def topspace_prod_topology by blast 
+  qed
+qed
+
+lemma connectedin_Times:
+   "connectedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow>
+        S = {} \<or> T = {} \<or> connectedin X S \<and> connectedin Y T"
+  by (force simp: connectedin_def subtopology_Times connected_space_prod_topology)
+
+
+subsection\<open>The notion of "separated between" (complement of "connected between)"\<close>
+
+definition separated_between 
+  where "separated_between X S T \<equiv>
+        \<exists>U V. openin X U \<and> openin X V \<and> U \<union> V = topspace X \<and> disjnt U V \<and> S \<subseteq> U \<and> T \<subseteq> V"
+
+lemma separated_between_alt:
+   "separated_between X S T \<longleftrightarrow>
+        (\<exists>U V. closedin X U \<and> closedin X V \<and> U \<union> V = topspace X \<and> disjnt U V \<and> S \<subseteq> U \<and> T \<subseteq> 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 \<longleftrightarrow>
+        (\<exists>U. closedin X U \<and> openin X U \<and> S \<subseteq> U \<and> T \<subseteq> 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:
+   "\<lbrakk>separated_between X S T; S' \<subseteq> S; T' \<subseteq> T\<rbrakk> \<Longrightarrow> separated_between X S' T'"
+  by (meson order.trans separated_between)
+
+lemma separated_between_refl:
+   "separated_between X S S \<longleftrightarrow> 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 \<longleftrightarrow> separated_between X T S"
+  by (metis disjnt_sym separated_between_alt sup_commute)
+
+lemma separated_between_imp_subset:
+   "separated_between X S T \<Longrightarrow> S \<subseteq> topspace X \<and> T \<subseteq> topspace X"
+  by (metis le_supI1 le_supI2 separated_between_def)
+
+lemma separated_between_empty: 
+  "(separated_between X {} S \<longleftrightarrow> S \<subseteq> topspace X) \<and> (separated_between X S {} \<longleftrightarrow> S \<subseteq> 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 \<union> U) \<longleftrightarrow> separated_between X S T \<and> separated_between X S U"
+  by (auto simp: separated_between)
+
+lemma separated_between_Un': 
+  "separated_between X (S \<union> T) U \<longleftrightarrow> separated_between X S U \<and> separated_between X T U"
+  by (simp add: separated_between_Un separated_between_sym)
+
+lemma separated_between_imp_disjoint:
+   "separated_between X S T \<Longrightarrow> disjnt S T"
+  by (meson disjnt_iff separated_between_def subsetD)
+
+lemma separated_between_imp_separatedin:
+   "separated_between X S T \<Longrightarrow> separatedin X S T"
+  by (meson separated_between_def separatedin_mono separatedin_open_sets)
+
+lemma separated_between_full:
+  assumes "S \<union> T = topspace X"
+  shows "separated_between X S T \<longleftrightarrow> disjnt S T \<and> closedin X S \<and> openin X S \<and> closedin X T \<and> openin X T"
+proof -
+  have "separated_between X S T \<longrightarrow> 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 \<union> T = topspace X \<Longrightarrow> (separated_between X S T \<longleftrightarrow> 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 \<longleftrightarrow>
+         (S = {} \<longrightarrow> T \<subseteq> topspace X) \<and> (\<forall>x \<in> 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 \<subseteq> topspace X"
+    by (meson equals0I separated_between_imp_subset)
+  show ?lhs
+  proof -
+    obtain U where U: "\<forall>x \<in> S. openin X (U x)"
+      "\<forall>x \<in> S. \<exists>V. openin X V \<and> U x \<union> V = topspace X \<and> disjnt (U x) V \<and> {x} \<subseteq> U x \<and> T \<subseteq> V"
+      using R unfolding separated_between_def by metis
+    then have "S \<subseteq> \<Union>(U ` S)"
+      by blast
+    then obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> (\<Union>i\<in>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 "\<And>x. x \<in> K \<Longrightarrow> closedin X (U x)"
+        by (smt (verit) \<open>K \<subseteq> S\<close> Diff_cancel U(2) Un_Diff Un_Diff_Int disjnt_def openin_closedin_eq subsetD)
+      then show "closedin X (\<Union> (U ` K))"
+        by (metis (mono_tags, lifting) \<open>finite K\<close> closedin_Union finite_imageI image_iff)
+      show "openin X (\<Union> (U ` K))"
+        using U(1) \<open>K \<subseteq> S\<close> by blast
+      show "S \<subseteq> \<Union> (U ` K)"
+        by (simp add: K)
+      have "\<And>x i. \<lbrakk>x \<in> T; i \<in> K; x \<in> U i\<rbrakk> \<Longrightarrow> False"
+        by (meson U(2) \<open>K \<subseteq> S\<close> disjnt_iff subsetD)
+      then show "T \<subseteq> topspace X - \<Union> (U ` K)"
+        using \<open>T \<subseteq> topspace X\<close> by auto
+    qed
+  qed
+qed
+
+lemma separated_between_pointwise_right:
+   "compactin X T
+        \<Longrightarrow> separated_between X S T \<longleftrightarrow> (T = {} \<longrightarrow> S \<subseteq> topspace X) \<and> (\<forall>y \<in> T. separated_between X S {y})"
+  by (meson separated_between_pointwise_left separated_between_sym)
+
+lemma separated_between_closure_of:
+  "S \<subseteq> topspace X \<Longrightarrow> separated_between X (X closure_of S) T \<longleftrightarrow> separated_between X S T"
+  by (meson closure_of_minimal_eq separated_between_alt)
+
+
+lemma separated_between_closure_of':
+ "T \<subseteq> topspace X \<Longrightarrow> separated_between X S (X closure_of T) \<longleftrightarrow> 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 \<longleftrightarrow> S \<subseteq> topspace X \<and> 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 \<longleftrightarrow> T \<subseteq> topspace X \<and> 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 \<longleftrightarrow>
+   T \<subseteq> topspace X \<and> disjnt S T \<and> 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 \<subseteq> U" "X closure_of T - X interior_of T \<subseteq> 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 \<subseteq> U - T" "T \<subseteq> topspace X - (U - T)"
+        using R U(3) by (force simp: disjnt_iff)+
+      have "T \<subseteq> 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 \<inter> 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 \<longleftrightarrow> S \<subseteq> topspace X \<and> disjnt S T \<and> 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:
+  "\<lbrakk>S \<subseteq> topspace X; disjnt S T\<rbrakk>
+   \<Longrightarrow> (separated_between X (X frontier_of S) T \<longleftrightarrow> separated_between X S T)"
+  using separated_between_frontier_of_eq by blast
+
+lemma separated_between_frontier_of':
+ "\<lbrakk>T \<subseteq> topspace X; disjnt S T\<rbrakk>
+   \<Longrightarrow> (separated_between X S (X frontier_of T) \<longleftrightarrow> separated_between X S T)"
+  using separated_between_frontier_of_eq' by auto
+
+lemma connected_space_separated_between:
+  "connected_space X \<longleftrightarrow> (\<forall>S T. separated_between X S T \<longrightarrow> S = {} \<or> 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
+        \<Longrightarrow> (separated_between X S T \<longleftrightarrow> S = {} \<and> T \<subseteq> topspace X \<or> S \<subseteq> topspace X \<and> T = {})"
+  by (metis connected_space_separated_between separated_between_empty)
+
+
+subsection\<open>Connected components\<close>
+
+lemma connected_component_of_subtopology_eq:
+   "connected_component_of (subtopology X U) a = connected_component_of X a \<longleftrightarrow>
+    connected_component_of_set X a \<subseteq> 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 \<in> connected_components_of X" "C \<subseteq> U"
+  shows "C \<in> connected_components_of (subtopology X U)"
+proof -
+  obtain a where a: "connected_component_of_set X a \<subseteq> U" and "a \<in> topspace X"
+             and Ceq: "C = connected_component_of_set X a"
+    using assms by (force simp: connected_components_of_def)
+  then have "a \<in> 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 \<open>a \<in> U\<close> \<open>a \<in> topspace X\<close> 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 \<in> 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 \<in> connected_components_of X" "C' \<in> connected_components_of X"
+    shows "(disjnt C C' \<longleftrightarrow> (C \<noteq> C'))"
+proof -
+  have "C \<noteq> {}"
+    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:
+   "\<lbrakk>C \<in> connected_components_of X; C' \<in> connected_components_of X\<rbrakk> \<Longrightarrow> C \<inter> C' \<noteq> {} \<longleftrightarrow> 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) \<Longrightarrow> finite(connected_components_of X)"
+  by (simp add: Union_connected_components_of finite_UnionD)
+
+lemma connected_component_of_unique:
+   "\<lbrakk>x \<in> C; connectedin X C; \<And>C'. x \<in> C' \<and> connectedin X C' \<Longrightarrow> C' \<subseteq> C\<rbrakk>
+        \<Longrightarrow> 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:
+   "\<lbrakk>C \<in> connected_components_of (subtopology X s); X closure_of C \<subseteq> s\<rbrakk> \<Longrightarrow> 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 \<in> 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) = (\<lambda>x. {x}) ` U"
+  by (simp add: connected_component_of_discrete_topology connected_components_of_def)
+
+lemma connected_component_of_continuous_image:
+   "\<lbrakk>continuous_map X Y f; connected_component_of X x y\<rbrakk>
+        \<Longrightarrow> 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 \<in> 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 " "\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x" 
+    "\<And>y. y \<in> topspace Y \<Longrightarrow> 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 \<times> connected_component_of_set Y y"
+proof (cases "x \<in> topspace X \<and> y \<in> topspace Y")
+  case True
+  show ?thesis
+  proof (rule connected_component_of_unique)
+    show "(x, y) \<in> connected_component_of_set X x \<times> 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 \<times> connected_component_of_set Y y)"
+      by (metis connectedin_Times connectedin_connected_component_of)
+    show "C \<subseteq> connected_component_of_set X x \<times> connected_component_of_set Y y"
+      if "(x, y) \<in> C \<and> 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 \<times> D |C D. C \<in> connected_components_of X \<and> D \<in> connected_components_of Y}" (is "?lhs=?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    apply (clarsimp simp: connected_components_of_def)
+    by (metis (no_types) connected_component_of_pair imageI)
+next
+  show "?rhs \<subseteq> ?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 \<in> extensional I then PiE I (\<lambda>i. connected_component_of_set (X i) (x i)) else {})"
+    (is "?lhs = If _ ?R _")    
+proof (cases "x \<in> topspace(product_topology X I)")
+  case True
+  have "?lhs = (\<Pi>\<^sub>E i\<in>I. connected_component_of_set (X i) (x i))"
+    if xX: "\<And>i. i\<in>I \<Longrightarrow> x i \<in> topspace (X i)" and ext: "x \<in> extensional I"
+  proof (rule connected_component_of_unique)
+    show "x \<in> ?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 \<subseteq> ?R"
+      if "x \<in> C \<and> connectedin (product_topology X I) C" for C 
+    proof -
+      have "C \<subseteq> extensional I"
+        using PiE_def connectedin_subset_topspace that by fastforce
+      have "\<And>y. y \<in> C \<Longrightarrow> y \<in> (\<Pi> i\<in>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 \<open>C \<subseteq> extensional I\<close> 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. \<forall>i \<in> I. B i \<in> connected_components_of(X i)}"  (is "?lhs=?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    by (auto simp: connected_components_of_def connected_component_of_product_topology PiE_iff)
+  show "?rhs \<subseteq> ?lhs"
+  proof
+    fix F
+    assume "F \<in> ?rhs"
+    then obtain B where Feq: "F = Pi\<^sub>E I B" and
+      "\<forall>i\<in>I. \<exists>x\<in>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: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i) \<and> B i = connected_component_of_set (X i) (f i)"
+      by metis
+    then have "(\<lambda>i\<in>I. f i) \<in> ((\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> extensional I)"
+      by simp
+    with f show "F \<in> ?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 \<open>Monotone maps (in the general topological sense)\<close>
+
+
+definition monotone_map 
+  where "monotone_map X Y f ==
+        f ` (topspace X) \<subseteq> topspace Y \<and>
+        (\<forall>y \<in> topspace Y. connectedin X {x \<in> topspace X. f x = y})"
+
+lemma monotone_map:
+  "monotone_map X Y f \<longleftrightarrow>
+   f ` (topspace X) \<subseteq> topspace Y \<and> (\<forall>y. connectedin X {x \<in> 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 \<longleftrightarrow> monotone_map X Y f \<and> f ` (topspace X) \<subseteq> S"
+  by (smt (verit, del_insts) le_inf_iff monotone_map topspace_subtopology)
+
+lemma monotone_map_from_subtopology:
+  assumes "monotone_map X Y f" 
+    "\<And>x y. x \<in> topspace X \<and> y \<in> topspace X \<and> x \<in> S \<and> f x = f y \<Longrightarrow> y \<in> 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 \<and> {x \<in> topspace X. f x \<in> v} = u
+        \<Longrightarrow> 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 \<subseteq> 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 \<in> topspace Y"
+  then have "{x \<in> topspace X. f x = y} = {} \<or> (\<exists>a \<in> topspace X. {x \<in> topspace X. f x = y} = {a})"
+    using assms(2) unfolding inj_on_def by blast
+  then show "connectedin X {x \<in> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 "\<not> connected_space X"
+  then obtain U V where "openin X U" "openin X V" "U \<inter> V = {}"
+    "U \<noteq> {}" "V \<noteq> {}" and topUV: "topspace X \<subseteq> U \<union> V"
+    by (auto simp: connected_space_def)
+  then have UVsub: "U \<subseteq> topspace X" "V \<subseteq> topspace X"
+    by (auto simp: openin_subset)
+  have "\<not> connected_space Y"
+    unfolding connected_space_def not_not
+  proof (intro exI conjI)
+    show "topspace Y \<subseteq> f`U \<union> f`V"
+      by (metis f(2) image_Un quotient_imp_surjective_map subset_Un_eq topUV)
+    show "f`U \<noteq> {}"
+      by (simp add: \<open>U \<noteq> {}\<close>)
+    show "(f`V) \<noteq> {}"
+      by (simp add: \<open>V \<noteq> {}\<close>)
+    have *: "y \<notin> f ` V" if "y \<in> f ` U" for y
+    proof -
+      have \<section>: "connectedin X {x \<in> topspace X. f x = y}"
+        using f(1) monotone_map by fastforce
+      show ?thesis
+        using connectedinD [OF \<section> \<open>openin X U\<close> \<open>openin X V\<close>] UVsub topUV \<open>U \<inter> V = {}\<close> that
+        by (force simp: disjoint_iff)
+    qed
+    then show "f`U \<inter> f`V = {}"
+      by blast
+    show "openin Y (f`U)"
+      using f \<open>openin X U\<close> topUV * unfolding quotient_map_saturated_open by force
+    show "openin Y (f`V)"
+      using f \<open>openin X V\<close> 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 \<or> closedin Y C"
+  shows "connectedin X {x \<in> topspace X. f x \<in> C}"
+proof -
+  have "connected_space (subtopology X {x \<in> topspace X. f x \<in> C})"
+  proof -
+    have "connected_space (subtopology Y C)"
+      using \<open>connectedin Y C\<close> connectedin_def by blast
+    moreover have "quotient_map (subtopology X {a \<in> topspace X. f a \<in> C}) (subtopology Y C) f"
+      by (simp add: assms quotient_map_restriction)
+    ultimately show ?thesis
+      using \<open>monotone_map X Y f\<close> 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 \<longleftrightarrow> (\<forall>C. connectedin Y C \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> C})"
+         (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+    unfolding connectedin_def
+  proof (intro strip conjI)
+    fix C
+    assume C: "C \<subseteq> topspace Y \<and> connected_space (subtopology Y C)"
+    show "connected_space (subtopology X {x \<in> topspace X. f x \<in> C})"
+    proof (rule connected_space_monotone_quotient_map_preimage)
+      show "monotone_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
+        by (simp add: L monotone_map_restriction)
+      show "quotient_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
+      proof (rule continuous_open_imp_quotient_map)
+        show "continuous_map (subtopology X {x \<in> topspace X. f x \<in> 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 "\<forall>y. connectedin Y {y} \<longrightarrow> connectedin X {x \<in> 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 \<longleftrightarrow> (\<forall>C. connectedin Y C \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> C})" 
+         (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+    unfolding connectedin_def
+  proof (intro strip conjI)
+    fix C
+    assume C: "C \<subseteq> topspace Y \<and> connected_space (subtopology Y C)"
+    show "connected_space (subtopology X {x \<in> topspace X. f x \<in> C})"
+    proof (rule connected_space_monotone_quotient_map_preimage)
+      show "monotone_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
+        by (simp add: L monotone_map_restriction)
+      show "quotient_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
+      proof (rule continuous_closed_imp_quotient_map)
+        show "continuous_map (subtopology X {x \<in> topspace X. f x \<in> 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 "\<forall>y. connectedin Y {y} \<longrightarrow> connectedin X {x \<in> topspace X. f x = y}"
+    by (smt (verit) Collect_cong singletonD singletonI)
+  then show ?lhs
+    by (simp add: fim monotone_map_def)
+qed
+
+subsection\<open>Other countability properties\<close>
+
+definition second_countable
+  where "second_countable X \<equiv>
+         \<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and>
+             (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U))"
+
+definition first_countable
+  where "first_countable X \<equiv>
+        \<forall>x \<in> topspace X.
+         \<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and>
+             (\<forall>U. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U))"
+
+definition separable_space
+  where "separable_space X \<equiv>
+        \<exists>C. countable C \<and> C \<subseteq> topspace X \<and> X closure_of C = topspace X"
+
+lemma second_countable:
+   "second_countable X \<longleftrightarrow>
+        (\<exists>\<B>. countable \<B> \<and> openin X = arbitrary union_of (\<lambda>x. x \<in> \<B>))"
+  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 \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V"
+    "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)"
+    using assms by (auto simp: second_countable_def)
+  show ?thesis
+    unfolding second_countable_def
+  proof (intro exI conjI)
+    show "\<forall>V\<in>((\<inter>)S) ` \<B>. openin (subtopology X S) V"
+      using openin_subtopology_Int2 \<B> by blast
+    show "\<forall>U x. openin (subtopology X S) U \<and> x \<in> U \<longrightarrow> (\<exists>V\<in>((\<inter>)S) ` \<B>. x \<in> V \<and> V \<subseteq> U)"
+      using \<B> unfolding openin_subtopology
+      by (smt (verit, del_insts) IntI image_iff inf_commute inf_le1 subset_iff)
+  qed (use \<B> in auto)
+qed
+
+
+lemma second_countable_discrete_topology:
+   "second_countable(discrete_topology U) \<longleftrightarrow> countable U" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  then
+  obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> V \<subseteq> U"
+    "\<And>W x. W \<subseteq> U \<and> x \<in> W \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> W)"
+    by (auto simp: second_countable_def)
+  then have "{x} \<in> \<B>" if "x \<in> U" for x
+    by (metis empty_subsetI insertCI insert_subset subset_antisym that)
+  then show ?rhs
+    by (smt (verit) countable_subset image_subsetI \<open>countable \<B>\<close> countable_image_inj_on [OF _ inj_singleton])
+next
+  assume ?rhs 
+  then show ?lhs
+    unfolding second_countable_def
+    by (rule_tac x="(\<lambda>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: "\<And>U. openin X U \<longrightarrow> openin Y (f ` U)"
+    using assms by (auto simp: open_map_def)
+  obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V"
+    and *: "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> 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 \<and> y \<in> V"
+    then obtain x where "x \<in> topspace X" and x: "f x = y"
+      by (metis fim image_iff openin_subset subsetD)
+
+    then obtain W where "W\<in>\<B>" "x \<in> W" "W \<subseteq> {x \<in> topspace X. f x \<in> V}"
+      using * [of "{x \<in> topspace X. f x \<in> V}" x] V assms openin_continuous_map_preimage 
+      by force
+    then show "\<exists>W \<in> (image f) ` \<B>. y \<in> W \<and> W \<subseteq> V"
+      using x by auto
+  qed (use \<B> openXYf in auto)
+qed
+
+lemma homeomorphic_space_second_countability:
+   "X homeomorphic_space Y \<Longrightarrow> (second_countable X \<longleftrightarrow> 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:
+   "\<lbrakk>retraction_map X Y r; second_countable X\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> 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 \<in> topspace (subtopology X S)"
+  then obtain \<B> where "countable \<B>" and \<B>: "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V"
+    "\<And>U. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)"
+    using assms first_countable_def by force
+  show "\<exists>\<B>. countable \<B> \<and> (\<forall>V\<in>\<B>. openin (subtopology X S) V) \<and> (\<forall>U. openin (subtopology X S) U \<and> x \<in> U \<longrightarrow> (\<exists>V\<in>\<B>. x \<in> V \<and> V \<subseteq> U))"
+  proof (intro exI conjI strip)
+    show "countable (((\<inter>)S) ` \<B>)"
+      using \<open>countable \<B>\<close> by blast
+    show "openin (subtopology X S) V" if "V \<in> ((\<inter>)S) ` \<B>" for V
+      using \<B> openin_subtopology_Int2 that by fastforce
+    show "\<exists>V\<in>((\<inter>)S) ` \<B>. x \<in> V \<and> V \<subseteq> U"
+      if "openin (subtopology X S) U \<and> x \<in> U" for U 
+      using that \<B>(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 \<in> U"
+  show "\<exists>\<B>. countable \<B> \<and> (\<forall>V\<in>\<B>. V \<subseteq> U) \<and> (\<forall>Ua. Ua \<subseteq> U \<and> x \<in> Ua \<longrightarrow> (\<exists>V\<in>\<B>. x \<in> V \<and> V \<subseteq> Ua))"
+    using \<open>x \<in> U\<close> 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 \<in> topspace Y"
+  have openXYf: "\<And>U. openin X U \<longrightarrow> openin Y (f ` U)"
+    using assms by (auto simp: open_map_def)
+  then obtain x where x: "x \<in> topspace X" "f x = y"
+    by (metis \<open>y \<in> topspace Y\<close> fim imageE)
+  obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V"
+    and *: "\<And>U. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)"
+    using assms x first_countable_def by force
+  show "\<exists>\<B>. countable \<B> \<and>
+              (\<forall>V\<in>\<B>. openin Y V) \<and> (\<forall>U. openin Y U \<and> y \<in> U \<longrightarrow> (\<exists>V\<in>\<B>. y \<in> V \<and> V \<subseteq> U))"
+  proof (intro exI conjI strip)
+    fix V assume "openin Y V \<and> y \<in> V"
+    then have "\<exists>W\<in>\<B>. x \<in> W \<and> W \<subseteq> {x \<in> topspace X. f x \<in> V}"
+      using * [of "{x \<in> topspace X. f x \<in> V}"] assms openin_continuous_map_preimage x 
+      by fastforce
+    then show "\<exists>V' \<in> (image f) ` \<B>. y \<in> V' \<and> V' \<subseteq> V"
+      using image_mono x by auto 
+  qed (use \<B> openXYf in force)+
+qed
+
+lemma homeomorphic_space_first_countability:
+  "X homeomorphic_space Y \<Longrightarrow> first_countable X \<longleftrightarrow> 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:
+   "\<lbrakk>retraction_map X Y r; first_countable X\<rbrakk> \<Longrightarrow> 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 \<subseteq> topspace X" "X closure_of C = topspace X"
+    by (meson assms separable_space_def)
+  then have "\<And>x T. \<lbrakk>x \<in> topspace X; x \<in> T; openin (subtopology X S) T\<rbrakk>
+           \<Longrightarrow> \<exists>y. y \<in> S \<and> y \<in> C \<and> y \<in> T"
+    by (smt (verit) \<open>openin X S\<close> in_closure_of openin_open_subtopology subsetD)
+  with C \<open>openin X S\<close> show ?thesis
+    unfolding separable_space_def
+    by (rule_tac x="S \<inter> 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: "\<And>S. f ` (X closure_of S) \<subseteq> Y closure_of f ` S"
+    by (simp add: assms continuous_map_image_closure_subset)
+  obtain C where C: "countable C" "C \<subseteq> 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:
+  "\<lbrakk>quotient_map X Y q; separable_space X\<rbrakk> \<Longrightarrow> separable_space Y"
+  by (meson quotient_imp_continuous_map quotient_imp_surjective_map separable_space_continuous_map_image)
+
+lemma separable_space_retraction_map_image:
+  "\<lbrakk>retraction_map X Y r; separable_space X\<rbrakk> \<Longrightarrow> separable_space Y"
+  using retraction_imp_quotient_map separable_space_quotient_map_image by blast
+
+lemma homeomorphic_separable_space:
+  "X homeomorphic_space Y \<Longrightarrow> (separable_space X \<longleftrightarrow> 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) \<longleftrightarrow> 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 \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V"
+    and *: "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)"
+    using assms by (auto simp: second_countable_def)
+  obtain c where c: "\<And>V. \<lbrakk>V \<in> \<B>; V \<noteq> {}\<rbrakk> \<Longrightarrow> c V \<in> V"
+    by (metis all_not_in_conv)
+  then have **: "\<And>x. x \<in> topspace X \<Longrightarrow> x \<in> X closure_of c ` (\<B> - {{}})"
+    using * by (force simp: closure_of_def)
+  show ?thesis
+    unfolding separable_space_def
+  proof (intro exI conjI)
+    show "countable (c ` (\<B>-{{}}))"
+      using \<B>(1) by blast
+    show "(c ` (\<B>-{{}})) \<subseteq> topspace X"
+      using \<B>(2) c openin_subset by fastforce
+    show "X closure_of (c ` (\<B>-{{}})) = 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 \<U>
+  assume "\<forall>U \<in> \<U>. openin X U" and UU: "\<Union>\<U> = topspace X"
+  obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V"
+    and *: "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)"
+    using assms by (auto simp: second_countable_def)
+  define \<B>' where "\<B>' = {B \<in> \<B>. \<exists>U. U \<in> \<U> \<and> B \<subseteq> U}"
+  have \<B>': "countable \<B>'" "\<Union>\<B>' = \<Union>\<U>"
+    using \<B> using "*" \<open>\<forall>U\<in>\<U>. openin X U\<close> by (fastforce simp: \<B>'_def)+
+  have "\<And>b. \<exists>U. b \<in> \<B>' \<longrightarrow> U \<in> \<U> \<and> b \<subseteq> U" 
+    by (simp add: \<B>'_def)
+  then obtain G where G: "\<And>b. b \<in> \<B>' \<longrightarrow> G b \<in> \<U> \<and> b \<subseteq> G b" 
+    by metis
+  with \<B>' UU show "\<exists>\<V>. countable \<V> \<and> \<V> \<subseteq> \<U> \<and> \<Union>\<V> = topspace X"
+    by (rule_tac x="G ` \<B>'" in exI) fastforce
+qed
+
+subsection \<open>Neigbourhood bases EXTRAS\<close>
+(* Neigbourhood bases (useful for "local" properties of various kind).       *)
+
+lemma openin_topology_neighbourhood_base_unique:
+   "openin X = arbitrary union_of P \<longleftrightarrow>
+        (\<forall>u. P u \<longrightarrow> openin X u) \<and> 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
+        \<Longrightarrow> (neighbourhood_base_at x P X \<longleftrightarrow>
+             (\<forall>w. b w \<and> x \<in> w \<longrightarrow> (\<exists>u v. openin X u \<and> P v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> 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 "\<And>S t. P S \<and> openin X t \<and> (t \<noteq> {}) \<and> t \<subseteq> S \<Longrightarrow> P t"
+  shows "neighbourhood_base_of P X \<longleftrightarrow>
+         (\<forall>x \<in> topspace X. \<exists>u v. openin X u \<and> P v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> 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) \<longleftrightarrow> x \<in> u \<Longrightarrow> 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) \<longleftrightarrow> (\<forall>x \<in> 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 \<longleftrightarrow> 
+  (\<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and> neighbourhood_base_of (\<lambda>A. A\<in>\<B>) X)"
+  by (metis (full_types) openin_topology_neighbourhood_base_unique second_countable)
+
+lemma first_countable_neighbourhood_base_alt:
+   "first_countable X \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. \<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and> neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) 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 \<longleftrightarrow>
+        (\<exists>\<B>. countable \<B> \<and> neighbourhood_base_of (\<lambda>V. V \<in> \<B>) X)" (is "?lhs=?rhs")
+proof
+  assume ?lhs 
+  then show ?rhs
+    using second_countable_neighbourhood_base_alt by blast
+next
+  assume ?rhs 
+  then obtain \<B> where "countable \<B>"
+    and \<B>: "\<And>W x. openin X W \<and> x \<in> W \<longrightarrow> (\<exists>U. openin X U \<and> (\<exists>V. V \<in> \<B> \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W))"
+    by (metis neighbourhood_base_of)
+  then show ?lhs
+    unfolding second_countable_neighbourhood_base_alt neighbourhood_base_of
+    apply (rule_tac x="(\<lambda>u. X interior_of u) ` \<B>" 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 \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. \<exists>\<B>. countable \<B> \<and> neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) 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 \<in> topspace X"
+    with R obtain \<B> where "countable \<B>" and \<B>: "neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) X"
+      by blast
+    then
+    show "\<exists>\<B>. countable \<B> \<and> Ball \<B> (openin X) \<and> neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) X"
+      unfolding neighbourhood_base_at_def
+      apply (rule_tac x="(\<lambda>u. X interior_of u) ` \<B>" in exI)
+      by (smt (verit, best) countable_image image_iff interior_of_eq interior_of_mono openin_interior_of)
+  qed
+qed
+
+
+subsection\<open>$T_0$ spaces and the Kolmogorov quotient\<close>
+
+definition t0_space where
+  "t0_space X \<equiv>
+     \<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x \<noteq> y \<longrightarrow> (\<exists>U. openin X U \<and> (x \<notin> U \<longleftrightarrow> y \<in> U))"
+
+lemma t0_space_expansive:
+   "\<lbrakk>topspace Y = topspace X; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> \<Longrightarrow> t0_space X \<Longrightarrow> t0_space Y"
+  by (metis t0_space_def)
+
+lemma t1_imp_t0_space: "t1_space X \<Longrightarrow> t0_space X"
+  by (metis t0_space_def t1_space_def)
+
+lemma t1_eq_symmetric_t0_space_alt:
+   "t1_space X \<longleftrightarrow>
+      t0_space X \<and>
+      (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x \<in> X closure_of {y} \<longleftrightarrow> y \<in> 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 \<longleftrightarrow> t0_space X \<and> (\<forall>x y. x \<in> X closure_of {y} \<longleftrightarrow> y \<in> X closure_of {x})"
+  by (auto simp: t1_eq_symmetric_t0_space_alt in_closure_of)
+
+lemma Hausdorff_imp_t0_space:
+   "Hausdorff_space X \<Longrightarrow> t0_space X"
+  by (simp add: Hausdorff_imp_t1_space t1_imp_t0_space)
+
+lemma t0_space:
+   "t0_space X \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x \<noteq> y \<longrightarrow> (\<exists>C. closedin X C \<and> (x \<notin> C \<longleftrightarrow> y \<in> 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 \<longleftrightarrow> 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 \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. X closure_of {x} = X closure_of {y} \<longrightarrow> 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 \<Longrightarrow> t0_space (subtopology X U)"
+  by (simp add: t0_space_def openin_subtopology) (metis Int_iff)
+
+lemma t0_space_retraction_map_image:
+   "\<lbrakk>retraction_map X Y r; t0_space X\<rbrakk> \<Longrightarrow> t0_space Y"
+  using hereditary_imp_retractive_property homeomorphic_t0_space t0_space_subtopology by blast
+
+lemma XY: "{x}\<times>{y} = {(x,y)}"
+  by simp
+
+lemma t0_space_prod_topologyI: "\<lbrakk>t0_space X; t0_space Y\<rbrakk> \<Longrightarrow> 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) \<longleftrightarrow> topspace (prod_topology X Y) = {} \<or> t0_space X \<and> 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) \<longleftrightarrow>
+        topspace(product_topology X I) = {} \<or> (\<forall>i \<in> 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 \<in> topspace (product_topology X I)"
+        and y: "y \<in> topspace (product_topology X I)"
+        and "x \<noteq> y"
+      then obtain i where "i \<in> I" "x i \<noteq> 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 \<notin> U \<longleftrightarrow> y i \<in> U)"
+        by (metis t0_space PiE_mem \<open>i \<in> I\<close> \<open>x i \<noteq> y i\<close> topspace_product_topology x y)
+      with \<open>i \<in> I\<close> x y show "\<exists>U. closedin (product_topology X I) U \<and> (x \<notin> U) = (y \<in> U)"
+        by (rule_tac x="PiE I (\<lambda>j. if j = i then U else topspace(X j))" in exI)
+          (simp add: closedin_product_topology PiE_iff)
+    qed
+  qed
+qed
+
+
+subsection \<open>Kolmogorov quotients\<close>
+
+definition Kolmogorov_quotient 
+  where "Kolmogorov_quotient X \<equiv> \<lambda>x. @y. \<forall>U. openin X U \<longrightarrow> (y \<in> U \<longleftrightarrow> x \<in> U)"
+
+lemma Kolmogorov_quotient_in_open:
+   "openin X U \<Longrightarrow> (Kolmogorov_quotient X x \<in> U \<longleftrightarrow> x \<in> U)"
+  by (smt (verit, ccfv_SIG) Kolmogorov_quotient_def someI_ex)
+
+lemma Kolmogorov_quotient_in_topspace:
+   "Kolmogorov_quotient X x \<in> topspace X \<longleftrightarrow> x \<in> topspace X"
+  by (simp add: Kolmogorov_quotient_in_open)
+
+lemma Kolmogorov_quotient_in_closed:
+  "closedin X C \<Longrightarrow> (Kolmogorov_quotient X x \<in> C \<longleftrightarrow> x \<in> 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 \<Longrightarrow> Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \<inter> 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 \<inter> U) = Kolmogorov_quotient X ` S \<inter> U"
+    using Kolmogorov_quotient_in_open [of X U] by auto
+  then show "\<exists>V. openin X V \<and> Kolmogorov_quotient X ` (S \<inter> U) = Kolmogorov_quotient X ` S \<inter> V"
+    using \<open>openin X U\<close> 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 \<Longrightarrow> Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \<inter> 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 \<longleftrightarrow>
+    (\<forall>U. openin X U \<longrightarrow> (x \<in> U \<longleftrightarrow> y \<in> 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 \<longleftrightarrow>
+    (\<forall>U. closedin X U \<longrightarrow> (x \<in> U \<longleftrightarrow> y \<in> 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 \<in> 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 \<Longrightarrow> x \<in> topspace X \<Longrightarrow> 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 \<subseteq> 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 \<subseteq> 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"
+              "\<And>x. x \<in> S \<Longrightarrow> g(Kolmogorov_quotient X x) = f x"
+proof -
+  have "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; Kolmogorov_quotient X x = Kolmogorov_quotient X y\<rbrakk>
+            \<Longrightarrow> 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 \<inter> Kolmogorov_quotient X ` S) = f ` S"
+    "\<And>x. x \<in> S \<Longrightarrow> 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\<open>Closed diagonals and graphs\<close>
+
+lemma Hausdorff_space_closedin_diagonal:
+  "Hausdorff_space X \<longleftrightarrow>
+        closedin (prod_topology X X) ((\<lambda>x. (x,x)) ` topspace X)"
+proof -
+  have \<section>: "((\<lambda>x. (x, x)) ` topspace X) \<subseteq> topspace X \<times> topspace X"
+    by auto
+  show ?thesis
+    apply (simp add: closedin_def openin_prod_topology_alt Hausdorff_space_def disjnt_iff \<section>)
+    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) (\<lambda>x. (x,x)) \<longleftrightarrow> Hausdorff_space X"
+proof -
+  have "section_map X (prod_topology X X) (\<lambda>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) (\<lambda>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 \<in> topspace X. f x = g x}"
+proof -
+  have \<section>:"{x \<in> topspace X. f x = g x} = {x \<in> topspace X. (f x,g x) \<in> ((\<lambda>y.(y,y)) ` topspace Y)}"
+    using f continuous_map_image_subset_topspace by fastforce
+  show ?thesis
+    unfolding \<section>
+  proof (intro closedin_continuous_map_preimage)
+    show "continuous_map X (prod_topology Y Y) (\<lambda>x. (f x, g x))"
+      by (simp add: continuous_map_pairedI f g)
+    show "closedin (prod_topology Y Y) ((\<lambda>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" "\<forall>x\<in>S. r x = x"
+    using assms by (meson retract_of_space_def)
+  then have \<section>: "S = {x \<in> topspace X. r x = x}"
+    using S retract_of_space_imp_subset by (force simp: continuous_map_def)
+  show ?thesis
+    unfolding \<section> 
+    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) ((\<lambda>x. (x, f x)) ` (topspace X)))
+         (\<lambda>x. (x, f x)) fst  \<longleftrightarrow>  continuous_map X Y f" 
+   (is "?lhs=?rhs")
+proof
+  assume ?lhs
+  then 
+  have h: "homeomorphic_map X (subtopology (prod_topology X Y) ((\<lambda>x. (x, f x)) ` topspace X)) (\<lambda>x. (x, f x))"
+    by (auto simp: homeomorphic_maps_map)
+  have "f = snd \<circ> (\<lambda>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 \<open> KC spaces, those where all compact sets are closed.\<close>
+
+definition kc_space 
+  where "kc_space X \<equiv> \<forall>S. compactin X S \<longrightarrow> 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:
+   "\<lbrakk>kc_space X; topspace Y = topspace X; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk>
+      \<Longrightarrow> kc_space Y"
+  by (meson compactin_contractive kc_space_def topology_finer_closedin)
+
+lemma compactin_imp_closedin_gen:
+   "\<lbrakk>kc_space X; compactin X S\<rbrakk> \<Longrightarrow> closedin X S"
+  using kc_space_def by blast
+
+lemma Hausdorff_imp_kc_space: "Hausdorff_space X \<Longrightarrow> kc_space X"
+  by (simp add: compactin_imp_closedin kc_space_def)
+
+lemma kc_imp_t1_space: "kc_space X \<Longrightarrow> t1_space X"
+  by (simp add: finite_imp_compactin kc_space_def t1_space_closedin_finite)
+
+lemma kc_space_subtopology:
+   "kc_space X \<Longrightarrow> 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 \<in> topspace X. f x \<in> 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" "\<And>x. x \<in> topspace Y \<Longrightarrow> 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 \<in> topspace Y. s x \<in> 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 \<Longrightarrow> kc_space X \<longleftrightarrow> 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 \<longleftrightarrow>
+         (\<forall>Y. topspace Y = topspace X \<and> (\<forall>S. openin X S \<longrightarrow> openin Y S) \<and> compact_space Y \<longrightarrow> 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 \<equiv> topology (arbitrary union_of (finite intersection_of (\<lambda>A. A = topspace X - S \<or> openin X A)
+                           relative_to (topspace X)))"
+    have "topspace Y = topspace X"
+      by (auto simp: Y_def)
+    have "openin X T \<longrightarrow> 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 "\<exists>\<F>'. finite \<F>' \<and> \<F>' \<subseteq> \<C> \<and> topspace X \<subseteq> \<Union> \<F>'" 
+        if \<C>: "\<C> \<subseteq> insert (topspace X - S) (Collect (openin X))" and sub: "topspace X \<subseteq> \<Union>\<C>" for \<C>
+      proof -
+        consider "\<C> \<subseteq> Collect (openin X)" | \<V> where "\<C> = insert (topspace X - S) \<V>" "\<V> \<subseteq> Collect (openin X)"
+          using \<C> 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 \<subseteq> \<Union>\<V>"
+            using S sub compactin_subset_topspace by blast
+          with 2 obtain \<F> where "finite \<F> \<and> \<F> \<subseteq> \<V> \<and> S \<subseteq> \<Union>\<F>"
+            using S unfolding compactin_def by (metis Ball_Collect)
+          with 2 show ?thesis
+            by (rule_tac x="insert (topspace X - S) \<F>" in exI) blast
+        qed
+      qed
+    qed (auto simp: Y_def)
+    have "Y = X"
+      using R \<open>\<And>S. openin X S \<longrightarrow> openin Y S\<close> \<open>compact_space Y\<close> \<open>topspace Y = topspace X\<close> 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:
+   "\<lbrakk>compact_space X; kc_space Y; continuous_map X Y f\<rbrakk> \<Longrightarrow> 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 \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> 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 \<subseteq> topspace X"
+      by (simp add: compactin_subset_topspace)
+    moreover have "X closure_of K \<subseteq> K"
+    proof
+      fix x
+      assume x: "x \<in> X closure_of K"
+      have "kc_space (subtopology X K)"
+        by (simp add: R \<open>compactin X K\<close>)
+      have "compactin X (insert x K)"
+        by (metis K x compactin_Un compactin_sing in_closure_of insert_is_Un)
+      then show "x \<in> 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) \<longleftrightarrow> 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) (\<lambda>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) \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> 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 \<equiv> fst ` L \<union> snd ` L"
+    have "L \<subseteq> K \<times> 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: \<open>compactin X K\<close> 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 \<open>L \<subseteq> K \<times> K\<close> 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) \<longleftrightarrow>
+        kc_space X \<and>
+        (\<forall>K. compactin X K \<longrightarrow> 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 \<subseteq> topspace X \<times> topspace Y"
+    using compactin_subset_topspace topspace_prod_topology by blast
+  moreover have "\<exists>T. openin (prod_topology X Y) T \<and> (a,b) \<in> T \<and> T \<subseteq> (topspace X \<times> topspace Y) - K"
+    if ab: "(a,b) \<in> (topspace X \<times> topspace Y) - K" for a b
+  proof - 
+    have "compactin Y {b}"
+      using that by force
+    moreover 
+    have "compactin Y {y \<in> topspace Y. (a,y) \<in> K}"
+    proof -
+      have "compactin (prod_topology X Y) (K \<inter> {a} \<times> 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 \<inter> {a} \<times> topspace Y)  homeomorphic_space 
+                     subtopology Y {y \<in> topspace Y. (a, y) \<in> 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 \<in> topspace Y. (a,y) \<in> K}"
+      using ab by force
+    ultimately obtain V U where VU: "openin Y V" "openin Y U" "{b} \<subseteq> V" "{y \<in> topspace Y. (a,y) \<in> K} \<subseteq> U" "disjnt V U"
+      using Hausdorff_space_compact_separation [OF Y] by blast
+    define V' where "V' \<equiv> topspace Y - U"
+    have W: "closedin Y V'" "{y \<in> topspace Y. (a,y) \<in> K} \<subseteq> topspace Y - V'" "disjnt V (topspace Y - V')"
+      using VU by (auto simp: V'_def disjnt_iff)
+    with VU obtain "V \<subseteq> topspace Y" "V' \<subseteq> topspace Y"
+      by (meson closedin_subset openin_closedin_eq)
+    then obtain "b \<in> V" "disjnt {y \<in> topspace Y. (a,y) \<in> K} V'" "V \<subseteq> V'"
+      using VU unfolding disjnt_iff V'_def by force
+    define C where "C \<equiv> image fst (K \<inter> {z \<in> topspace(prod_topology X Y). snd z \<in> V'})"
+    have "closedin (prod_topology X Y) {z \<in> topspace (prod_topology X Y). snd z \<in> V'}"
+        using closedin_continuous_map_preimage \<open>closedin Y V'\<close> 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) \<times> V)"
+        by (simp add: VU \<open>closedin X C\<close> openin_diff openin_prod_Times_iff)
+      have "a \<notin> C"
+        using VU by (auto simp: C_def V'_def)
+      then show "(a, b) \<in> (topspace X - C) \<times> V"
+        using \<open>a \<notin> C\<close> \<open>b \<in> V\<close> that by blast
+      show "(topspace X - C) \<times> V \<subseteq> topspace X \<times> topspace Y - K"
+        using \<open>V \<subseteq> V'\<close> \<open>V \<subseteq> topspace Y\<close> 
+        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:
+   "\<lbrakk>Hausdorff_space X; kc_space Y\<rbrakk> \<Longrightarrow> kc_space (prod_topology X Y)"
+  using kc_space_prod_topology_left homeomorphic_kc_space homeomorphic_space_prod_topology_swap by blast
+
+
+subsection \<open>Regular spaces\<close>
+
+text \<open>Regular spaces are *not* a priori assumed to be Hausdorff or $T_1$\<close>
+
+definition regular_space 
+  where "regular_space X \<equiv>
+        \<forall>C a. closedin X C \<and> a \<in> topspace X - C
+                \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> C \<subseteq> V \<and> 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: "(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> 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 \<in> topspace Y" and "a \<notin> C"
+    then obtain "closedin X (g ` C)" "g a \<in> topspace X" "g a \<notin> g ` C"
+      using \<open>closedin Y C\<close> 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 \<in> S" "g`C \<subseteq> 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 \<in> f`S \<and> C \<subseteq> 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 "\<exists>U V. openin Y U \<and> openin Y V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V"
+      by metis
+  qed
+qed
+
+lemma homeomorphic_regular_space:
+   "X homeomorphic_space Y
+        \<Longrightarrow> (regular_space X \<longleftrightarrow> regular_space Y)"
+  by (meson homeomorphic_regular_space_aux homeomorphic_space_sym)
+
+lemma regular_space:
+   "regular_space X \<longleftrightarrow>
+        (\<forall>C a. closedin X C \<and> a \<in> topspace X - C
+              \<longrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> 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 \<and> a \<in> topspace X - C"
+  show "(\<exists>V. openin X U \<and> openin X V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V)  
+    \<longleftrightarrow> (openin X U \<and> a \<in> U \<and> 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 \<subseteq> 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 \<longleftrightarrow> regular_space X" (is "?lhs=?rhs")
+proof -
+  have "?lhs \<longleftrightarrow> (\<forall>W x. openin X W \<and> x \<in> W \<longrightarrow>
+                  (\<exists>U. openin X U \<and> (\<exists>V. closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W)))"
+    by (simp add: neighbourhood_base_of)
+  also have "\<dots> \<longleftrightarrow> (\<forall>W x. closedin X W \<and> x \<in> topspace X - W \<longrightarrow>
+                     (\<exists>U. openin X U \<and> (\<exists>V. closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> topspace X - W)))"
+    by (smt (verit) Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq)
+  also have "\<dots> \<longleftrightarrow> ?rhs"
+  proof -
+    have \<section>: "(\<exists>V. closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> topspace X - W) 
+         \<longleftrightarrow> (\<exists>V. openin X V \<and> x \<in> U \<and> W \<subseteq> V \<and> disjnt U V)" (is "?lhs=?rhs")
+      if "openin X U"  "closedin X W" "x \<in> topspace X" "x \<notin> W" for W U x
+    proof
+      assume ?lhs with \<open>closedin X W\<close> show ?rhs
+        unfolding closedin_def by (smt (verit) Diff_mono disjnt_Diff1 double_diff subset_eq)
+    next
+      assume ?rhs with \<open>openin X U\<close> 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 \<section> 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 \<Longrightarrow> 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:
+   "\<lbrakk>retraction_map X Y r; regular_space X\<rbrakk> \<Longrightarrow> regular_space Y"
+  using hereditary_imp_retractive_property homeomorphic_regular_space regular_space_subtopology by blast
+
+lemma regular_t0_imp_Hausdorff_space:
+   "\<lbrakk>regular_space X; t0_space X\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> (t0_space X \<longleftrightarrow> Hausdorff_space X)"
+  using Hausdorff_imp_t0_space regular_t0_imp_Hausdorff_space by blast
+
+lemma regular_t1_imp_Hausdorff_space:
+   "\<lbrakk>regular_space X; t1_space X\<rbrakk> \<Longrightarrow> Hausdorff_space X"
+  by (simp add: regular_t0_imp_Hausdorff_space t1_imp_t0_space)
+
+lemma regular_t1_eq_Hausdorff_space:
+   "regular_space X \<Longrightarrow> t1_space X \<longleftrightarrow> 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 \<in> topspace X" and "a \<notin> S"
+  then show "\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> S \<subseteq> V \<and> 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 = {} \<Longrightarrow> 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 \<and> Hausdorff_space X \<longleftrightarrow>
+    neighbourhood_base_of (\<lambda>C. closedin X C \<and> 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 \<Longrightarrow> kc_space X \<longleftrightarrow> 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 "\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> 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 "\<exists>U V. x \<in> S \<longrightarrow> openin X U \<and> openin X V \<and> x \<in> U \<and> T \<subseteq> V \<and> 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: "\<And>x. x \<in> S \<Longrightarrow> openin X (U x) \<and> openin X (V x) \<and> x \<in> (U x) \<and> T \<subseteq> (V x) \<and> disjnt (U x) (V x)" 
+    by metis
+  then obtain \<F> where "finite \<F>" "\<F> \<subseteq> U ` S" "S \<subseteq> \<Union> \<F>"
+    using S unfolding compactin_def by (smt (verit) UN_iff image_iff subsetI)
+  then obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> \<Union>(U ` K)"
+    by (metis finite_subset_image)
+  show ?thesis 
+  proof (intro exI conjI)
+    show "openin X (\<Union>(U ` K))"
+      using \<open>K \<subseteq> S\<close> UV by blast
+    show "openin X (\<Inter>(V ` K))"
+      using False K UV \<open>K \<subseteq> S\<close> \<open>finite K\<close> by blast
+    show "S \<subseteq> \<Union>(U ` K)"
+      by (simp add: K)
+    show "T \<subseteq> \<Inter>(V ` K)"
+      using UV \<open>K \<subseteq> S\<close> by blast
+    show "disjnt (\<Union>(U ` K)) (\<Inter>(V ` K))"
+      by (smt (verit) Inter_iff UN_E UV \<open>K \<subseteq> S\<close> disjnt_iff image_eqI subset_iff)
+  qed
+qed
+
+lemma regular_space_compact_closed_sets:
+   "regular_space X \<longleftrightarrow>
+        (\<forall>S T. compactin X S \<and> closedin X T \<and> disjnt S T
+           \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> 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 \<in> topspace X" and "x \<notin> S"
+    then obtain U V where "openin X U \<and> openin X V \<and> {x} \<subseteq> U \<and> S \<subseteq> V \<and> disjnt U V"
+      by (metis R compactin_sing disjnt_empty1 disjnt_insert1)
+    then show "\<exists>U. openin X U \<and> x \<in> U \<and> 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) \<longleftrightarrow>
+        topspace X = {} \<or> topspace Y = {} \<or> regular_space X \<and> 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 = {} \<or> 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) \<in> W"
+      then obtain U V where U: "openin X U" "x \<in> U" and V: "openin Y V" "y \<in> V" 
+        and "U \<times> V \<subseteq> W"
+        by (metis openin_prod_topology_alt)
+      obtain D1 C1 where 1: "openin X D1" "closedin X C1" "x \<in> D1" "D1 \<subseteq> C1" "C1 \<subseteq> U"
+        by (metis \<open>regular_space X\<close> U neighbourhood_base_of neighbourhood_base_of_closedin)
+      obtain D2 C2 where 2: "openin Y D2" "closedin Y C2" "y \<in> D2" "D2 \<subseteq> C2" "C2 \<subseteq> V"
+        by (metis \<open>regular_space Y\<close> V neighbourhood_base_of neighbourhood_base_of_closedin)
+      show "\<exists>U V. openin (prod_topology X Y) U \<and> closedin (prod_topology X Y) V \<and>
+                  (x,y) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
+      proof (intro conjI exI)
+        show "openin (prod_topology X Y) (D1 \<times> D2)"
+          by (simp add: 1 2 openin_prod_Times_iff)
+        show "closedin (prod_topology X Y) (C1 \<times> C2)"
+          by (simp add: 1 2 closedin_prod_Times_iff)
+      qed (use 1 2 \<open>U \<times> V \<subseteq> W\<close> in auto)
+    qed
+  qed
+qed
+
+lemma regular_space_product_topology:
+   "regular_space (product_topology X I) \<longleftrightarrow>
+    topspace (product_topology X I) = {} \<or> (\<forall>i \<in> 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 \<in> topspace (product_topology X I)"
+      by blast
+    define \<F> where "\<F> \<equiv> {Pi\<^sub>E I U |U. finite {i \<in> I. U i \<noteq> topspace (X i)}
+                        \<and> (\<forall>i\<in>I. openin (X i) (U i))}"
+    have oo: "openin (product_topology X I) = arbitrary union_of (\<lambda>W. W \<in> \<F>)"
+      by (simp add: \<F>_def openin_product_topology product_topology_base_alt)
+    have "\<exists>U V. openin (product_topology X I) U \<and> closedin (product_topology X I) V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> Pi\<^sub>E I W"
+      if fin: "finite {i \<in> I. W i \<noteq> topspace (X i)}" 
+        and opeW: "\<And>k. k \<in> I \<Longrightarrow> openin (X k) (W k)" and x: "x \<in> PiE I W" for W x
+    proof -
+      have "\<And>i. i \<in> I \<Longrightarrow> \<exists>U V. openin (X i) U \<and> closedin (X i) V \<and> x i \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W i"
+        by (metis False PiE_iff R neighbourhood_base_of neighbourhood_base_of_closedin opeW x)
+      then obtain U C where UC: 
+        "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (U i) \<and> closedin (X i) (C i) \<and> x i \<in> U i \<and> U i \<subseteq> C i \<and> C i \<subseteq> W i"
+        by metis
+      define PI where "PI \<equiv> \<lambda>V. PiE I (\<lambda>i. if W i = topspace(X i) then topspace(X i) else V i)"
+      show ?thesis
+      proof (intro conjI exI)
+        have "\<forall>i\<in>I. W i \<noteq> topspace (X i) \<longrightarrow> 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 \<in> PI U"
+          using UC x by (fastforce simp: PI_def)
+        show "PI U \<subseteq> PI C"
+          by (smt (verit) UC Orderings.order_eq_iff PiE_mono PI_def)
+        show "PI C \<subseteq> 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: \<F>_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: "\<And>y. y \<in> topspace X \<Longrightarrow> closedin Z {x \<in> topspace Z. f x = y}"
+    and contg: "continuous_map Z Y g"
+  shows "closed_map Z (prod_topology X Y) (\<lambda>x. (f x, g x))"
+  unfolding closed_map_def
+proof (intro strip)
+  fix C assume "closedin Z C"
+  then have "C \<subseteq> topspace Z"
+    by (simp add: closedin_subset)
+  have "f ` topspace Z \<subseteq> topspace X" "g ` topspace Z \<subseteq> topspace Y"
+    by (simp_all add: assms closed_map_imp_subset_topspace)
+  show "closedin (prod_topology X Y) ((\<lambda>x. (f x, g x)) ` C)"
+    unfolding closedin_def topspace_prod_topology
+  proof (intro conjI)
+    have "closedin Y (g ` C)"
+      using \<open>closedin Z C\<close> assms(3) closed_map_def by blast
+    with assms show "(\<lambda>x. (f x, g x)) ` C \<subseteq> topspace X \<times> topspace Y"
+      using \<open>C \<subseteq> topspace Z\<close> \<open>f ` topspace Z \<subseteq> topspace X\<close> continuous_map_closedin subsetD by fastforce
+    have *: "\<exists>T. openin (prod_topology X Y) T \<and> (y1,y2) \<in> T \<and> T \<subseteq> topspace X \<times> topspace Y - (\<lambda>x. (f x, g x)) ` C"
+      if "(y1,y2) \<notin> (\<lambda>x. (f x, g x)) ` C" and y1: "y1 \<in> topspace X" and y2: "y2 \<in> topspace Y"
+      for y1 y2
+    proof -
+      define A where "A \<equiv> topspace Z - (C \<inter> {x \<in> topspace Z. f x = y1})"
+      have A: "openin Z A" "{x \<in> topspace Z. g x = y2} \<subseteq> A"
+        using that \<open>closedin Z C\<close> clo that(2) by (auto simp: A_def)
+      obtain V0 where "openin Y V0 \<and> y2 \<in> V0" and UA: "{x \<in> topspace Z. g x \<in> V0} \<subseteq> A"
+        using g A y2 unfolding closed_map_fibre_neighbourhood by blast
+      then obtain V V' where VV: "openin Y V \<and> closedin Y V' \<and> y2 \<in> V \<and> V \<subseteq> V'" and "V' \<subseteq> V0"
+        by (metis (no_types, lifting) \<open>regular_space Y\<close> neighbourhood_base_of neighbourhood_base_of_closedin)
+      with UA have subA: "{x \<in> topspace Z. g x \<in> V'} \<subseteq> A"
+        by blast
+      show ?thesis
+      proof -
+        define B where "B \<equiv> topspace Z - (C \<inter> {x \<in> topspace Z. g x \<in> V'})"
+        have "openin Z B"
+          using VV \<open>closedin Z C\<close> contg by (fastforce simp: B_def continuous_map_closedin)
+        have "{x \<in> topspace Z. f x = y1} \<subseteq> B"
+          using A_def subA by (auto simp: A_def B_def)
+        then obtain U where "openin X U" "y1 \<in> U" and subB: "{x \<in> topspace Z. f x \<in> U} \<subseteq> B"
+          using \<open>openin Z B\<close> y1 f unfolding closed_map_fibre_neighbourhood by meson
+        show ?thesis
+        proof (intro conjI exI)
+          show "openin (prod_topology X Y) (U \<times> V)"
+            by (metis VV \<open>openin X U\<close> openin_prod_Times_iff)
+          show "(y1, y2) \<in> U \<times> V"
+            by (simp add: VV \<open>y1 \<in> U\<close>)
+          show "U \<times> V \<subseteq> topspace X \<times> topspace Y - (\<lambda>x. (f x, g x)) ` C"
+            using VV \<open>C \<subseteq> topspace Z\<close> \<open>openin X U\<close> 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 \<times> topspace Y - (\<lambda>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 \<and> continuous_map Z X f \<and> (\<forall>z \<in> topspace Y. closedin Z {x \<in> topspace Z. g x = z})
+         \<or> regular_space Y \<and> continuous_map Z Y g \<and> (\<forall>y \<in> topspace X. closedin Z {x \<in> topspace Z. f x = y}))"
+         (is "?RSX \<or> ?RSY")
+       shows "closed_map Z (prod_topology X Y) (\<lambda>x. (f x, g x))"
+  using D
+proof
+  assume RSX: ?RSX
+  have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x,y). (y,x)) \<circ> (\<lambda>x. (g x, f x))"
+    by auto
+  show ?thesis
+    unfolding eq
+  proof (rule closed_map_compose)
+    show "closed_map Z (prod_topology Y X) (\<lambda>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) (\<lambda>(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 \<and> regular_space Y \<or> regular_space X \<and> t1_space Y"
+  shows "closed_map Z (prod_topology X Y) (\<lambda>x. (f x, g x))"
+proof (rule closed_map_paired_gen)
+  show "regular_space X \<and> continuous_map Z X f \<and> (\<forall>z\<in>topspace Y. closedin Z {x \<in> topspace Z. g x = z}) \<or> regular_space Y \<and> continuous_map Z Y g \<and> (\<forall>y\<in>topspace X. closedin Z {x \<in> 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 \<circ> f)" "continuous_map Z X (fst \<circ> f)"
+    "closed_map Z Y (snd \<circ> f)" "continuous_map Z Y (snd \<circ> f)"
+    "t1_space X \<and> regular_space Y \<or> regular_space X \<and> t1_space Y"
+  shows "closed_map Z (prod_topology X Y) f"
+proof -
+  have "closed_map Z (prod_topology X Y) (\<lambda>a. ((fst \<circ> f) a, (snd \<circ> 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 \<in> topspace X" and subW: "{x} \<times> C \<subseteq> W"
+  shows "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> C \<subseteq> V \<and> U \<times> V \<subseteq> W"
+proof (cases "C = {}")
+  case True
+  with x show ?thesis by auto
+next
+  case False
+  have "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W" 
+    if "y \<in> C" for y
+    using W openin_prod_topology_alt subW subsetD that by fastforce
+  then obtain U V where UV: "\<And>y. y \<in> C \<Longrightarrow> openin X (U y) \<and> openin Y (V y) \<and> x \<in> U y \<and> y \<in> V y \<and> U y \<times> V y \<subseteq> W" 
+    by metis
+  then obtain D where D: "finite D" "D \<subseteq> C" "C \<subseteq> \<Union> (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 (\<Inter> (U ` D))" "openin Y (\<Union> (V ` D))"
+      using D False UV by blast+
+    show "x \<in> \<Inter> (U ` D)" "C \<subseteq> \<Union> (V ` D)" "\<Inter> (U ` D) \<times> \<Union> (V ` D) \<subseteq> 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 \<in> topspace X \<times> topspace Y. fst x \<in> U} = U \<times> topspace Y"
+    if "U \<subseteq> topspace X" for U
+    using that by force
+  have **: "\<And>U y. \<lbrakk>openin (prod_topology X Y) U; y \<in> topspace X;
+            {x \<in> topspace X \<times> topspace Y. fst x = y} \<subseteq> U\<rbrakk>
+           \<Longrightarrow> \<exists>V. openin X V \<and> y \<in> V \<and> V \<times> topspace Y \<subseteq> 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:
+   "\<lbrakk>closed_map X Y f; regular_space X;
+     \<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}\<rbrakk>
+    \<Longrightarrow> closed_map X (prod_topology X Y) (\<lambda>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"
+    "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
+  shows "closed_map X (prod_topology Y X) (\<lambda>x. (f x, x))"
+proof -
+  have eq: "(\<lambda>x. (f x, x)) = (\<lambda>(x,y). (y,x)) \<circ> (\<lambda>x. (x, f x))"
+    by auto
+  show ?thesis
+    unfolding eq
+  proof (rule closed_map_compose)
+    show "closed_map X (prod_topology X Y) (\<lambda>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) (\<lambda>(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"
+          "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
+  shows "closedin (prod_topology X Y) ((\<lambda>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"
+    "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
+  shows "proper_map X (prod_topology X Y) (\<lambda>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"
+    "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
+  shows "proper_map X (prod_topology Y X) (\<lambda>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 \<in> topspace Y" and "y \<notin> C"
+  have "closed_map X Y f" "(\<forall>y \<in> topspace Y. compactin X {x \<in> topspace X. f x = y})"
+    using pmapf proper_map_def by force+
+  moreover have "closedin X {z \<in> topspace X. f z \<in> C}"
+    using \<open>closedin Y C\<close> contf continuous_map_closedin by fastforce
+  moreover have "disjnt {z \<in> topspace X. f z = y} {z \<in> topspace X. f z \<in> C}"
+    using \<open>y \<notin> C\<close> disjnt_iff by blast
+  ultimately
+  obtain U V where UV: "openin X U" "openin X V" "{z \<in> topspace X. f z = y} \<subseteq> U" "{z \<in> topspace X. f z \<in> C} \<subseteq> V"
+                  and dUV: "disjnt U V"
+    using \<open>y \<in> topspace Y\<close> \<open>regular_space X\<close> unfolding regular_space_compact_closed_sets
+    by meson
+
+  have *: "\<And>U T. openin X U \<and> T \<subseteq> topspace Y \<and> {x \<in> topspace X. f x \<in> T} \<subseteq> U \<longrightarrow>
+         (\<exists>V. openin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U)"
+   using \<open>closed_map X Y f\<close> unfolding closed_map_preimage_neighbourhood by blast
+  obtain V1 where V1: "openin Y V1 \<and> y \<in> V1" and sub1: "{x \<in> topspace X. f x \<in> V1} \<subseteq> U"
+    using * [of U "{y}"] UV \<open>y \<in> topspace Y\<close> by auto
+  moreover
+  obtain V2 where "openin Y V2 \<and> C \<subseteq> V2" and sub2: "{x \<in> topspace X. f x \<in> V2} \<subseteq> V"
+    by (smt (verit, ccfv_SIG) * UV \<open>closedin Y C\<close> closedin_subset mem_Collect_eq subset_iff)
+  moreover have "disjnt V1 V2"
+  proof -
+    have "\<And>x. \<lbrakk>\<forall>x. x \<in> U \<longrightarrow> x \<notin> V; x \<in> V1; x \<in> V2\<rbrakk> \<Longrightarrow> 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 "\<exists>U V. openin Y U \<and> openin Y V \<and> y \<in> U \<and> C \<subseteq> V \<and> disjnt U V"
+    by meson
+qed
+
+lemma regular_space_perfect_map_image:
+   "\<lbrakk>regular_space X; perfect_map X Y f\<rbrakk> \<Longrightarrow> 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 \<longleftrightarrow> 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: "(\<forall>y \<in> topspace Y. compactin X {x \<in> 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 \<in> topspace X" and "x \<notin> C"
+    obtain U1 U2 where "openin X U1" "openin X U2" "{x} \<subseteq> U1" and "disjnt U1 U2"
+      and subV: "C \<inter> {z \<in> topspace X. f z = f x} \<subseteq> U2"
+    proof (rule Hausdorff_space_compact_separation [of X "{x}" "C \<inter> {z \<in> topspace X. f z = f x}", OF \<open>Hausdorff_space X\<close>])
+      show "compactin X {x}"
+        by (simp add: \<open>x \<in> topspace X\<close>)
+      show "compactin X (C \<inter> {z \<in> topspace X. f z = f x})"
+        using \<open>closedin X C\<close> fim \<open>x \<in> topspace X\<close> closed_Int_compactin preYf by fastforce
+      show "disjnt {x} (C \<inter> {z \<in> topspace X. f z = f x})"
+        using \<open>x \<notin> C\<close> by force
+    qed
+    have "closedin Y (f ` (C - U2))"
+      using \<open>closed_map X Y f\<close> \<open>closedin X C\<close> \<open>openin X U2\<close> closed_map_def by blast
+    moreover
+    have "f x \<in> topspace Y - f ` (C - U2)"
+      using \<open>closedin X C\<close> \<open>continuous_map X Y f\<close> \<open>x \<in> topspace X\<close> closedin_subset continuous_map_def subV by fastforce
+    ultimately
+    obtain V1 V2 where VV: "openin Y V1" "openin Y V2" "f x \<in> V1" 
+                and subV2: "f ` (C - U2) \<subseteq> V2" and "disjnt V1 V2"
+      by (meson R regular_space_def)
+    show "\<exists>U U'. openin X U \<and> openin X U' \<and> x \<in> U \<and> C \<subseteq> U' \<and> disjnt U U'"
+    proof (intro exI conjI)
+      show "openin X (U1 \<inter> {x \<in> topspace X. f x \<in> V1})"
+        using VV(1) \<open>continuous_map X Y f\<close> \<open>openin X U1\<close> continuous_map by fastforce
+      show "openin X (U2 \<union> {x \<in> topspace X. f x \<in> V2})"
+        using VV(2) \<open>continuous_map X Y f\<close> \<open>openin X U2\<close> continuous_map by fastforce
+      show "x \<in> U1 \<inter> {x \<in> topspace X. f x \<in> V1}"
+        using VV(3) \<open>x \<in> topspace X\<close> \<open>{x} \<subseteq> U1\<close> by auto
+      show "C \<subseteq> U2 \<union> {x \<in> topspace X. f x \<in> V2}"
+        using \<open>closedin X C\<close> closedin_subset subV2 by auto
+      show "disjnt (U1 \<inter> {x \<in> topspace X. f x \<in> V1}) (U2 \<union> {x \<in> topspace X. f x \<in> V2})"
+        using \<open>disjnt U1 U2\<close> \<open>disjnt V1 V2\<close> by (auto simp: disjnt_iff)
+    qed
+  qed
+qed
+
+
+
+subsection\<open>Locally compact spaces\<close>
+
+definition locally_compact_space 
+  where "locally_compact_space X \<equiv>
+    \<forall>x \<in> topspace X. \<exists>U K. openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> 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 "\<exists>V K. openin Y V \<and> compactin Y K \<and> f x \<in> V \<and> V \<subseteq> K"
+    if "x \<in> topspace X" "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> 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 \<longleftrightarrow> 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" "\<forall>x\<in>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 \<in> topspace X" "x \<in> T"
+    obtain U K where UK: "openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K"
+      by (meson X \<open>x \<in> topspace X\<close> locally_compact_space_def)
+    then have "compactin (subtopology X T) (r ` K) \<and> T \<inter> U \<subseteq> r ` K"
+      by (smt (verit) IntD1 image_compactin image_iff inf_le2 r subset_iff)
+    then show "\<exists>U. openin X U \<and> (\<exists>K. compactin (subtopology X T) K \<and> x \<in> U \<and> T \<inter> U \<subseteq> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<in> W"
+  then obtain U K where "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K"
+    by (metis loc locally_compact_space_def openin_subset subsetD)
+  moreover have "openin X (U \<inter> W) \<and> x \<in> U \<inter> W"
+    using \<open>openin X W\<close> \<open>x \<in> W\<close> \<open>openin X U\<close> \<open>x \<in> U\<close> by blast
+  then have "\<exists>u' v. openin X u' \<and> closedin X v \<and> x \<in> u' \<and> u' \<subseteq> v \<and> v \<subseteq> U \<and> v \<subseteq> W"
+    using reg
+    by (metis le_infE neighbourhood_base_of neighbourhood_base_of_closedin)
+  then show "\<exists>U V. openin X U \<and> compactin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
+    by (meson \<open>U \<subseteq> K\<close> \<open>compactin X K\<close> closed_compactin subset_trans)
+qed
+
+lemma Hausdorff_regular: "\<lbrakk>Hausdorff_space X; neighbourhood_base_of (compactin X) X\<rbrakk> \<Longrightarrow> 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 \<in> W"
+  then have "x \<in> topspace X"
+    using openin_subset by blast 
+  then obtain U K where "openin X U" "compactin X K" and UK: "x \<in> U" "U \<subseteq> K"
+    by (meson loc locally_compact_space_def)
+  with \<open>Hausdorff_space X\<close> have "regular_space (subtopology X K)"
+    using Hausdorff_space_subtopology compact_Hausdorff_imp_regular_space compact_space_subtopology by blast
+  then have "\<exists>U' V'. openin (subtopology X K) U' \<and> closedin (subtopology X K) V' \<and> x \<in> U' \<and> U' \<subseteq> V' \<and> V' \<subseteq> K \<inter> W"
+    unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of
+    by (meson IntI \<open>U \<subseteq> K\<close> \<open>openin X W\<close> \<open>x \<in> U\<close> \<open>x \<in> W\<close> openin_subtopology_Int2 subsetD)
+  then obtain V C where "openin X V" "closedin X C" and VC: "x \<in> K \<inter> V" "K \<inter> V \<subseteq> K \<inter> C" "K \<inter> C \<subseteq> K \<inter> W"
+    by (metis Int_commute closedin_subtopology openin_subtopology)
+  show "\<exists>U V. openin X U \<and> closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
+  proof (intro conjI exI)
+    show "openin X (U \<inter> V)"
+      using \<open>openin X U\<close> \<open>openin X V\<close> by blast
+    show "closedin X (K \<inter> C)"
+      using \<open>closedin X C\<close> \<open>compactin X K\<close> compactin_imp_closedin \<open>Hausdorff_space X\<close> by blast
+  qed (use UK VC in auto)
+qed
+
+lemma locally_compact_space_neighbourhood_base:
+  "Hausdorff_space X \<or> regular_space X
+        \<Longrightarrow> locally_compact_space X \<longleftrightarrow> 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 \<and> (Hausdorff_space X \<or> regular_space X) \<longleftrightarrow> locally_compact_space X \<and> regular_space X"
+  using locally_compact_Hausdorff_imp_regular_space by blast
+
+lemma locally_compact_space_compact_closedin:
+  assumes  "Hausdorff_space X \<or> regular_space X"
+  shows "locally_compact_space X \<longleftrightarrow>
+         (\<forall>x \<in> topspace X. \<exists>U K. openin X U \<and> compactin X K \<and> closedin X K \<and> x \<in> U \<and> U \<subseteq> 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 \<or> regular_space X"
+  shows "locally_compact_space X \<longleftrightarrow>
+         (\<forall>x \<in> topspace X. \<exists>U. openin X U \<and> compactin X (X closure_of U) \<and> x \<in> 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 \<or> regular_space X"
+  shows "locally_compact_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>C. compactin X C \<and> 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 \<open>regular_space X\<close> 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 \<or> regular_space X"
+  shows "locally_compact_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>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 (\<lambda>A. compactin X A \<and> 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 \<or> regular_space X"
+  shows "locally_compact_space X \<longleftrightarrow> 
+             neighbourhood_base_of (\<lambda>U. openin X U \<and> 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 (\<lambda>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 \<or> regular_space X"
+  shows "locally_compact_space X \<longleftrightarrow>
+         (\<forall>K. compactin X K
+              \<longrightarrow> (\<exists>U L. openin X U \<and> compactin X L \<and> closedin X L \<and> K \<subseteq> U \<and> U \<subseteq> L))"
+         (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  then obtain U L where UL: "\<forall>x \<in> topspace X. openin X (U x) \<and> compactin X (L x) \<and> closedin X (L x) \<and> x \<in> U x \<and> U x \<subseteq> 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 \<subseteq> topspace X"
+      by (simp add: compactin_subset_topspace)
+    then have *: "(\<forall>U\<in>U ` K. openin X U) \<and> K \<subseteq> \<Union> (U ` K)"
+      using UL by blast
+    with \<open>compactin X K\<close> obtain KF where KF: "finite KF" "KF \<subseteq> K" "K \<subseteq> \<Union>(U ` KF)"
+      by (metis compactinD finite_subset_image)
+    show "\<exists>U L. openin X U \<and> compactin X L \<and> closedin X L \<and> K \<subseteq> U \<and> U \<subseteq> L"
+    proof (intro conjI exI)
+      show "openin X (\<Union> (U ` KF))"
+        using "*" \<open>KF \<subseteq> K\<close> by fastforce
+      show "compactin X (\<Union> (L ` KF))"
+        by (smt (verit) UL \<open>K \<subseteq> topspace X\<close> KF compactin_Union finite_imageI imageE subset_iff)
+      show "closedin X (\<Union> (L ` KF))"
+        by (smt (verit) UL \<open>K \<subseteq> topspace X\<close> KF closedin_Union finite_imageI imageE subsetD)
+    qed (use UL \<open>K \<subseteq> topspace X\<close> 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 \<and> regular_space X \<longleftrightarrow>
+        neighbourhood_base_of (\<lambda>C. compactin X C \<and> 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 \<and> kc_space X \<longleftrightarrow>
+        locally_compact_space X \<and> 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 \<and> kc_space X \<longleftrightarrow>
+        locally_compact_space X \<and> Hausdorff_space X \<and> regular_space X"
+  using Hausdorff_regular locally_compact_kc_space by blast
+
+lemma locally_compact_kc_imp_regular_space:
+   "\<lbrakk>neighbourhood_base_of (compactin X) X; kc_space X\<rbrakk> \<Longrightarrow> regular_space X"
+  using Hausdorff_regular locally_compact_imp_kc_eq_Hausdorff_space by blast
+
+lemma kc_locally_compact_space:
+   "kc_space X
+    \<Longrightarrow> neighbourhood_base_of (compactin X) X \<longleftrightarrow> locally_compact_space X \<and> Hausdorff_space X \<and> 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 \<in> topspace X" "x \<in> S"
+  then obtain U K where UK: "openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K"
+    by (meson loc locally_compact_space_def)
+  show "\<exists>U. openin (subtopology X S) U \<and> 
+            (\<exists>K. compactin (subtopology X S) K \<and> x \<in> U \<and> U \<subseteq> K)"
+  proof (intro conjI exI)
+    show "openin (subtopology X S) (S \<inter> U)"
+      by (simp add: UK openin_subtopology_Int2)
+    show "compactin (subtopology X S) (S \<inter> 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 \<in> topspace X" "x \<in> S"
+  then obtain U K where UK: "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K"
+    by (meson loc locally_compact_space_def)
+  have "openin X (U \<inter> S)"
+    by (simp add: UK \<open>openin X S\<close> openin_Int)
+  with UK reg x obtain V C 
+      where VC: "openin X V" "closedin X C" "x \<in> V" "V \<subseteq> C" "C \<subseteq> U" "C \<subseteq> S"
+    by (metis IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin)
+  show "\<exists>U. openin (subtopology X S) U \<and> 
+            (\<exists>K. compactin (subtopology X S) K \<and> x \<in> U \<and> U \<subseteq> K)"
+  proof (intro conjI exI)
+    show "openin (subtopology X S) V"
+      using VC by (meson \<open>openin X S\<close> openin_open_subtopology order_trans)
+    show "compactin (subtopology X S) (C \<inter> 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:
+  "\<lbrakk>continuous_map X X' f; open_map X X' f;
+    f ` topspace X = topspace X'; locally_compact_space X\<rbrakk> \<Longrightarrow> 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 \<subseteq> 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 \<in> S"
+  then obtain T K where *: "openin X T" "compactin X K" "K \<subseteq> S" "a \<in> S" "a \<in> T" "S \<inter> T \<subseteq> K"
+    using loc unfolding locally_compact_space_def
+  by (metis IntE S compactin_subtopology inf_commute openin_subtopology topspace_subtopology_subset)
+  have "T \<inter> X closure_of S \<subseteq> X closure_of (T \<inter> S)"
+    by (simp add: "*"(1) openin_Int_closure_of_subset)
+  also have "... \<subseteq> S"
+    using * \<open>Hausdorff_space X\<close> by (metis closure_of_minimal compactin_imp_closedin order.trans inf_commute)
+  finally have "T \<inter> X closure_of S \<subseteq> T \<inter> S" by simp 
+  then have "openin (subtopology X (X closure_of S)) (T \<inter> S)"
+    unfolding openin_subtopology using \<open>openin X T\<close> S closure_of_subset by fastforce
+  with * show "\<exists>T. openin (subtopology X (X closure_of S)) T \<and> a \<in> T \<and> T \<subseteq> S"
+    by blast
+qed
+
+lemma locally_compact_subspace_closed_Int_openin:
+   "\<lbrakk>Hausdorff_space X \<and> S \<subseteq> topspace X \<and> locally_compact_space(subtopology X S)\<rbrakk>
+        \<Longrightarrow> \<exists>C U. closedin X C \<and> openin X U \<and> C \<inter> 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 \<longleftrightarrow> S \<subseteq> topspace X \<and> locally_compact_space(subtopology X S)" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  then obtain "S \<subseteq> 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 "(\<exists>C U. closedin X C \<and> openin X U \<and> C \<inter> U = S) \<longleftrightarrow> S \<subseteq> topspace X \<and> 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 \<inter> U"
+    by blast
+  then have "C \<subseteq> topspace X"
+    by (simp add: closedin_subset)
+  have "locally_compact_space (subtopology (subtopology X C) (topspace (subtopology X C) \<inter> U))"
+    proof (rule locally_compact_space_open_subset)
+  show "regular_space (subtopology X C)"
+    by (simp add: \<open>Hausdorff_space X\<close> loc locally_compact_Hausdorff_imp_regular_space regular_space_subtopology)
+  show "locally_compact_space (subtopology X C)"
+    by (simp add: \<open>closedin X C\<close> loc locally_compact_space_closed_subset)
+  show "openin (subtopology X C) (topspace (subtopology X C) \<inter> U)"
+    by (simp add: \<open>openin X U\<close> Int_left_commute inf_commute openin_Int openin_subtopology_Int2)
+qed
+    then show ?rhs
+      by (metis Seq \<open>C \<subseteq> topspace X\<close> 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:
+   "\<lbrakk>Hausdorff_space X; S \<subseteq> topspace X; X closure_of S = topspace X;
+     locally_compact_space (subtopology X S)\<rbrakk> \<Longrightarrow> 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) \<longleftrightarrow>
+        topspace (prod_topology X Y) = {} \<or>
+        locally_compact_space X \<and> 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 \<in> topspace X" "z \<in> 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 \<in> topspace X" and "y \<in> topspace Y"
+      obtain U C where "openin X U" "compactin X C" "x \<in> U" "U \<subseteq> C"
+        by (meson False R \<open>x \<in> topspace X\<close> locally_compact_space_def)
+      obtain V D where "openin Y V" "compactin Y D" "y \<in> V" "V \<subseteq> D"
+        by (meson False R \<open>y \<in> topspace Y\<close> locally_compact_space_def)
+      show "\<exists>U. openin (prod_topology X Y) U \<and> (\<exists>K. compactin (prod_topology X Y) K \<and> (x, y) \<in> U \<and> U \<subseteq> K)"
+      proof (intro exI conjI)
+        show "openin (prod_topology X Y) (U \<times> V)"
+          by (simp add: \<open>openin X U\<close> \<open>openin Y V\<close> openin_prod_Times_iff)
+        show "compactin (prod_topology X Y) (C \<times> D)"
+          by (simp add: \<open>compactin X C\<close> \<open>compactin Y D\<close> compactin_Times)
+        show "(x, y) \<in> U \<times> V"
+          by (simp add: \<open>x \<in> U\<close> \<open>y \<in> V\<close>)
+        show "U \<times> V \<subseteq> C \<times> D"
+          by (simp add: Sigma_mono \<open>U \<subseteq> C\<close> \<open>V \<subseteq> D\<close>)
+      qed
+    qed
+  qed
+qed
+
+lemma locally_compact_space_product_topology:
+   "locally_compact_space(product_topology X I) \<longleftrightarrow>
+        topspace(product_topology X I) = {} \<or>
+        finite {i \<in> I. \<not> compact_space(X i)} \<and> (\<forall>i \<in> 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 \<in> 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 \<in> U" "U \<subseteq> C"
+      by (meson locally_compact_space_def)
+    then obtain V where finV: "finite {i \<in> I. V i \<noteq> topspace (X i)}" and "\<forall>i \<in> I. openin (X i) (V i)" 
+                    and "z \<in> PiE I V" "PiE I V \<subseteq> U"
+      by (auto simp: openin_product_topology_alt)
+    have "compact_space (X i)" if "i \<in> I" "V i = topspace (X i)" for i
+    proof -
+      have "compactin (X i) ((\<lambda>x. x i) ` C)"
+        using \<open>compactin (product_topology X I) C\<close> image_compactin
+        by (metis continuous_map_product_projection \<open>i \<in> I\<close>)
+      moreover have "V i \<subseteq> (\<lambda>x. x i) ` C"
+      proof -
+        have "V i \<subseteq> (\<lambda>x. x i) ` Pi\<^sub>E I V"
+          by (metis \<open>z \<in> Pi\<^sub>E I V\<close> empty_iff image_projection_PiE order_refl \<open>i \<in> I\<close>)
+        also have "\<dots> \<subseteq> (\<lambda>x. x i) ` C"
+          using \<open>U \<subseteq> C\<close> \<open>Pi\<^sub>E I V \<subseteq> U\<close> 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 \<in> I. \<not> compact_space (X i)}"
+      by (metis (mono_tags, lifting) mem_Collect_eq finite_subset subsetI)
+    moreover have "locally_compact_space (X i)" if "i \<in> 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 \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
+      have "\<exists>U C. openin (X i) U \<and> compactin (X i) C \<and> z i \<in> U \<and> U \<subseteq> C \<and>
+                    (compact_space(X i) \<longrightarrow> U = topspace(X i) \<and> C = topspace(X i))" 
+        if "i \<in> 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: "\<And>i. i \<in> I \<Longrightarrow> 
+             openin (X i) (U i) \<and> compactin (X i) (C i) \<and> z i \<in> U i \<and> U i \<subseteq> C i \<and>
+                    (compact_space(X i) \<longrightarrow> U i = topspace(X i) \<and> C i = topspace(X i))"
+        by metis
+      show "\<exists>U. openin (product_topology X I) U \<and> (\<exists>K. compactin (product_topology X I) K \<and> z \<in> U \<and> U \<subseteq> 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) \<longleftrightarrow> (\<forall>i \<in> 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 \<in> I" and y: "y \<in> topspace (X i)"
+    then obtain U K where UK: "openin (X i) U" "compactin (X i) K" "y \<in> U" "U \<subseteq> K"
+      using R by (fastforce simp: locally_compact_space_def)
+    then show "\<exists>U. openin (sum_topology X I) U \<and> (\<exists>K. compactin (sum_topology X I) K \<and> (i, y) \<in> U \<and> U \<subseteq> K)"
+      by (metis \<open>i \<in> I\<close> 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 \<or> regular_space Z" 
+    and f: "quotient_map X Y f"
+  shows "quotient_map (prod_topology Z X) (prod_topology Z Y) (\<lambda>(x,y). (x,f y))"
+proof -
+  define h where "h \<equiv> (\<lambda>(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 \<in> topspace Z \<times> topspace X. h u \<in> W}
+   \<longleftrightarrow> openin (prod_topology Z Y) W"   (is "?lhs=?rhs")
+    if W: "W \<subseteq> topspace Z \<times> topspace Y" for W
+  proof
+    define S where "S \<equiv> {u \<in> topspace Z \<times> topspace X. h u \<in> W}"
+    assume ?lhs 
+    then have L: "openin (prod_topology Z X) S"
+      using S_def by blast
+    have "\<exists>T. openin (prod_topology Z Y) T \<and> (x0, z0) \<in> T \<and> T \<subseteq> W"
+      if \<section>: "(x0,z0) \<in> W" for x0 z0 
+    proof -
+      have x0: "x0 \<in> topspace Z"
+        using W that by blast
+      obtain y0 where y0: "y0 \<in> topspace X" "f y0 = z0"
+        by (metis W fim imageE insert_absorb insert_subset mem_Sigma_iff \<section>)
+      then have "(x0, y0) \<in> S"
+        by (simp add: S_def h_def that x0)
+      have "continuous_map Z (prod_topology Z X) (\<lambda>x. (x, y0))"
+        by (simp add: continuous_map_paired y0)
+      with openin_continuous_map_preimage [OF _ L] 
+      have ope_ZS: "openin Z {x \<in> topspace Z. (x,y0) \<in> S}"
+        by blast
+      obtain U U' where "openin Z U" "compactin Z U'" "closedin Z U'" 
+        "x0 \<in> U"  "U \<subseteq> U'" "U' \<subseteq> {x \<in> topspace Z. (x,y0) \<in> S}"
+        using loc ope_ZS x0 \<open>(x0, y0) \<in> S\<close>
+        by (force simp: locally_compact_space_neighbourhood_base_closedin [OF reg] 
+            neighbourhood_base_of)
+      then have D: "U' \<times> {y0} \<subseteq> S"
+        by (auto simp: )
+      define V where "V \<equiv> {z \<in> topspace Y. U' \<times> {y \<in> topspace X. f y = z} \<subseteq> S}"
+      have "z0 \<in> V"
+        using D y0 Int_Collect fim by (fastforce simp: h_def V_def S_def)
+      have "openin X {x \<in> topspace X. f x \<in> V} \<Longrightarrow> 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 \<in> topspace X. f x \<in> V}"
+      proof -
+        let ?Z = "subtopology Z U'"
+        have *: "{x \<in> topspace X. f x \<in> V} = topspace X - snd ` (U' \<times> topspace X - S)"
+          by (force simp: V_def S_def h_def simp flip: fim)
+        have "compact_space ?Z"
+          using \<open>compactin Z U'\<close> compactin_subspace by auto
+        moreover have "closedin (prod_topology ?Z X) (U' \<times> topspace X - S)"
+          by (simp add: L \<open>closedin Z U'\<close> 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 \<times> V)"
+          by (simp add: openin_prod_Times_iff \<open>openin Z U\<close> \<open>openin Y V\<close>)
+        show "(x0, z0) \<in> U \<times> V"
+          by (simp add: \<open>x0 \<in> U\<close> \<open>z0 \<in> V\<close>)
+        show "U \<times> V \<subseteq> W"
+          using \<open>U \<subseteq> U'\<close> 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 \<or> regular_space Z" 
+    and f: "quotient_map X Y f"
+  shows "quotient_map (prod_topology X Z) (prod_topology Y Z) (\<lambda>(x,y). (f x,y))"
+proof -
+  have "(\<lambda>(x,y). (f x,y)) = prod.swap \<circ> (\<lambda>(x,y). (x,f y)) \<circ> 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) (\<lambda>(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 \<in> topspace X"
+  then obtain U K where "openin X' U" "compactin X' K" "f x \<in> U" "U \<subseteq> K"
+    using assms unfolding locally_compact_space_def perfect_map_def
+    by (metis (no_types, lifting) continuous_map_closedin)
+  show "\<exists>U K. openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K"
+  proof (intro exI conjI)
+    have "continuous_map X X' f"
+      using f perfect_map_def by blast
+    then show "openin X {x \<in> topspace X. f x \<in> U}"
+      by (simp add: \<open>openin X' U\<close> continuous_map)
+    show "compactin X {x \<in> topspace X. f x \<in> K}"
+      using \<open>compactin X' K\<close> f perfect_imp_proper_map proper_map_alt by blast
+  qed (use x \<open>f x \<in> U\<close> \<open>U \<subseteq> K\<close> in auto)
+qed
+
+
+subsection\<open>Special characterizations of classes of functions into and out of R\<close>
+
+lemma monotone_map_into_euclideanreal_alt:
+  assumes "continuous_map X euclideanreal f" 
+  shows "(\<forall>k. is_interval k \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> k}) \<longleftrightarrow>
+         connected_space X \<and> 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 \<in> topspace X. f x \<in> {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 \<noteq> {}" "V \<noteq> {}" "disjnt U V"
+         and UV: "{x \<in> topspace X. f x \<in> {a..b}} = U \<union> V" 
+         and dis: "disjnt U {x \<in> topspace X. f x = b}" "disjnt V {x \<in> topspace X. f x = a}" 
+       for a b U V
+  proof -
+    define E1 where "E1 \<equiv> U \<union> {x \<in> topspace X. f x \<in> {c. c \<le> a}}"
+    define E2 where "E2 \<equiv> V \<union> {x \<in> topspace X. f x \<in> {c. b \<le> c}}"
+    have "closedin X {x \<in> topspace X. f x \<le> a}" "closedin X {x \<in> topspace X. b \<le> 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 \<inter> E2 = {}"
+      unfolding E1_def E2_def using \<open>a<b\<close> \<open>disjnt U V\<close> dis UV
+      by (simp add: disjnt_def set_eq_iff) (smt (verit))
+    have "topspace X \<subseteq> E1 \<union> E2"
+      unfolding E1_def E2_def using UV by fastforce
+    have "E1 = {} \<or> E2 = {}"
+      using R connected_space_closedin
+      using \<open>E1 \<inter> E2 = {}\<close> \<open>closedin X E1\<close> \<open>closedin X E2\<close> \<open>topspace X \<subseteq> E1 \<union> E2\<close> by blast
+    then show False
+      using E1_def E2_def \<open>U \<noteq> {}\<close> \<open>V \<noteq> {}\<close> by fastforce
+  qed
+  show ?lhs
+  proof (intro strip)
+    fix K :: "real set"
+    assume "is_interval K"
+    have False
+      if "a \<in> K" "b \<in> K" and clo: "closedin X U" "closedin X V" 
+         and UV: "{x. x \<in> topspace X \<and> f x \<in> K} \<subseteq> U \<union> V"
+                 "U \<inter> V \<inter> {x. x \<in> topspace X \<and> f x \<in> K} = {}" 
+         and nondis: "\<not> disjnt U {x. x \<in> topspace X \<and> f x = a}"
+                     "\<not> disjnt V {x. x \<in> topspace X \<and> f x = b}" 
+     for a b U V
+    proof -
+      have "\<forall>y. connectedin X {x. x \<in> topspace X \<and> f x = y}"
+        using R monotone_map by fastforce
+      then have **: False if "p \<in> U \<and> q \<in> V \<and> f p = f q \<and> f q \<in> K" for p q
+        unfolding connectedin_closedin
+        using \<open>a \<in> K\<close> \<open>b \<in> K\<close> 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 \<equiv> {x \<in> topspace X. f x \<in> {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 \<inter> W" "V \<inter> W"])
+          show "closedin X (U \<inter> W)" "closedin X (V \<inter> W)"
+            using \<open>closedin X W\<close> clo by auto
+          show "U \<inter> W \<noteq> {}" "V \<inter> W \<noteq> {}"
+            using nondis 1 by (auto simp: disjnt_iff W_def)
+          show "disjnt (U \<inter> W) (V \<inter> W)"
+            using \<open>is_interval K\<close> unfolding is_interval_1 disjnt_iff W_def
+            by (metis (mono_tags, lifting) \<open>a \<in> K\<close> \<open>b \<in> K\<close> ** Int_Collect atLeastAtMost_iff)
+          have "\<And>x. \<lbrakk>x \<in> topspace X; a \<le> f x; f x \<le> b\<rbrakk> \<Longrightarrow> x \<in> U \<or> x \<in> V"
+            using \<open>a \<in> K\<close> \<open>b \<in> K\<close> \<open>is_interval K\<close> UV unfolding is_interval_1 disjnt_iff
+            by blast
+          then show "{x \<in> topspace X. f x \<in> {a..b}} = U \<inter> W \<union> V \<inter> W"
+            by (auto simp: W_def)
+          show "disjnt (U \<inter> W) {x \<in> topspace X. f x = b}" "disjnt (V \<inter> W) {x \<in> topspace X. f x = a}"
+            using ** \<open>a \<in> K\<close> \<open>b \<in> K\<close> nondis by (force simp: disjnt_iff)+
+        qed
+      next
+        case 2
+        then show ?thesis
+          using ** nondis \<open>b \<in> K\<close> by (force simp add: disjnt_iff)
+      next
+        case 3
+        define W where "W \<equiv> {x \<in> topspace X. f x \<in> {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 \<inter> W" "U \<inter> W"])
+          show "closedin X (U \<inter> W)" "closedin X (V \<inter> W)"
+            using \<open>closedin X W\<close> clo by auto
+          show "U \<inter> W \<noteq> {}" "V \<inter> W \<noteq> {}"
+            using nondis 3 by (auto simp: disjnt_iff W_def)
+          show "disjnt (V \<inter> W) (U \<inter> W)"
+            using \<open>is_interval K\<close> unfolding is_interval_1 disjnt_iff W_def
+            by (metis (mono_tags, lifting) \<open>a \<in> K\<close> \<open>b \<in> K\<close> ** Int_Collect atLeastAtMost_iff)
+          have "\<And>x. \<lbrakk>x \<in> topspace X; b \<le> f x; f x \<le> a\<rbrakk> \<Longrightarrow> x \<in> U \<or> x \<in> V"
+            using \<open>a \<in> K\<close> \<open>b \<in> K\<close> \<open>is_interval K\<close> UV unfolding is_interval_1 disjnt_iff
+            by blast
+          then show "{x \<in> topspace X. f x \<in> {b..a}} = V \<inter> W \<union> U \<inter> W"
+            by (auto simp: W_def)
+          show "disjnt (V \<inter> W) {x \<in> topspace X. f x = a}" "disjnt (U \<inter> W) {x \<in> topspace X. f x = b}"
+            using ** \<open>a \<in> K\<close> \<open>b \<in> K\<close> nondis by (force simp: disjnt_iff)+
+        qed      
+      qed
+    qed
+    then show "connectedin X {x \<in> topspace X. f x \<in> K}"
+      unfolding connectedin_closedin disjnt_iff by blast
+  qed
+qed
+
+lemma monotone_map_into_euclideanreal:
+   "\<lbrakk>connected_space X; continuous_map X euclideanreal f\<rbrakk>
+    \<Longrightarrow> monotone_map X euclideanreal f \<longleftrightarrow>
+        (\<forall>k. is_interval k \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> k})"
+  by (simp add: monotone_map_into_euclideanreal_alt)
+
+lemma monotone_map_euclideanreal_alt:
+   "(\<forall>I::real set. is_interval I \<longrightarrow> is_interval {x::real. x \<in> S \<and> f x \<in> I}) \<longleftrightarrow>
+    is_interval S \<and> (mono_on S f \<or> 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 \<in> S" "b \<in> S" "c \<in> S" "a<b" "b<c" and d: "f a < f b \<and> f c < f b \<or> f a > f b \<and> f c > f b" for a b c
+      using d
+    proof
+      assume "f a < f b \<and> 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 \<and> 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 \<or> monotone_on S (\<le>) (\<ge>) 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
+   "\<lbrakk>is_interval S; continuous_on S f\<rbrakk> \<Longrightarrow> 
+    monotone_map (top_of_set S) euclideanreal f \<longleftrightarrow> (mono_on S f \<or> monotone_on S (\<le>) (\<ge>) 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 \<Rightarrow> real"
+  assumes "is_interval S" "continuous_on S f"
+  shows "inj_on f S \<longleftrightarrow> strict_mono_on S f \<or> 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\<open>Normal spaces including Urysohn's lemma and the Tietze extension theorem\<close>
+
+definition normal_space 
+  where "normal_space X \<equiv>
+        \<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T 
+              \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> 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 \<in> topspace X. r x \<in> S}" "closedin X {x \<in> topspace X. r x \<in> T}"
+    using closedin_continuous_map_preimage \<open>closedin Y S\<close> \<open>closedin Y T\<close> r'
+    by (auto simp: retraction_maps_def)
+  moreover
+  have "disjnt {x \<in> topspace X. r x \<in> S} {x \<in> topspace X. r x \<in> T}"
+    using \<open>disjnt S T\<close> by (auto simp: disjnt_def)
+  ultimately
+  obtain U V where UV: "openin X U \<and> openin X V \<and> {x \<in> topspace X. r x \<in> S} \<subseteq> U \<and> {x \<in> topspace X. r x \<in> T} \<subseteq> V" "disjnt U V"
+    by (meson X normal_space_def)
+  show "\<exists>U V. openin Y U \<and> openin Y V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
+  proof (intro exI conjI)
+    show "openin Y {x \<in> topspace Y. r' x \<in> U}" "openin Y {x \<in> topspace Y. r' x \<in> V}"
+      using openin_continuous_map_preimage UV r'
+      by (auto simp: retraction_maps_def)
+    show "S \<subseteq> {x \<in> topspace Y. r' x \<in> U}" "T \<subseteq> {x \<in> topspace Y. r' x \<in> V}"
+      using openin_continuous_map_preimage UV r' \<open>closedin Y S\<close> \<open>closedin Y T\<close> 
+      by (auto simp add: closedin_def continuous_map_closedin retraction_maps_def subset_iff)
+    show "disjnt {x \<in> topspace Y. r' x \<in> U} {x \<in> topspace Y. r' x \<in> V}"
+      using \<open>disjnt U V\<close> by (auto simp: disjnt_def)
+  qed
+qed
+
+lemma homeomorphic_normal_space:
+   "X homeomorphic_space Y \<Longrightarrow> normal_space X \<longleftrightarrow> 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 \<longleftrightarrow>
+    (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T
+          \<longrightarrow> (\<exists>U. openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U)))"
+proof -
+  have "(\<exists>V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V) \<longleftrightarrow> openin X U \<and> S \<subseteq> U \<and> 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 \<Longrightarrow> ?rhs"
+      by (smt (verit, best) disjnt_iff in_closure_of subsetD)
+    assume R: ?rhs
+    then have "(U \<union> S) \<inter> (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 \<subseteq> 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 \<longleftrightarrow>
+    (\<forall>S U. closedin X S \<and> openin X U \<and> S \<subseteq> U \<longrightarrow> (\<exists>V. openin X V \<and> S \<subseteq> V \<and> X closure_of V \<subseteq> U))"
+proof -
+  have "\<exists>V. openin X V \<and> S \<subseteq> V \<and> X closure_of V \<subseteq> U"
+    if "\<And>T. closedin X T \<longrightarrow> disjnt S T \<longrightarrow> (\<exists>U. openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U))"
+       "closedin X S" "openin X U" "S \<subseteq> 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 "\<exists>U. openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U)"
+    if "\<And>U. openin X U \<and> S \<subseteq> U \<longrightarrow> (\<exists>V. openin X V \<and> S \<subseteq> V \<and> X closure_of V \<subseteq> 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 \<longleftrightarrow>
+    (\<forall>S T. S \<subseteq> topspace X \<and> T \<subseteq> topspace X \<and>
+              disjnt (X closure_of S) (X closure_of T)
+              \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V))" 
+   (is "?lhs=?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (meson closedin_closure_of closure_of_subset normal_space_def order.trans)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (metis closedin_subset closure_of_eq normal_space_def)
+qed
+
+lemma normal_space_disjoint_closures:
+   "normal_space X \<longleftrightarrow>
+    (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T
+          \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and>
+                    disjnt (X closure_of U) (X closure_of V)))"
+   (is "?lhs=?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (metis closedin_closure_of normal_space)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (smt (verit) closure_of_subset disjnt_iff normal_space openin_subset subset_eq)
+qed
+
+lemma normal_space_dual:
+   "normal_space X \<longleftrightarrow>
+    (\<forall>U V. openin X U \<longrightarrow> openin X V \<and> U \<union> V = topspace X
+          \<longrightarrow> (\<exists>S T. closedin X S \<and> closedin X T \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> S \<union> T = topspace X))"
+   (is "_ = ?rhs")
+proof -
+  have "normal_space X \<longleftrightarrow>
+        (\<forall>U V. closedin X U \<longrightarrow> closedin X V \<longrightarrow> disjnt U V \<longrightarrow>
+              (\<exists>S T. \<not> (openin X S \<and> openin X T \<longrightarrow>
+                         \<not> (U \<subseteq> S \<and> V \<subseteq> T \<and> disjnt S T))))"
+    unfolding normal_space_def by meson
+  also have "... \<longleftrightarrow> (\<forall>U V. openin X U \<longrightarrow> openin X V \<and> disjnt (topspace X - U) (topspace X - V) \<longrightarrow>
+              (\<exists>S T. \<not> (openin X S \<and> openin X T \<longrightarrow>
+                         \<not> (topspace X - U \<subseteq> S \<and> topspace X - V \<subseteq> T \<and> disjnt S T))))"
+    by (auto simp: all_closedin)
+  also have "... \<longleftrightarrow> ?rhs"
+  proof -
+    have *: "disjnt (topspace X - U) (topspace X - V) \<longleftrightarrow> U \<union> V = topspace X"
+      if "U \<subseteq> topspace X" "V \<subseteq> 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 \<in> topspace X" "y \<in> topspace X" "x \<noteq> y"
+  then have "disjnt {x} {y}"
+    by (auto simp: disjnt_iff)
+  then show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> 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 \<Longrightarrow> t1_space X \<longleftrightarrow> Hausdorff_space X"
+  using normal_t1_imp_Hausdorff_space t1_or_Hausdorff_space by blast
+
+lemma normal_t1_imp_regular_space:
+   "\<lbrakk>normal_space X; t1_space X\<rbrakk> \<Longrightarrow> 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:
+   "\<lbrakk>compact_space X; Hausdorff_space X \<or> regular_space X\<rbrakk>
+        \<Longrightarrow> 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 \<longleftrightarrow>
+    (\<forall>S T. fsigma_in X S \<and> fsigma_in X T \<and> separatedin X S T
+           \<longrightarrow> (\<exists>U B. openin X U \<and> openin X B \<and> S \<subseteq> U \<and> T \<subseteq> B \<and> 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: "\<And>n. closedin X (C n)" "\<And>n. C n \<subseteq> C (Suc n)" "\<Union> (range C) = S"
+      by (meson fsigma_in_ascending)
+    assume "fsigma_in X T" 
+    then obtain D where D: "\<And>n. closedin X (D n)" "\<And>n. D n \<subseteq> D (Suc n)" "\<Union> (range D) = T"
+      by (meson fsigma_in_ascending)
+    assume "separatedin X S T"
+    have "\<And>n. disjnt (D n) (X closure_of S)"
+      by (metis D(3) \<open>separatedin X S T\<close> disjnt_Union1 disjnt_def rangeI separatedin_def)
+    then have "\<And>n. \<exists>V V'. openin X V \<and> openin X V' \<and> D n \<subseteq> V \<and> X closure_of S \<subseteq> V' \<and> disjnt V V'"
+      by (metis D(1) L closedin_closure_of normal_space_def)
+    then obtain V V' where V: "\<And>n. openin X (V n)" and "\<And>n. openin X (V' n)" "\<And>n. disjnt (V n) (V' n)"
+          and DV:  "\<And>n. D n \<subseteq> V n" 
+          and subV': "\<And>n. X closure_of S \<subseteq> V' n"
+      by metis
+    then have VV: "V' n \<inter> 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 "\<And>n. disjnt (C n) (X closure_of T)"
+      by (metis C(3) \<open>separatedin X S T\<close> disjnt_Union1 disjnt_def rangeI separatedin_def)
+    then have "\<And>n. \<exists>U U'. openin X U \<and> openin X U' \<and> C n \<subseteq> U \<and> X closure_of T \<subseteq> U' \<and> disjnt U U'"
+      by (metis C(1) L closedin_closure_of normal_space_def)
+    then obtain U U' where U: "\<And>n. openin X (U n)" and "\<And>n. openin X (U' n)" "\<And>n. disjnt (U n) (U' n)"
+          and CU:  "\<And>n. C n \<subseteq> U n" 
+          and subU': "\<And>n. X closure_of T \<subseteq> U' n"
+      by metis
+    then have UU: "U' n \<inter> 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 "\<exists>U B. openin X U \<and> openin X B \<and> S \<subseteq> U \<and> T \<subseteq> B \<and> disjnt U B"
+    proof (intro conjI exI)
+      have "\<And>S n. closedin X (\<Union>m\<le>n. X closure_of V m)"
+        by (force intro: closedin_Union)
+      then show "openin X (\<Union>n. U n - (\<Union>m\<le>n. X closure_of V m))"
+        using U by blast
+      have "\<And>S n. closedin X (\<Union>m\<le>n. X closure_of U m)"
+        by (force intro: closedin_Union)
+      then show "openin X (\<Union>n. V n - (\<Union>m\<le>n. X closure_of U m))"
+        using V by blast
+      have "S \<subseteq> topspace X"
+        by (simp add: \<open>fsigma_in X S\<close> fsigma_in_subset)
+      then show "S \<subseteq> (\<Union>n. U n - (\<Union>m\<le>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 \<subseteq> topspace X"
+        by (simp add: \<open>fsigma_in X T\<close> fsigma_in_subset)
+      then show "T \<subseteq> (\<Union>n. V n - (\<Union>m\<le>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 "\<And>x m n. \<lbrakk>x \<in> U n; x \<in> V m; \<forall>k\<le>m. x \<notin> X closure_of U k\<rbrakk> \<Longrightarrow> \<exists>k\<le>n. x \<in> X closure_of V k"
+        by (meson U V closure_of_subset nat_le_linear openin_subset subsetD)
+      then show "disjnt (\<Union>n. U n - (\<Union>m\<le>n. X closure_of V m)) (\<Union>n. V n - (\<Union>m\<le>n. X closure_of U m))"
+        by (force simp: disjnt_iff)
+    qed
+  qed
+next
+  show "?rhs \<Longrightarrow> ?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 \<and> openin X B \<and> T \<subseteq> A \<and> U \<subseteq> B \<and> disjnt A B"
+    by (metis assms fsigma_in_fsigma_subtopology normal_space_fsigmas separatedin_subtopology)
+  then
+  show "\<exists>A B. openin (subtopology X S) A \<and> openin (subtopology X S) B \<and> T \<subseteq> A \<and>
+   U \<subseteq> B \<and> 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 \<in> topspace X. f x \<in> S}" "closedin X {x \<in> topspace X. f x \<in> T}"
+    using \<open>closedin Y S\<close> \<open>closedin Y T\<close> closedin_continuous_map_preimage contf by auto
+  moreover
+  have "disjnt {x \<in> topspace X. f x \<in> S} {x \<in> topspace X. f x \<in> T}"
+    using \<open>disjnt S T\<close> by (auto simp: disjnt_iff)
+  ultimately
+  obtain U V where "closedin X U" "closedin X V" 
+    and subXU: "{x \<in> topspace X. f x \<in> S} \<subseteq> topspace X - U" 
+    and subXV: "{x \<in> topspace X. f x \<in> T} \<subseteq> topspace X - V" 
+    and dis: "disjnt (topspace X - U) (topspace X -V)"
+    using \<open>normal_space X\<close> by (force simp add: normal_space_def ex_openin)
+  have "closedin Y (f ` U)" "closedin Y (f ` V)"
+    using \<open>closedin X U\<close> \<open>closedin X V\<close> clof closed_map_def by blast+
+  moreover have "S \<subseteq> topspace Y - f ` U"
+    using \<open>closedin Y S\<close> \<open>closedin X U\<close> subXU by (force dest: closedin_subset)
+  moreover have "T \<subseteq> topspace Y - f ` V"
+    using \<open>closedin Y T\<close> \<open>closedin X V\<close> 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 "\<exists>U V. openin Y U \<and> openin Y V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
+    by (force simp add: ex_openin)
+qed
+
+
+subsection \<open>Hereditary topological properties\<close>
+
+definition hereditarily 
+  where "hereditarily P X \<equiv>
+        \<forall>S. S \<subseteq> topspace X \<longrightarrow> P(subtopology X S)"
+
+lemma hereditarily:
+   "hereditarily P X \<longleftrightarrow> (\<forall>S. P(subtopology X S))"
+  by (metis Int_lower1 hereditarily_def subtopology_restrict)
+
+lemma hereditarily_mono:
+   "\<lbrakk>hereditarily P X; \<And>x. P x \<Longrightarrow> Q x\<rbrakk> \<Longrightarrow> hereditarily Q X"
+  by (simp add: hereditarily)
+
+lemma hereditarily_inc:
+   "hereditarily P X \<Longrightarrow> P X"
+  by (metis hereditarily subtopology_topspace)
+
+lemma hereditarily_subtopology:
+   "hereditarily P X \<Longrightarrow> 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 \<subseteq> topspace Y"
+  then have nx: "normal_space (subtopology X {x \<in> topspace X. f x \<in> T})"
+    by (meson X hereditarily)
+  moreover have "continuous_map (subtopology X {x \<in> topspace X. f x \<in> 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 \<in> topspace X. f x \<in> 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
+      \<Longrightarrow> (hereditarily normal_space X \<longleftrightarrow> 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:
+   "\<lbrakk>retraction_map X Y r; hereditarily normal_space X\<rbrakk> \<Longrightarrow> hereditarily normal_space Y"
+  by (smt (verit) hereditarily_subtopology hereditary_imp_retractive_property homeomorphic_hereditarily_normal_space)
+
+subsection\<open>Limits in a topological space\<close>
+
+lemma limitin_const_iff:
+  assumes "t1_space X" "\<not> trivial_limit F"
+  shows "limitin X (\<lambda>k. a) l F \<longleftrightarrow> l \<in> topspace X \<and> 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 \<sigma> l sequentially" and "S \<subseteq> range \<sigma>" and SX: "S \<subseteq> topspace X"
+  shows "compactin X (insert l S)"
+unfolding compactin_def
+proof (intro conjI strip)
+  show "insert l S \<subseteq> topspace X"
+    by (meson SX insert_subset lim limitin_topspace)
+  fix \<U>
+  assume \<section>: "Ball \<U> (openin X) \<and> insert l S \<subseteq> \<Union> \<U>"
+  have "\<exists>V. finite V \<and> V \<subseteq> \<U> \<and> (\<exists>t \<in> V. l \<in> t) \<and> S \<subseteq> \<Union> V"
+    if *: "\<forall>x \<in> S. \<exists>T \<in> \<U>. x \<in> T" and "T \<in> \<U>" "l \<in> T" for T
+  proof -
+    obtain V where V: "\<And>x. x \<in> S \<Longrightarrow> V x \<in> \<U> \<and> x \<in> V x"
+      using * by metis
+    obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> \<sigma> n \<in> T"
+      by (meson "\<section>" \<open>T \<in> \<U>\<close> \<open>l \<in> T\<close> lim limitin_sequentially)
+    show ?thesis
+    proof (intro conjI exI)
+      have "x \<in> T"
+        if "x \<in> S" and "\<forall>A. (\<forall>x \<in> S. (\<forall>n\<le>N. x \<noteq> \<sigma> n) \<or> A \<noteq> V x) \<or> x \<notin> A" for x 
+        by (metis (no_types) N V that assms(2) imageE nle_le subsetD)
+      then show "S \<subseteq> \<Union> (insert T (V ` (S \<inter> \<sigma> ` {0..N})))"
+        by force
+    qed (use V that in auto)
+  qed
+  then show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> insert l S \<subseteq> \<Union> \<F>"
+    by (smt (verit, best) Union_iff \<section> insert_subset subsetD)
+qed
+
+lemma limitin_Hausdorff_unique:
+  assumes "limitin X f l1 F" "limitin X f l2 F" "\<not> trivial_limit F" "Hausdorff_space X"
+  shows "l1 = l2"
+proof (rule ccontr)
+  assume "l1 \<noteq> l2"
+  with assms obtain U V where "openin X U" "openin X V" "l1 \<in> U" "l2 \<in> V" "disjnt U V"
+    by (metis Hausdorff_space_def limitin_topspace)
+  then have "eventually (\<lambda>x. f x \<in> U) F" "eventually (\<lambda>x. f x \<in> V) F"
+    using assms by (fastforce simp: limitin_def)+
+  then have "\<exists>x. f x \<in> U \<and> f x \<in> V"
+    using assms eventually_elim2 filter_eq_iff by fastforce
+  with assms \<open>disjnt U V\<close> 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 \<noteq> l2"
+  define A where "A \<equiv> insert l1 (range f - {l2})"
+  have "l1 \<in> topspace X"
+    using lim1 limitin_def by fastforce
+  moreover have "compactin X (insert l1 (topspace X \<inter> (range f - {l2})))"
+    by (meson Diff_subset compactin_sequence_with_limit inf_le1 inf_le2 lim1 subset_trans)
+  ultimately have "compactin X (topspace X \<inter> A)"
+    by (simp add: A_def)
+  then have OXA: "openin X (topspace X - A)"
+    by (metis Diff_Diff_Int Diff_subset \<open>kc_space X\<close> kc_space_def openin_closedin_eq)
+  have "l2 \<in> topspace X - A"
+    using \<open>l1 \<noteq> l2\<close> A_def lim2 limitin_topspace by fastforce
+  then have "\<forall>\<^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 \<open>l1 \<noteq> l2\<close> \<open>kc_space X\<close> by fastforce
+qed
+
+lemma limitin_closedin:
+  assumes lim: "limitin X f l F" 
+    and "closedin X S" and ev: "eventually (\<lambda>x. f x \<in> S) F" "\<not> trivial_limit F"
+  shows "l \<in> S"
+proof (rule ccontr)
+  assume "l \<notin> S"
+  have "\<forall>\<^sub>F x in F. f x \<in> topspace X - S"
+    by (metis Diff_iff \<open>l \<notin> S\<close> \<open>closedin X S\<close> closedin_def lim limitin_def)
+  with ev eventually_elim2 trivial_limit_def show False
+    by force
+qed
+
+end
+
--- a/src/HOL/Analysis/Abstract_Topology.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Abstract_Topology.thy	Sun May 07 22:51:23 2023 +0200
@@ -1,5 +1,4 @@
-(*  Author:     L C Paulson, University of Cambridge [ported from HOL Light]
-*)
+(*  Author:     L C Paulson, University of Cambridge [ported from HOL Light] *)
 
 section \<open>Operators involving abstract topology\<close>
 
@@ -174,6 +173,18 @@
   shows "closedin U (S - T)"
   by (metis Int_Diff cT closedin_Int closedin_subset inf.orderE oS openin_closedin_eq)
 
+lemma all_openin: "(\<forall>U. openin X U \<longrightarrow> P U) \<longleftrightarrow> (\<forall>U. closedin X U \<longrightarrow> P (topspace X - U))"
+  by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq)
+
+lemma all_closedin: "(\<forall>U. closedin X U \<longrightarrow> P U) \<longleftrightarrow> (\<forall>U. openin X U \<longrightarrow> P (topspace X - U))"
+  by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq)
+
+lemma ex_openin: "(\<exists>U. openin X U \<and> P U) \<longleftrightarrow> (\<exists>U. closedin X U \<and> P (topspace X - U))"
+  by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq)
+
+lemma ex_closedin: "(\<exists>U. closedin X U \<and> P U) \<longleftrightarrow> (\<exists>U. openin X U \<and> P (topspace X - U))"
+  by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq)
+
 
 subsection\<open>The discrete topology\<close>
 
@@ -285,6 +296,10 @@
   unfolding closedin_def topspace_subtopology
   by (auto simp: openin_subtopology)
 
+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 \<longleftrightarrow> V \<subseteq> topspace U"
   unfolding openin_subtopology
   by auto (metis IntD1 in_mono openin_subset)
@@ -368,6 +383,10 @@
      "closedin X S \<Longrightarrow> (closedin (subtopology X S) T \<longleftrightarrow> closedin X T \<and> T \<subseteq> S)"
   by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE)
 
+lemma closedin_trans_full:
+   "\<lbrakk>closedin (subtopology X U) S; closedin X U\<rbrakk> \<Longrightarrow> closedin X S"
+  using closedin_closed_subtopology by blast
+
 lemma openin_subtopology_Un:
     "\<lbrakk>openin (subtopology X T) S; openin (subtopology X U) S\<rbrakk>
      \<Longrightarrow> openin (subtopology X (T \<union> U)) S"
@@ -1371,7 +1390,7 @@
   by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure)
 
 lemma closure_of_Union_subset: "\<Union>((\<lambda>S. X closure_of S) ` \<A>) \<subseteq> X closure_of (\<Union>\<A>)"
-  by clarify (meson Union_upper closure_of_mono subsetD)
+  by (simp add: SUP_le_iff Sup_upper closure_of_mono)
 
 lemma closure_of_locally_finite_Union:
   assumes "locally_finite_in X \<A>" 
@@ -1867,6 +1886,250 @@
     by (metis \<open>closedin X T\<close> closed_map_def closedin_subtopology f)
 qed
 
+lemma closed_map_closure_of_image:
+   "closed_map X Y f \<longleftrightarrow>
+        f ` topspace X \<subseteq> topspace Y \<and>
+        (\<forall>S. S \<subseteq> topspace X \<longrightarrow> Y closure_of (f ` S) \<subseteq> 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 \<longleftrightarrow> (\<forall>S. image f (X interior_of S) \<subseteq> 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 \<longleftrightarrow> (\<forall>S\<subseteq>topspace X. f ` (X interior_of S) \<subseteq> 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 \<longleftrightarrow>
+       (f ` topspace X \<subseteq> topspace Y \<and> (\<forall>S. f ` (X interior_of S) \<subseteq> 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 \<longleftrightarrow>
+    (f ` topspace X \<subseteq> topspace Y \<and>
+     (\<forall>U T. closedin X U \<and> T \<subseteq> topspace Y \<and>
+            {x \<in> topspace X. f x \<in> T} \<subseteq> U \<longrightarrow>
+            (\<exists>V. closedin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U)))" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs
+  show ?rhs
+  proof (intro conjI strip)
+    show "f ` topspace X \<subseteq> topspace Y"
+      by (simp add: \<open>open_map X Y f\<close> open_map_imp_subset_topspace)
+  next
+    fix U T
+    assume UT: "closedin X U \<and> T \<subseteq> topspace Y \<and> {x \<in> topspace X. f x \<in> T} \<subseteq> 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 "\<exists>V. closedin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> 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 \<in> topspace X. f x \<in> topspace Y - f ` U} \<subseteq> topspace X - U"
+      by blast
+    then obtain V where V: "closedin Y V" 
+      and sub: "topspace Y - f ` U \<subseteq> V" "{x \<in> topspace X. f x \<in> V} \<subseteq> topspace X - U"
+      using R \<open>openin X U\<close> by (meson Diff_subset openin_closedin_eq) 
+    then have "f ` U \<subseteq> topspace Y - V"
+      using R \<open>openin X U\<close> 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 \<longleftrightarrow>
+        image f (topspace X) \<subseteq> topspace Y \<and>
+        (\<forall>U T. openin X U \<and> T \<subseteq> topspace Y \<and>
+              {x \<in> topspace X. f x \<in> T} \<subseteq> U
+              \<longrightarrow> (\<exists>V. openin Y V \<and> T \<subseteq> V \<and>
+                      {x \<in> topspace X. f x \<in> V} \<subseteq> U))" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs
+  show ?rhs
+  proof (intro conjI strip)
+    show "f ` topspace X \<subseteq> topspace Y"
+      by (simp add: L closed_map_imp_subset_topspace)
+  next
+    fix U T
+    assume UT: "openin X U \<and> T \<subseteq> topspace Y \<and> {x \<in> topspace X. f x \<in> T} \<subseteq> 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 "\<exists>V. openin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> 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 \<in> topspace X. f x \<in> topspace Y - f ` U} \<subseteq> topspace X - U"
+      by blast
+    then obtain V where V: "openin Y V" 
+      and sub: "topspace Y - f ` U \<subseteq> V" "{x \<in> topspace X. f x \<in> V} \<subseteq> topspace X - U"
+      using R Diff_subset \<open>closedin X U\<close> unfolding closedin_def
+      by (smt (verit, ccfv_threshold) Collect_mem_eq Collect_mono_iff)
+    then have "f ` U \<subseteq> topspace Y - V"
+      using R \<open>closedin X U\<close> 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 \<longleftrightarrow>
+     f ` (topspace X) \<subseteq> topspace Y \<and>
+     (\<forall>U y. openin X U \<and> y \<in> topspace Y \<and> {x \<in> topspace X. f x = y} \<subseteq> U
+     \<longrightarrow> (\<exists>V. openin Y V \<and> y \<in> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U))"
+  unfolding closed_map_preimage_neighbourhood
+proof (intro conj_cong refl all_cong1)
+  fix U
+  assume "f ` topspace X \<subseteq> topspace Y"
+  show "(\<forall>T. openin X U \<and> T \<subseteq> topspace Y \<and> {x \<in> topspace X. f x \<in> T} \<subseteq> U 
+         \<longrightarrow> (\<exists>V. openin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U))
+      = (\<forall>y. openin X U \<and> y \<in> topspace Y \<and> {x \<in> topspace X. f x = y} \<subseteq> U 
+         \<longrightarrow> (\<exists>V. openin Y V \<and> y \<in> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U))" 
+    (is "(\<forall>T. ?P T) \<longleftrightarrow> (\<forall>y. ?Q y)")
+  proof
+    assume L [rule_format]: "\<forall>T. ?P T"
+    show "\<forall>y. ?Q y"
+    proof
+      fix y show "?Q y"
+        using L [of "{y}"] by blast
+    qed
+  next
+    assume R: "\<forall>y. ?Q y"
+    show "\<forall>T. ?P T"
+    proof (cases "openin X U")
+      case True
+      note [[unify_search_bound=3]]
+      obtain V where V: "\<And>y. \<lbrakk>y \<in> topspace Y; {x \<in> topspace X. f x = y} \<subseteq> U\<rbrakk> \<Longrightarrow>
+                       openin Y (V y) \<and> y \<in> V y \<and> {x \<in> topspace X. f x \<in> V y} \<subseteq> U"
+        using R by (simp add: True) meson
+      show ?thesis
+      proof clarify
+        fix T
+        assume "openin X U" and "T \<subseteq> topspace Y" and "{x \<in> topspace X. f x \<in> T} \<subseteq> U"
+        with V show "\<exists>V. openin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U"
+          by (rule_tac x="\<Union>y\<in>T. V y" in exI) fastforce
+      qed
+    qed auto
+  qed
+qed
+
+lemma open_map_in_subtopology:
+   "openin Y S
+        \<Longrightarrow> (open_map X (subtopology Y S) f \<longleftrightarrow> open_map X Y f \<and> f ` (topspace X) \<subseteq> 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:
+   "\<lbrakk>openin Y S; open_map X (subtopology Y S) f\<rbrakk> \<Longrightarrow> open_map X Y f"
+  using open_map_in_subtopology by blast
+
+lemma closed_map_in_subtopology:
+   "closedin Y S
+        \<Longrightarrow> closed_map X (subtopology Y S) f \<longleftrightarrow> (closed_map X Y f \<and> f ` topspace X \<subseteq> 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:
+   "\<lbrakk>closedin Y S; closed_map X (subtopology Y S) f\<rbrakk> \<Longrightarrow> 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 \<circ> 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 \<in> topspace X. f x \<in> U}"
+    using contf closedin_continuous_map_preimage by blast
+  then have "closedin Z ((g \<circ> f) ` {x \<in> topspace X. f x \<in> U})"
+    using cmf closed_map_def by blast
+  moreover 
+  have "\<And>y. y \<in> U \<Longrightarrow> \<exists>x \<in> topspace X. f x \<in> U \<and> g y = g (f x)"
+    by (smt (verit, ccfv_SIG) \<open>closedin Y U\<close> closedin_subset fim image_iff subsetD)
+  then have "(g \<circ> f) ` {x \<in> topspace X. f x \<in> U} = g`U" by auto
+  ultimately show "closedin Z (g ` U)"
+    by metis
+qed
+
+text \<open>identical proof as the above\<close>
+lemma open_map_from_composition_left:
+  assumes cmf: "open_map X Z (g \<circ> 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 \<in> topspace X. f x \<in> U}"
+    using contf openin_continuous_map_preimage by blast
+  then have "openin Z ((g \<circ> f) ` {x \<in> topspace X. f x \<in> U})"
+    using cmf open_map_def by blast
+  moreover 
+  have "\<And>y. y \<in> U \<Longrightarrow> \<exists>x \<in> topspace X. f x \<in> U \<and> g y = g (f x)"
+    by (smt (verit, ccfv_SIG) \<open>openin Y U\<close> openin_subset fim image_iff subsetD)
+  then have "(g \<circ> f) ` {x \<in> topspace X. f x \<in> 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 \<circ> f)" "f ` topspace X \<subseteq> 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 "\<And>y c. \<lbrakk>y \<in> topspace Y; g y = g (f c); c \<in> C\<rbrakk> \<Longrightarrow> y \<in> f ` C"
+    using \<open>closedin X C\<close> assms closedin_subset inj_onD by fastforce
+  then
+  have "f ` C = {x \<in> topspace Y. g x \<in> (g \<circ> f) ` C}"
+    using \<open>closedin X C\<close> assms(2) closedin_subset by fastforce
+  moreover have "closedin Z ((g \<circ> f) ` C)"
+    using \<open>closedin X C\<close> cmf closed_map_def by blast
+  ultimately show "closedin Y (f ` C)"
+    using assms(3) closedin_continuous_map_preimage by fastforce
+qed
+
+text \<open>identical proof as the above\<close>
+lemma open_map_from_composition_right:
+  assumes "open_map X Z (g \<circ> f)" "f ` topspace X \<subseteq> 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 "\<And>y c. \<lbrakk>y \<in> topspace Y; g y = g (f c); c \<in> C\<rbrakk> \<Longrightarrow> y \<in> f ` C"
+    using \<open>openin X C\<close> assms openin_subset inj_onD by fastforce
+  then
+  have "f ` C = {x \<in> topspace Y. g x \<in> (g \<circ> f) ` C}"
+    using \<open>openin X C\<close> assms(2) openin_subset by fastforce
+  moreover have "openin Z ((g \<circ> f) ` C)"
+    using \<open>openin X C\<close> assms(1) open_map_def by blast
+  ultimately show "openin Y (f ` C)"
+    using assms(3) openin_continuous_map_preimage by fastforce
+qed
+
 subsection\<open>Quotient maps\<close>
                                       
 definition quotient_map where
@@ -2151,6 +2414,84 @@
   qed
 qed
 
+lemma quotient_map_saturated_closed:
+     "quotient_map X Y f \<longleftrightarrow>
+        continuous_map X Y f \<and> f ` (topspace X) = topspace Y \<and>
+        (\<forall>U. closedin X U \<and> {x \<in> topspace X. f x \<in> f ` U} \<subseteq> U \<longrightarrow> closedin Y (f ` U))"
+     (is "?lhs = ?rhs")
+proof
+  assume L: ?lhs
+  then obtain fim: "f ` topspace X = topspace Y"
+    and Y: "\<And>U. U \<subseteq> topspace Y \<Longrightarrow> closedin Y U = closedin X {x \<in> topspace X. f x \<in> 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 \<and> {x \<in> topspace X. f x \<in> f ` U} \<subseteq> U"
+    then have sub:  "f ` U \<subseteq> topspace Y" and eq: "{x \<in> topspace X. f x \<in> 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: "\<And>U. closedin Y U \<Longrightarrow> closedin X {x \<in> topspace X. f x \<in> U}"
+    and fim: "f ` topspace X = topspace Y"
+    and XY: "\<And>U. \<lbrakk>closedin X U; {x \<in> topspace X. f x \<in> f ` U} \<subseteq> U\<rbrakk> \<Longrightarrow> 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 \<subseteq> topspace Y" and X: "closedin X {x \<in> topspace X. f x \<in> U}"
+    have feq: "f ` {x \<in> topspace X. f x \<in> U} = U"
+      using \<open>U \<subseteq> topspace Y\<close> fim by auto
+    show "closedin Y U"
+      using XY [OF X] by (simp add: feq)
+  next
+    fix U :: "'b set"
+    assume "U \<subseteq> topspace Y" and Y: "closedin Y U"
+    show "closedin X {x \<in> topspace X. f x \<in> U}"
+      by (metis YX [OF Y])
+  qed
+qed
+
+lemma quotient_map_onto_image:
+  assumes "f ` topspace X \<subseteq> topspace Y" and U: "\<And>U. U \<subseteq> topspace Y \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> 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 \<subseteq> topspace Y \<inter> f ` topspace X"
+  with U have "openin X {x \<in> topspace X. f x \<in> U} \<Longrightarrow> \<exists>x. openin Y x \<and> U = f ` topspace X \<inter> x"
+    by fastforce
+  moreover have "\<exists>x. openin Y x \<and> U = f ` topspace X \<inter> x \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> U}"
+    by (metis (mono_tags, lifting) Collect_cong IntE IntI U image_eqI openin_subset)
+  ultimately show "openin X {x \<in> topspace X. f x \<in> 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 "\<And>x y. \<lbrakk>x \<in> topspace X; y \<in> topspace X; f x = f y\<rbrakk> \<Longrightarrow> h x = h y"
+  obtains g where "continuous_map Y Z g" "g ` topspace Y = h ` topspace X"
+                  "\<And>x. x \<in> topspace X \<Longrightarrow> g(f x) = h x"
+proof -
+  obtain g where g: "\<And>x. x \<in> topspace X \<Longrightarrow> h x = g(f x)"
+    using function_factors_left_gen[of "\<lambda>x. x \<in> 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\<open> Separated Sets\<close>
 
 definition separatedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
@@ -2245,6 +2586,11 @@
   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 \<union> T = topspace X
+   \<Longrightarrow> separatedin X S T \<longleftrightarrow> disjnt S T \<and> closedin X S \<and> openin X S \<and> closedin X T \<and> openin X T"
+  by (metis separatedin_open_sets separation_closedin_Un_gen separation_openin_Un_gen subtopology_topspace)
+
 
 subsection\<open>Homeomorphisms\<close>
 text\<open>(1-way and 2-way versions may be useful in places)\<close>
@@ -2766,6 +3112,10 @@
     by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset)
 qed
 
+lemma connectedinD:
+     "\<lbrakk>connectedin X S; openin X E1; openin X E2; S \<subseteq> E1 \<union> E2; E1 \<inter> E2 \<inter> S = {}; E1 \<inter> S \<noteq> {}; E2 \<inter> S \<noteq> {}\<rbrakk> \<Longrightarrow> False"
+  by (meson connectedin)
+
 lemma connectedin_iff_connected [simp]: "connectedin euclidean S \<longleftrightarrow> connected S"
   by (simp add: connected_def connectedin)
 
@@ -2966,7 +3316,7 @@
     by (metis closure_of_eq)
 qed
 
-lemma connectedin_inter_frontier_of:
+lemma connectedin_Int_frontier_of:
   assumes "connectedin X S" "S \<inter> T \<noteq> {}" "S - T \<noteq> {}"
   shows "S \<inter> X frontier_of T \<noteq> {}"
 proof -
@@ -3068,7 +3418,7 @@
     moreover have "connectedin (discrete_topology U) S \<longleftrightarrow> (\<exists>a. S = {a})"
     proof
       show "connectedin (discrete_topology U) S \<Longrightarrow> \<exists>a. S = {a}"
-        using False connectedin_inter_frontier_of insert_Diff by fastforce
+        using False connectedin_Int_frontier_of insert_Diff by fastforce
     qed (use True in auto)
     ultimately show ?thesis
       by auto
@@ -3360,7 +3710,7 @@
 corollary compact_space_imp_nest:
   fixes C :: "nat \<Rightarrow> 'a set"
   assumes "compact_space X" and clo: "\<And>n. closedin X (C n)"
-    and ne: "\<And>n. C n \<noteq> {}" and inc: "\<And>m n. m \<le> n \<Longrightarrow> C n \<subseteq> C m"
+    and ne: "\<And>n. C n \<noteq> {}" and dec: "decseq C"
   shows "(\<Inter>n. C n) \<noteq> {}"
 proof -
   let ?\<U> = "range (\<lambda>n. \<Inter>m \<le> n. C m)"
@@ -3370,8 +3720,8 @@
   proof -
     obtain n where "\<And>k. k \<in> K \<Longrightarrow> k \<le> n"
       using Max.coboundedI \<open>finite K\<close> by blast
-    with inc have "C n \<subseteq> (\<Inter>n\<in>K. \<Inter>m \<le> n. C m)"
-    by blast
+    with dec have "C n \<subseteq> (\<Inter>n\<in>K. \<Inter>m \<le> n. C m)"
+      unfolding decseq_def by blast
   with ne [of n] show ?thesis
     by blast
   qed
@@ -3556,6 +3906,10 @@
   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 \<Longrightarrow> (closed_map X Y f \<longleftrightarrow> closedin Y (f ` topspace X))"
+  using closed_map_def embedding_imp_closed_map by blast
+
 
 subsection\<open>Retraction and section maps\<close>
 
@@ -3688,6 +4042,7 @@
   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>\<open>tag unimportant\<close> \<open>Half-global and completely global cases\<close>
 
 lemma continuous_openin_preimage_gen:
@@ -4131,6 +4486,188 @@
   using assms continuous_on_generated_topo_iff by blast
 
 
+subsection\<open>Continuity via bases/subbases, hence upper and lower semicontinuity\<close>
+
+lemma continuous_map_into_topology_base:
+  assumes P: "openin Y = arbitrary union_of P"
+    and f: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y"
+    and ope: "\<And>U. P U \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> U}"
+  shows "continuous_map X Y f"
+proof -
+  have *: "\<And>\<U>. (\<And>t. t \<in> \<U> \<Longrightarrow> P t) \<Longrightarrow> openin X {x \<in> topspace X. \<exists>U\<in>\<U>. f x \<in> 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 \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and> (\<forall>U. P U \<longrightarrow> openin X {x \<in> topspace X. f x \<in> U})"
+ (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  then have "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y"
+    by (meson continuous_map_def)
+  moreover have "\<And>U. P U \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> 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 \<equiv> topology(arbitrary union_of (finite intersection_of P relative_to U))"
+  assumes f: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y"
+    and ope: "\<And>U. P U \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> 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 \<in> topspace X. f x \<in> V}"
+    if \<section>: "(finite intersection_of P relative_to U) V" for V 
+  proof -
+    define finv where "finv \<equiv> \<lambda>V. {x \<in> topspace X. f x \<in> V}"
+    obtain \<U> where \<U>: "finite \<U>" "\<And>V. V \<in> \<U> \<Longrightarrow> P V"
+                      "{x \<in> topspace X. f x \<in> V} = (\<Inter>V \<in> insert U \<U>. finv V)"
+      using \<section> by (fastforce simp: finv_def intersection_of_def relative_to_def)
+    show ?thesis
+      unfolding \<U>
+    proof (intro openin_Inter ope)
+      have U: "U = topspace Y"
+        unfolding Y_def using topspace_subbase by fastforce
+      fix V
+      assume V: "V \<in> finv ` insert U \<U>"
+      with U f have "openin X {x \<in> topspace X. f x \<in> U}"
+        by (auto simp: openin_subopen [of X "Collect _"])
+      then show "openin X V"
+        using V \<U>(2) ope by (fastforce simp: finv_def)
+    qed (use \<open>finite \<U>\<close> 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 \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and> (\<forall>U. P U \<longrightarrow> openin X {x \<in> topspace X. f x \<in> U})"
+   (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+  proof (intro conjI strip)
+    show "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> 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 \<in> topspace X. f x \<in> V} = {x \<in> topspace X. f x \<in> U \<inter> V}"
+      using L by (auto simp: continuous_map)
+    have "openin Y (U \<inter> V)"
+      unfolding assms openin_subbase
+      by (meson \<open>P V\<close> arbitrary_union_of_inc finite_intersection_of_inc relative_to_inc)
+    show "openin X {x \<in> topspace X. f x \<in> V}"
+      using L \<open>openin Y (U \<inter> V)\<close> continuous_map eq by fastforce
+  qed
+next
+  show "?rhs \<Longrightarrow> ?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 (\<lambda>x. x \<in> range greaterThan \<union> range lessThan) relative_to U))
+ = subtopology euclidean U"
+proof -
+  have "\<exists>V. (finite intersection_of (\<lambda>x. x \<in> range greaterThan \<or> x \<in> range lessThan)) V \<and> x \<in> V \<and> V \<subseteq> W"
+    if "open W" "x \<in> W" for W and x::'a
+    using \<open>open W\<close> [unfolded open_generated_order] \<open>x \<in> W\<close>
+  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 "\<And>V::'a set. (finite intersection_of (\<lambda>x. x \<in> range greaterThan \<or> x \<in> range lessThan)) V \<Longrightarrow> open V"
+    by (force simp: intersection_of_def subset_iff)
+  ultimately have *: "openin (euclidean::'a topology) = 
+           (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> range greaterThan \<or> x \<in> 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 \<longleftrightarrow>
+         (\<forall>x \<in> topspace X. f x \<in> U) \<and>
+         (\<forall>a. openin X {x \<in> topspace X. f x > a}) \<and>
+         (\<forall>a. openin X {x \<in> 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 \<Rightarrow> 'b::order_topology"
+  shows "continuous_map X euclidean f \<longleftrightarrow>
+         (\<forall>a. openin X {x \<in> topspace X. f x > a}) \<and>
+         (\<forall>a. openin X {x \<in> topspace X. f x < a})"
+  using continuous_map_upper_lower_semicontinuous_lt_gen [where U=UNIV]
+  by auto
+
+lemma Int_Collect_imp_eq: "A \<inter> {x. x\<in>A \<longrightarrow> P x} = {x\<in>A. P x}"
+  by blast
+
+lemma continuous_map_upper_lower_semicontinuous_le_gen:
+  shows
+    "continuous_map X (subtopology euclideanreal U) f \<longleftrightarrow>
+         (\<forall>x \<in> topspace X. f x \<in> U) \<and>
+         (\<forall>a. closedin X {x \<in> topspace X. f x \<ge> a}) \<and>
+         (\<forall>a. closedin X {x \<in> topspace X. f x \<le> 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 \<longleftrightarrow>
+         (\<forall>a. closedin X {x \<in> topspace X. f x \<ge> a}) \<and>
+         (\<forall>a. closedin X {x \<in> topspace X. f x \<le> 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 \<longleftrightarrow>
+         (\<forall>x \<in> topspace X. f x \<in> U) \<and>
+         (\<forall>a. openin X {x \<in> topspace X. f x < a}) \<and>
+         (\<forall>a. closedin X {x \<in> topspace X. f x \<le> 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 \<longleftrightarrow>
+         (\<forall>a. openin X {x \<in> topspace X. f x < a}) \<and>
+         (\<forall>a. closedin X {x \<in> topspace X. f x \<le> a})"
+  using continuous_map_upper_lower_semicontinuous_lte_gen [where U=UNIV]
+  by auto
+
+
 subsection\<^marker>\<open>tag important\<close> \<open>Pullback topology\<close>
 
 text \<open>Pulling back a topology by map gives again a topology. \<open>subtopology\<close> is
--- a/src/HOL/Analysis/Abstract_Topology_2.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Abstract_Topology_2.thy	Sun May 07 22:51:23 2023 +0200
@@ -347,7 +347,7 @@
       continuous_on S f; continuous_on T f\<rbrakk>
      \<Longrightarrow> continuous_on (S \<union> T) f"
   unfolding continuous_on closedin_limpt
-  by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within)
+  by (metis Lim_trivial_limit Lim_within_Un Un_iff trivial_limit_within)
 
 lemma continuous_on_cases_local:
      "\<lbrakk>closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T;
@@ -1138,6 +1138,20 @@
     by (force simp: retract_of_space_def)
 qed
 
+lemma retract_of_space_trans:
+  assumes "S retract_of_space X"  "T retract_of_space subtopology X S"
+  shows "T retract_of_space X"
+  using assms
+  apply (simp add: retract_of_space_retraction_maps)
+  by (metis id_comp inf.absorb_iff2 retraction_maps_compose subtopology_subtopology)
+
+lemma retract_of_space_subtopology:
+  assumes "S retract_of_space X"  "S \<subseteq> U"
+  shows "S retract_of_space subtopology X U"
+  using assms
+  apply (clarsimp simp: retract_of_space_def)
+  by (metis continuous_map_from_subtopology inf.absorb2 subtopology_subtopology)
+
 lemma retract_of_space_clopen:
   assumes "openin X S" "closedin X S" "S = {} \<Longrightarrow> topspace X = {}"
   shows "S retract_of_space X"
@@ -1194,6 +1208,13 @@
   using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map
         section_map_def by blast
 
+lemma hereditary_imp_retractive_property:
+  assumes "\<And>X S. P X \<Longrightarrow> P(subtopology X S)" 
+          "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')"
+  assumes "retraction_map X X' r" "P X"
+  shows "Q X'"
+  by (meson assms retraction_map_def retraction_maps_section_image2)
+
 subsection\<open>Paths and path-connectedness\<close>
 
 definition pathin :: "'a topology \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where
@@ -1591,4 +1612,31 @@
     connected_component_of_set X x = connected_component_of_set X y"
   by (meson connected_component_of_nonoverlap)
 
+subsection\<open>Combining theorems for continuous functions into the reals\<close>
+
+text \<open>The homeomorphism between the real line and the open interval $(-1,1)$\<close>
+
+lemma continuous_map_real_shrink:
+  "continuous_map euclideanreal (top_of_set {-1<..<1}) (\<lambda>x. x / (1 + \<bar>x\<bar>))"
+proof -
+  have "continuous_on UNIV (\<lambda>x::real. x / (1 + \<bar>x\<bar>))"
+    by (intro continuous_intros) auto
+  then show ?thesis
+    by (auto simp: continuous_map_in_subtopology divide_simps)
+qed
+
+lemma continuous_on_real_grow:
+  "continuous_on {-1<..<1} (\<lambda>x::real. x / (1 - \<bar>x\<bar>))"
+  by (intro continuous_intros) auto
+
+lemma real_grow_shrink:
+  fixes x::real 
+  shows "x / (1 + \<bar>x\<bar>) / (1 - \<bar>x / (1 + \<bar>x\<bar>)\<bar>) = x"
+  by (force simp: divide_simps)
+
+lemma homeomorphic_maps_real_shrink:
+  "homeomorphic_maps euclideanreal (subtopology euclideanreal {-1<..<1}) 
+     (\<lambda>x. x / (1 + \<bar>x\<bar>))  (\<lambda>y. y / (1 - \<bar>y\<bar>))"
+  by (force simp: homeomorphic_maps_def continuous_map_real_shrink continuous_on_real_grow divide_simps)
+
 end
\ No newline at end of file
--- a/src/HOL/Analysis/Analysis.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Analysis.thy	Sun May 07 22:51:23 2023 +0200
@@ -4,6 +4,9 @@
   Convex
   Determinants
   (* Topology *)
+  FSigma
+  Sum_Topology
+  Abstract_Topological_Spaces
   Connected
   Abstract_Limits
   Isolated
--- a/src/HOL/Analysis/Arcwise_Connected.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Arcwise_Connected.thy	Sun May 07 22:51:23 2023 +0200
@@ -533,6 +533,216 @@
   by (auto simp: field_split_simps)
 
 
+lemma padic_rational_approximation_straddle:
+  assumes "\<epsilon> > 0" "p > 1"
+  obtains n q r 
+    where "of_int q / p^n < x" "x < of_int r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>"
+proof -
+  obtain n where n: "2 / \<epsilon> < p ^ n"
+    using \<open>p>1\<close> real_arch_pow by blast
+  define q where "q \<equiv> \<lfloor>p ^ n * x\<rfloor> - 1"
+  show thesis
+    proof
+      show "q / p ^ n < x" "x < real_of_int (q+2) / p ^ n"
+        using assms by (simp_all add: q_def divide_simps floor_less_cancel mult.commute)
+      show "\<bar>q / p ^ n - real_of_int (q+2) / p ^ n\<bar> < \<epsilon>"
+        using assms n by (simp add: q_def divide_simps mult.commute)
+    qed
+qed
+
+lemma padic_rational_approximation_straddle_pos:
+  assumes "\<epsilon> > 0" "p > 1" "x > 0"
+  obtains n q r 
+    where "of_nat q / p^n < x" "x < of_nat r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>"
+proof -
+  obtain n q r 
+    where *: "of_int q / p^n < x" "x < of_int r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>"
+    using padic_rational_approximation_straddle assms by metis
+  then have "r \<ge> 0"
+    using assms by (smt (verit, best) divide_nonpos_pos of_int_0_le_iff zero_less_power)
+  show thesis
+  proof
+    show "real (max 0 (nat q)) / p ^ n < x"
+      using * by (metis assms(3) div_0 max_nat.left_neutral nat_eq_iff2 of_nat_0 of_nat_nat) 
+    show "x < real (nat r) / p ^ n"
+      using \<open>r \<ge> 0\<close> * by force
+    show "\<bar>real (max 0 (nat q)) / p ^ n - real (nat r) / p ^ n\<bar> < \<epsilon>"
+      using * assms by (simp add: divide_simps)
+  qed
+qed
+
+lemma padic_rational_approximation_straddle_pos_le:
+  assumes "\<epsilon> > 0" "p > 1" "x \<ge> 0"
+  obtains n q r 
+    where "of_nat q / p^n \<le> x" "x < of_nat r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>"
+proof -
+  obtain n q r 
+    where *: "of_int q / p^n < x" "x < of_int r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>"
+    using padic_rational_approximation_straddle assms by metis
+  then have "r \<ge> 0"
+    using assms by (smt (verit, best) divide_nonpos_pos of_int_0_le_iff zero_less_power)
+  show thesis
+  proof
+    show "real (max 0 (nat q)) / p ^ n \<le> x"
+      using * assms(3) nle_le by fastforce
+    show "x < real (nat r) / p ^ n"
+      using \<open>r \<ge> 0\<close> * by force
+    show "\<bar>real (max 0 (nat q)) / p ^ n - real (nat r) / p ^ n\<bar> < \<epsilon>"
+       using * assms by (simp add: divide_simps)
+  qed
+qed
+
+
+subsubsection \<open>Definition by recursion on dyadic rationals in [0,1]\<close>
+
+lemma recursion_on_dyadic_fractions:
+  assumes base: "R a b"
+    and step: "\<And>x y. R x y \<Longrightarrow> \<exists>z. R x z \<and> R z y" and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z"
+  shows "\<exists>f :: real \<Rightarrow> 'a. f 0 = a \<and> f 1 = b \<and>
+               (\<forall>x \<in> dyadics \<inter> {0..1}. \<forall>y \<in> dyadics \<inter> {0..1}. x < y \<longrightarrow> R (f x) (f y))"
+proof -
+  obtain mid where mid: "R x y \<Longrightarrow> R x (mid x y)" "R x y \<Longrightarrow> R (mid x y) y" for x y 
+    using step by metis
+  define g where "g \<equiv> rec_nat (\<lambda>k. if k = 0 then a else b) (\<lambda>n r k. if even k then r (k div 2) else mid (r ((k - 1) div 2)) (r ((Suc k) div 2)))"
+  have g0 [simp]: "g 0 = (\<lambda>k. if k = 0 then a else b)"
+    by (simp add: g_def)
+  have gSuc [simp]: "\<And>n. g(Suc n) = (\<lambda>k. if even k then g n (k div 2) else mid (g n ((k - 1) div 2)) (g n ((Suc k) div 2)))"
+    by (auto simp: g_def)
+  have g_eq_g: "2 ^ d * k = k' \<Longrightarrow> g n k = g (n + d) k'" for n d k k'
+    by (induction d arbitrary: k k') auto
+  have "g n k = g n' k'" if "real k / 2^n = real k' / 2^n'" "n' \<le> n" for k n k' n'
+  proof -
+    have "real k = real k' * 2 ^ (n-n')"
+      using that by (simp add: power_diff divide_simps)
+    then have "k = k' * 2 ^ (n-n')"
+      using of_nat_eq_iff by fastforce
+    with g_eq_g show ?thesis
+      by (metis le_add_diff_inverse mult.commute that(2))
+  qed
+  then have g_eq_g: "g n k = g n' k'" if "real k / 2 ^ n = real k' / 2 ^ n'" for k n k' n'
+    by (metis nat_le_linear that)
+  then obtain f where "(\<lambda>(k,n). g n k) = f \<circ> (\<lambda>(k,n). k / 2 ^ n)"
+    using function_factors_left by (smt (verit, del_insts) case_prod_beta')
+  then have f_eq_g: "\<And>k n. f(real k / 2 ^ n) = g n k"
+    by (simp add: fun_eq_iff)
+  show ?thesis
+  proof (intro exI conjI strip)
+    show "f 0 = a"
+      by (metis f_eq_g g0 div_0 of_nat_0)
+    show "f 1 = b"
+      by (metis f_eq_g g0 div_by_1 of_nat_1_eq_iff power_0 zero_neq_one)
+    show "R (f x) (f y)" 
+      if x: "x \<in> dyadics \<inter> {0..1}" and y: "y \<in> dyadics \<inter> {0..1}" and "x < y" for x y
+    proof -
+      obtain n1 k1 where xeq: "x = real k1 / 2^n1" "k1 \<le> 2^n1"
+        using x by (auto simp: dyadics_def)
+      obtain n2 k2 where yeq: "y = real k2 / 2^n2" "k2 \<le> 2^n2"
+        using y by (auto simp: dyadics_def)
+      have xcommon: "x = real(2^n2 * k1) / 2 ^ (n1+n2)"
+        using xeq by (simp add: power_add)
+      have ycommon: "y = real(2^n1 * k2) / 2 ^ (n1+n2)"
+        using yeq by (simp add: power_add)
+      have *: "R (g n j) (g n k)" if "j < k" "k \<le> 2^n" for n j k
+        using that
+      proof (induction n arbitrary: j k)
+        case 0
+        then show ?case
+          by (simp add: base)
+      next
+        case (Suc n)
+        show ?case
+        proof (cases "even j")
+          case True
+          then obtain a where [simp]: "j = 2*a"
+            by blast
+          show ?thesis 
+          proof (cases "even k")
+            case True
+            with Suc show ?thesis
+              by (auto elim!: evenE)
+          next
+            case False
+            then obtain b where [simp]: "k = Suc (2*b)"
+              using oddE by fastforce
+            show ?thesis
+              using Suc
+              apply simp
+              by (smt (verit, ccfv_SIG) less_Suc_eq linorder_not_le local.trans mid(1) nat_mult_less_cancel1 pos2)
+          qed
+        next
+          case False
+          then obtain a where [simp]: "j = Suc (2*a)"
+            using oddE by fastforce
+          show ?thesis
+          proof (cases "even k")
+            case True
+            then obtain b where [simp]: "k = 2*b"
+              by blast
+            show ?thesis
+              using Suc 
+              apply simp
+              by (smt (verit, ccfv_SIG) Suc_leI Suc_lessD le_trans lessI linorder_neqE_nat linorder_not_le local.trans mid(2) nat_mult_less_cancel1 pos2)
+          next
+            case False
+            then obtain b where [simp]: "k = Suc (2*b)"
+              using oddE by fastforce
+            show ?thesis
+              using Suc 
+              apply simp
+              by (smt (verit) Suc_leI le_trans lessI less_or_eq_imp_le linorder_neqE_nat linorder_not_le local.trans mid(1) mid(2) nat_mult_less_cancel1 pos2)
+            qed
+        qed
+      qed
+      show ?thesis
+        unfolding xcommon ycommon f_eq_g
+      proof (rule *)
+        show "2 ^ n2 * k1 < 2 ^ n1 * k2"
+          using of_nat_less_iff \<open>x < y\<close> by (fastforce simp: xeq yeq field_simps)
+        show "2 ^ n1 * k2 \<le> 2 ^ (n1 + n2)"
+          by (simp add: power_add yeq)
+      qed
+    qed
+  qed
+qed
+
+lemma dyadics_add:
+  assumes "x \<in> dyadics" "y \<in> dyadics"
+  shows "x+y \<in> dyadics"
+proof -
+  obtain i j m n where x: "x = of_nat i / 2 ^ m" and y: "y = of_nat j / 2 ^ n"
+    using assms by (auto simp: dyadics_def)
+  have xcommon: "x = of_nat(2^n * i) / 2 ^ (m+n)"
+    using x by (simp add: power_add)
+  moreover
+  have ycommon: "y = of_nat(2^m * j) / 2 ^ (m+n)"
+    using y by (simp add: power_add)
+  ultimately have "x+y = (of_nat(2^n * i + 2^m * j)) / 2 ^ (m+n)"
+    by (simp add: field_simps)
+  then show ?thesis
+    unfolding dyadics_def by blast
+qed
+
+lemma dyadics_diff:
+  fixes x :: "'a::linordered_field"
+  assumes "x \<in> dyadics" "y \<in> dyadics" "y \<le> x"
+  shows "x-y \<in> dyadics"
+proof -
+  obtain i j m n where x: "x = of_nat i / 2 ^ m" and y: "y = of_nat j / 2 ^ n"
+    using assms by (auto simp: dyadics_def)
+  have j_le_i: "j * 2 ^ m \<le> i * 2 ^ n"
+    using of_nat_le_iff \<open>y \<le> x\<close> unfolding x y by (fastforce simp add: divide_simps)
+  have xcommon: "x = of_nat(2^n * i) / 2 ^ (m+n)"
+    using x by (simp add: power_add)
+  moreover
+  have ycommon: "y = of_nat(2^m * j) / 2 ^ (m+n)"
+    using y by (simp add: power_add)
+  ultimately have "x-y = (of_nat(2^n * i - 2^m * j)) / 2 ^ (m+n)"
+    by (simp add: xcommon ycommon field_simps j_le_i of_nat_diff)
+  then show ?thesis
+    unfolding dyadics_def by blast
+qed
+
+
 
 theorem homeomorphic_monotone_image_interval:
   fixes f :: "real \<Rightarrow> 'a::{real_normed_vector,complete_space}"
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy	Sun May 07 22:51:23 2023 +0200
@@ -175,39 +175,6 @@
   finally show "norm (A *v x) \<le> CARD('m) * real (CARD('n)) * B * norm x" .
 qed
 
-lemma rational_approximation:
-  assumes "e > 0"
-  obtains r::real where "r \<in> \<rat>" "\<bar>r - x\<bar> < e"
-  using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto
-
-lemma Rats_closure_real: "closure \<rat> = (UNIV::real set)"
-proof -
-  have "\<And>x::real. x \<in> closure \<rat>"
-    by (metis closure_approachable dist_real_def rational_approximation)
-  then show ?thesis by auto
-qed
-
-proposition matrix_rational_approximation:
-  fixes A :: "real^'n^'m"
-  assumes "e > 0"
-  obtains B where "\<And>i j. B$i$j \<in> \<rat>" "onorm(\<lambda>x. (A - B) *v x) < e"
-proof -
-  have "\<forall>i j. \<exists>q \<in> \<rat>. \<bar>q - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
-    using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"])
-  then obtain B where B: "\<And>i j. B$i$j \<in> \<rat>" and Bclo: "\<And>i j. \<bar>B$i$j - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
-    by (auto simp: lambda_skolem Bex_def)
-  show ?thesis
-  proof
-    have "onorm ((*v) (A - B)) \<le> real CARD('m) * real CARD('n) *
-    (e / (2 * real CARD('m) * real CARD('n)))"
-      apply (rule onorm_le_matrix_component)
-      using Bclo by (simp add: abs_minus_commute less_imp_le)
-    also have "\<dots> < e"
-      using \<open>0 < e\<close> by (simp add: field_split_simps)
-    finally show "onorm ((*v) (A - B)) < e" .
-  qed (use B in auto)
-qed
-
 lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
   unfolding inner_simps scalar_mult_eq_scaleR by auto
 
@@ -478,6 +445,42 @@
 qed
 *)
 
+subsection\<open>Arbitrarily good rational approximations\<close>
+
+lemma rational_approximation:
+  assumes "e > 0"
+  obtains r::real where "r \<in> \<rat>" "\<bar>r - x\<bar> < e"
+  using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto
+
+lemma Rats_closure_real: "closure \<rat> = (UNIV::real set)"
+proof -
+  have "\<And>x::real. x \<in> closure \<rat>"
+    by (metis closure_approachable dist_real_def rational_approximation)
+  then show ?thesis by auto
+qed
+
+proposition matrix_rational_approximation:
+  fixes A :: "real^'n^'m"
+  assumes "e > 0"
+  obtains B where "\<And>i j. B$i$j \<in> \<rat>" "onorm(\<lambda>x. (A - B) *v x) < e"
+proof -
+  have "\<forall>i j. \<exists>q \<in> \<rat>. \<bar>q - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
+    using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"])
+  then obtain B where B: "\<And>i j. B$i$j \<in> \<rat>" and Bclo: "\<And>i j. \<bar>B$i$j - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
+    by (auto simp: lambda_skolem Bex_def)
+  show ?thesis
+  proof
+    have "onorm ((*v) (A - B)) \<le> real CARD('m) * real CARD('n) *
+    (e / (2 * real CARD('m) * real CARD('n)))"
+      apply (rule onorm_le_matrix_component)
+      using Bclo by (simp add: abs_minus_commute less_imp_le)
+    also have "\<dots> < e"
+      using \<open>0 < e\<close> by (simp add: field_split_simps)
+    finally show "onorm ((*v) (A - B)) < e" .
+  qed (use B in auto)
+qed
+
+
 subsection "Derivative"
 
 definition\<^marker>\<open>tag important\<close> "jacobian f net = matrix(frechet_derivative f net)"
@@ -547,8 +550,7 @@
 lemma interval_split_cart:
   "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
   "cbox a b \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
-  apply (rule_tac[!] set_eqI)
-  unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart
+  unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart set_eq_iff
   unfolding vec_lambda_beta
   by auto
 
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Sun May 07 22:51:23 2023 +0200
@@ -1716,6 +1716,11 @@
   shows "connected S \<longleftrightarrow> convex S"
   by (metis is_interval_convex convex_connected is_interval_connected_1)
 
+lemma connected_space_iff_is_interval_1 [iff]:
+  fixes S :: "real set"
+  shows "connected_space (top_of_set S) \<longleftrightarrow> is_interval S"
+  using connectedin_topspace is_interval_connected_1 by force
+
 lemma connected_convex_1_gen:
   fixes S :: "'a :: euclidean_space set"
   assumes "DIM('a) = 1"
--- a/src/HOL/Analysis/Elementary_Topology.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Elementary_Topology.thy	Sun May 07 22:51:23 2023 +0200
@@ -1323,32 +1323,6 @@
 lemma at_within_eq_bot_iff: "at c within A = bot \<longleftrightarrow> c \<notin> closure (A - {c})"
   using not_trivial_limit_within[of c A] by blast
 
-text \<open>Some property holds "sufficiently close" to the limit point.\<close>
-
-lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
-  by simp
-
-lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
-  by (simp add: filter_eq_iff)
-
-lemma Lim_topological:
-  "(f \<longlongrightarrow> l) net \<longleftrightarrow>
-    trivial_limit net \<or> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)"
-  unfolding tendsto_def trivial_limit_eq by auto
-
-lemma eventually_within_Un:
-  "eventually P (at x within (s \<union> t)) \<longleftrightarrow>
-    eventually P (at x within s) \<and> eventually P (at x within t)"
-  unfolding eventually_at_filter
-  by (auto elim!: eventually_rev_mp)
-
-lemma Lim_within_union:
- "(f \<longlongrightarrow> l) (at x within (s \<union> t)) \<longleftrightarrow>
-  (f \<longlongrightarrow> l) (at x within s) \<and> (f \<longlongrightarrow> l) (at x within t)"
-  unfolding tendsto_def
-  by (auto simp: eventually_within_Un)
-
-
 subsection \<open>Limits\<close>
 
 text \<open>The expected monotonicity property.\<close>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/FSigma.thy	Sun May 07 22:51:23 2023 +0200
@@ -0,0 +1,236 @@
+(*  Author:     L C Paulson, University of Cambridge [ported from HOL Light, metric.ml] *)
+
+section \<open>$F$-Sigma and $G$-Delta sets in a Topological Space\<close>
+
+text \<open>An $F$-sigma set is a countable union of closed sets; a $G$-delta set is the dual.\<close>
+
+theory FSigma
+  imports Abstract_Topology
+begin
+
+
+definition fsigma_in 
+  where "fsigma_in X \<equiv> countable union_of closedin X"
+
+definition gdelta_in 
+  where "gdelta_in X \<equiv> (countable intersection_of openin X) relative_to topspace X"
+
+lemma fsigma_in_ascending:
+   "fsigma_in X S \<longleftrightarrow> (\<exists>C. (\<forall>n. closedin X (C n)) \<and> (\<forall>n. C n \<subseteq> C(Suc n)) \<and> \<Union> (range C) = S)"
+  unfolding fsigma_in_def
+  by (metis closedin_Un countable_union_of_ascending closedin_empty)
+
+lemma gdelta_in_alt:
+   "gdelta_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> (countable intersection_of openin X) S"
+proof -
+  have "(countable intersection_of openin X) (topspace X)"
+    by (simp add: countable_intersection_of_inc)
+  then show ?thesis
+    unfolding gdelta_in_def
+    by (metis countable_intersection_of_inter relative_to_def relative_to_imp_subset relative_to_subset)
+qed
+
+lemma fsigma_in_subset: "fsigma_in X S \<Longrightarrow> S \<subseteq> topspace X"
+  using closedin_subset by (fastforce simp: fsigma_in_def union_of_def subset_iff)
+
+lemma gdelta_in_subset: "gdelta_in X S \<Longrightarrow> S \<subseteq> topspace X"
+  by (simp add: gdelta_in_alt)
+
+lemma closed_imp_fsigma_in: "closedin X S \<Longrightarrow> fsigma_in X S"
+  by (simp add: countable_union_of_inc fsigma_in_def)
+
+lemma open_imp_gdelta_in: "openin X S \<Longrightarrow> gdelta_in X S"
+  by (simp add: countable_intersection_of_inc gdelta_in_alt openin_subset)
+
+lemma fsigma_in_empty [iff]: "fsigma_in X {}"
+  by (simp add: closed_imp_fsigma_in)
+
+lemma gdelta_in_empty [iff]: "gdelta_in X {}"
+  by (simp add: open_imp_gdelta_in)
+
+lemma fsigma_in_topspace [iff]: "fsigma_in X (topspace X)"
+  by (simp add: closed_imp_fsigma_in)
+
+lemma gdelta_in_topspace [iff]: "gdelta_in X (topspace X)"
+  by (simp add: open_imp_gdelta_in)
+
+lemma fsigma_in_Union:
+   "\<lbrakk>countable T; \<And>S. S \<in> T \<Longrightarrow> fsigma_in X S\<rbrakk> \<Longrightarrow> fsigma_in X (\<Union> T)"
+  by (simp add: countable_union_of_Union fsigma_in_def)
+
+lemma fsigma_in_Un:
+   "\<lbrakk>fsigma_in X S; fsigma_in X T\<rbrakk> \<Longrightarrow> fsigma_in X (S \<union> T)"
+  by (simp add: countable_union_of_Un fsigma_in_def)
+
+lemma fsigma_in_Int:
+   "\<lbrakk>fsigma_in X S; fsigma_in X T\<rbrakk> \<Longrightarrow> fsigma_in X (S \<inter> T)"
+  by (simp add: closedin_Int countable_union_of_Int fsigma_in_def)
+
+lemma gdelta_in_Inter:
+   "\<lbrakk>countable T; T\<noteq>{}; \<And>S. S \<in> T \<Longrightarrow> gdelta_in X S\<rbrakk> \<Longrightarrow> gdelta_in X (\<Inter> T)"
+  by (simp add: Inf_less_eq countable_intersection_of_Inter gdelta_in_alt)
+
+lemma gdelta_in_Int:
+   "\<lbrakk>gdelta_in X S; gdelta_in X T\<rbrakk> \<Longrightarrow> gdelta_in X (S \<inter> T)"
+  by (simp add: countable_intersection_of_inter gdelta_in_alt le_infI2)
+
+lemma gdelta_in_Un:
+   "\<lbrakk>gdelta_in X S; gdelta_in X T\<rbrakk> \<Longrightarrow> gdelta_in X (S \<union> T)"
+  by (simp add: countable_intersection_of_union gdelta_in_alt openin_Un)
+
+lemma fsigma_in_diff:
+  assumes "fsigma_in X S" "gdelta_in X T"
+  shows "fsigma_in X (S - T)"
+proof -
+  have [simp]: "S - (S \<inter> T) = S - T" for S T :: "'a set"
+    by auto
+  have [simp]: "topspace X - \<Inter>\<T> = (\<Union>T\<in>\<T>. topspace X - T)" for \<T>
+    by auto
+  have "\<And>\<T>. \<lbrakk>countable \<T>; \<T> \<subseteq> Collect (openin X)\<rbrakk> \<Longrightarrow>
+             (countable union_of closedin X) (\<Union> ((-) (topspace X) ` \<T>))"
+    by (metis Ball_Collect countable_union_of_UN countable_union_of_inc openin_closedin_eq)
+  then have "\<forall>S. gdelta_in X S \<longrightarrow> fsigma_in X (topspace X - S)"
+    by (simp add: fsigma_in_def gdelta_in_def all_relative_to all_intersection_of del: UN_simps)
+  then show ?thesis
+    by (metis Diff_Int2 Diff_Int_distrib2 assms fsigma_in_Int fsigma_in_subset inf.absorb_iff2)
+qed
+
+lemma gdelta_in_diff:
+  assumes "gdelta_in X S" "fsigma_in X T"
+  shows "gdelta_in X (S - T)"
+proof -
+  have [simp]: "topspace X - \<Union>\<T> = topspace X \<inter> (\<Inter>T\<in>\<T>. topspace X - T)" for \<T>
+    by auto
+  have "\<And>\<T>. \<lbrakk>countable \<T>; \<T> \<subseteq> Collect (closedin X)\<rbrakk>
+         \<Longrightarrow> (countable intersection_of openin X relative_to topspace X)
+              (topspace X \<inter> \<Inter> ((-) (topspace X) ` \<T>))"
+    by (metis Ball_Collect closedin_def countable_intersection_of_INT countable_intersection_of_inc relative_to_inc)
+  then have "\<forall>S. fsigma_in X S \<longrightarrow> gdelta_in X (topspace X - S)"
+    by (simp add: fsigma_in_def gdelta_in_def all_union_of del: INT_simps)
+  then show ?thesis
+    by (metis Diff_Int2 Diff_Int_distrib2 assms gdelta_in_Int gdelta_in_alt inf.orderE inf_commute)
+qed
+
+lemma gdelta_in_fsigma_in:
+   "gdelta_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> fsigma_in X (topspace X - S)"
+  by (metis double_diff fsigma_in_diff fsigma_in_topspace gdelta_in_alt gdelta_in_diff gdelta_in_topspace)
+
+lemma fsigma_in_gdelta_in:
+   "fsigma_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> gdelta_in X (topspace X - S)"
+  by (metis Diff_Diff_Int fsigma_in_subset gdelta_in_fsigma_in inf.absorb_iff2)
+
+lemma gdelta_in_descending:
+   "gdelta_in X S \<longleftrightarrow> (\<exists>C. (\<forall>n. openin X (C n)) \<and> (\<forall>n. C(Suc n) \<subseteq> C n) \<and> \<Inter>(range C) = S)" (is "?lhs=?rhs")
+proof
+  assume ?lhs
+  then obtain C where C: "S \<subseteq> topspace X" "\<And>n. closedin X (C n)" 
+                         "\<And>n. C n \<subseteq> C(Suc n)" "\<Union>(range C) = topspace X - S"
+    by (meson fsigma_in_ascending gdelta_in_fsigma_in)
+  define D where "D \<equiv> \<lambda>n. topspace X - C n"
+  have "\<And>n. openin X (D n) \<and> D (Suc n) \<subseteq> D n"
+    by (simp add: Diff_mono C D_def openin_diff)
+  moreover have "\<Inter>(range D) = S"
+    by (simp add: Diff_Diff_Int Int_absorb1 C D_def)
+  ultimately show ?rhs
+    by metis
+next
+  assume ?rhs
+  then obtain C where "S \<subseteq> topspace X" 
+                and C: "\<And>n. openin X (C n)" "\<And>n. C(Suc n) \<subseteq> C n" "\<Inter>(range C) = S"
+    using openin_subset by fastforce
+  define D where "D \<equiv> \<lambda>n. topspace X - C n"
+  have "\<And>n. closedin X (D n) \<and> D n \<subseteq> D(Suc n)"
+    by (simp add: Diff_mono C closedin_diff D_def)
+  moreover have "\<Union>(range D) = topspace X - S"
+    using C D_def by blast
+  ultimately show ?lhs
+    by (metis \<open>S \<subseteq> topspace X\<close> fsigma_in_ascending gdelta_in_fsigma_in)
+qed
+
+lemma homeomorphic_map_fsigmaness:
+  assumes f: "homeomorphic_map X Y f" and "U \<subseteq> topspace X"
+  shows "fsigma_in Y (f ` U) \<longleftrightarrow> fsigma_in X U"  (is "?lhs=?rhs")
+proof -
+  obtain g where g: "homeomorphic_maps X Y f g" and g: "homeomorphic_map Y X g"
+    and 1: "(\<forall>x \<in> topspace X. g(f x) = x)" and 2: "(\<forall>y \<in> topspace Y. f(g y) = y)"
+    using assms homeomorphic_map_maps by (metis homeomorphic_maps_map)
+  show ?thesis
+  proof
+    assume ?lhs
+    then obtain \<V> where "countable \<V>" and \<V>: "\<V> \<subseteq> Collect (closedin Y)" "\<Union>\<V> = f`U"
+      by (force simp: fsigma_in_def union_of_def)
+    define \<U> where "\<U> \<equiv> image (image g) \<V>"
+    have "countable \<U>"
+      using \<U>_def \<open>countable \<V>\<close> by blast
+    moreover have "\<U> \<subseteq> Collect (closedin X)"
+      using f g homeomorphic_map_closedness_eq \<V> unfolding \<U>_def by blast
+    moreover have "\<Union>\<U> \<subseteq> U"
+      unfolding \<U>_def  by (smt (verit) assms 1 \<V> image_Union image_iff in_mono subsetI)
+    moreover have "U \<subseteq> \<Union>\<U>"
+      unfolding \<U>_def using assms \<V>
+      by (smt (verit, del_insts) "1" imageI image_Union subset_iff)
+    ultimately show ?rhs
+      by (metis fsigma_in_def subset_antisym union_of_def)
+  next
+    assume ?rhs
+    then obtain \<V> where "countable \<V>" and \<V>: "\<V> \<subseteq> Collect (closedin X)" "\<Union>\<V> = U"
+      by (auto simp: fsigma_in_def union_of_def)
+    define \<U> where "\<U> \<equiv> image (image f) \<V>"
+    have "countable \<U>"
+      using \<U>_def \<open>countable \<V>\<close> by blast
+    moreover have "\<U> \<subseteq> Collect (closedin Y)"
+      using f g homeomorphic_map_closedness_eq \<V> unfolding \<U>_def by blast
+    moreover have "\<Union>\<U> = f`U"
+      unfolding \<U>_def using \<V> by blast
+    ultimately show ?lhs
+      by (metis fsigma_in_def union_of_def)
+  qed
+qed
+
+lemma homeomorphic_map_fsigmaness_eq:
+   "homeomorphic_map X Y f
+        \<Longrightarrow> (fsigma_in X U \<longleftrightarrow> U \<subseteq> topspace X \<and> fsigma_in Y (f ` U))"
+  by (metis fsigma_in_subset homeomorphic_map_fsigmaness)
+
+lemma homeomorphic_map_gdeltaness:
+  assumes "homeomorphic_map X Y f" "U \<subseteq> topspace X"
+  shows "gdelta_in Y (f ` U) \<longleftrightarrow> gdelta_in X U"
+proof -
+  have "topspace Y - f ` U = f ` (topspace X - U)"
+    by (metis (no_types, lifting) Diff_subset assms homeomorphic_eq_everything_map inj_on_image_set_diff)
+  then show ?thesis
+    using assms homeomorphic_imp_surjective_map
+    by (fastforce simp: gdelta_in_fsigma_in homeomorphic_map_fsigmaness_eq)
+qed
+
+lemma homeomorphic_map_gdeltaness_eq:
+   "homeomorphic_map X Y f
+    \<Longrightarrow> gdelta_in X U \<longleftrightarrow> U \<subseteq> topspace X \<and> gdelta_in Y (f ` U)"
+  by (meson gdelta_in_subset homeomorphic_map_gdeltaness)
+
+lemma fsigma_in_relative_to:
+   "(fsigma_in X relative_to S) = fsigma_in (subtopology X S)"
+  by (simp add: fsigma_in_def countable_union_of_relative_to closedin_relative_to)
+
+lemma fsigma_in_subtopology:
+   "fsigma_in (subtopology X U) S \<longleftrightarrow> (\<exists>T. fsigma_in X T \<and> S = T \<inter> U)"
+  by (metis fsigma_in_relative_to inf_commute relative_to_def)
+
+lemma gdelta_in_relative_to:
+   "(gdelta_in X relative_to S) = gdelta_in (subtopology X S)"
+  apply (simp add: gdelta_in_def)
+  by (metis countable_intersection_of_relative_to openin_relative_to subtopology_restrict)
+
+lemma gdelta_in_subtopology:
+   "gdelta_in (subtopology X U) S \<longleftrightarrow> (\<exists>T. gdelta_in X T \<and> S = T \<inter> U)"
+  by (metis gdelta_in_relative_to inf_commute relative_to_def)
+
+lemma fsigma_in_fsigma_subtopology:
+   "fsigma_in X S \<Longrightarrow> (fsigma_in (subtopology X S) T \<longleftrightarrow> fsigma_in X T \<and> T \<subseteq> S)"
+  by (metis fsigma_in_Int fsigma_in_gdelta_in fsigma_in_subtopology inf.orderE topspace_subtopology_subset)
+
+lemma gdelta_in_gdelta_subtopology:
+   "gdelta_in X S \<Longrightarrow> (gdelta_in (subtopology X S) T \<longleftrightarrow> gdelta_in X T \<and> T \<subseteq> S)"
+  by (metis gdelta_in_Int gdelta_in_subset gdelta_in_subtopology inf.orderE topspace_subtopology_subset)
+
+end
--- a/src/HOL/Analysis/Further_Topology.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Further_Topology.thy	Sun May 07 22:51:23 2023 +0200
@@ -969,9 +969,6 @@
 
 subsection\<open> Special cases and corollaries involving spheres\<close>
 
-lemma disjnt_Diff1: "X \<subseteq> Y' \<Longrightarrow> disjnt (X - Y) (X' - Y')"
-  by (auto simp: disjnt_def)
-
 proposition extend_map_affine_to_sphere_cofinite_simple:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "compact S" "convex U" "bounded U"
--- a/src/HOL/Analysis/Homotopy.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Homotopy.thy	Sun May 07 22:51:23 2023 +0200
@@ -156,7 +156,7 @@
   qed
   show ?thesis
     using assms
-    apply (clarsimp simp add: homotopic_with_def)
+    apply (clarsimp simp: homotopic_with_def)
     subgoal for h
       by (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *])
     done
@@ -211,7 +211,7 @@
     show "\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x))"
     proof 
       fix t show "t\<in>{0..1} \<Longrightarrow> P (\<lambda>x. k (t, x))"
-        by (cases "t \<le> 1/2") (auto simp add: k_def P)
+        by (cases "t \<le> 1/2") (auto simp: k_def P)
     qed
   qed
 qed
@@ -578,7 +578,7 @@
 
 proposition homotopic_paths_join:
     "\<lbrakk>homotopic_paths S p p'; homotopic_paths S q q'; pathfinish p = pathstart q\<rbrakk> \<Longrightarrow> homotopic_paths S (p +++ q) (p' +++ q')"
-  apply (clarsimp simp add: homotopic_paths_def homotopic_with_def)
+  apply (clarsimp simp: homotopic_paths_def homotopic_with_def)
   apply (rename_tac k1 k2)
   apply (rule_tac x="(\<lambda>y. ((k1 \<circ> Pair (fst y)) +++ (k2 \<circ> Pair (fst y))) (snd y))" in exI)
   apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def)
@@ -639,8 +639,8 @@
     show "continuous_on {x \<in> ?A. snd x \<le> 1/2} (\<lambda>t. p (fst t * (2 * snd t)))"
          "continuous_on {x \<in> ?A. 1/2 \<le> snd x} (\<lambda>t. p (fst t * (1 - (2 * snd t - 1))))"
          "continuous_on ?A snd"
-      by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp add: mult_le_one)+
-  qed (auto simp add: algebra_simps)
+      by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp: mult_le_one)+
+  qed (auto simp: algebra_simps)
   then show ?thesis
     using assms
     apply (subst homotopic_paths_sym_eq)
@@ -1004,7 +1004,7 @@
         by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2)
       show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \<in> {0..1}" for t 
         using assms
-        unfolding joinpaths_def subpath_def by (auto simp add: divide_simps add.commute mult.commute mult.left_commute)
+        unfolding joinpaths_def subpath_def by (auto simp: divide_simps add.commute mult.commute mult.left_commute)
     qed (use False in auto)
   qed
   then show ?thesis
@@ -1062,26 +1062,24 @@
   then have "continuous_on ({0..1} \<times> {0..1}) (g \<circ> fst)"
     by (fastforce intro!: continuous_intros)+
   with g show ?thesis
-    by (auto simp add: homotopic_loops_def homotopic_with_def path_defs image_subset_iff)
+    by (auto simp: homotopic_loops_def homotopic_with_def path_defs image_subset_iff)
 qed
 
 lemma homotopic_loops_imp_path_component_value:
-   "\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk>
-        \<Longrightarrow> path_component S (p t) (q t)"
-apply (clarsimp simp add: homotopic_loops_def homotopic_with_def path_defs)
+   "\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> path_component S (p t) (q t)"
+apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs)
 apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI)
 apply (fastforce elim!: continuous_on_subset intro!: continuous_intros)
 done
 
 lemma homotopic_points_eq_path_component:
    "homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow> path_component S a b"
-by (auto simp: path_component_imp_homotopic_points
-         dest: homotopic_loops_imp_path_component_value [where t=1])
+  using homotopic_loops_imp_path_component_value path_component_imp_homotopic_points by fastforce
 
 lemma path_connected_eq_homotopic_points:
-    "path_connected S \<longleftrightarrow>
+  "path_connected S \<longleftrightarrow>
       (\<forall>a b. a \<in> S \<and> b \<in> S \<longrightarrow> homotopic_loops S (linepath a a) (linepath b b))"
-by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component)
+  by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component)
 
 
 subsection\<open>Simply connected sets\<close>
@@ -1134,16 +1132,9 @@
   then show ?rhs
   using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected)
 next
-  assume r: ?rhs
-  note pa = r [THEN conjunct2, rule_format]
-  show ?lhs
-  proof (clarsimp simp add: simply_connected_eq_contractible_loop_any)
-    fix p a
-    assume "path p" and "path_image p \<subseteq> S" "pathfinish p = pathstart p"
-      and "a \<in> S"
-    with pa [of p] show "homotopic_loops S p (linepath a a)"
-      using homotopic_loops_trans path_connected_eq_homotopic_points r by blast
-  qed
+  assume ?rhs
+  then show ?lhs
+    by (meson homotopic_loops_trans path_connected_eq_homotopic_points simply_connected_eq_contractible_loop_any)
 qed
 
 lemma simply_connected_eq_contractible_loop_all:
@@ -1158,19 +1149,8 @@
 next
   case False
   then obtain a where "a \<in> S" by blast
-  show ?thesis
-  proof
-    assume "simply_connected S"
-    then show ?rhs
-      using \<open>a \<in> S\<close> \<open>simply_connected S\<close> simply_connected_eq_contractible_loop_any
-      by blast
-  next
-    assume ?rhs
-    then show "simply_connected S"
-      unfolding simply_connected_eq_contractible_loop_any 
-      by (meson False homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans 
-          path_component_imp_homotopic_points path_component_refl)
-  qed
+  then show ?thesis
+    by (meson False homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any)
 qed
 
 lemma simply_connected_eq_contractible_path:
@@ -1212,19 +1192,17 @@
            "path_image q \<subseteq> S" "pathstart q = pathstart p"
            "pathfinish q = pathfinish p" for p q
   proof -
-    have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))"
-      by (simp add: homotopic_paths_rid homotopic_paths_sym that)
-    also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p))
-                                 (p +++ reversepath q +++ q)"
+    have "homotopic_paths S p (p +++ reversepath q +++ q)"
       using that
-      by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath)
+      by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym 
+          homotopic_paths_trans pathstart_linepath)
     also have "homotopic_paths S (p +++ reversepath q +++ q)
                                  ((p +++ reversepath q) +++ q)"
       by (simp add: that homotopic_paths_assoc)
     also have "homotopic_paths S ((p +++ reversepath q) +++ q)
                                  (linepath (pathstart q) (pathstart q) +++ q)"
       using * [of "p +++ reversepath q"] that
-      by (simp add: homotopic_paths_join path_image_join)
+      by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join)
     also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q"
       using that homotopic_paths_lid by blast
     finally show ?thesis .
@@ -1291,12 +1269,12 @@
   have "\<forall>p. path p \<and>
             path_image p \<subseteq> S \<and> pathfinish p = pathstart p \<longrightarrow>
             homotopic_loops S p (linepath a a)"
-    using a apply (clarsimp simp add: homotopic_loops_def homotopic_with_def path_defs)
+    using a apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs)
     apply (rule_tac x="(h \<circ> (\<lambda>y. (fst y, (p \<circ> snd) y)))" in exI)
     apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset)
     done
   with \<open>a \<in> S\<close> show ?thesis
-    by (auto simp add: simply_connected_eq_contractible_loop_all False)
+    by (auto simp: simply_connected_eq_contractible_loop_all False)
 qed
 
 corollary contractible_imp_connected:
@@ -1369,20 +1347,20 @@
 lemma homotopic_into_contractible:
   fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set"
   assumes f: "continuous_on S f" "f ` S \<subseteq> T"
-      and g: "continuous_on S g" "g ` S \<subseteq> T"
-      and T: "contractible T"
-    shows "homotopic_with_canon (\<lambda>h. True) S T f g"
-using homotopic_through_contractible [of S f T id T g id]
-by (simp add: assms contractible_imp_path_connected)
+    and g: "continuous_on S g" "g ` S \<subseteq> T"
+    and T: "contractible T"
+  shows "homotopic_with_canon (\<lambda>h. True) S T f g"
+  using homotopic_through_contractible [of S f T id T g id]
+  by (simp add: assms contractible_imp_path_connected)
 
 lemma homotopic_from_contractible:
   fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set"
   assumes f: "continuous_on S f" "f ` S \<subseteq> T"
-      and g: "continuous_on S g" "g ` S \<subseteq> T"
-      and "contractible S" "path_connected T"
-    shows "homotopic_with_canon (\<lambda>h. True) S T f g"
-using homotopic_through_contractible [of S id S f T id g]
-by (simp add: assms contractible_imp_path_connected)
+    and g: "continuous_on S g" "g ` S \<subseteq> T"
+    and "contractible S" "path_connected T"
+  shows "homotopic_with_canon (\<lambda>h. True) S T f g"
+  using homotopic_through_contractible [of S id S f T id g]
+  by (simp add: assms contractible_imp_path_connected)
 
 subsection\<open>Starlike sets\<close>
 
@@ -1402,8 +1380,7 @@
 proof -
   have "rel_interior S \<noteq> {}"
     by (simp add: assms rel_interior_eq_empty)
-  then obtain a where a: "a \<in> rel_interior S"  by blast
-  with ST have "a \<in> T"  by blast
+  with ST obtain a where a: "a \<in> rel_interior S" and "a \<in> T" by blast
   have "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
     by (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) (use assms in auto)
   then have "\<forall>x\<in>T. a \<in> T \<and> open_segment a x \<subseteq> T"
@@ -1439,7 +1416,7 @@
 lemma starlike_imp_contractible:
   fixes S :: "'a::real_normed_vector set"
   shows "starlike S \<Longrightarrow> contractible S"
-using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def)
+  using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def)
 
 lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)"
   by (simp add: starlike_imp_contractible)
@@ -1447,27 +1424,27 @@
 lemma starlike_imp_simply_connected:
   fixes S :: "'a::real_normed_vector set"
   shows "starlike S \<Longrightarrow> simply_connected S"
-by (simp add: contractible_imp_simply_connected starlike_imp_contractible)
+  by (simp add: contractible_imp_simply_connected starlike_imp_contractible)
 
 lemma convex_imp_simply_connected:
   fixes S :: "'a::real_normed_vector set"
   shows "convex S \<Longrightarrow> simply_connected S"
-using convex_imp_starlike starlike_imp_simply_connected by blast
+  using convex_imp_starlike starlike_imp_simply_connected by blast
 
 lemma starlike_imp_path_connected:
   fixes S :: "'a::real_normed_vector set"
   shows "starlike S \<Longrightarrow> path_connected S"
-by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected)
+  by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected)
 
 lemma starlike_imp_connected:
   fixes S :: "'a::real_normed_vector set"
   shows "starlike S \<Longrightarrow> connected S"
-by (simp add: path_connected_imp_connected starlike_imp_path_connected)
+  by (simp add: path_connected_imp_connected starlike_imp_path_connected)
 
 lemma is_interval_simply_connected_1:
   fixes S :: "real set"
   shows "is_interval S \<longleftrightarrow> simply_connected S"
-using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto
+  by (meson convex_imp_simply_connected is_interval_connected_1 is_interval_convex_1 simply_connected_imp_connected)
 
 lemma contractible_empty [simp]: "contractible {}"
   by (simp add: contractible_def homotopic_on_emptyI)
@@ -1476,15 +1453,8 @@
   fixes S :: "'a::euclidean_space set"
   assumes "convex S" and TS: "rel_interior S \<subseteq> T" "T \<subseteq> closure S"
   shows "contractible T"
-proof (cases "S = {}")
-  case True
-  with assms show ?thesis
-    by (simp add: subsetCE)
-next
-  case False
-  show ?thesis
-    by (meson False assms starlike_convex_tweak_boundary_points starlike_imp_contractible)
-qed
+  by (metis assms closure_eq_empty contractible_empty empty_subsetI 
+      starlike_convex_tweak_boundary_points starlike_imp_contractible subset_antisym)
 
 lemma convex_imp_contractible:
   fixes S :: "'a::real_normed_vector set"
@@ -1494,13 +1464,13 @@
 lemma contractible_sing [simp]:
   fixes a :: "'a::real_normed_vector"
   shows "contractible {a}"
-by (rule convex_imp_contractible [OF convex_singleton])
+  by (rule convex_imp_contractible [OF convex_singleton])
 
 lemma is_interval_contractible_1:
   fixes S :: "real set"
   shows  "is_interval S \<longleftrightarrow> contractible S"
-using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1
-      is_interval_simply_connected_1 by auto
+  using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1
+    is_interval_simply_connected_1 by auto
 
 lemma contractible_Times:
   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
@@ -1556,9 +1526,7 @@
 lemma locally_open_subset:
   assumes "locally P S" "openin (top_of_set S) t"
     shows "locally P t"
-  using assms
-  unfolding locally_def
-  by (elim all_forward) (meson dual_order.trans openin_imp_subset openin_subset_trans openin_trans)
+  by (smt (verit, ccfv_SIG) assms order.trans locally_def openin_imp_subset openin_subset_trans openin_trans)
 
 lemma locally_diff_closed:
     "\<lbrakk>locally P S; closedin (top_of_set S) t\<rbrakk> \<Longrightarrow> locally P (S - t)"
@@ -1575,16 +1543,13 @@
     using zero_less_one by blast
   then show ?thesis
     unfolding locally_def
-    by (auto simp add: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR)
+    by (auto simp: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR)
 qed
 
 lemma locally_iff:
     "locally P S \<longleftrightarrow>
      (\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T)))"
-  apply (simp add: le_inf_iff locally_def openin_open, safe)
-   apply (metis IntE IntI le_inf_iff)
-  apply (metis IntI Int_subset_iff)
-  done
+  by (smt (verit) locally_def openin_open)
 
 lemma locally_Int:
   assumes S: "locally P S" and T: "locally P T"
@@ -1649,34 +1614,29 @@
     using hom by (auto simp: homeomorphism_def)
   have gw: "g ` W = S \<inter> f -` W"
     using \<open>W \<subseteq> T\<close> g by force
-  have \<circ>: "openin (top_of_set S) (g ` W)"
-  proof -
-    have "continuous_on S f"
-      using f(3) by blast
-    then show "openin (top_of_set S) (g ` W)"
-      by (simp add: gw Collect_conj_eq \<open>openin (top_of_set T) W\<close> continuous_on_open f(2))
-  qed
+  have "openin (top_of_set S) (g ` W)"
+    using \<open>openin (top_of_set T) W\<close> continuous_on_open f gw by auto
   then obtain U V
     where osu: "openin (top_of_set S) U" and uv: "P V" "g y \<in> U" "U \<subseteq> V" "V \<subseteq> g ` W"
-    using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \<open>y \<in> W\<close> by force
+    by (metis S \<open>y \<in> W\<close> image_eqI locallyE)
   have "V \<subseteq> S" using uv by (simp add: gw)
   have fv: "f ` V = T \<inter> {x. g x \<in> V}"
     using \<open>f ` S = T\<close> f \<open>V \<subseteq> S\<close> by auto
+  have contvf: "continuous_on V f"
+    using \<open>V \<subseteq> S\<close> continuous_on_subset f(3) by blast
   have "f ` V \<subseteq> W"
     using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto
-  have contvf: "continuous_on V f"
-    using \<open>V \<subseteq> S\<close> continuous_on_subset f(3) by blast
-  have contvg: "continuous_on (f ` V) g"
-    using \<open>f ` V \<subseteq> W\<close> \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast
+  then have contvg: "continuous_on (f ` V) g"
+    using \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast
   have "V \<subseteq> g ` f ` V"
     by (metis \<open>V \<subseteq> S\<close> hom homeomorphism_def homeomorphism_of_subsets order_refl)
   then have homv: "homeomorphism V (f ` V) f g"
-    using \<open>V \<subseteq> S\<close> f by (auto simp add: homeomorphism_def contvf contvg)
+    using \<open>V \<subseteq> S\<close> f by (auto simp: homeomorphism_def contvf contvg)
   have "openin (top_of_set (g ` T)) U"
     using \<open>g ` T = S\<close> by (simp add: osu)
-  then have 1: "openin (top_of_set T) (T \<inter> g -` U)"
+  then have "openin (top_of_set T) (T \<inter> g -` U)"
     using \<open>continuous_on T g\<close> continuous_on_open [THEN iffD1] by blast
-  have 2: "\<exists>V. Q V \<and> y \<in> (T \<inter> g -` U) \<and> (T \<inter> g -` U) \<subseteq> V \<and> V \<subseteq> W"
+  moreover have "\<exists>V. Q V \<and> y \<in> (T \<inter> g -` U) \<and> (T \<inter> g -` U) \<subseteq> V \<and> V \<subseteq> W"
   proof (intro exI conjI)
     show "Q (f ` V)"
       using Q homv \<open>P V\<close> by blast
@@ -1687,44 +1647,28 @@
     show "f ` V \<subseteq> W"
       using \<open>f ` V \<subseteq> W\<close> by blast
   qed
-  show "\<exists>U. openin (top_of_set T) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
-    by (meson 1 2)
+  ultimately show "\<exists>U. openin (top_of_set T) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
+    by meson
 qed
 
 lemma homeomorphism_locally:
   fixes f:: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  assumes hom: "homeomorphism S T f g"
-      and eq: "\<And>S T. homeomorphism S T f g \<Longrightarrow> (P S \<longleftrightarrow> Q T)"
+  assumes "homeomorphism S T f g"
+      and "\<And>S T. homeomorphism S T f g \<Longrightarrow> (P S \<longleftrightarrow> Q T)"
     shows "locally P S \<longleftrightarrow> locally Q T"
-     (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then show ?rhs
-    using eq hom homeomorphism_locally_imp by blast
-next
-  assume ?rhs
-  then show ?lhs
-    using eq homeomorphism_sym homeomorphism_symD [OF hom] 
-    by (blast intro: homeomorphism_locally_imp) 
-qed
+  by (smt (verit) assms homeomorphism_locally_imp homeomorphism_symD)
 
 lemma homeomorphic_locally:
   fixes S:: "'a::metric_space set" and T:: "'b::metric_space set"
   assumes hom: "S homeomorphic T"
           and iff: "\<And>X Y. X homeomorphic Y \<Longrightarrow> (P X \<longleftrightarrow> Q Y)"
     shows "locally P S \<longleftrightarrow> locally Q T"
-proof -
-  obtain f g where hom: "homeomorphism S T f g"
-    using assms by (force simp: homeomorphic_def)
-  then show ?thesis
-    using homeomorphic_def local.iff
-    by (blast intro!: homeomorphism_locally)
-qed
+  by (smt (verit, ccfv_SIG) hom homeomorphic_def homeomorphism_locally homeomorphism_locally_imp iff)
 
 lemma homeomorphic_local_compactness:
   fixes S:: "'a::metric_space set" and T:: "'b::metric_space set"
   shows "S homeomorphic T \<Longrightarrow> locally compact S \<longleftrightarrow> locally compact T"
-by (simp add: homeomorphic_compactness homeomorphic_locally)
+  by (simp add: homeomorphic_compactness homeomorphic_locally)
 
 lemma locally_translation:
   fixes P :: "'a :: real_normed_vector set \<Rightarrow> bool"
@@ -1746,7 +1690,7 @@
       and oo: "\<And>T. openin (top_of_set S) T \<Longrightarrow> openin (top_of_set (f ` S)) (f ` T)"
       and Q: "\<And>T. \<lbrakk>T \<subseteq> S; P T\<rbrakk> \<Longrightarrow> Q(f ` T)"
     shows "locally Q (f ` S)"
-proof (clarsimp simp add: locally_def)
+proof (clarsimp simp: locally_def)
   fix W y
   assume oiw: "openin (top_of_set (f ` S)) W" and "y \<in> W"
   then have "W \<subseteq> f ` S" by (simp add: openin_euclidean_subtopology_iff)
@@ -1756,8 +1700,7 @@
     using \<open>W \<subseteq> f ` S\<close> \<open>y \<in> W\<close> by blast
   then obtain U V
     where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W"
-    using P [unfolded locally_def, rule_format, of "(S \<inter> f -` W)" x] oivf \<open>y \<in> W\<close>
-    by auto
+    by (metis IntI P \<open>y \<in> W\<close> locallyE oivf vimageI)
   then have "openin (top_of_set (f ` S)) (f ` U)"
     by (simp add: oo)
   then show "\<exists>X. openin (top_of_set (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)"
@@ -1778,17 +1721,16 @@
 proof -
   let ?A = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}"
   let ?B = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}"
-  have 1: "openin (top_of_set S) ?A"
-    by (subst openin_subopen, blast)
-  have 2: "openin (top_of_set S) ?B"
-    by (subst openin_subopen, blast)
-  have \<section>: "?A \<inter> ?B = {}"
+  have "?A \<inter> ?B = {}"
     by (clarsimp simp: set_eq_iff) (metis (no_types, opaque_lifting) Int_iff opD openin_Int)
-  have *: "S \<subseteq> ?A \<union> ?B"
+  moreover have "S \<subseteq> ?A \<union> ?B"
     by clarsimp (meson opI)
-  have "?A = {} \<or> ?B = {}"
-    using \<open>connected S\<close> [unfolded connected_openin, simplified, rule_format, OF 1 \<section> * 2] 
-    by blast
+  moreover have "openin (top_of_set S) ?A"
+    by (subst openin_subopen, blast)
+  moreover have "openin (top_of_set S) ?B"
+    by (subst openin_subopen, blast)
+  ultimately have "?A = {} \<or> ?B = {}"
+    by (metis (no_types, lifting) \<open>connected S\<close> connected_openin)
   then show ?thesis
     by clarsimp (meson opI etc)
 qed
@@ -1851,15 +1793,11 @@
   shows "locally (\<lambda>U. f constant_on U) S \<longleftrightarrow> f constant_on S" (is "?lhs = ?rhs")
 proof
   assume ?lhs
-  then have "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x\<in>T. f x = f a)"
-    unfolding locally_def
-    by (metis (mono_tags, opaque_lifting) constant_on_def constant_on_subset openin_subtopology_self)
   then show ?rhs
-    using assms
-    by (simp add: locally_constant_imp_constant)
+    by (smt (verit, del_insts) assms constant_on_def locally_constant_imp_constant locally_def openin_subtopology_self subset_iff)
 next
   assume ?rhs then show ?lhs
-    using assms by (metis constant_on_subset locallyI openin_imp_subset order_refl)
+    by (metis constant_on_subset locallyI openin_imp_subset order_refl)
 qed
 
 
@@ -1938,10 +1876,8 @@
   proof
     fix x
     assume "x \<in> S"
-    then obtain e where "e>0" and e: "closed (cball x e \<inter> S)"
-      using R by blast
-    then have "compact (cball x e \<inter> S)"
-      by (simp add: bounded_Int compact_eq_bounded_closed)
+    then obtain e where "e>0" and "compact (cball x e \<inter> S)"
+      by (metis Int_commute compact_Int_closed compact_cball inf.right_idem R)
     moreover have "\<forall>y\<in>ball x e \<inter> S. \<exists>\<epsilon>>0. cball y \<epsilon> \<inter> S \<subseteq> ball x e"
       by (meson Elementary_Metric_Spaces.open_ball IntD1 le_infI1 open_contains_cball_eq)
     moreover have "openin (top_of_set S) (ball x e \<inter> S)"
@@ -1984,10 +1920,8 @@
     have "openin (top_of_set S) (\<Union> (u ` T))"
       using T that uv by fastforce
     moreover
-    have "compact (\<Union> (v ` T))"
-      by (meson T compact_UN subset_eq that(1) uv)
-    moreover have "\<Union> (v ` T) \<subseteq> S"
-      by (metis SUP_least T(1) subset_eq that(1) uv)
+    obtain "compact (\<Union> (v ` T))" "\<Union> (v ` T) \<subseteq> S"
+      by (metis T UN_subset_iff \<open>K \<subseteq> S\<close> compact_UN subset_iff uv)
     ultimately show ?thesis
       using T by auto 
   qed
@@ -1996,7 +1930,7 @@
 next
   assume ?rhs
   then show ?lhs
-    apply (clarsimp simp add: locally_compact)
+    apply (clarsimp simp: locally_compact)
     apply (drule_tac x="{x}" in spec, simp)
     done
 qed
@@ -2014,14 +1948,7 @@
     have ope: "openin (top_of_set S) (ball x e)"
       by (meson e open_ball ball_subset_cball dual_order.trans open_subset)
     show ?thesis
-    proof (intro exI conjI)
-      let ?U = "ball x e"
-      let ?V = "cball x e"
-      show "x \<in> ?U" "?U \<subseteq> ?V" "?V \<subseteq> S" "compact ?V"
-        using \<open>e > 0\<close> e by auto
-      show "openin (top_of_set S) ?U"
-        using ope by blast
-    qed
+      by (meson \<open>0 < e\<close> ball_subset_cball centre_in_ball compact_cball e ope)
   qed
   show ?thesis
     unfolding locally_compact by (blast intro: *)
@@ -2045,13 +1972,13 @@
 
 lemma locally_compact_Int:
   fixes S :: "'a :: t2_space set"
-  shows "\<lbrakk>locally compact S; locally compact t\<rbrakk> \<Longrightarrow> locally compact (S \<inter> t)"
-by (simp add: compact_Int locally_Int)
+  shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<inter> T)"
+  by (simp add: compact_Int locally_Int)
 
 lemma locally_compact_closedin:
   fixes S :: "'a :: heine_borel set"
-  shows "\<lbrakk>closedin (top_of_set S) t; locally compact S\<rbrakk>
-        \<Longrightarrow> locally compact t"
+  shows "\<lbrakk>closedin (top_of_set S) T; locally compact S\<rbrakk>
+        \<Longrightarrow> locally compact T"
   unfolding closedin_closed
   using closed_imp_locally_compact locally_compact_Int by blast
 
@@ -2130,12 +2057,7 @@
     obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \<inter> T)"
       using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
     moreover have "closed (cball x (min e1 e2) \<inter> (S \<union> T))"
-    proof -
-      have "closed (cball x e1 \<inter> (cball x e2 \<inter> S))"
-        by (metis closed_Int closed_cball e1 inf_left_commute)
-      then show ?thesis
-        by (simp add: Int_Un_distrib cball_min_Int closed_Int closed_Un e2 inf_assoc)
-    qed
+      by (smt (verit) Int_Un_distrib2 Int_commute cball_min_Int closed_Int closed_Un closed_cball e1 e2 inf_left_commute)
     ultimately show ?thesis
       by (rule_tac x="min e1 e2" in exI) linarith
   qed
@@ -2298,7 +2220,7 @@
             by (simp_all add: closedin_closed_Int)
           moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2"
             using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12
-            by (auto simp add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+)
+            by (auto simp: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+)
           ultimately have cloDV1: "closedin (top_of_set D) (D \<inter> V1)"
                       and cloDV2:  "closedin (top_of_set D) (D \<inter> V2)"
             by metis+
@@ -2721,7 +2643,7 @@
   fixes S :: "'a:: real_normed_vector set"
   assumes "convex S"
   shows "locally path_connected S"
-proof (clarsimp simp add: locally_path_connected)
+proof (clarsimp simp: locally_path_connected)
   fix V x
   assume "openin (top_of_set S) V" and "x \<in> V"
   then obtain T e where  "V = S \<inter> T" "x \<in> S" "0 < e" "ball x e \<subseteq> T"
@@ -3824,7 +3746,7 @@
       show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S)
                            (\<lambda>(t,x). (1 - t) * x + t * z)"
         using  \<open>z \<in> S\<close> 
-        by (auto simp add: case_prod_unfold intro!: continuous_intros \<section>)
+        by (auto simp: case_prod_unfold intro!: continuous_intros \<section>)
     qed auto
   qed (simp add: contractible_space_empty)
 qed
@@ -3946,7 +3868,7 @@
     by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto)
   moreover have "homotopic_with_canon (\<lambda>x. True) T U (f \<circ> id) (f \<circ> (h \<circ> k))"
     by (rule homotopic_with_compose_continuous_left [where Y=T])
-       (use f in \<open>auto simp add: hom homotopic_with_symD\<close>)
+       (use f in \<open>auto simp: hom homotopic_with_symD\<close>)
   ultimately show ?thesis
     using that homotopic_with_trans by (fastforce simp add: o_def)
 qed
@@ -3987,7 +3909,7 @@
     by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto)
   moreover have "homotopic_with_canon (\<lambda>x. True) U T (id \<circ> f) ((h \<circ> k) \<circ> f)"
     by (rule homotopic_with_compose_continuous_right [where X=T])
-       (use f in \<open>auto simp add: hom homotopic_with_symD\<close>)
+       (use f in \<open>auto simp: hom homotopic_with_symD\<close>)
   ultimately show ?thesis
     using homotopic_with_trans by (fastforce simp add: o_def)
 qed
@@ -4123,7 +4045,7 @@
   fixes U :: "'a::euclidean_space set"
   assumes "convex U" "\<not> collinear U" "countable S"
     shows "path_connected(U - S)"
-proof (clarsimp simp add: path_connected_def)
+proof (clarsimp simp: path_connected_def)
   fix a b
   assume "a \<in> U" "a \<notin> S" "b \<in> U" "b \<notin> S"
   let ?m = "midpoint a b"
@@ -4245,7 +4167,7 @@
   next
     assume ge2: "aff_dim S \<ge> 2"
     then have "\<not> collinear S"
-    proof (clarsimp simp add: collinear_affine_hull)
+    proof (clarsimp simp: collinear_affine_hull)
       fix u v
       assume "S \<subseteq> affine hull {u, v}"
       then have "aff_dim S \<le> aff_dim {u, v}"
@@ -4282,7 +4204,7 @@
   assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S"
       and "\<not> collinear S" "countable T"
     shows "path_connected(S - T)"
-proof (clarsimp simp add: path_connected_component)
+proof (clarsimp simp: path_connected_component)
   fix x y
   assume xy: "x \<in> S" "x \<notin> T" "y \<in> S" "y \<notin> T"
   show "path_component (S - T) x y"
@@ -4483,7 +4405,7 @@
       by blast
   next
     show "cball a r \<inter> T \<subseteq> f ` (cball a r \<inter> T)"
-    proof (clarsimp simp add: dist_norm norm_minus_commute)
+    proof (clarsimp simp: dist_norm norm_minus_commute)
       fix x
       assume x: "norm (x - a) \<le> r" and "x \<in> T"
       have "\<exists>v \<in> {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \<bullet> 1 = 0"
@@ -4575,7 +4497,7 @@
       then show "continuous_on S gg"
         by (rule continuous_on_subset) (use ST in auto)
       show "ff ` S \<subseteq> S"
-      proof (clarsimp simp add: ff_def)
+      proof (clarsimp simp: ff_def)
         fix x
         assume "x \<in> S" and x: "dist a x < r" and "x \<in> T"
         then have "f x \<in> cball a r \<inter> T"
@@ -4584,7 +4506,7 @@
           using ST(1) \<open>x \<in> T\<close> gid hom homeomorphism_def x by fastforce
       qed
       show "gg ` S \<subseteq> S"
-      proof (clarsimp simp add: gg_def)
+      proof (clarsimp simp: gg_def)
         fix x
         assume "x \<in> S" and x: "dist a x < r" and "x \<in> T"
         then have "g x \<in> cball a r \<inter> T"
--- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy	Sun May 07 22:51:23 2023 +0200
@@ -2,7 +2,7 @@
 
 theory Ordered_Euclidean_Space
 imports
-  Convex_Euclidean_Space
+  Convex_Euclidean_Space Abstract_Limits
   "HOL-Library.Product_Order"
 begin
 
@@ -157,6 +157,32 @@
   shows "((\<lambda>i. inf (X i) (Y i)) \<longlongrightarrow> inf x y) net"
    unfolding inf_min eucl_inf by (intro assms tendsto_intros)
 
+lemma tendsto_Inf:
+  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::ordered_euclidean_space"
+  assumes "finite K" "\<And>i. i \<in> K \<Longrightarrow> ((\<lambda>x. f x i) \<longlongrightarrow> l i) F"
+  shows "((\<lambda>x. Inf (f x ` K)) \<longlongrightarrow> Inf (l ` K)) F"
+  using assms
+  by (induction K rule: finite_induct) (auto simp: cInf_insert_If tendsto_inf)
+
+lemma tendsto_Sup:
+  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::ordered_euclidean_space"
+  assumes "finite K" "\<And>i. i \<in> K \<Longrightarrow> ((\<lambda>x. f x i) \<longlongrightarrow> l i) F"
+  shows "((\<lambda>x. Sup (f x ` K)) \<longlongrightarrow> Sup (l ` K)) F"
+  using assms
+  by (induction K rule: finite_induct) (auto simp: cSup_insert_If tendsto_sup)
+
+lemma continuous_map_Inf:
+  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::ordered_euclidean_space"
+  assumes "finite K" "\<And>i. i \<in> K \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x i)"
+  shows "continuous_map X euclidean (\<lambda>x. INF i\<in>K. f x i)"
+  using assms by (simp add: continuous_map_atin tendsto_Inf)
+
+lemma continuous_map_Sup:
+  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::ordered_euclidean_space"
+  assumes "finite K" "\<And>i. i \<in> K \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x i)"
+  shows "continuous_map X euclidean (\<lambda>x. SUP i\<in>K. f x i)"
+  using assms by (simp add: continuous_map_atin tendsto_Sup)
+
 lemma tendsto_componentwise_max:
   assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F"
   shows "((\<lambda>x. (\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)) \<longlongrightarrow> (\<Sum>i\<in>Basis. max (l \<bullet> i) (m \<bullet> i) *\<^sub>R i)) F"
@@ -167,10 +193,6 @@
   shows "((\<lambda>x. (\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)) \<longlongrightarrow> (\<Sum>i\<in>Basis. min (l \<bullet> i) (m \<bullet> i) *\<^sub>R i)) F"
   by (intro tendsto_intros assms)
 
-lemma (in order) atLeastatMost_empty'[simp]:
-  "(\<not> a \<le> b) \<Longrightarrow> {a..b} = {}"
-  by (auto)
-
 instance real :: ordered_euclidean_space
   by standard auto
 
--- a/src/HOL/Analysis/Path_Connected.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy	Sun May 07 22:51:23 2023 +0200
@@ -2313,6 +2313,174 @@
   using assms homeomorphic_map_path_component_of by fastforce
 
 
+subsection\<open>Paths and path-connectedness\<close>
+
+lemma path_connected_space_quotient_map_image:
+   "\<lbrakk>quotient_map X Y q; path_connected_space X\<rbrakk> \<Longrightarrow> path_connected_space Y"
+  by (metis path_connectedin_continuous_map_image path_connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map)
+
+lemma path_connected_space_retraction_map_image:
+   "\<lbrakk>retraction_map X Y r; path_connected_space X\<rbrakk> \<Longrightarrow> path_connected_space Y"
+  using path_connected_space_quotient_map_image retraction_imp_quotient_map by blast
+
+lemma path_connected_space_prod_topology:
+  "path_connected_space(prod_topology X Y) \<longleftrightarrow>
+        topspace(prod_topology X Y) = {} \<or> path_connected_space X \<and> path_connected_space Y"
+proof (cases "topspace(prod_topology X Y) = {}")
+  case True
+  then show ?thesis
+    by (simp add: path_connected_space_topspace_empty)
+next
+  case False
+  have "path_connected_space (prod_topology X Y)" 
+    if X: "path_connected_space X" and Y: "path_connected_space Y"
+  proof (clarsimp simp: path_connected_space_def)
+    fix x y x' y'
+    assume "x \<in> topspace X" and "y \<in> topspace Y" and "x' \<in> topspace X" and "y' \<in> topspace Y"
+    obtain f where "pathin X f" "f 0 = x" "f 1 = x'"
+      by (meson X \<open>x \<in> topspace X\<close> \<open>x' \<in> topspace X\<close> path_connected_space_def)
+    obtain g where "pathin Y g" "g 0 = y" "g 1 = y'"
+      by (meson Y \<open>y \<in> topspace Y\<close> \<open>y' \<in> topspace Y\<close> path_connected_space_def)
+    show "\<exists>h. pathin (prod_topology X Y) h \<and> h 0 = (x,y) \<and> h 1 = (x',y')"
+    proof (intro exI conjI)
+      show "pathin (prod_topology X Y) (\<lambda>t. (f t, g t))"
+        using \<open>pathin X f\<close> \<open>pathin Y g\<close> by (simp add: continuous_map_paired pathin_def)
+      show "(\<lambda>t. (f t, g t)) 0 = (x, y)"
+        using \<open>f 0 = x\<close> \<open>g 0 = y\<close> by blast
+      show "(\<lambda>t. (f t, g t)) 1 = (x', y')"
+        using \<open>f 1 = x'\<close> \<open>g 1 = y'\<close> by blast
+    qed
+  qed
+  then show ?thesis
+    by (metis False Sigma_empty1 Sigma_empty2 topspace_prod_topology path_connected_space_retraction_map_image
+        retraction_map_fst  retraction_map_snd) 
+qed
+
+lemma path_connectedin_Times:
+   "path_connectedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow>
+        S = {} \<or> T = {} \<or> path_connectedin X S \<and> path_connectedin Y T"
+  by (auto simp add: path_connectedin_def subtopology_Times path_connected_space_prod_topology)
+
+
+subsection\<open>Path components\<close>
+
+lemma path_component_of_subtopology_eq:
+  "path_component_of (subtopology X U) x = path_component_of X x \<longleftrightarrow> path_component_of_set X x \<subseteq> U"  
+  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (metis path_connectedin_path_component_of path_connectedin_subtopology)
+next
+  show "?rhs \<Longrightarrow> ?lhs"
+    unfolding fun_eq_iff
+    by (metis path_connectedin_subtopology path_component_of path_component_of_aux path_component_of_mono)
+qed
+
+lemma path_components_of_subtopology:
+  assumes "C \<in> path_components_of X" "C \<subseteq> U"
+  shows "C \<in> path_components_of (subtopology X U)"
+  using assms path_component_of_refl path_component_of_subtopology_eq topspace_subtopology
+  by (smt (verit) imageE path_component_in_path_components_of path_components_of_def)
+
+lemma path_imp_connected_component_of:
+   "path_component_of X x y \<Longrightarrow> connected_component_of X x y"
+  by (metis in_mono mem_Collect_eq path_component_subset_connected_component_of)
+
+lemma finite_path_components_of_finite:
+   "finite(topspace X) \<Longrightarrow> finite(path_components_of X)"
+  by (simp add: Union_path_components_of finite_UnionD)
+
+lemma path_component_of_continuous_image:
+  "\<lbrakk>continuous_map X X' f; path_component_of X x y\<rbrakk> \<Longrightarrow> path_component_of X' (f x) (f y)"
+  by (meson image_eqI path_component_of path_connectedin_continuous_map_image)
+
+lemma path_component_of_pair [simp]:
+   "path_component_of_set (prod_topology X Y) (x,y) =
+    path_component_of_set X x \<times> path_component_of_set Y y"   (is "?lhs = ?rhs")
+proof (cases "?lhs = {}")
+  case True
+  then show ?thesis
+    by (metis Sigma_empty1 Sigma_empty2 mem_Sigma_iff path_component_of_eq_empty topspace_prod_topology) 
+next
+  case False
+  then have "path_component_of X x x" "path_component_of Y y y"
+    using path_component_of_eq_empty path_component_of_refl by fastforce+
+  moreover
+  have "path_connectedin (prod_topology X Y) (path_component_of_set X x \<times> path_component_of_set Y y)"
+    by (metis path_connectedin_Times path_connectedin_path_component_of)
+  moreover have "path_component_of X x a" "path_component_of Y y b"
+    if "(x, y) \<in> C'" "(a,b) \<in> C'" and "path_connectedin (prod_topology X Y) C'" for C' a b
+    by (smt (verit, best) that continuous_map_fst continuous_map_snd fst_conv snd_conv path_component_of path_component_of_continuous_image)+
+  ultimately show ?thesis
+    by (intro path_component_of_unique) auto
+qed
+
+lemma path_components_of_prod_topology:
+   "path_components_of (prod_topology X Y) =
+    (\<lambda>(C,D). C \<times> D) ` (path_components_of X \<times> path_components_of Y)"
+  by (force simp add: image_iff path_components_of_def)
+
+lemma path_components_of_prod_topology':
+   "path_components_of (prod_topology X Y) =
+    {C \<times> D |C D. C \<in> path_components_of X \<and> D \<in> path_components_of Y}"
+  by (auto simp: path_components_of_prod_topology)
+
+lemma path_component_of_product_topology:
+   "path_component_of_set (product_topology X I) f =
+    (if f \<in> extensional I then PiE I (\<lambda>i. path_component_of_set (X i) (f i)) else {})"
+    (is "?lhs = ?rhs")
+proof (cases "path_component_of_set (product_topology X I) f = {}")
+  case True
+  then show ?thesis
+    by (smt (verit) PiE_eq_empty_iff PiE_iff path_component_of_eq_empty topspace_product_topology)
+next
+  case False
+  then have [simp]: "f \<in> extensional I"
+    by (auto simp: path_component_of_eq_empty PiE_iff path_component_of_equiv)
+  show ?thesis
+  proof (intro path_component_of_unique)
+    show "f \<in> ?rhs"
+      using False path_component_of_eq_empty path_component_of_refl by force
+    show "path_connectedin (product_topology X I) (if f \<in> extensional I then \<Pi>\<^sub>E i\<in>I. path_component_of_set (X i) (f i) else {})"
+      by (simp add: path_connectedin_PiE path_connectedin_path_component_of)
+    fix C'
+    assume "f \<in> C'" and C': "path_connectedin (product_topology X I) C'" 
+    show "C' \<subseteq> ?rhs"
+    proof -
+      have "C' \<subseteq> extensional I"
+        using PiE_def C' path_connectedin_subset_topspace by fastforce
+      with \<open>f \<in> C'\<close> C' show ?thesis
+        apply (clarsimp simp: PiE_iff subset_iff)
+        by (smt (verit, ccfv_threshold) continuous_map_product_projection path_component_of path_component_of_continuous_image)
+    qed   
+  qed
+qed
+
+lemma path_components_of_product_topology:
+  "path_components_of (product_topology X I) =
+    {PiE I B |B. \<forall>i \<in> I. B i \<in> path_components_of(X i)}"  (is "?lhs=?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    apply (simp add: path_components_of_def image_subset_iff)
+    by (smt (verit, best) PiE_iff image_eqI path_component_of_product_topology)
+next
+  show "?rhs \<subseteq> ?lhs"
+  proof
+    fix F
+    assume "F \<in> ?rhs"
+    then obtain B where B: "F = Pi\<^sub>E I B"
+      and "\<forall>i\<in>I. \<exists>x\<in>topspace (X i). B i = path_component_of_set (X i) x"
+      by (force simp add: path_components_of_def image_iff)
+    then obtain f where ftop: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i)"
+      and BF: "\<And>i. i \<in> I \<Longrightarrow> B i = path_component_of_set (X i) (f i)"
+      by metis
+    then have "F = path_component_of_set (product_topology X I) (restrict f I)"
+      by (metis (mono_tags, lifting) B PiE_cong path_component_of_product_topology restrict_apply' restrict_extensional)
+    then show "F \<in> ?lhs"
+      by (simp add: ftop path_component_in_path_components_of)
+  qed
+qed
+
 subsection \<open>Sphere is path-connected\<close>
 
 lemma path_connected_punctured_universe:
--- a/src/HOL/Analysis/Product_Topology.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/Product_Topology.thy	Sun May 07 22:51:23 2023 +0200
@@ -1,7 +1,7 @@
 section\<open>The binary product topology\<close>
 
 theory Product_Topology
-imports Function_Topology
+  imports Function_Topology
 begin
 
 section \<open>Product Topology\<close>
@@ -464,6 +464,10 @@
    "homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(x,y). (y,x))"
   using homeomorphic_map_maps homeomorphic_maps_swap by metis
 
+lemma homeomorphic_space_prod_topology_swap:
+   "(prod_topology X Y) homeomorphic_space (prod_topology Y X)"
+  using homeomorphic_map_swap homeomorphic_space by blast
+
 lemma embedding_map_graph:
    "embedding_map X (prod_topology X Y) (\<lambda>x. (x, f x)) \<longleftrightarrow> continuous_map X Y f"
     (is "?lhs = ?rhs")
@@ -472,8 +476,7 @@
   have "snd \<circ> (\<lambda>x. (x, f x)) = f"
     by force
   moreover have "continuous_map X Y (snd \<circ> (\<lambda>x. (x, f x)))"
-    using L
-    unfolding embedding_map_def
+    using L unfolding embedding_map_def
     by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map)
   ultimately show ?rhs
     by simp
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/Sum_Topology.thy	Sun May 07 22:51:23 2023 +0200
@@ -0,0 +1,146 @@
+section\<open>Disjoint sum of arbitarily many spaces\<close>
+
+theory Sum_Topology
+  imports Abstract_Topology
+begin
+
+
+definition sum_topology :: "('a \<Rightarrow> 'b topology) \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'b) topology" where
+  "sum_topology X I \<equiv>
+    topology (\<lambda>U. U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i \<in> I. openin (X i) {x. (i,x) \<in> U}))"
+
+lemma is_sum_topology: "istopology (\<lambda>U. U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i\<in>I. openin (X i) {x. (i, x) \<in> U}))"
+proof -
+  have 1: "{x. (i, x) \<in> S \<inter> T} = {x. (i, x) \<in> S} \<inter> {x::'b. (i, x) \<in> T}" for S T and i::'a
+    by auto
+  have 2: "{x. (i, x) \<in> \<Union> \<K>} = (\<Union>K\<in>\<K>. {x::'b. (i, x) \<in> K})" for \<K> and i::'a
+    by auto
+  show ?thesis
+    unfolding istopology_def 1 2 by blast
+qed
+
+lemma openin_sum_topology:
+   "openin (sum_topology X I) U \<longleftrightarrow>
+        U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i \<in> I. openin (X i) {x. (i,x) \<in> U})"
+  by (auto simp: sum_topology_def is_sum_topology)
+
+lemma openin_disjoint_union:
+   "openin (sum_topology X I) (Sigma I U) \<longleftrightarrow> (\<forall>i \<in> I. openin (X i) (U i))"
+  using openin_subset by (force simp: openin_sum_topology)
+
+lemma topspace_sum_topology [simp]:
+   "topspace(sum_topology X I) = Sigma I (topspace \<circ> X)"
+  by (metis comp_apply openin_disjoint_union openin_subset openin_sum_topology openin_topspace subset_antisym)
+
+lemma openin_sum_topology_alt:
+   "openin (sum_topology X I) U \<longleftrightarrow> (\<exists>T. U = Sigma I T \<and> (\<forall>i \<in> I. openin (X i) (T i)))"
+  by (bestsimp simp: openin_sum_topology dest: openin_subset)
+
+lemma forall_openin_sum_topology:
+   "(\<forall>U. openin (sum_topology X I) U \<longrightarrow> P U) \<longleftrightarrow> (\<forall>T. (\<forall>i \<in> I. openin (X i) (T i)) \<longrightarrow> P(Sigma I T))"
+  by (auto simp: openin_sum_topology_alt)
+
+lemma exists_openin_sum_topology:
+   "(\<exists>U. openin (sum_topology X I) U \<and> P U) \<longleftrightarrow>
+    (\<exists>T. (\<forall>i \<in> I. openin (X i) (T i)) \<and> P(Sigma I T))"
+  by (auto simp: openin_sum_topology_alt)
+
+lemma closedin_sum_topology:
+   "closedin (sum_topology X I) U \<longleftrightarrow> U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i \<in> I. closedin (X i) {x. (i,x) \<in> U})"
+     (is "?lhs = ?rhs")
+proof
+  assume L: ?lhs
+  then have U: "U \<subseteq> Sigma I (topspace \<circ> X)"
+    using closedin_subset by fastforce
+  then have "\<forall>i\<in>I. {x. (i, x) \<in> U} \<subseteq> topspace (X i)"
+    by fastforce
+  moreover have "openin (X i) (topspace (X i) - {x. (i, x) \<in> U})" if "i\<in>I" for i
+    apply (subst openin_subopen)
+    using L that unfolding closedin_def openin_sum_topology
+    by (bestsimp simp: o_def subset_iff)
+  ultimately show ?rhs
+    by (simp add: U closedin_def)
+next
+  assume R: ?rhs
+  then have "openin (X i) {x. (i, x) \<in> topspace (sum_topology X I) - U}" if "i\<in>I" for i
+    apply (subst openin_subopen)
+    using that unfolding closedin_def openin_sum_topology
+    by (bestsimp simp: o_def subset_iff)
+  then show ?lhs
+    by (simp add: R closedin_def openin_sum_topology)
+qed
+
+lemma closedin_disjoint_union:
+   "closedin (sum_topology X I) (Sigma I U) \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) (U i))"
+  using closedin_subset by (force simp: closedin_sum_topology)
+
+lemma closedin_sum_topology_alt:
+   "closedin (sum_topology X I) U \<longleftrightarrow> (\<exists>T. U = Sigma I T \<and> (\<forall>i \<in> I. closedin (X i) (T i)))"
+  using closedin_subset unfolding closedin_sum_topology by auto blast
+
+lemma forall_closedin_sum_topology:
+   "(\<forall>U. closedin (sum_topology X I) U \<longrightarrow> P U) \<longleftrightarrow>
+        (\<forall>t. (\<forall>i \<in> I. closedin (X i) (t i)) \<longrightarrow> P(Sigma I t))"
+  by (metis closedin_sum_topology_alt)
+
+lemma exists_closedin_sum_topology:
+   "(\<exists>U. closedin (sum_topology X I) U \<and> P U) \<longleftrightarrow>
+    (\<exists>T. (\<forall>i \<in> I. closedin (X i) (T i)) \<and> P(Sigma I T))"
+  by (simp add: closedin_sum_topology_alt) blast
+
+lemma open_map_component_injection:
+   "i \<in> I \<Longrightarrow> open_map (X i) (sum_topology X I) (\<lambda>x. (i,x))"
+  unfolding open_map_def openin_sum_topology
+  using openin_subset openin_subopen by (force simp: image_iff)
+
+lemma closed_map_component_injection:
+  assumes "i \<in> I"
+  shows "closed_map(X i) (sum_topology X I) (\<lambda>x. (i,x))"
+proof -
+  have "closedin (X j) {x. j = i \<and> x \<in> U}"
+    if "\<And>U S. closedin U S \<Longrightarrow> S \<subseteq> topspace U" and "closedin (X i) U" and "j \<in> I" for U j
+    using that by (cases "j=i") auto
+  then show ?thesis
+    using closedin_subset assms by (force simp: closed_map_def closedin_sum_topology image_iff)
+qed
+
+lemma continuous_map_component_injection:
+   "i \<in> I \<Longrightarrow> continuous_map(X i) (sum_topology X I) (\<lambda>x. (i,x))"
+  apply (clarsimp simp: continuous_map_def openin_sum_topology)
+  by (smt (verit, best) Collect_cong mem_Collect_eq openin_subset subsetD)
+
+lemma subtopology_sum_topology:
+  "subtopology (sum_topology X I) (Sigma I S) =
+        sum_topology (\<lambda>i. subtopology (X i) (S i)) I"
+  unfolding topology_eq
+proof (intro strip iffI)
+  fix T
+  assume *: "openin (subtopology (sum_topology X I) (Sigma I S)) T"
+  then obtain U where U: "\<forall>i\<in>I. openin (X i) (U i) \<and> T = Sigma I U \<inter> Sigma I S" 
+    by (auto simp: openin_subtopology openin_sum_topology_alt)
+  have "T = (SIGMA i:I. U i \<inter> S i)"
+  proof
+    show "T \<subseteq> (SIGMA i:I. U i \<inter> S i)"
+      by (metis "*" SigmaE Sigma_Int_distrib2 U openin_imp_subset subset_iff)
+    show "(SIGMA i:I. U i \<inter> S i) \<subseteq> T"
+      using U by fastforce
+  qed
+  then show "openin (sum_topology (\<lambda>i. subtopology (X i) (S i)) I) T"
+    by (simp add: U openin_disjoint_union openin_subtopology_Int)
+next
+  fix T
+  assume *: "openin (sum_topology (\<lambda>i. subtopology (X i) (S i)) I) T"
+  then obtain U where "\<forall>i\<in>I. \<exists>T. openin (X i) T \<and> U i = T \<inter> S i" and Teq: "T = Sigma I U"
+    by (auto simp: openin_subtopology openin_sum_topology_alt)
+  then obtain B where "\<And>i. i\<in>I \<Longrightarrow> openin (X i) (B i) \<and> U i = B i \<inter> S i"
+    by metis
+  then show "openin (subtopology (sum_topology X I) (Sigma I S)) T"
+    by (auto simp: openin_subtopology openin_sum_topology_alt Teq)
+qed
+
+lemma embedding_map_component_injection:
+   "i \<in> I \<Longrightarrow> embedding_map (X i) (sum_topology X I) (\<lambda>x. (i,x))"
+  by (metis injective_open_imp_embedding_map continuous_map_component_injection
+            open_map_component_injection inj_onI prod.inject)
+
+end
--- a/src/HOL/Analysis/T1_Spaces.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Analysis/T1_Spaces.thy	Sun May 07 22:51:23 2023 +0200
@@ -49,6 +49,9 @@
   apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton)
   using t1_space_alt by auto
 
+lemma t1_space_euclidean: "t1_space (euclidean :: 'a::metric_space topology)"
+  by (simp add: t1_space_closedin_singleton)
+
 lemma closedin_t1_singleton:
    "\<lbrakk>t1_space X; a \<in> topspace X\<rbrakk> \<Longrightarrow> closedin X {a}"
   by (simp add: t1_space_closedin_singleton)
@@ -380,8 +383,7 @@
 lemma Hausdorff_space_discrete_topology [simp]:
    "Hausdorff_space (discrete_topology U)"
   unfolding Hausdorff_space_def
-  apply safe
-  by (metis discrete_topology_unique_alt disjnt_empty2 disjnt_insert2 insert_iff mk_disjoint_insert topspace_discrete_topology)
+  by (metis Hausdorff_space_compact_sets Hausdorff_space_def compactin_discrete_topology equalityE openin_discrete_topology)
 
 lemma compactin_Int:
    "\<lbrakk>Hausdorff_space X; compactin X S; compactin X T\<rbrakk> \<Longrightarrow> compactin X (S \<inter> T)"
@@ -395,6 +397,10 @@
    "\<lbrakk>Hausdorff_space X; finite S\<rbrakk> \<Longrightarrow> X derived_set_of S = {}"
   using Hausdorff_imp_t1_space t1_space_derived_set_of_finite by auto
 
+lemma infinite_perfect_set:
+   "\<lbrakk>Hausdorff_space X; S \<subseteq> X derived_set_of S; S \<noteq> {}\<rbrakk> \<Longrightarrow> infinite S"
+  using derived_set_of_finite by blast
+
 lemma derived_set_of_singleton:
    "Hausdorff_space X \<Longrightarrow> X derived_set_of {x} = {}"
   by (simp add: derived_set_of_finite)
@@ -698,4 +704,47 @@
   qed
 qed
 
+lemma Hausdorff_space_closed_neighbourhood:
+   "Hausdorff_space X \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. \<exists>U C. openin X U \<and> closedin X C \<and>
+                      Hausdorff_space(subtopology X C) \<and> x \<in> U \<and> U \<subseteq> C)" (is "_ = ?rhs")
+proof
+  assume R: ?rhs
+  show "Hausdorff_space X"
+    unfolding Hausdorff_space_def
+  proof clarify
+    fix x y
+    assume x: "x \<in> topspace X" and y: "y \<in> topspace X" and "x \<noteq> y"
+    obtain T C where *: "openin X T" "closedin X C" "x \<in> T" "T \<subseteq> C"
+                 and C: "Hausdorff_space (subtopology X C)"
+      by (meson R \<open>x \<in> topspace X\<close>)
+    show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V"
+    proof (cases "y \<in> C")
+      case True
+      with * C obtain U V where U: "openin (subtopology X C) U"
+                        and V: "openin (subtopology X C) V"
+                        and "x \<in> U" "y \<in> V" "disjnt U V"
+        unfolding Hausdorff_space_def
+        by (smt (verit, best) \<open>x \<noteq> y\<close> closedin_subset subsetD topspace_subtopology_subset)
+      then obtain U' V' where UV': "U = U' \<inter> C" "openin X U'" "V = V' \<inter> C" "openin X V'"
+        by (meson openin_subtopology)
+      have "disjnt (T \<inter> U') V'"
+        using \<open>disjnt U V\<close> UV' \<open>T \<subseteq> C\<close> by (force simp: disjnt_iff)
+      with \<open>T \<subseteq> C\<close> have "disjnt (T \<inter> U') (V' \<union> (topspace X - C))"
+        unfolding disjnt_def by blast
+      moreover
+      have "openin X (T \<inter> U')"
+        by (simp add: \<open>openin X T\<close> \<open>openin X U'\<close> openin_Int)
+      moreover have "openin X (V' \<union> (topspace X - C))"
+        using \<open>closedin X C\<close> \<open>openin X V'\<close> by auto
+      ultimately show ?thesis
+        using UV' \<open>x \<in> T\<close> \<open>x \<in> U\<close> \<open>y \<in> V\<close> by blast
+    next
+      case False
+      with * y show ?thesis
+        by (force simp: closedin_def disjnt_def)
+    qed
+  qed
+qed fastforce
+
 end
--- a/src/HOL/Archimedean_Field.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Archimedean_Field.thy	Sun May 07 22:51:23 2023 +0200
@@ -16,15 +16,10 @@
 proof -
   have "Sup (uminus ` S) = - (Inf S)"
   proof (rule antisym)
-    show "- (Inf S) \<le> Sup (uminus ` S)"
-      apply (subst minus_le_iff)
-      apply (rule cInf_greatest [OF \<open>S \<noteq> {}\<close>])
-      apply (subst minus_le_iff)
-      apply (rule cSup_upper)
-       apply force
-      using bdd
-      apply (force simp: abs_le_iff bdd_above_def)
-      done
+    have "\<And>x. x \<in> S \<Longrightarrow> bdd_above (uminus ` S)"
+      using bdd by (force simp: abs_le_iff bdd_above_def)
+  then show "- (Inf S) \<le> Sup (uminus ` S)"
+    by (meson cInf_greatest [OF \<open>S \<noteq> {}\<close>] cSUP_upper minus_le_iff)
   next
     have *: "\<And>x. x \<in> S \<Longrightarrow> Inf S \<le> x"
       by (meson abs_le_iff bdd bdd_below_def cInf_lower minus_le_iff)
--- a/src/HOL/Conditionally_Complete_Lattices.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Conditionally_Complete_Lattices.thy	Sun May 07 22:51:23 2023 +0200
@@ -481,6 +481,12 @@
   assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
 qed (intro cSup_eq_non_empty assms)
 
+lemma cSup_unique:
+  fixes b :: "'a :: {conditionally_complete_lattice, no_bot}"
+  assumes "\<And>c. (\<forall>x \<in> s. x \<le> c) \<longleftrightarrow> b \<le> c"
+  shows "Sup s = b"
+  by (metis assms cSup_eq order.refl)
+
 lemma cInf_eq:
   fixes a :: "'a :: {conditionally_complete_lattice, no_top}"
   assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
@@ -490,6 +496,12 @@
   assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
 qed (intro cInf_eq_non_empty assms)
 
+lemma cInf_unique:
+  fixes b :: "'a :: {conditionally_complete_lattice, no_top}"
+  assumes "\<And>c. (\<forall>x \<in> s. x \<ge> c) \<longleftrightarrow> b \<ge> c"
+  shows "Inf s = b"
+  by (meson assms cInf_eq order.refl)
+
 class conditionally_complete_linorder = conditionally_complete_lattice + linorder
 begin
 
--- a/src/HOL/Data_Structures/Leftist_Heap.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Data_Structures/Leftist_Heap.thy	Sun May 07 22:51:23 2023 +0200
@@ -158,6 +158,8 @@
 
 subsection "Complexity"
 
+text \<open>We count only the calls of the only recursive function: @{const merge}\<close>
+
 text\<open>Explicit termination argument: sum of sizes\<close>
 
 fun T_merge :: "'a::ord lheap \<Rightarrow> 'a lheap \<Rightarrow> nat" where
@@ -168,11 +170,11 @@
    else T_merge t1 r2) + 1"
 
 definition T_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat" where
-"T_insert x t = T_merge (Node Leaf (x, 1) Leaf) t + 1"
+"T_insert x t = T_merge (Node Leaf (x, 1) Leaf) t"
 
 fun T_del_min :: "'a::ord lheap \<Rightarrow> nat" where
-"T_del_min Leaf = 1" |
-"T_del_min (Node l _ r) = T_merge l r + 1"
+"T_del_min Leaf = 0" |
+"T_del_min (Node l _ r) = T_merge l r"
 
 lemma T_merge_min_height: "ltree l \<Longrightarrow> ltree r \<Longrightarrow> T_merge l r \<le> min_height l + min_height r + 1"
 proof(induction l r rule: merge.induct)
@@ -185,7 +187,7 @@
   le_log2_of_power[OF min_height_size1[of r]] T_merge_min_height[of l r] assms
 by linarith
 
-corollary T_insert_log: "ltree t \<Longrightarrow> T_insert x t \<le> log 2 (size1 t) + 3"
+corollary T_insert_log: "ltree t \<Longrightarrow> T_insert x t \<le> log 2 (size1 t) + 2"
 using T_merge_log[of "Node Leaf (x, 1) Leaf" t]
 by(simp add: T_insert_def split: tree.split)
 
@@ -203,15 +205,15 @@
 qed
 
 corollary T_del_min_log: assumes "ltree t"
-  shows "T_del_min t \<le> 2 * log 2 (size1 t) + 1"
+  shows "T_del_min t \<le> 2 * log 2 (size1 t)"
 proof(cases t rule: tree2_cases)
   case Leaf thus ?thesis using assms by simp
 next
   case [simp]: (Node l _ _ r)
-  have "T_del_min t = T_merge l r + 1" by simp
-  also have "\<dots> \<le> log 2 (size1 l) + log 2 (size1 r) + 2"
+  have "T_del_min t = T_merge l r" by simp
+  also have "\<dots> \<le> log 2 (size1 l) + log 2 (size1 r) + 1"
     using \<open>ltree t\<close> T_merge_log[of l r] by (auto simp del: T_merge.simps)
-  also have "\<dots> \<le> 2 * log 2 (size1 t) + 1"
+  also have "\<dots> \<le> 2 * log 2 (size1 t)"
     using ld_ld_1_less[of "size1 l" "size1 r"] by (simp)
   finally show ?thesis .
 qed
--- a/src/HOL/Filter.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Filter.thy	Sun May 07 22:51:23 2023 +0200
@@ -458,6 +458,12 @@
 lemma False_imp_not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> \<not> trivial_limit net \<Longrightarrow> \<not> eventually (\<lambda>x. P x) net"
   by (simp add: eventually_False)
 
+lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
+  by simp
+
+lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
+  by (simp add: filter_eq_iff)
+
 lemma eventually_Inf: "eventually P (Inf B) \<longleftrightarrow> (\<exists>X\<subseteq>B. finite X \<and> eventually P (Inf X))"
 proof -
   let ?F = "\<lambda>P. \<exists>X\<subseteq>B. finite X \<and> eventually P (Inf X)"
--- a/src/HOL/Fun.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Fun.thy	Sun May 07 22:51:23 2023 +0200
@@ -1031,6 +1031,12 @@
 abbreviation strict_mono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
   where "strict_mono_on A \<equiv> monotone_on A (<) (<)"
 
+abbreviation antimono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
+  where "antimono_on A \<equiv> monotone_on A (\<le>) (\<ge>)"
+
+abbreviation strict_antimono_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b :: ord) \<Rightarrow> bool"
+  where "strict_antimono_on A \<equiv> monotone_on A (<) (>)"
+
 lemma mono_on_def[no_atp]: "mono_on A f \<longleftrightarrow> (\<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s)"
   by (auto simp add: monotone_on_def)
 
@@ -1265,6 +1271,32 @@
   "strict_mono_on A (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) \<Longrightarrow> mono_on A f"
   by (rule mono_onI, rule strict_mono_on_leD)
 
+lemma mono_imp_strict_mono:
+  fixes f :: "'a::order \<Rightarrow> 'b::order"
+  shows "\<lbrakk>mono_on S f; inj_on f S\<rbrakk> \<Longrightarrow> strict_mono_on S f"
+  by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff)
+
+lemma strict_mono_iff_mono:
+  fixes f :: "'a::linorder \<Rightarrow> 'b::order"
+  shows "strict_mono_on S f \<longleftrightarrow> mono_on S f \<and> inj_on f S"
+proof
+  show "strict_mono_on S f \<Longrightarrow> mono_on S f \<and> inj_on f S"
+    by (simp add: strict_mono_on_imp_inj_on strict_mono_on_imp_mono_on)
+qed (auto intro: mono_imp_strict_mono)
+
+lemma antimono_imp_strict_antimono:
+  fixes f :: "'a::order \<Rightarrow> 'b::order"
+  shows "\<lbrakk>antimono_on S f; inj_on f S\<rbrakk> \<Longrightarrow> strict_antimono_on S f"
+  by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff)
+
+lemma strict_antimono_iff_antimono:
+  fixes f :: "'a::linorder \<Rightarrow> 'b::order"
+  shows "strict_antimono_on S f \<longleftrightarrow> antimono_on S f \<and> inj_on f S"
+proof
+  show "strict_antimono_on S f \<Longrightarrow> antimono_on S f \<and> inj_on f S"
+    by (force simp add: monotone_on_def intro: linorder_inj_onI)
+qed (auto intro: antimono_imp_strict_antimono)
+
 lemma mono_compose: "mono Q \<Longrightarrow> mono (\<lambda>i x. Q i (f x))"
   unfolding mono_def le_fun_def by auto
 
--- a/src/HOL/Library/Set_Idioms.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Library/Set_Idioms.thy	Sun May 07 22:51:23 2023 +0200
@@ -557,4 +557,162 @@
     by blast
 qed
 
+lemma countable_union_of_empty [simp]: "(countable union_of P) {}"
+  by (simp add: union_of_empty)
+
+lemma countable_intersection_of_empty [simp]: "(countable intersection_of P) UNIV"
+  by (simp add: intersection_of_empty)
+
+lemma countable_union_of_inc: "P S \<Longrightarrow> (countable union_of P) S"
+  by (simp add: union_of_inc)
+
+lemma countable_intersection_of_inc: "P S \<Longrightarrow> (countable intersection_of P) S"
+  by (simp add: intersection_of_inc)
+
+lemma countable_union_of_complement:
+  "(countable union_of P) S \<longleftrightarrow> (countable intersection_of (\<lambda>S. P(-S))) (-S)" 
+  (is "?lhs=?rhs")
+proof
+  assume ?lhs
+  then obtain \<U> where "countable \<U>" and \<U>: "\<U> \<subseteq> Collect P" "\<Union>\<U> = S"
+    by (metis union_of_def)
+  define \<U>' where "\<U>' \<equiv> (\<lambda>C. -C) ` \<U>"
+  have "\<U>' \<subseteq> {S. P (- S)}" "\<Inter>\<U>' = -S"
+    using \<U>'_def \<U> by auto
+  then show ?rhs
+    unfolding intersection_of_def by (metis \<U>'_def \<open>countable \<U>\<close> countable_image)
+next
+  assume ?rhs
+  then obtain \<U> where "countable \<U>" and \<U>: "\<U> \<subseteq> {S. P (- S)}" "\<Inter>\<U> = -S"
+    by (metis intersection_of_def)
+  define \<U>' where "\<U>' \<equiv> (\<lambda>C. -C) ` \<U>"
+  have "\<U>' \<subseteq> Collect P" "\<Union> \<U>' = S"
+    using \<U>'_def \<U> by auto
+  then show ?lhs
+    unfolding union_of_def
+    by (metis \<U>'_def \<open>countable \<U>\<close> countable_image)
+qed
+
+lemma countable_intersection_of_complement:
+   "(countable intersection_of P) S \<longleftrightarrow> (countable union_of (\<lambda>S. P(- S))) (- S)"
+  by (simp add: countable_union_of_complement)
+
+lemma countable_union_of_explicit:
+  assumes "P {}"
+  shows "(countable union_of P) S \<longleftrightarrow>
+         (\<exists>T. (\<forall>n::nat. P(T n)) \<and> \<Union>(range T) = S)" (is "?lhs=?rhs")
+proof
+  assume ?lhs
+  then obtain \<U> where "countable \<U>" and \<U>: "\<U> \<subseteq> Collect P" "\<Union>\<U> = S"
+    by (metis union_of_def)
+  then show ?rhs
+    by (metis SUP_bot Sup_empty assms from_nat_into mem_Collect_eq range_from_nat_into subsetD)
+next
+  assume ?rhs
+  then show ?lhs
+    by (metis countableI_type countable_image image_subset_iff mem_Collect_eq union_of_def)
+qed
+
+lemma countable_union_of_ascending:
+  assumes empty: "P {}" and Un: "\<And>T U. \<lbrakk>P T; P U\<rbrakk> \<Longrightarrow> P(T \<union> U)"
+  shows "(countable union_of P) S \<longleftrightarrow>
+         (\<exists>T. (\<forall>n. P(T n)) \<and> (\<forall>n. T n \<subseteq> T(Suc n)) \<and> \<Union>(range T) = S)" (is "?lhs=?rhs")
+proof
+  assume ?lhs
+  then obtain T where T: "\<And>n::nat. P(T n)" "\<Union>(range T) = S"
+    by (meson empty countable_union_of_explicit)
+  have "P (\<Union> (T ` {..n}))" for n
+    by (induction n) (auto simp: atMost_Suc Un T)
+  with T show ?rhs
+    by (rule_tac x="\<lambda>n. \<Union>k\<le>n. T k" in exI) force
+next
+  assume ?rhs
+  then show ?lhs
+    using empty countable_union_of_explicit by auto
+qed
+
+lemma countable_union_of_idem [simp]:
+  "countable union_of countable union_of P = countable union_of P"  (is "?lhs=?rhs")
+proof
+  fix S
+  show "(countable union_of countable union_of P) S = (countable union_of P) S"
+  proof
+    assume L: "?lhs S"
+    then obtain \<U> where "countable \<U>" and \<U>: "\<U> \<subseteq> Collect (countable union_of P)" "\<Union>\<U> = S"
+      by (metis union_of_def)
+    then have "\<forall>U \<in> \<U>. \<exists>\<V>. countable \<V> \<and> \<V> \<subseteq> Collect P \<and> U = \<Union>\<V>"
+      by (metis Ball_Collect union_of_def)
+    then obtain \<F> where \<F>: "\<forall>U \<in> \<U>. countable (\<F> U) \<and> \<F> U \<subseteq> Collect P \<and> U = \<Union>(\<F> U)"
+      by metis
+    have "countable (\<Union> (\<F> ` \<U>))"
+      using \<F> \<open>countable \<U>\<close> by blast
+    moreover have "\<Union> (\<F> ` \<U>) \<subseteq> Collect P"
+      by (simp add: Sup_le_iff \<F>)
+    moreover have "\<Union> (\<Union> (\<F> ` \<U>)) = S"
+      by auto (metis Union_iff \<F> \<U>(2))+
+    ultimately show "?rhs S"
+      by (meson union_of_def)
+  qed (simp add: countable_union_of_inc)
+qed
+
+lemma countable_intersection_of_idem [simp]:
+   "countable intersection_of countable intersection_of P =
+        countable intersection_of P"
+  by (force simp: countable_intersection_of_complement)
+
+lemma countable_union_of_Union:
+   "\<lbrakk>countable \<U>; \<And>S. S \<in> \<U> \<Longrightarrow> (countable union_of P) S\<rbrakk>
+        \<Longrightarrow> (countable union_of P) (\<Union> \<U>)"
+  by (metis Ball_Collect countable_union_of_idem union_of_def)
+
+lemma countable_union_of_UN:
+   "\<lbrakk>countable I; \<And>i. i \<in> I \<Longrightarrow> (countable union_of P) (U i)\<rbrakk>
+        \<Longrightarrow> (countable union_of P) (\<Union>i\<in>I. U i)"
+  by (metis (mono_tags, lifting) countable_image countable_union_of_Union imageE)
+
+lemma countable_union_of_Un:
+  "\<lbrakk>(countable union_of P) S; (countable union_of P) T\<rbrakk>
+           \<Longrightarrow> (countable union_of P) (S \<union> T)"
+  by (smt (verit) Union_Un_distrib countable_Un le_sup_iff union_of_def)
+
+lemma countable_intersection_of_Inter:
+   "\<lbrakk>countable \<U>; \<And>S. S \<in> \<U> \<Longrightarrow> (countable intersection_of P) S\<rbrakk>
+        \<Longrightarrow> (countable intersection_of P) (\<Inter> \<U>)"
+  by (metis countable_intersection_of_idem intersection_of_def mem_Collect_eq subsetI)
+
+lemma countable_intersection_of_INT:
+   "\<lbrakk>countable I; \<And>i. i \<in> I \<Longrightarrow> (countable intersection_of P) (U i)\<rbrakk>
+        \<Longrightarrow> (countable intersection_of P) (\<Inter>i\<in>I. U i)"
+  by (metis (mono_tags, lifting) countable_image countable_intersection_of_Inter imageE)
+
+lemma countable_intersection_of_inter:
+   "\<lbrakk>(countable intersection_of P) S; (countable intersection_of P) T\<rbrakk>
+           \<Longrightarrow> (countable intersection_of P) (S \<inter> T)"
+  by (simp add: countable_intersection_of_complement countable_union_of_Un)
+
+lemma countable_union_of_Int:
+  assumes S: "(countable union_of P) S" and T: "(countable union_of P) T"
+    and Int: "\<And>S T. P S \<and> P T \<Longrightarrow> P(S \<inter> T)"
+  shows "(countable union_of P) (S \<inter> T)"
+proof -
+  obtain \<U> where "countable \<U>" and \<U>: "\<U> \<subseteq> Collect P" "\<Union>\<U> = S"
+    using S by (metis union_of_def)
+  obtain \<V> where "countable \<V>" and \<V>: "\<V> \<subseteq> Collect P" "\<Union>\<V> = T"
+    using T by (metis union_of_def)
+  have "\<And>U V. \<lbrakk>U \<in> \<U>; V \<in> \<V>\<rbrakk> \<Longrightarrow> (countable union_of P) (U \<inter> V)"
+    using \<U> \<V> by (metis Ball_Collect countable_union_of_inc local.Int)
+  then have "(countable union_of P) (\<Union>U\<in>\<U>. \<Union>V\<in>\<V>. U \<inter> V)"
+    by (meson \<open>countable \<U>\<close> \<open>countable \<V>\<close> countable_union_of_UN)
+  moreover have "S \<inter> T = (\<Union>U\<in>\<U>. \<Union>V\<in>\<V>. U \<inter> V)"
+    by (simp add: \<U> \<V>)
+  ultimately show ?thesis
+    by presburger
+qed
+
+lemma countable_intersection_of_union:
+  assumes S: "(countable intersection_of P) S" and T: "(countable intersection_of P) T"
+    and Un: "\<And>S T. P S \<and> P T \<Longrightarrow> P(S \<union> T)"
+  shows "(countable intersection_of P) (S \<union> T)"
+  by (metis (mono_tags, lifting) Compl_Int S T Un compl_sup countable_intersection_of_complement countable_union_of_Int)
+
 end
--- a/src/HOL/Limits.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Limits.thy	Sun May 07 22:51:23 2023 +0200
@@ -3251,4 +3251,21 @@
   for f :: "real \<Rightarrow> real"
   using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff)
 
+lemma Lim_topological:
+  "(f \<longlongrightarrow> l) net \<longleftrightarrow>
+    trivial_limit net \<or> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)"
+  unfolding tendsto_def trivial_limit_eq by auto
+
+lemma eventually_within_Un:
+  "eventually P (at x within (s \<union> t)) \<longleftrightarrow>
+    eventually P (at x within s) \<and> eventually P (at x within t)"
+  unfolding eventually_at_filter
+  by (auto elim!: eventually_rev_mp)
+
+lemma Lim_within_Un:
+ "(f \<longlongrightarrow> l) (at x within (s \<union> t)) \<longleftrightarrow>
+  (f \<longlongrightarrow> l) (at x within s) \<and> (f \<longlongrightarrow> l) (at x within t)"
+  unfolding tendsto_def
+  by (auto simp: eventually_within_Un)
+
 end
--- a/src/HOL/Real.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Real.thy	Sun May 07 22:51:23 2023 +0200
@@ -1096,6 +1096,27 @@
   by (metis reals_Archimedean3 dual_order.order_iff_strict le0 le_less_trans not_le x0 xc)
 
 
+lemma inverse_Suc: "inverse (Suc n) > 0"
+  using of_nat_0_less_iff positive_imp_inverse_positive zero_less_Suc by blast
+
+lemma Archimedean_eventually_inverse:
+  fixes \<epsilon>::real shows "(\<forall>\<^sub>F n in sequentially. inverse (real (Suc n)) < \<epsilon>) \<longleftrightarrow> 0 < \<epsilon>"
+  (is "?lhs=?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    unfolding eventually_at_top_dense using inverse_Suc order_less_trans by blast
+next
+  assume ?rhs
+  then obtain N where "inverse (Suc N) < \<epsilon>"
+    using reals_Archimedean by blast
+  moreover have "inverse (Suc n) \<le> inverse (Suc N)" if "n \<ge> N" for n
+    using inverse_Suc that by fastforce
+  ultimately show ?lhs
+    unfolding eventually_sequentially
+    using order_le_less_trans by blast
+qed
+
 subsection \<open>Rationals\<close>
 
 lemma Rats_abs_iff[simp]:
@@ -1381,10 +1402,34 @@
 lemma forall_pos_mono_1:
   "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
     (\<And>n. P (inverse (real (Suc n)))) \<Longrightarrow> 0 < e \<Longrightarrow> P e"
-  apply (rule forall_pos_mono)
-  apply auto
-  apply (metis Suc_pred of_nat_Suc)
-  done
+  using reals_Archimedean by blast
+
+lemma Archimedean_eventually_pow:
+  fixes x::real
+  assumes "1 < x"
+  shows "\<forall>\<^sub>F n in sequentially. b < x ^ n"
+proof -
+  obtain N where "\<And>n. n\<ge>N \<Longrightarrow> b < x ^ n"
+    by (metis assms le_less order_less_trans power_strict_increasing_iff real_arch_pow)
+  then show ?thesis
+    using eventually_sequentially by blast
+qed
+
+lemma Archimedean_eventually_pow_inverse:
+  fixes x::real
+  assumes "\<bar>x\<bar> < 1" "\<epsilon> > 0"
+  shows "\<forall>\<^sub>F n in sequentially. \<bar>x^n\<bar> < \<epsilon>"
+proof (cases "x = 0")
+  case True
+  then show ?thesis
+    by (simp add: assms eventually_at_top_dense zero_power)
+next
+  case False
+  then have "\<forall>\<^sub>F n in sequentially. inverse \<epsilon> < inverse \<bar>x\<bar> ^ n"
+    by (simp add: Archimedean_eventually_pow assms(1) one_less_inverse)
+  then show ?thesis
+    by eventually_elim (metis \<open>\<epsilon> > 0\<close> inverse_less_imp_less power_abs power_inverse)
+qed
 
 
 subsection \<open>Floor and Ceiling Functions from the Reals to the Integers\<close>
--- a/src/HOL/Set.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Set.thy	Sun May 07 22:51:23 2023 +0200
@@ -1946,6 +1946,9 @@
 lemma disjnt_Un2 [simp]: "disjnt C (A \<union> B) \<longleftrightarrow> disjnt C A \<and> disjnt C B"
   by (auto simp: disjnt_def)
 
+lemma disjnt_Diff1: "disjnt (X-Y) (U-V)" and disjnt_Diff2: "disjnt (U-V) (X-Y)" if "X \<subseteq> V"
+  using that by (auto simp: disjnt_def)
+
 lemma disjoint_image_subset: "\<lbrakk>pairwise disjnt \<A>; \<And>X. X \<in> \<A> \<Longrightarrow> f X \<subseteq> X\<rbrakk> \<Longrightarrow> pairwise disjnt (f `\<A>)"
   unfolding disjnt_def pairwise_def by fast
 
--- a/src/HOL/Set_Interval.thy	Sun May 07 14:25:41 2023 +0200
+++ b/src/HOL/Set_Interval.thy	Sun May 07 22:51:23 2023 +0200
@@ -280,8 +280,8 @@
 context order
 begin
 
-lemma atLeastatMost_empty[simp]:
-  "b < a \<Longrightarrow> {a..b} = {}"
+lemma atLeastatMost_empty[simp]: "b < a \<Longrightarrow> {a..b} = {}" 
+  and atLeastatMost_empty'[simp]: "\<not> a \<le> b \<Longrightarrow> {a..b} = {}"
   by(auto simp: atLeastAtMost_def atLeast_def atMost_def)
 
 lemma atLeastLessThan_empty[simp]:
--- a/src/Pure/Isar/code.ML	Sun May 07 14:25:41 2023 +0200
+++ b/src/Pure/Isar/code.ML	Sun May 07 22:51:23 2023 +0200
@@ -758,29 +758,24 @@
 
 fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy);
 
-fun expand_eta thy k thm =
-  let
-    val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
-    val (_, args) = strip_comb lhs;
-    val l = if k = ~1
-      then (length o fst o strip_abs) rhs
-      else Int.max (0, k - length args);
-    val (raw_vars, _) = Term.strip_abs_eta l rhs;
-    val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs [])))
-      raw_vars;
-    fun expand (v, ty) thm = Drule.fun_cong_rule thm
-      (Thm.global_cterm_of thy (Var ((v, 0), ty)));
-  in
-    thm
-    |> fold expand vars
-    |> Conv.fconv_rule Drule.beta_eta_conversion
-  end;
-
 fun same_arity thy thms =
   let
-    val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
-    val k = fold (Integer.max o num_args_of o Thm.prop_of) thms 0;
-  in map (expand_eta thy k) thms end;
+    val lhs_rhss = map (Logic.dest_equals o Thm.plain_prop_of) thms;
+    val k = fold (Integer.max o length o snd o strip_comb o fst) lhs_rhss 0;
+    fun expand_eta (lhs, rhs) thm =
+      let
+        val l = k - length (snd (strip_comb lhs));
+        val (raw_vars, _) = Term.strip_abs_eta l rhs;
+        val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs [])))
+          raw_vars;
+        fun expand (v, ty) thm = Drule.fun_cong_rule thm
+          (Thm.global_cterm_of thy (Var ((v, 0), ty)));
+      in
+        thm
+        |> fold expand vars
+        |> Conv.fconv_rule Drule.beta_eta_conversion
+      end;
+  in map2 expand_eta lhs_rhss thms end;
 
 fun mk_desymbolization pre post mk vs =
   let
--- a/src/Tools/Code/code_thingol.ML	Sun May 07 14:25:41 2023 +0200
+++ b/src/Tools/Code/code_thingol.ML	Sun May 07 22:51:23 2023 +0200
@@ -587,25 +587,6 @@
   | is_undefined_clause ctxt _ =
       false;
 
-fun satisfied_app wanted (ty, ts) =
-  let
-    val given = length ts;
-    val delta = wanted - given;
-    val rty = (drop wanted o binder_types) ty ---> body_type ty;
-  in
-    if delta = 0 then
-      (([], (ts, rty)), [])
-    else if delta < 0 then
-      let
-        val (ts1, ts2) = chop wanted ts
-      in (([], (ts1, rty)), ts2) end
-    else
-      let
-        val tys = (take delta o drop given o binder_types) ty;
-        val vs_tys = invent_params ((fold o fold_aterms) Term.declare_term_frees ts) tys;
-      in ((vs_tys, (ts @ map Free vs_tys, rty)), []) end
-  end
-
 fun ensure_tyco ctxt algbr eqngr permissive tyco =
   let
     val thy = Proof_Context.theory_of ctxt;
@@ -751,41 +732,34 @@
         then translation_error ctxt permissive some_thm deps
           "Abstraction violation" ("constant " ^ Code.string_of_const thy c)
       else ()
-  in translate_const_proper ctxt algbr eqngr permissive some_thm (c, ty) (deps, program) end
-and translate_const_proper ctxt algbr eqngr permissive some_thm (c, ty) =
-  let
-    val thy = Proof_Context.theory_of ctxt;
     val (annotate, ty') = dest_tagged_type ty;
     val typargs = Sign.const_typargs thy (c, ty');
     val sorts = Code_Preproc.sortargs eqngr c;
     val (dom, range) = Term.strip_type ty';
   in
-    ensure_const ctxt algbr eqngr permissive c
-    ##>> fold_map (translate_typ ctxt algbr eqngr permissive) typargs
-    ##>> fold_map (translate_dicts ctxt algbr eqngr permissive some_thm) (typargs ~~ sorts)
-    ##>> fold_map (translate_typ ctxt algbr eqngr permissive) (range :: dom)
-    #>> (fn (((c, typargs), dss), range :: dom) =>
+    (deps, program)
+    |> ensure_const ctxt algbr eqngr permissive c
+    ||>> fold_map (translate_typ ctxt algbr eqngr permissive) typargs
+    ||>> fold_map (translate_dicts ctxt algbr eqngr permissive some_thm) (typargs ~~ sorts)
+    ||>> fold_map (translate_typ ctxt algbr eqngr permissive) (range :: dom)
+    |>> (fn (((c, typargs), dss), range :: dom) =>
       { sym = Constant c, typargs = typargs, dicts = dss,
         dom = dom, range = range, annotation =
           if annotate then SOME (dom `--> range) else NONE })
   end
-and translate_case ctxt algbr eqngr permissive some_thm (t_pos, []) (c_ty, ts) =
+and translate_case ctxt algbr eqngr permissive some_thm (t_pos, []) _ (const as { dom, ... }, ts) =
       let
         fun project_term xs = nth xs t_pos;
         val project_clause = the_single o nth_drop t_pos;
-        val ty_case = project_term (binder_types (snd c_ty));
-        fun distill_clauses ty_case t =
+        val ty_case = project_term dom;
+        fun distill_clauses t =
           map (fn ([pat], body) => (pat, body)) (distill_minimized_clause [ty_case] t)
       in
-        translate_const ctxt algbr eqngr permissive some_thm NONE c_ty
-        ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts
-        ##>> translate_typ ctxt algbr eqngr permissive ty_case
-        #>> (fn ((const, ts), ty_case) =>
-            ICase { term = project_term ts, typ = ty_case,
-              clauses = (filter_out (is_undefined_clause ctxt) o distill_clauses ty_case o project_clause) ts,
-              primitive = IConst const `$$ ts })
+        pair (ICase { term = project_term ts, typ = ty_case,
+          clauses = (filter_out (is_undefined_clause ctxt) o distill_clauses o project_clause) ts,
+          primitive = IConst const `$$ ts })
       end
-  | translate_case ctxt algbr eqngr permissive some_thm (t_pos, case_pats) (c_ty, ts) =
+  | translate_case ctxt algbr eqngr permissive some_thm (t_pos, case_pats) ty (const as { dom, ... }, ts) =
       let
         fun project_term xs = nth xs t_pos;
         fun project_cases xs =
@@ -793,40 +767,39 @@
           |> nth_drop t_pos
           |> curry (op ~~) case_pats
           |> map_filter (fn (NONE, _) => NONE | (SOME _, x) => SOME x);
-        val ty_case = project_term (binder_types (snd c_ty));
-        val constrs = map_filter I case_pats ~~ project_cases ts
-          |> map (fn ((c, n), t) =>
-            ((c, (take n o binder_types o fastype_of_tagged_term) t ---> ty_case), n));
+        val tys = take (length case_pats + 1) (binder_types ty);
+        val ty_case = project_term tys;
+        val ty_case' = project_term dom;
+        val constrs = map_filter I case_pats ~~ project_cases tys
+          |> map (fn ((c, n), ty) =>
+            ((c, (take n o binder_types) ty ---> ty_case), n));
         fun distill_clauses constrs ts_clause =
           maps (fn ((constr as { dom = tys, ... }, n), t) =>
             map (fn (pat_args, body) => (IConst constr `$$ pat_args, body))
               (distill_minimized_clause (take n tys) t))
                 (constrs ~~ ts_clause);
       in
-        translate_const ctxt algbr eqngr permissive some_thm NONE c_ty
-        ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts
-        ##>> translate_typ ctxt algbr eqngr permissive ty_case
-        ##>> fold_map (fn (c_ty, n) =>
+        fold_map (fn (c_ty, n) =>
           translate_const ctxt algbr eqngr permissive some_thm NONE c_ty #>> rpair n) constrs
-        #>> (fn (((const, ts), ty_case), constrs) =>
-            ICase { term = project_term ts, typ = ty_case,
+        #>> (fn constrs =>
+            ICase { term = project_term ts, typ = ty_case',
               clauses = (filter_out (is_undefined_clause ctxt) o distill_clauses constrs o project_cases) ts,
               primitive = IConst const `$$ ts })
       end
-and translate_app_case ctxt algbr eqngr permissive some_thm pattern_schema c_ty ((vs_tys, (ts1, rty)), ts2) =
-  fold_map (fn (v, ty) => translate_typ ctxt algbr eqngr permissive ty #>> pair (SOME v)) vs_tys
-  ##>> translate_case ctxt algbr eqngr permissive some_thm pattern_schema (c_ty, ts1)
-  ##>> translate_typ ctxt algbr eqngr permissive rty
-  ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts2
-  #>> (fn (((vs_tys, t), rty), ts) => (vs_tys `|==> (t, rty)) `$$ ts)
 and translate_app ctxt algbr eqngr permissive some_thm some_abs (c_ty as (c, ty), ts) =
-  case Code.get_case_schema (Proof_Context.theory_of ctxt) c of
-    SOME (wanted, pattern_schema) =>
-      translate_app_case ctxt algbr eqngr permissive some_thm pattern_schema c_ty (satisfied_app wanted (ty, ts))
-  | NONE =>
-      translate_const ctxt algbr eqngr permissive some_thm some_abs c_ty
-      ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts
-      #>> (fn (const, ts) => IConst const `$$ ts)
+  translate_const ctxt algbr eqngr permissive some_thm some_abs c_ty
+  ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm NONE) ts
+  #-> (fn (const, ts) => case Code.get_case_schema (Proof_Context.theory_of ctxt) c of
+      SOME (wanted, pattern_schema) =>
+        let
+          val ((vs_tys, (ts1, rty)), ts2) = satisfied_application wanted (const, ts);
+          val (_, ty') = dest_tagged_type ty;
+        in
+          translate_case ctxt algbr eqngr permissive some_thm pattern_schema ty' (const, ts1)
+          #>> (fn t => (vs_tys `|==> (t, rty)) `$$ ts2)
+        end
+    | NONE =>
+        pair (IConst const `$$ ts))
 and translate_tyvar_sort ctxt (algbr as (proj_sort, _)) eqngr permissive (v, sort) =
   fold_map (ensure_class ctxt algbr eqngr permissive) (proj_sort sort)
   #>> (fn sort => (unprefix "'" v, sort))