merged
authordesharna
Fri, 26 May 2023 09:56:20 +0200
changeset 78108 250785900816
parent 78107 0366e49dab85 (current diff)
parent 78101 300537844bb7 (diff)
child 78109 5c6db3d1b602
merged
--- a/Admin/components/components.sha1	Fri May 26 09:49:45 2023 +0200
+++ b/Admin/components/components.sha1	Fri May 26 09:56:20 2023 +0200
@@ -374,6 +374,7 @@
 7056b285af67902b32f5049349a064f073f05860 polyml-5.9-cc80e2b43c38.tar.gz
 0c396bd6b46ff11a2432b91aab2be0248bd9b0a4 polyml-5.9.tar.gz
 f399ab9ee4a586fddeb6e73ca3605af65a89f969 polyml-5e9c8155ea96.tar.gz
+2cea4dd48bb8b171bc04c9793a55c7fa4c2d96f1 polyml-a5d5fba90286.tar.gz
 49f1adfacdd6d29fa9f72035d94a31eaac411a97 polyml-test-0a6ebca445fc.tar.gz
 2a8c4421e0a03c0d6ad556b3c36c34eb11568adb polyml-test-1236652ebd55.tar.gz
 8e83fb5088cf265902b8da753a8eac5fe3f6a14b polyml-test-159dc81efc3b.tar.gz
--- a/Admin/components/main	Fri May 26 09:49:45 2023 +0200
+++ b/Admin/components/main	Fri May 26 09:56:20 2023 +0200
@@ -25,7 +25,7 @@
 nunchaku-0.5
 opam-2.0.7
 pdfjs-2.14.305
-polyml-5e9c8155ea96
+polyml-a5d5fba90286
 postgresql-42.5.0
 prismjs-1.29.0
 rsync-3.2.7
--- a/src/Doc/Implementation/Proof.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/Doc/Implementation/Proof.thy	Fri May 26 09:56:20 2023 +0200
@@ -270,7 +270,8 @@
   cterm list -> Proof.context -> thm list * Proof.context"} \\
   @{define_ML Assumption.add_assumes: "
   cterm list -> Proof.context -> thm list * Proof.context"} \\
-  @{define_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
+  @{define_ML Assumption.export: "Proof.context -> Proof.context -> thm -> thm"} \\
+  @{define_ML Assumption.export_goal: "Proof.context -> Proof.context -> thm -> thm"} \\
   @{define_ML Assumption.export_term: "Proof.context -> Proof.context -> term -> term"} \\
   \end{mldecls}
 
@@ -290,11 +291,12 @@
   \<^ML>\<open>Assumption.add_assms\<close> where the export rule performs \<open>\<Longrightarrow>\<hyphen>intro\<close> or
   \<open>#\<Longrightarrow>\<hyphen>intro\<close>, depending on goal mode.
 
-  \<^descr> \<^ML>\<open>Assumption.export\<close>~\<open>is_goal inner outer thm\<close> exports result \<open>thm\<close>
-  from the \<open>inner\<close> context back into the \<open>outer\<close> one; \<open>is_goal = true\<close> means
-  this is a goal context. The result is in HHF normal form. Note that
-  \<^ML>\<open>Proof_Context.export\<close> combines \<^ML>\<open>Variable.export\<close> and
-  \<^ML>\<open>Assumption.export\<close> in the canonical way.
+  \<^descr> \<^ML>\<open>Assumption.export\<close>~\<open>inner outer thm\<close> exports result \<open>thm\<close> from the
+  \<open>inner\<close> context back into the \<open>outer\<close> one; \<^ML>\<open>Assumption.export_goal\<close>
+  does the same in a goal context (\<^theory_text>\<open>fix/assume/show\<close> in Isabelle/Isar). The
+  result is always in HHF normal form. Note that \<^ML>\<open>Proof_Context.export\<close>
+  combines \<^ML>\<open>Variable.export\<close> and \<^ML>\<open>Assumption.export\<close> in the
+  canonical way.
 
   \<^descr> \<^ML>\<open>Assumption.export_term\<close>~\<open>inner outer t\<close> exports term \<open>t\<close> from the
   \<open>inner\<close> context back into the \<open>outer\<close> one. This is analogous to
@@ -318,7 +320,7 @@
   val eq' = Thm.symmetric eq;
 
   (*back to original context -- discharges assumption*)
-  val r = Assumption.export false ctxt1 ctxt0 eq';
+  val r = Assumption.export ctxt1 ctxt0 eq';
 \<close>
 
 text \<open>
--- a/src/Doc/Isar_Ref/Generic.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/Doc/Isar_Ref/Generic.thy	Fri May 26 09:56:20 2023 +0200
@@ -840,9 +840,9 @@
 
 (*<*)experiment begin(*>*)
 simproc_setup unit ("x::unit") =
-  \<open>fn _ => fn _ => fn ct =>
+  \<open>K (K (fn ct =>
     if HOLogic.is_unit (Thm.term_of ct) then NONE
-    else SOME (mk_meta_eq @{thm unit_eq})\<close>
+    else SOME (mk_meta_eq @{thm unit_eq})))\<close>
 (*<*)end(*>*)
 
 text \<open>
--- a/src/Doc/Isar_Ref/Spec.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/Doc/Isar_Ref/Spec.thy	Fri May 26 09:56:20 2023 +0200
@@ -417,7 +417,7 @@
     @@{command declare} (@{syntax thms} + @'and')
   \<close>
 
-  \<^descr> \<^theory_text>\<open>declaration d\<close> adds the declaration function \<open>d\<close> of ML type \<^ML_type>\<open>declaration\<close>, to the current local theory under construction. In later
+  \<^descr> \<^theory_text>\<open>declaration d\<close> adds the declaration function \<open>d\<close> of ML type \<^ML_type>\<open>Morphism.declaration\<close>, to the current local theory under construction. In later
   application contexts, the function is transformed according to the morphisms
   being involved in the interpretation hierarchy.
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/Abstract_Metric_Spaces.thy	Fri May 26 09:56:20 2023 +0200
@@ -0,0 +1,2634 @@
+section \<open>Abstract Metric Spaces\<close>
+
+theory Abstract_Metric_Spaces
+  imports Elementary_Metric_Spaces Abstract_Limits Abstract_Topological_Spaces
+begin
+
+(*Avoid a clash with the existing metric_space locale (from the type class)*)
+locale Metric_space =
+  fixes M :: "'a set" and d :: "'a \<Rightarrow> 'a \<Rightarrow> real"
+  assumes nonneg [simp]: "\<And>x y. 0 \<le> d x y"
+  assumes commute: "\<And>x y. d x y = d y x"
+  assumes zero [simp]: "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> d x y = 0 \<longleftrightarrow> x=y"
+  assumes triangle: "\<And>x y z. \<lbrakk>x \<in> M; y \<in> M; z \<in> M\<rbrakk> \<Longrightarrow> d x z \<le> d x y + d y z"
+
+text \<open>Link with the type class version\<close>
+interpretation Met_TC: Metric_space UNIV dist
+  by (simp add: dist_commute dist_triangle Metric_space.intro)
+
+context Metric_space
+begin
+
+lemma subspace: "M' \<subseteq> M \<Longrightarrow> Metric_space M' d"
+  by (simp add: commute in_mono Metric_space.intro triangle)
+
+lemma abs_mdist [simp] : "\<bar>d x y\<bar> = d x y"
+  by simp
+
+lemma mdist_pos_less: "\<lbrakk>x \<noteq> y; x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> 0 < d x y"
+  by (metis less_eq_real_def nonneg zero)
+
+lemma mdist_zero [simp]: "x \<in> M \<Longrightarrow> d x x = 0"
+  by simp
+
+lemma mdist_pos_eq [simp]: "\<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> 0 < d x y \<longleftrightarrow> x \<noteq> y"
+  using mdist_pos_less zero by fastforce
+
+lemma triangle': "\<lbrakk>x \<in> M; y \<in> M; z \<in> M\<rbrakk> \<Longrightarrow> d x z \<le> d x y + d z y"
+  by (simp add: commute triangle)
+
+lemma triangle'': "\<lbrakk>x \<in> M; y \<in> M; z \<in> M\<rbrakk> \<Longrightarrow> d x z \<le> d y x + d y z"
+  by (simp add: commute triangle)
+
+lemma mdist_reverse_triangle: "\<lbrakk>x \<in> M; y \<in> M; z \<in> M\<rbrakk> \<Longrightarrow> \<bar>d x y - d y z\<bar> \<le> d x z"
+  by (smt (verit) commute triangle)
+
+text\<open> Open and closed balls                                                                \<close>
+
+definition mball where "mball x r \<equiv> {y. x \<in> M \<and> y \<in> M \<and> d x y < r}"
+definition mcball where "mcball x r \<equiv> {y. x \<in> M \<and> y \<in> M \<and> d x y \<le> r}"
+
+lemma in_mball [simp]: "y \<in> mball x r \<longleftrightarrow> x \<in> M \<and> y \<in> M \<and> d x y < r"
+  by (simp add: local.Metric_space_axioms Metric_space.mball_def)
+
+lemma centre_in_mball_iff [iff]: "x \<in> mball x r \<longleftrightarrow> x \<in> M \<and> 0 < r"
+  using in_mball mdist_zero by force
+
+lemma mball_subset_mspace: "mball x r \<subseteq> M"
+  by auto
+
+lemma mball_eq_empty: "mball x r = {} \<longleftrightarrow> (x \<notin> M) \<or> r \<le> 0"
+  by (smt (verit, best) Collect_empty_eq centre_in_mball_iff mball_def nonneg)
+
+lemma mball_subset: "\<lbrakk>d x y + a \<le> b; y \<in> M\<rbrakk> \<Longrightarrow> mball x a \<subseteq> mball y b"
+  by (smt (verit) commute in_mball subsetI triangle)
+
+lemma disjoint_mball: "r + r' \<le> d x x' \<Longrightarrow> disjnt (mball x r) (mball x' r')"
+  by (smt (verit) commute disjnt_iff in_mball triangle)
+
+lemma mball_subset_concentric: "r \<le> s \<Longrightarrow> mball x r \<subseteq> mball x s"
+  by auto
+
+lemma in_mcball [simp]: "y \<in> mcball x r \<longleftrightarrow> x \<in> M \<and> y \<in> M \<and> d x y \<le> r"
+  by (simp add: local.Metric_space_axioms Metric_space.mcball_def)
+
+lemma centre_in_mcball_iff [iff]: "x \<in> mcball x r \<longleftrightarrow> x \<in> M \<and> 0 \<le> r"
+  using mdist_zero by force
+
+lemma mcball_eq_empty: "mcball x r = {} \<longleftrightarrow> (x \<notin> M) \<or> r < 0"
+  by (smt (verit, best) Collect_empty_eq centre_in_mcball_iff empty_iff mcball_def nonneg)
+
+lemma mcball_subset_mspace: "mcball x r \<subseteq> M"
+  by auto
+
+lemma mball_subset_mcball: "mball x r \<subseteq> mcball x r"
+  by auto
+
+lemma mcball_subset: "\<lbrakk>d x y + a \<le> b; y \<in> M\<rbrakk> \<Longrightarrow> mcball x a \<subseteq> mcball y b"
+  by (smt (verit) in_mcball mdist_reverse_triangle subsetI)
+
+lemma mcball_subset_concentric: "r \<le> s \<Longrightarrow> mcball x r \<subseteq> mcball x s"
+  by force
+
+lemma mcball_subset_mball: "\<lbrakk>d x y + a < b; y \<in> M\<rbrakk> \<Longrightarrow> mcball x a \<subseteq> mball y b"
+  by (smt (verit) commute in_mball in_mcball subsetI triangle)
+
+lemma mcball_subset_mball_concentric: "a < b \<Longrightarrow> mcball x a \<subseteq> mball x b"
+  by force
+
+end
+
+
+
+subsection \<open>Metric topology                                                           \<close>
+
+context Metric_space
+begin
+
+definition mopen where 
+  "mopen U \<equiv> U \<subseteq> M \<and> (\<forall>x. x \<in> U \<longrightarrow> (\<exists>r>0. mball x r \<subseteq> U))"
+
+definition mtopology :: "'a topology" where 
+  "mtopology \<equiv> topology mopen"
+
+lemma is_topology_metric_topology [iff]: "istopology mopen"
+proof -
+  have "\<And>S T. \<lbrakk>mopen S; mopen T\<rbrakk> \<Longrightarrow> mopen (S \<inter> T)"
+    by (smt (verit, del_insts) Int_iff in_mball mopen_def subset_eq)
+  moreover have "\<And>\<K>. (\<forall>K\<in>\<K>. mopen K) \<longrightarrow> mopen (\<Union>\<K>)"
+    using mopen_def by fastforce
+  ultimately show ?thesis
+    by (simp add: istopology_def)
+qed
+
+lemma openin_mtopology: "openin mtopology U \<longleftrightarrow> U \<subseteq> M \<and> (\<forall>x. x \<in> U \<longrightarrow> (\<exists>r>0. mball x r \<subseteq> U))"
+  by (simp add: mopen_def mtopology_def)
+
+lemma topspace_mtopology [simp]: "topspace mtopology = M"
+  by (meson order.refl mball_subset_mspace openin_mtopology openin_subset openin_topspace subset_antisym zero_less_one)
+
+lemma subtopology_mspace [simp]: "subtopology mtopology M = mtopology"
+  by (metis subtopology_topspace topspace_mtopology)
+
+lemma open_in_mspace [iff]: "openin mtopology M"
+  by (metis openin_topspace topspace_mtopology)
+
+lemma closedin_mspace [iff]: "closedin mtopology M"
+  by (metis closedin_topspace topspace_mtopology)
+
+lemma openin_mball [iff]: "openin mtopology (mball x r)"
+proof -
+  have "\<And>y. \<lbrakk>x \<in> M; d x y < r\<rbrakk> \<Longrightarrow> \<exists>s>0. mball y s \<subseteq> mball x r"
+    by (metis add_diff_cancel_left' add_diff_eq commute less_add_same_cancel1 mball_subset order_refl)
+  then show ?thesis
+    by (auto simp: openin_mtopology)
+qed
+
+lemma mcball_eq_cball [simp]: "Met_TC.mcball = cball"
+  by force
+
+lemma mball_eq_ball [simp]: "Met_TC.mball = ball"
+  by force
+
+lemma mopen_eq_open [simp]: "Met_TC.mopen = open"
+  by (force simp: open_contains_ball Met_TC.mopen_def)
+
+lemma limitin_iff_tendsto [iff]: "limitin Met_TC.mtopology \<sigma> x F = tendsto \<sigma> x F"
+  by (simp add: Met_TC.mtopology_def)
+
+lemma mtopology_is_euclideanreal [simp]: "Met_TC.mtopology = euclideanreal"
+  by (simp add: Met_TC.mtopology_def)
+
+(*
+lemma metric_injective_image:
+   "\<And>f m s.
+        f ` s \<subseteq> M \<and>
+        (\<forall>x y. x \<in> s \<and> y \<in> s \<and> f x = f y \<Longrightarrow> x = y)
+        \<Longrightarrow> (mspace(metric(s,\<lambda>(x,y). d (f x) (f y))) = s) \<and>
+            (d(metric(s,\<lambda>(x,y). d (f x) (f y))) =
+             \<lambda>(x,y). d (f x) (f y))"
+oops
+  REWRITE_TAC[\<subseteq>; FORALL_IN_IMAGE; INJECTIVE_ON_ALT] THEN
+  REPEAT GEN_TAC THEN STRIP_TAC THEN
+  REWRITE_TAC[mspace; d; GSYM PAIR_EQ] THEN
+  REWRITE_TAC[GSYM(CONJUNCT2 metric_tybij); is_metric_space] THEN
+  REWRITE_TAC[GSYM mspace; GSYM d] THEN
+  ASM_SIMP_TAC[MDIST_POS_LE; MDIST_TRIANGLE; MDIST_0] THEN
+  ASM_MESON_TAC[MDIST_SYM]);;
+*)
+
+lemma mtopology_base:
+   "mtopology = topology(arbitrary union_of (\<lambda>U. \<exists>x \<in> M. \<exists>r>0. U = mball x r))"
+proof -
+  have "\<And>S. \<exists>x r. x \<in> M \<and> 0 < r \<and> S = mball x r \<Longrightarrow> openin mtopology S"
+    using openin_mball by blast
+  moreover have "\<And>U x. \<lbrakk>openin mtopology U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>B. (\<exists>x r. x \<in> M \<and> 0 < r \<and> B = mball x r) \<and> x \<in> B \<and> B \<subseteq> U"
+    by (metis centre_in_mball_iff in_mono openin_mtopology)
+  ultimately show ?thesis
+    by (smt (verit) topology_base_unique)
+qed
+
+lemma closedin_metric:
+   "closedin mtopology C \<longleftrightarrow> C \<subseteq> M \<and> (\<forall>x. x \<in> M - C \<longrightarrow> (\<exists>r>0. disjnt C (mball x r)))"  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    unfolding closedin_def openin_mtopology
+    by (metis Diff_disjoint disjnt_def disjnt_subset2 topspace_mtopology)
+  show "?rhs \<Longrightarrow> ?lhs"
+    unfolding closedin_def openin_mtopology disjnt_def
+    by (metis Diff_subset Diff_triv Int_Diff Int_commute inf.absorb_iff2 mball_subset_mspace topspace_mtopology)
+qed
+
+lemma closedin_mcball [iff]: "closedin mtopology (mcball x r)"
+proof -
+  have "\<exists>ra>0. disjnt (mcball x r) (mball y ra)" if "x \<notin> M" for y
+    by (metis disjnt_empty1 gt_ex mcball_eq_empty that)
+  moreover have "disjnt (mcball x r) (mball y (d x y - r))" if "y \<in> M" "d x y > r" for y
+    using that disjnt_iff in_mball in_mcball mdist_reverse_triangle by force
+  ultimately show ?thesis
+    using closedin_metric mcball_subset_mspace by fastforce
+qed
+
+lemma mball_iff_mcball: "(\<exists>r>0. mball x r \<subseteq> U) = (\<exists>r>0. mcball x r \<subseteq> U)"
+  by (meson dense mball_subset_mcball mcball_subset_mball_concentric order_trans)
+
+lemma openin_mtopology_mcball:
+  "openin mtopology U \<longleftrightarrow> U \<subseteq> M \<and> (\<forall>x. x \<in> U \<longrightarrow> (\<exists>r. 0 < r \<and> mcball x r \<subseteq> U))"
+  using mball_iff_mcball openin_mtopology by presburger
+
+lemma metric_derived_set_of:
+  "mtopology derived_set_of S = {x \<in> M. \<forall>r>0. \<exists>y\<in>S. y\<noteq>x \<and> y \<in> mball x r}" (is "?lhs=?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    unfolding openin_mtopology derived_set_of_def 
+    by clarsimp (metis in_mball openin_mball openin_mtopology zero)
+  show "?rhs \<subseteq> ?lhs"
+    unfolding openin_mtopology derived_set_of_def 
+    by clarify (metis subsetD topspace_mtopology)
+qed
+
+lemma metric_closure_of:
+   "mtopology closure_of S = {x \<in> M. \<forall>r>0. \<exists>y \<in> S. y \<in> mball x r}"
+proof -
+  have "\<And>x r. \<lbrakk>0 < r; x \<in> mtopology closure_of S\<rbrakk> \<Longrightarrow> \<exists>y\<in>S. y \<in> mball x r"
+    by (metis centre_in_mball_iff in_closure_of openin_mball topspace_mtopology)
+  moreover have "\<And>x T. \<lbrakk>x \<in> M; \<forall>r>0. \<exists>y\<in>S. y \<in> mball x r\<rbrakk> \<Longrightarrow> x \<in> mtopology closure_of S"
+    by (smt (verit) in_closure_of in_mball openin_mtopology subsetD topspace_mtopology)
+  ultimately show ?thesis
+    by (auto simp: in_closure_of)
+qed
+
+lemma metric_closure_of_alt:
+  "mtopology closure_of S = {x \<in> M. \<forall>r>0. \<exists>y \<in> S. y \<in> mcball x r}"
+proof -
+  have "\<And>x r. \<lbrakk>\<forall>r>0. x \<in> M \<and> (\<exists>y\<in>S. y \<in> mcball x r); 0 < r\<rbrakk> \<Longrightarrow> \<exists>y\<in>S. y \<in> M \<and> d x y < r"
+    by (meson dense in_mcball le_less_trans)
+  then show ?thesis
+    by (fastforce simp: metric_closure_of in_closure_of)
+qed
+
+lemma metric_interior_of:
+   "mtopology interior_of S = {x \<in> M. \<exists>\<epsilon>>0. mball x \<epsilon> \<subseteq> S}" (is "?lhs=?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    using interior_of_maximal_eq openin_mtopology by fastforce
+  show "?rhs \<subseteq> ?lhs"
+    using interior_of_def openin_mball by fastforce
+qed
+
+lemma metric_interior_of_alt:
+   "mtopology interior_of S = {x \<in> M. \<exists>\<epsilon>>0. mcball x \<epsilon> \<subseteq> S}"
+  by (fastforce simp: mball_iff_mcball metric_interior_of)
+
+lemma in_interior_of_mball:
+   "x \<in> mtopology interior_of S \<longleftrightarrow> x \<in> M \<and> (\<exists>\<epsilon>>0. mball x \<epsilon> \<subseteq> S)"
+  using metric_interior_of by force
+
+lemma in_interior_of_mcball:
+   "x \<in> mtopology interior_of S \<longleftrightarrow> x \<in> M \<and> (\<exists>\<epsilon>>0. mcball x \<epsilon> \<subseteq> S)"
+  using metric_interior_of_alt by force
+
+lemma Hausdorff_space_mtopology: "Hausdorff_space mtopology"
+  unfolding Hausdorff_space_def
+proof clarify
+  fix x y
+  assume x: "x \<in> topspace mtopology" and y: "y \<in> topspace mtopology" and "x \<noteq> y"
+  then have gt0: "d x y / 2 > 0"
+    by auto
+  have "disjnt (mball x (d x y / 2)) (mball y (d x y / 2))"
+    by (simp add: disjoint_mball)
+  then show "\<exists>U V. openin mtopology U \<and> openin mtopology V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V"
+    by (metis centre_in_mball_iff gt0 openin_mball topspace_mtopology x y)
+qed
+
+
+
+subsection\<open>Bounded sets\<close>
+
+definition mbounded where "mbounded S \<longleftrightarrow> (\<exists>x B. S \<subseteq> mcball x B)"
+
+lemma mbounded_pos: "mbounded S \<longleftrightarrow> (\<exists>x B. 0 < B \<and> S \<subseteq> mcball x B)"
+proof -
+  have "\<exists>x' r'. 0 < r' \<and> S \<subseteq> mcball x' r'" if "S \<subseteq> mcball x r" for x r
+    by (metis gt_ex less_eq_real_def linorder_not_le mcball_subset_concentric order_trans that)
+  then show ?thesis
+    by (auto simp: mbounded_def)
+qed
+
+lemma mbounded_alt:
+  "mbounded S \<longleftrightarrow> S \<subseteq> M \<and> (\<exists>B. \<forall>x \<in> S. \<forall>y \<in> S. d x y \<le> B)"
+proof -
+  have "\<And>x B. S \<subseteq> mcball x B \<Longrightarrow> \<forall>x\<in>S. \<forall>y\<in>S. d x y \<le> 2 * B"
+    by (smt (verit, best) commute in_mcball subsetD triangle)
+  then show ?thesis
+    apply (auto simp: mbounded_def subset_iff)
+     apply blast+
+    done
+qed
+
+
+lemma mbounded_alt_pos:
+  "mbounded S \<longleftrightarrow> S \<subseteq> M \<and> (\<exists>B>0. \<forall>x \<in> S. \<forall>y \<in> S. d x y \<le> B)"
+  by (smt (verit, del_insts) gt_ex mbounded_alt)
+
+lemma mbounded_subset: "\<lbrakk>mbounded T; S \<subseteq> T\<rbrakk> \<Longrightarrow> mbounded S"
+  by (meson mbounded_def order_trans)
+
+lemma mbounded_subset_mspace: "mbounded S \<Longrightarrow> S \<subseteq> M"
+  by (simp add: mbounded_alt)
+
+lemma mbounded:
+   "mbounded S \<longleftrightarrow> S = {} \<or> (\<forall>x \<in> S. x \<in> M) \<and> (\<exists>y B. y \<in> M \<and> (\<forall>x \<in> S. d y x \<le> B))"
+  by (meson all_not_in_conv in_mcball mbounded_def subset_iff)
+
+lemma mbounded_empty [iff]: "mbounded {}"
+  by (simp add: mbounded)
+
+lemma mbounded_mcball: "mbounded (mcball x r)"
+  using mbounded_def by auto
+
+lemma mbounded_mball [iff]: "mbounded (mball x r)"
+  by (meson mball_subset_mcball mbounded_def)
+
+lemma mbounded_insert: "mbounded (insert a S) \<longleftrightarrow> a \<in> M \<and> mbounded S"
+proof -
+  have "\<And>y B. \<lbrakk>y \<in> M; \<forall>x\<in>S. d y x \<le> B\<rbrakk>
+           \<Longrightarrow> \<exists>y. y \<in> M \<and> (\<exists>B \<ge> d y a. \<forall>x\<in>S. d y x \<le> B)"
+    by (metis order.trans nle_le)
+  then show ?thesis
+    by (auto simp: mbounded)
+qed
+
+lemma mbounded_Int: "mbounded S \<Longrightarrow> mbounded (S \<inter> T)"
+  by (meson inf_le1 mbounded_subset)
+
+lemma mbounded_Un: "mbounded (S \<union> T) \<longleftrightarrow> mbounded S \<and> mbounded T" (is "?lhs=?rhs")
+proof
+  assume R: ?rhs
+  show ?lhs
+  proof (cases "S={} \<or> T={}")
+    case True then show ?thesis
+      using R by auto
+  next
+    case False
+    obtain x y B C where "S \<subseteq> mcball x B" "T \<subseteq> mcball y C" "B > 0" "C > 0" "x \<in> M" "y \<in> M"
+      using R mbounded_pos
+      by (metis False mcball_eq_empty subset_empty)
+    then have "S \<union> T \<subseteq> mcball x (B + C + d x y)"
+      by (smt (verit) commute dual_order.trans le_supI mcball_subset mdist_pos_eq)
+    then show ?thesis
+      using mbounded_def by blast
+  qed
+next
+  show "?lhs \<Longrightarrow> ?rhs"
+    using mbounded_def by auto
+qed
+
+lemma mbounded_Union:
+  "\<lbrakk>finite \<F>; \<And>X. X \<in> \<F> \<Longrightarrow> mbounded X\<rbrakk> \<Longrightarrow> mbounded (\<Union>\<F>)"
+  by (induction \<F> rule: finite_induct) (auto simp: mbounded_Un)
+
+lemma mbounded_closure_of:
+   "mbounded S \<Longrightarrow> mbounded (mtopology closure_of S)"
+  by (meson closedin_mcball closure_of_minimal mbounded_def)
+
+lemma mbounded_closure_of_eq:
+   "S \<subseteq> M \<Longrightarrow> (mbounded (mtopology closure_of S) \<longleftrightarrow> mbounded S)"
+  by (metis closure_of_subset mbounded_closure_of mbounded_subset topspace_mtopology)
+
+
+lemma maxdist_thm:
+  assumes "mbounded S"
+      and "x \<in> S"
+      and "y \<in> S"
+    shows  "d x y = (SUP z\<in>S. \<bar>d x z - d z y\<bar>)"
+proof -
+  have "\<bar>d x z - d z y\<bar> \<le> d x y" if "z \<in> S" for z
+    by (metis all_not_in_conv assms mbounded mdist_reverse_triangle that) 
+  moreover have "d x y \<le> r"
+    if "\<And>z. z \<in> S \<Longrightarrow> \<bar>d x z - d z y\<bar> \<le> r" for r :: real
+    using that assms mbounded_subset_mspace mdist_zero by fastforce
+  ultimately show ?thesis
+    by (intro cSup_eq [symmetric]) auto
+qed
+
+
+lemma metric_eq_thm: "\<lbrakk>S \<subseteq> M; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> (x = y) = (\<forall>z\<in>S. d x z = d y z)"
+  by (metis commute  subset_iff zero)
+
+lemma compactin_imp_mbounded:
+  assumes "compactin mtopology S"
+  shows "mbounded S"
+proof -
+  have "S \<subseteq> M"
+    and com: "\<And>\<U>. \<lbrakk>\<forall>U\<in>\<U>. openin mtopology U; S \<subseteq> \<Union>\<U>\<rbrakk> \<Longrightarrow> \<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> S \<subseteq> \<Union>\<F>"
+    using assms by (auto simp: compactin_def mbounded_def)
+  show ?thesis
+  proof (cases "S = {}")
+    case False
+    with \<open>S \<subseteq> M\<close> obtain a where "a \<in> S" "a \<in> M"
+      by blast
+    with \<open>S \<subseteq> M\<close> gt_ex have "S \<subseteq> \<Union>(range (mball a))"
+      by force
+    moreover have "\<forall>U \<in> range (mball a). openin mtopology U"
+      by (simp add: openin_mball)
+    ultimately obtain \<F> where "finite \<F>" "\<F> \<subseteq> range (mball a)" "S \<subseteq> \<Union>\<F>"
+      by (meson com)
+  then show ?thesis
+      using mbounded_Union mbounded_subset by fastforce
+  qed auto
+qed
+
+end
+
+
+subsection\<open>Subspace of a metric space\<close>
+
+locale submetric = Metric_space + 
+  fixes A
+  assumes subset: "A \<subseteq> M"
+
+sublocale submetric \<subseteq> sub: Metric_space A d
+  by (simp add: subset subspace)
+
+context submetric
+begin 
+
+lemma mball_submetric_eq: "sub.mball a r = (if a \<in> A then A \<inter> mball a r else {})"
+and mcball_submetric_eq: "sub.mcball a r = (if a \<in> A then A \<inter> mcball a r else {})"
+  using subset by force+
+
+lemma mtopology_submetric: "sub.mtopology = subtopology mtopology A"
+  unfolding topology_eq
+proof (intro allI iffI)
+  fix S
+  assume "openin sub.mtopology S"
+  then have "\<exists>T. openin (subtopology mtopology A) T \<and> x \<in> T \<and> T \<subseteq> S" if "x \<in> S" for x
+    by (metis mball_submetric_eq openin_mball openin_subtopology_Int2 sub.centre_in_mball_iff sub.openin_mtopology subsetD that)
+  then show "openin (subtopology mtopology A) S"
+    by (meson openin_subopen)
+next
+  fix S
+  assume "openin (subtopology mtopology A) S"
+  then obtain T where "openin mtopology T" "S = T \<inter> A"
+    by (meson openin_subtopology)
+  then have "mopen T"
+    by (simp add: mopen_def openin_mtopology)
+  then have "sub.mopen (T \<inter> A)"
+    unfolding sub.mopen_def mopen_def
+    by (metis inf.coboundedI2 mball_submetric_eq Int_iff \<open>S = T \<inter> A\<close> inf.bounded_iff subsetI)
+  then show "openin sub.mtopology S"
+    using \<open>S = T \<inter> A\<close> sub.mopen_def sub.openin_mtopology by force
+qed
+
+lemma mbounded_submetric: "sub.mbounded T \<longleftrightarrow> mbounded T \<and> T \<subseteq> A"
+  by (meson mbounded_alt sub.mbounded_alt subset subset_trans)
+
+end
+
+lemma (in Metric_space) submetric_empty [iff]: "submetric M d {}"
+  by (simp add: Metric_space_axioms submetric.intro submetric_axioms_def)
+
+
+subsection\<open>The discrete metric\<close>
+
+locale discrete_metric =
+  fixes M :: "'a set"
+
+definition (in discrete_metric) dd :: "'a \<Rightarrow> 'a \<Rightarrow> real"
+  where "dd \<equiv> \<lambda>x y::'a. if x=y then 0 else 1"
+
+lemma metric_M_dd: "Metric_space M discrete_metric.dd"
+  by (simp add: discrete_metric.dd_def Metric_space.intro)
+
+sublocale discrete_metric \<subseteq> disc: Metric_space M dd
+  by (simp add: metric_M_dd)
+
+
+lemma (in discrete_metric) mopen_singleton:
+  assumes "x \<in> M" shows "disc.mopen {x}"
+proof -
+  have "disc.mball x (1/2) \<subseteq> {x}"
+    by (smt (verit) dd_def disc.in_mball less_divide_eq_1_pos singleton_iff subsetI)
+  with assms show ?thesis
+    using disc.mopen_def half_gt_zero_iff zero_less_one by blast
+qed
+
+lemma (in discrete_metric) mtopology_discrete_metric:
+   "disc.mtopology = discrete_topology M"
+proof -
+  have "\<And>x. x \<in> M \<Longrightarrow> openin disc.mtopology {x}"
+    by (simp add: disc.mtopology_def mopen_singleton)
+  then show ?thesis
+    by (metis disc.topspace_mtopology discrete_topology_unique)
+qed
+
+lemma (in discrete_metric) discrete_ultrametric:
+   "dd x z \<le> max (dd x y) (dd y z)"
+  by (simp add: dd_def)
+
+
+lemma (in discrete_metric) dd_le1: "dd x y \<le> 1"
+  by (simp add: dd_def)
+
+lemma (in discrete_metric) mbounded_discrete_metric: "disc.mbounded S \<longleftrightarrow> S \<subseteq> M"
+  by (meson dd_le1 disc.mbounded_alt)
+
+
+
+subsection\<open>Metrizable spaces\<close>
+
+definition metrizable_space where
+  "metrizable_space X \<equiv> \<exists>M d. Metric_space M d \<and> X = Metric_space.mtopology M d"
+
+lemma (in Metric_space) metrizable_space_mtopology: "metrizable_space mtopology"
+  using local.Metric_space_axioms metrizable_space_def by blast
+
+lemma openin_mtopology_eq_open [simp]: "openin Met_TC.mtopology = open"
+  by (simp add: Met_TC.mtopology_def)
+
+lemma closedin_mtopology_eq_closed [simp]: "closedin Met_TC.mtopology = closed"
+proof -
+  have "(euclidean::'a topology) = Met_TC.mtopology"
+    by (simp add: Met_TC.mtopology_def)
+  then show ?thesis
+    using closed_closedin by fastforce
+qed
+
+lemma compactin_mtopology_eq_compact [simp]: "compactin Met_TC.mtopology = compact"
+  by (simp add: compactin_def compact_eq_Heine_Borel fun_eq_iff) meson
+
+lemma metrizable_space_discrete_topology:
+   "metrizable_space(discrete_topology U)"
+  by (metis discrete_metric.mtopology_discrete_metric metric_M_dd metrizable_space_def)
+
+lemma metrizable_space_subtopology:
+  assumes "metrizable_space X"
+  shows "metrizable_space(subtopology X S)"
+proof -
+  obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d"
+    using assms metrizable_space_def by blast
+  then interpret submetric M d "M \<inter> S"
+    by (simp add: submetric.intro submetric_axioms_def)
+  show ?thesis
+    unfolding metrizable_space_def
+    by (metis X mtopology_submetric sub.Metric_space_axioms subtopology_restrict topspace_mtopology)
+qed
+
+lemma homeomorphic_metrizable_space_aux:
+  assumes "X homeomorphic_space Y" "metrizable_space X"
+  shows "metrizable_space Y"
+proof -
+  obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d"
+    using assms by (auto simp: metrizable_space_def)
+  then interpret m: Metric_space M d 
+    by simp
+  obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g"
+    and fg: "(\<forall>x \<in> M. 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
+  define d' where "d' x y \<equiv> d (g x) (g y)" for x y
+  interpret m': Metric_space "topspace Y" "d'"
+    unfolding d'_def
+  proof
+    show "(d (g x) (g y) = 0) = (x = y)" if "x \<in> topspace Y" "y \<in> topspace Y" for x y
+      by (metis fg X hmg homeomorphic_imp_surjective_map imageI m.topspace_mtopology m.zero that)
+    show "d (g x) (g z) \<le> d (g x) (g y) + d (g y) (g z)"
+      if "x \<in> topspace Y" and "y \<in> topspace Y" and "z \<in> topspace Y" for x y z
+      by (metis X that hmg homeomorphic_eq_everything_map imageI m.topspace_mtopology m.triangle)
+  qed (auto simp: m.nonneg m.commute)
+  have "Y = Metric_space.mtopology (topspace Y) d'"
+    unfolding topology_eq
+  proof (intro allI)
+    fix S
+    have "openin m'.mtopology S" if S: "S \<subseteq> topspace Y" and "openin X (g ` S)"
+      unfolding m'.openin_mtopology
+    proof (intro conjI that strip)
+      fix y
+      assume "y \<in> S"
+      then obtain r where "r>0" and r: "m.mball (g y) r \<subseteq> g ` S" 
+        using X \<open>openin X (g ` S)\<close> m.openin_mtopology using \<open>y \<in> S\<close> by auto
+      then have "g ` m'.mball y r \<subseteq> m.mball (g y) r"
+        using X d'_def hmg homeomorphic_imp_surjective_map by fastforce
+      with S fg have "m'.mball y r \<subseteq> S"
+        by (smt (verit, del_insts) image_iff m'.in_mball r subset_iff)
+      then show "\<exists>r>0. m'.mball y r \<subseteq> S"
+        using \<open>0 < r\<close> by blast 
+    qed
+    moreover have "openin X (g ` S)" if ope': "openin m'.mtopology S"
+    proof -
+      have "\<exists>r>0. m.mball (g y) r \<subseteq> g ` S" if "y \<in> S" for y
+      proof -
+        have y: "y \<in> topspace Y"
+          using m'.openin_mtopology ope' that by blast
+        obtain r where "r > 0" and r: "m'.mball y r \<subseteq> S"
+          using ope' by (meson \<open>y \<in> S\<close> m'.openin_mtopology)
+        moreover have "\<And>x. \<lbrakk>x \<in> M; d (g y) x < r\<rbrakk> \<Longrightarrow> \<exists>u. u \<in> topspace Y \<and> d' y u < r \<and> x = g u"
+          using fg X d'_def hmf homeomorphic_imp_surjective_map by fastforce
+        ultimately have "m.mball (g y) r \<subseteq> g ` m'.mball y r"
+          using y by (force simp: m'.openin_mtopology)
+        then show ?thesis
+          using \<open>0 < r\<close> r by blast
+      qed
+      then show ?thesis
+        using X hmg homeomorphic_imp_surjective_map m.openin_mtopology ope' openin_subset by fastforce
+    qed
+    ultimately have "(S \<subseteq> topspace Y \<and> openin X (g ` S)) = openin m'.mtopology S"
+      using m'.topspace_mtopology openin_subset by blast
+    then show "openin Y S = openin m'.mtopology S"
+      by (simp add: m'.mopen_def homeomorphic_map_openness_eq [OF hmg])
+  qed
+  then show ?thesis
+    using m'.metrizable_space_mtopology by force
+qed
+
+lemma homeomorphic_metrizable_space:
+  assumes "X homeomorphic_space Y"
+  shows "metrizable_space X \<longleftrightarrow> metrizable_space Y"
+  using assms homeomorphic_metrizable_space_aux homeomorphic_space_sym by metis
+
+lemma metrizable_space_retraction_map_image:
+   "retraction_map X Y r \<and> metrizable_space X
+        \<Longrightarrow> metrizable_space Y"
+  using hereditary_imp_retractive_property metrizable_space_subtopology homeomorphic_metrizable_space
+  by blast
+
+
+lemma metrizable_imp_Hausdorff_space:
+   "metrizable_space X \<Longrightarrow> Hausdorff_space X"
+  by (metis Metric_space.Hausdorff_space_mtopology metrizable_space_def)
+
+(**
+lemma metrizable_imp_kc_space:
+   "metrizable_space X \<Longrightarrow> kc_space X"
+oops
+  MESON_TAC[METRIZABLE_IMP_HAUSDORFF_SPACE; HAUSDORFF_IMP_KC_SPACE]);;
+
+lemma kc_space_mtopology:
+   "kc_space mtopology"
+oops
+  REWRITE_TAC[GSYM FORALL_METRIZABLE_SPACE; METRIZABLE_IMP_KC_SPACE]);;
+**)
+
+lemma metrizable_imp_t1_space:
+   "metrizable_space X \<Longrightarrow> t1_space X"
+  by (simp add: Hausdorff_imp_t1_space metrizable_imp_Hausdorff_space)
+
+lemma closed_imp_gdelta_in:
+  assumes X: "metrizable_space X" and S: "closedin X S"
+  shows "gdelta_in X S"
+proof -
+  obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d"
+    using X metrizable_space_def by blast
+  then interpret M: Metric_space M d
+    by blast
+  have "S \<subseteq> M"
+    using M.closedin_metric \<open>X = M.mtopology\<close> S by blast
+  show ?thesis
+  proof (cases "S = {}")
+    case True
+    then show ?thesis
+      by simp
+  next
+    case False
+    have "\<exists>y\<in>S. d x y < inverse (1 + real n)" if "x \<in> S" for x n
+      using \<open>S \<subseteq> M\<close> M.mdist_zero [of x] that by force
+    moreover
+    have "x \<in> S" if "x \<in> M" and \<section>: "\<And>n. \<exists>y\<in>S. d x y < inverse(Suc n)" for x
+    proof -
+      have *: "\<exists>y\<in>S. d x y < \<epsilon>" if "\<epsilon> > 0" for \<epsilon>
+        by (metis \<section> that not0_implies_Suc order_less_le order_less_le_trans real_arch_inverse)
+      have "closedin M.mtopology S"
+        using S by (simp add: Xeq)
+      then show ?thesis
+        apply (simp add: M.closedin_metric)
+        by (metis * \<open>x \<in> M\<close> M.in_mball disjnt_insert1 insert_absorb subsetD)
+    qed
+    ultimately have Seq: "S = \<Inter>(range (\<lambda>n. {x\<in>M. \<exists>y\<in>S. d x y < inverse(Suc n)}))"
+      using \<open>S \<subseteq> M\<close> by force
+    have "openin M.mtopology {xa \<in> M. \<exists>y\<in>S. d xa y < inverse (1 + real n)}" for n
+    proof (clarsimp simp: M.openin_mtopology)
+      fix x y
+      assume "x \<in> M" "y \<in> S" and dxy: "d x y < inverse (1 + real n)"
+      then have "\<And>z. \<lbrakk>z \<in> M; d x z < inverse (1 + real n) - d x y\<rbrakk> \<Longrightarrow> \<exists>y\<in>S. d z y < inverse (1 + real n)"
+        by (smt (verit) M.commute M.triangle \<open>S \<subseteq> M\<close> in_mono)
+      with dxy show "\<exists>r>0. M.mball x r \<subseteq> {z \<in> M. \<exists>y\<in>S. d z y < inverse (1 + real n)}"
+        by (rule_tac x="inverse(Suc n) - d x y" in exI) auto
+    qed
+    then show ?thesis
+      apply (subst Seq)
+      apply (force simp: Xeq intro: gdelta_in_Inter open_imp_gdelta_in)
+      done
+  qed
+qed
+
+lemma open_imp_fsigma_in:
+   "\<lbrakk>metrizable_space X; openin X S\<rbrakk> \<Longrightarrow> fsigma_in X S"
+  by (meson closed_imp_gdelta_in fsigma_in_gdelta_in openin_closedin openin_subset)
+
+(*NEEDS first_countable
+lemma first_countable_mtopology:
+   "first_countable mtopology"
+oops
+  GEN_TAC THEN REWRITE_TAC[first_countable; TOPSPACE_MTOPOLOGY] THEN
+  X_GEN_TAC `x::A` THEN DISCH_TAC THEN
+  EXISTS_TAC `{ mball m (x::A,r) | rational r \<and> 0 < r}` THEN
+  REWRITE_TAC[FORALL_IN_GSPEC; OPEN_IN_MBALL; EXISTS_IN_GSPEC] THEN
+  ONCE_REWRITE_TAC[SET_RULE
+   `{f x | S x \<and> Q x} = f ` {x \<in> S. Q x}`] THEN
+  SIMP_TAC[COUNTABLE_IMAGE; COUNTABLE_RATIONAL; COUNTABLE_RESTRICT] THEN
+  REWRITE_TAC[OPEN_IN_MTOPOLOGY] THEN
+  X_GEN_TAC `U::A=>bool` THEN STRIP_TAC THEN
+  FIRST_X_ASSUM(MP_TAC \<circ> SPEC `x::A`) THEN
+  ASM_REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+  X_GEN_TAC `r::real` THEN STRIP_TAC THEN FIRST_ASSUM
+   (MP_TAC \<circ> SPEC `r::real` \<circ> MATCH_MP RATIONAL_APPROXIMATION_BELOW) THEN
+  MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `q::real` THEN
+  REWRITE_TAC[REAL_SUB_REFL] THEN STRIP_TAC THEN
+  ASM_SIMP_TAC[CENTRE_IN_MBALL] THEN
+  TRANS_TAC SUBSET_TRANS `mball m (x::A,r)` THEN
+  ASM_SIMP_TAC[MBALL_SUBSET_CONCENTRIC; REAL_LT_IMP_LE]);;
+
+lemma metrizable_imp_first_countable:
+   "metrizable_space X \<Longrightarrow> first_countable X"
+oops
+  REWRITE_TAC[FORALL_METRIZABLE_SPACE; FIRST_COUNTABLE_MTOPOLOGY]);;
+*)
+
+lemma mball_eq_ball [simp]: "Met_TC.mball = ball"
+  by force
+
+lemma mopen_eq_open [simp]: "Met_TC.mopen = open"
+  by (force simp: open_contains_ball Met_TC.mopen_def)
+
+lemma metrizable_space_euclidean:
+  "metrizable_space (euclidean :: 'a::metric_space topology)"
+  unfolding metrizable_space_def
+  by (metis Met_TC.Metric_space_axioms Met_TC.mtopology_def mopen_eq_open)
+
+lemma (in Metric_space) regular_space_mtopology:
+   "regular_space mtopology"
+unfolding regular_space_def
+proof clarify
+  fix C a
+  assume C: "closedin mtopology C" and a: "a \<in> topspace mtopology" and "a \<notin> C"
+  have "openin mtopology (topspace mtopology - C)"
+    by (simp add: C openin_diff)
+  then obtain r where "r>0" and r: "mball a r \<subseteq> topspace mtopology - C"
+    unfolding openin_mtopology using \<open>a \<notin> C\<close> a by auto
+  show "\<exists>U V. openin mtopology U \<and> openin mtopology V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V"
+  proof (intro exI conjI)
+    show "a \<in> mball a (r/2)"
+      using \<open>0 < r\<close> a by force
+    show "C \<subseteq> topspace mtopology - mcball a (r/2)"
+      using C \<open>0 < r\<close> r by (fastforce simp: closedin_metric)
+  qed (auto simp: openin_mball closedin_mcball openin_diff disjnt_iff)
+qed
+
+lemma metrizable_imp_regular_space:
+   "metrizable_space X \<Longrightarrow> regular_space X"
+  by (metis Metric_space.regular_space_mtopology metrizable_space_def)
+
+lemma regular_space_euclidean:
+ "regular_space (euclidean :: 'a::metric_space topology)"
+  by (simp add: metrizable_imp_regular_space metrizable_space_euclidean)
+
+
+subsection\<open>Limits at a point in a topological space\<close>
+
+lemma (in Metric_space) eventually_atin_metric:
+   "eventually P (atin mtopology a) \<longleftrightarrow>
+        (a \<in> M \<longrightarrow> (\<exists>\<delta>>0. \<forall>x. x \<in> M \<and> 0 < d x a \<and> d x a < \<delta> \<longrightarrow> P x))"  (is "?lhs=?rhs")
+proof (cases "a \<in> M")
+  case True
+  show ?thesis
+  proof
+    assume L: ?lhs 
+    with True obtain U where "openin mtopology U" "a \<in> U" and U: "\<forall>x\<in>U - {a}. P x"
+      by (auto simp: eventually_atin)
+    then obtain r where "r>0" and "mball a r \<subseteq> U"
+      by (meson openin_mtopology)
+    with U show ?rhs
+      by (smt (verit, ccfv_SIG) commute in_mball insert_Diff_single insert_iff subset_iff)
+  next
+    assume ?rhs 
+    then obtain \<delta> where "\<delta>>0" and \<delta>: "\<forall>x. x \<in> M \<and> 0 < d x a \<and> d x a < \<delta> \<longrightarrow> P x"
+      using True by blast
+    then have "\<forall>x \<in> mball a \<delta> - {a}. P x"
+      by (simp add: commute)
+    then show ?lhs
+      unfolding eventually_atin openin_mtopology
+      by (metis True \<open>0 < \<delta>\<close> centre_in_mball_iff openin_mball openin_mtopology) 
+  qed
+qed auto
+
+subsection \<open>Normal spaces and metric spaces\<close>
+
+lemma (in Metric_space) normal_space_mtopology:
+   "normal_space mtopology"
+  unfolding normal_space_def
+proof clarify
+  fix S T
+  assume "closedin mtopology S"
+  then have "\<And>x. x \<in> M - S \<Longrightarrow> (\<exists>r>0. mball x r \<subseteq> M - S)"
+    by (simp add: closedin_def openin_mtopology)
+  then obtain \<delta> where d0: "\<And>x. x \<in> M - S \<Longrightarrow> \<delta> x > 0 \<and> mball x (\<delta> x) \<subseteq> M - S"
+    by metis
+  assume "closedin mtopology T"
+  then have "\<And>x. x \<in> M - T \<Longrightarrow> (\<exists>r>0. mball x r \<subseteq> M - T)"
+    by (simp add: closedin_def openin_mtopology)
+  then obtain \<epsilon> where e: "\<And>x. x \<in> M - T \<Longrightarrow> \<epsilon> x > 0 \<and> mball x (\<epsilon> x) \<subseteq> M - T"
+    by metis
+  assume "disjnt S T"
+  have "S \<subseteq> M" "T \<subseteq> M"
+    using \<open>closedin mtopology S\<close> \<open>closedin mtopology T\<close> closedin_metric by blast+
+  have \<delta>: "\<And>x. x \<in> T \<Longrightarrow> \<delta> x > 0 \<and> mball x (\<delta> x) \<subseteq> M - S"
+    by (meson DiffI \<open>T \<subseteq> M\<close> \<open>disjnt S T\<close> d0 disjnt_iff subsetD)
+  have \<epsilon>: "\<And>x. x \<in> S \<Longrightarrow> \<epsilon> x > 0 \<and> mball x (\<epsilon> x) \<subseteq> M - T"
+    by (meson Diff_iff \<open>S \<subseteq> M\<close> \<open>disjnt S T\<close> disjnt_iff e subsetD)
+  show "\<exists>U V. openin mtopology U \<and> openin mtopology V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
+  proof (intro exI conjI)
+    show "openin mtopology (\<Union>x\<in>S. mball x (\<epsilon> x / 2))" "openin mtopology (\<Union>x\<in>T. mball x (\<delta> x / 2))"
+      by force+
+    show "S \<subseteq> (\<Union>x\<in>S. mball x (\<epsilon> x / 2))"
+      using \<epsilon> \<open>S \<subseteq> M\<close> by force
+    show "T \<subseteq> (\<Union>x\<in>T. mball x (\<delta> x / 2))"
+      using \<delta> \<open>T \<subseteq> M\<close> by force
+    show "disjnt (\<Union>x\<in>S. mball x (\<epsilon> x / 2)) (\<Union>x\<in>T. mball x (\<delta> x / 2))"
+      using \<epsilon> \<delta> 
+      apply (clarsimp simp: disjnt_iff subset_iff)
+      by (smt (verit, ccfv_SIG) field_sum_of_halves triangle')
+  qed 
+qed
+
+lemma metrizable_imp_normal_space:
+   "metrizable_space X \<Longrightarrow> normal_space X"
+  by (metis Metric_space.normal_space_mtopology metrizable_space_def)
+
+subsection\<open>Topological limitin in metric spaces\<close>
+
+
+lemma (in Metric_space) limitin_mspace:
+   "limitin mtopology f l F \<Longrightarrow> l \<in> M"
+  using limitin_topspace by fastforce
+
+lemma (in Metric_space) limitin_metric_unique:
+   "\<lbrakk>limitin mtopology f l1 F; limitin mtopology f l2 F; F \<noteq> bot\<rbrakk> \<Longrightarrow> l1 = l2"
+  by (meson Hausdorff_space_mtopology limitin_Hausdorff_unique)
+
+lemma (in Metric_space) limitin_metric:
+   "limitin mtopology f l F \<longleftrightarrow> l \<in> M \<and> (\<forall>\<epsilon>>0. eventually (\<lambda>x. f x \<in> M \<and> d (f x) l < \<epsilon>) F)"  
+   (is "?lhs=?rhs")
+proof
+  assume L: ?lhs
+  show ?rhs
+    unfolding limitin_def
+  proof (intro conjI strip)
+    show "l \<in> M"
+      using L limitin_mspace by blast
+    fix \<epsilon>::real
+    assume "\<epsilon>>0"
+    then have "\<forall>\<^sub>F x in F. f x \<in> mball l \<epsilon>"
+      using L openin_mball by (fastforce simp: limitin_def)
+    then show "\<forall>\<^sub>F x in F. f x \<in> M \<and> d (f x) l < \<epsilon>"
+      using commute eventually_mono by fastforce
+  qed
+next
+  assume R: ?rhs 
+  then show ?lhs
+    by (force simp: limitin_def commute openin_mtopology subset_eq elim: eventually_mono)
+qed
+
+lemma (in Metric_space) limit_metric_sequentially:
+   "limitin mtopology f l sequentially \<longleftrightarrow>
+     l \<in> M \<and> (\<forall>\<epsilon>>0. \<exists>N. \<forall>n\<ge>N. f n \<in> M \<and> d (f n) l < \<epsilon>)"
+  by (auto simp: limitin_metric eventually_sequentially)
+
+lemma (in submetric) limitin_submetric_iff:
+   "limitin sub.mtopology f l F \<longleftrightarrow>
+     l \<in> A \<and> eventually (\<lambda>x. f x \<in> A) F \<and> limitin mtopology f l F" (is "?lhs=?rhs")
+  by (simp add: limitin_subtopology mtopology_submetric)
+
+lemma (in Metric_space) metric_closedin_iff_sequentially_closed:
+   "closedin mtopology S \<longleftrightarrow>
+    S \<subseteq> M \<and> (\<forall>\<sigma> l. range \<sigma> \<subseteq> S \<and> limitin mtopology \<sigma> l sequentially \<longrightarrow> l \<in> S)" (is "?lhs=?rhs")
+proof
+  assume ?lhs then show ?rhs
+    by (force simp: closedin_metric limitin_closedin range_subsetD)
+next
+  assume R: ?rhs
+  show ?lhs
+    unfolding closedin_metric
+  proof (intro conjI strip)
+    show "S \<subseteq> M"
+      using R by blast
+    fix x
+    assume "x \<in> M - S"
+    have False if "\<forall>r>0. \<exists>y. y \<in> M \<and> y \<in> S \<and> d x y < r"
+    proof -
+      have "\<forall>n. \<exists>y. y \<in> M \<and> y \<in> S \<and> d x y < inverse(Suc n)"
+        using that by auto
+      then obtain \<sigma> where \<sigma>: "\<And>n. \<sigma> n \<in> M \<and> \<sigma> n \<in> S \<and> d x (\<sigma> n) < inverse(Suc n)"
+        by metis
+      then have "range \<sigma> \<subseteq> M"
+        by blast
+      have "\<exists>N. \<forall>n\<ge>N. d x (\<sigma> n) < \<epsilon>" if "\<epsilon>>0" for \<epsilon>
+      proof -
+        have "real (Suc (nat \<lceil>inverse \<epsilon>\<rceil>)) \<ge> inverse \<epsilon>"
+          by linarith
+        then have "\<forall>n \<ge> nat \<lceil>inverse \<epsilon>\<rceil>. d x (\<sigma> n) < \<epsilon>"
+          by (metis \<sigma> inverse_inverse_eq inverse_le_imp_le nat_ceiling_le_eq nle_le not_less_eq_eq order.strict_trans2 that)
+        then show ?thesis ..
+      qed
+      with \<sigma> have "limitin mtopology \<sigma> x sequentially"
+        using \<open>x \<in> M - S\<close> commute limit_metric_sequentially by auto
+      then show ?thesis
+        by (metis R DiffD2 \<sigma> image_subset_iff \<open>x \<in> M - S\<close>)
+    qed
+    then show "\<exists>r>0. disjnt S (mball x r)"
+      by (meson disjnt_iff in_mball)
+  qed
+qed
+
+lemma (in Metric_space) limit_atin_metric:
+   "limitin X f y (atin mtopology x) \<longleftrightarrow>
+      y \<in> topspace X \<and>
+      (x \<in> M
+       \<longrightarrow> (\<forall>V. openin X V \<and> y \<in> V
+               \<longrightarrow> (\<exists>\<delta>>0.  \<forall>x'. x' \<in> M \<and> 0 < d x' x \<and> d x' x < \<delta> \<longrightarrow> f x' \<in> V)))"
+  by (force simp: limitin_def eventually_atin_metric)
+
+lemma (in Metric_space) limitin_metric_dist_null:
+   "limitin mtopology f l F \<longleftrightarrow> l \<in> M \<and> eventually (\<lambda>x. f x \<in> M) F \<and> ((\<lambda>x. d (f x) l) \<longlongrightarrow> 0) F"
+  by (simp add: limitin_metric tendsto_iff eventually_conj_iff all_conj_distrib imp_conjR gt_ex)
+
+
+subsection\<open>Cauchy sequences and complete metric spaces\<close>
+
+context Metric_space
+begin
+
+definition MCauchy :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool"
+  where "MCauchy \<sigma> \<equiv> range \<sigma> \<subseteq> M \<and> (\<forall>\<epsilon>>0. \<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>)"
+
+definition mcomplete
+  where "mcomplete \<equiv> (\<forall>\<sigma>. MCauchy \<sigma> \<longrightarrow> (\<exists>x. limitin mtopology \<sigma> x sequentially))"
+
+lemma mcomplete_empty [iff]: "Metric_space.mcomplete {} d"
+  by (simp add: Metric_space.MCauchy_def Metric_space.mcomplete_def subspace)
+
+lemma MCauchy_imp_MCauchy_suffix: "MCauchy \<sigma> \<Longrightarrow> MCauchy (\<sigma> \<circ> (+)n)"
+  unfolding MCauchy_def image_subset_iff comp_apply
+  by (metis UNIV_I add.commute trans_le_add1) 
+
+lemma mcomplete:
+   "mcomplete \<longleftrightarrow>
+    (\<forall>\<sigma>. (\<forall>\<^sub>F n in sequentially. \<sigma> n \<in> M) \<and>
+     (\<forall>\<epsilon>>0. \<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>) \<longrightarrow>
+     (\<exists>x. limitin mtopology \<sigma> x sequentially))" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+  proof clarify
+    fix \<sigma>
+    assume "\<forall>\<^sub>F n in sequentially. \<sigma> n \<in> M"
+      and \<sigma>: "\<forall>\<epsilon>>0. \<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>"
+    then obtain N where "\<And>n. n\<ge>N \<Longrightarrow> \<sigma> n \<in> M"
+      by (auto simp: eventually_sequentially)
+    with \<sigma> have "MCauchy (\<sigma> \<circ> (+)N)"
+      unfolding MCauchy_def image_subset_iff comp_apply by (meson le_add1 trans_le_add2)
+    then obtain x where "limitin mtopology (\<sigma> \<circ> (+)N) x sequentially"
+      using L MCauchy_imp_MCauchy_suffix mcomplete_def by blast
+    then have "limitin mtopology \<sigma> x sequentially"
+      unfolding o_def by (auto simp: add.commute limitin_sequentially_offset_rev)
+    then show "\<exists>x. limitin mtopology \<sigma> x sequentially" ..
+  qed
+qed (simp add: mcomplete_def MCauchy_def image_subset_iff)
+
+lemma mcomplete_empty_mspace: "M = {} \<Longrightarrow> mcomplete"
+  using MCauchy_def mcomplete_def by blast
+
+lemma MCauchy_const [simp]: "MCauchy (\<lambda>n. a) \<longleftrightarrow> a \<in> M"
+  using MCauchy_def mdist_zero by auto
+
+lemma convergent_imp_MCauchy:
+  assumes "range \<sigma> \<subseteq> M" and lim: "limitin mtopology \<sigma> l sequentially"
+  shows "MCauchy \<sigma>"
+  unfolding MCauchy_def image_subset_iff
+proof (intro conjI strip)
+  fix \<epsilon>::real
+  assume "\<epsilon> > 0"
+  then have "\<forall>\<^sub>F n in sequentially. \<sigma> n \<in> M \<and> d (\<sigma> n) l < \<epsilon>/2"
+    using half_gt_zero lim limitin_metric by blast
+  then obtain N where "\<And>n. n\<ge>N \<Longrightarrow> \<sigma> n \<in> M \<and> d (\<sigma> n) l < \<epsilon>/2"
+    by (force simp: eventually_sequentially)
+  then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>"
+    by (smt (verit) Metric_space.limitin_mspace Metric_space.mdist_reverse_triangle Metric_space_axioms field_sum_of_halves lim)
+qed (use assms in blast)
+
+
+lemma mcomplete_alt:
+   "mcomplete \<longleftrightarrow> (\<forall>\<sigma>. MCauchy \<sigma> \<longleftrightarrow> range \<sigma> \<subseteq> M \<and> (\<exists>x. limitin mtopology \<sigma> x sequentially))"
+  using MCauchy_def convergent_imp_MCauchy mcomplete_def by blast
+
+lemma MCauchy_subsequence:
+  assumes "strict_mono r" "MCauchy \<sigma>"
+  shows "MCauchy (\<sigma> \<circ> r)"
+proof -
+  have "d (\<sigma> (r n)) (\<sigma> (r n')) < \<epsilon>"
+    if "N \<le> n" "N \<le> n'" "strict_mono r" "\<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>"
+    for \<epsilon> N n n'
+    using that by (meson le_trans strict_mono_imp_increasing)
+  moreover have "range (\<lambda>x. \<sigma> (r x)) \<subseteq> M"
+    using MCauchy_def assms by blast
+  ultimately show ?thesis
+    using assms by (simp add: MCauchy_def) metis
+qed
+
+lemma MCauchy_offset:
+  assumes cau: "MCauchy (\<sigma> \<circ> (+)k)" and \<sigma>: "\<And>n. n < k \<Longrightarrow> \<sigma> n \<in> M" 
+  shows "MCauchy \<sigma>"
+  unfolding MCauchy_def image_subset_iff
+proof (intro conjI strip)
+  fix n
+  show "\<sigma> n \<in> M"
+    using assms
+    unfolding MCauchy_def image_subset_iff
+    by (metis UNIV_I comp_apply le_iff_add linorder_not_le)
+next
+  fix \<epsilon> :: real
+  assume "\<epsilon> > 0"
+  obtain N where "\<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d ((\<sigma> \<circ> (+)k) n) ((\<sigma> \<circ> (+)k) n') < \<epsilon>"
+    using cau \<open>\<epsilon> > 0\<close> by (fastforce simp: MCauchy_def)
+  then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>"
+    unfolding o_def
+    apply (rule_tac x="k+N" in exI)
+    by (smt (verit, del_insts) add.assoc le_add1 less_eqE)
+qed
+
+lemma MCauchy_convergent_subsequence:
+  assumes cau: "MCauchy \<sigma>" and "strict_mono r" 
+     and lim: "limitin mtopology (\<sigma> \<circ> r) a sequentially"
+  shows "limitin mtopology \<sigma> a sequentially"
+  unfolding limitin_metric
+proof (intro conjI strip)
+  show "a \<in> M"
+    by (meson assms limitin_mspace)
+  fix \<epsilon> :: real
+  assume "\<epsilon> > 0"
+  then obtain N1 where N1: "\<And>n n'. \<lbrakk>n\<ge>N1; n'\<ge>N1\<rbrakk> \<Longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>/2"
+    using cau unfolding MCauchy_def by (meson half_gt_zero)
+  obtain N2 where N2: "\<And>n. n \<ge> N2 \<Longrightarrow> (\<sigma> \<circ> r) n \<in> M \<and> d ((\<sigma> \<circ> r) n) a < \<epsilon>/2"
+    by (metis (no_types, lifting) lim \<open>\<epsilon> > 0\<close> half_gt_zero limit_metric_sequentially)
+  have "\<sigma> n \<in> M \<and> d (\<sigma> n) a < \<epsilon>" if "n \<ge> max N1 N2" for n
+  proof (intro conjI)
+    show "\<sigma> n \<in> M"
+      using MCauchy_def cau by blast
+    have "N1 \<le> r n"
+      by (meson \<open>strict_mono r\<close> le_trans max.cobounded1 strict_mono_imp_increasing that)
+    then show "d (\<sigma> n) a < \<epsilon>"
+      using N1[of n "r n"] N2[of n] \<open>\<sigma> n \<in> M\<close> \<open>a \<in> M\<close> triangle that by fastforce
+  qed
+  then show "\<forall>\<^sub>F n in sequentially. \<sigma> n \<in> M \<and> d (\<sigma> n) a < \<epsilon>"
+    using eventually_sequentially by blast
+qed
+
+lemma MCauchy_interleaving_gen:
+  "MCauchy (\<lambda>n. if even n then x(n div 2) else y(n div 2)) \<longleftrightarrow>
+    (MCauchy x \<and> MCauchy y \<and> (\<lambda>n. d (x n) (y n)) \<longlonglongrightarrow> 0)" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs
+  have evens: "strict_mono (\<lambda>n::nat. 2 * n)" and odds: "strict_mono (\<lambda>n::nat. Suc (2 * n))"
+    by (auto simp: strict_mono_def)
+  show ?rhs
+  proof (intro conjI)
+    show "MCauchy x" "MCauchy y"
+      using MCauchy_subsequence [OF evens L] MCauchy_subsequence [OF odds L] by (auto simp: o_def)
+    show "(\<lambda>n. d (x n) (y n)) \<longlonglongrightarrow> 0"
+      unfolding LIMSEQ_iff
+    proof (intro strip)
+      fix \<epsilon> :: real
+      assume "\<epsilon> > 0"
+      then obtain N where N: 
+        "\<And>n n'. \<lbrakk>n\<ge>N; n'\<ge>N\<rbrakk> \<Longrightarrow> d (if even n then x (n div 2) else y (n div 2))
+                                   (if even n' then x (n' div 2) else y (n' div 2))  < \<epsilon>"
+        using L MCauchy_def by fastforce
+      have "d (x n) (y n) < \<epsilon>" if "n\<ge>N" for n
+        using N [of "2*n" "Suc(2*n)"] that by auto
+      then show "\<exists>N. \<forall>n\<ge>N. norm (d (x n) (y n) - 0) < \<epsilon>"
+        by auto
+    qed
+  qed
+next
+  assume R: ?rhs
+  show ?lhs
+    unfolding MCauchy_def
+  proof (intro conjI strip)
+    show "range (\<lambda>n. if even n then x (n div 2) else y (n div 2)) \<subseteq> M"
+      using R by (auto simp: MCauchy_def)
+    fix \<epsilon> :: real
+    assume "\<epsilon> > 0"
+    obtain Nx where Nx: "\<And>n n'. \<lbrakk>n\<ge>Nx; n'\<ge>Nx\<rbrakk> \<Longrightarrow> d (x n) (x n')  < \<epsilon>/2"
+      by (meson half_gt_zero MCauchy_def R \<open>\<epsilon> > 0\<close>)
+    obtain Ny where Ny: "\<And>n n'. \<lbrakk>n\<ge>Ny; n'\<ge>Ny\<rbrakk> \<Longrightarrow> d (y n) (y n')  < \<epsilon>/2"
+      by (meson half_gt_zero MCauchy_def R \<open>\<epsilon> > 0\<close>)
+    obtain Nxy where Nxy: "\<And>n. n\<ge>Nxy \<Longrightarrow> d (x n) (y n) < \<epsilon>/2"
+      using R \<open>\<epsilon> > 0\<close> half_gt_zero unfolding LIMSEQ_iff
+      by (metis abs_mdist diff_zero real_norm_def)
+    define N where "N \<equiv> 2 * Max{Nx,Ny,Nxy}"
+    show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \<epsilon>"
+    proof (intro exI strip)
+      fix n n'
+      assume "N \<le> n" and "N \<le> n'"
+      then have "n div 2 \<ge> Nx" "n div 2 \<ge> Ny" "n div 2 \<ge> Nxy" "n' div 2 \<ge> Nx" "n' div 2 \<ge> Ny" 
+        by (auto simp: N_def)
+      then have dxyn: "d (x (n div 2)) (y (n div 2)) < \<epsilon>/2" 
+            and dxnn': "d (x (n div 2)) (x (n' div 2)) < \<epsilon>/2"
+            and dynn': "d (y (n div 2)) (y (n' div 2)) < \<epsilon>/2"
+        using Nx Ny Nxy by blast+
+      have inM: "x (n div 2) \<in> M" "x (n' div 2) \<in> M""y (n div 2) \<in> M" "y (n' div 2) \<in> M"
+        using Metric_space.MCauchy_def Metric_space_axioms R by blast+
+      show "d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \<epsilon>"
+      proof (cases "even n")
+        case nt: True
+        show ?thesis
+        proof (cases "even n'")
+          case True
+          with \<open>\<epsilon> > 0\<close> nt dxnn' show ?thesis by auto
+        next
+          case False
+          with nt dxyn dynn' inM triangle show ?thesis
+            by fastforce
+        qed
+      next
+        case nf: False
+        show ?thesis 
+        proof (cases "even n'")
+          case True
+          then show ?thesis
+            by (smt (verit) \<open>\<epsilon> > 0\<close> dxyn dxnn' triangle commute inM field_sum_of_halves)
+        next
+          case False
+          with \<open>\<epsilon> > 0\<close> nf dynn' show ?thesis by auto
+        qed
+      qed
+    qed
+  qed
+qed
+
+lemma MCauchy_interleaving:
+   "MCauchy (\<lambda>n. if even n then \<sigma>(n div 2) else a) \<longleftrightarrow>
+    range \<sigma> \<subseteq> M \<and> limitin mtopology \<sigma> a sequentially"  (is "?lhs=?rhs")
+proof -
+  have "?lhs \<longleftrightarrow> (MCauchy \<sigma> \<and> a \<in> M \<and> (\<lambda>n. d (\<sigma> n) a) \<longlonglongrightarrow> 0)"
+    by (simp add: MCauchy_interleaving_gen [where y = "\<lambda>n. a"])
+  also have "... = ?rhs"
+    by (metis MCauchy_def always_eventually convergent_imp_MCauchy limitin_metric_dist_null range_subsetD)
+  finally show ?thesis .
+qed
+
+lemma mcomplete_nest:
+   "mcomplete \<longleftrightarrow>
+      (\<forall>C::nat \<Rightarrow>'a set. (\<forall>n. closedin mtopology (C n)) \<and>
+          (\<forall>n. C n \<noteq> {}) \<and> decseq C \<and> (\<forall>\<epsilon>>0. \<exists>n a. C n \<subseteq> mcball a \<epsilon>)
+          \<longrightarrow> \<Inter> (range C) \<noteq> {})" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+    unfolding imp_conjL
+  proof (intro strip)
+    fix C :: "nat \<Rightarrow> 'a set"
+    assume clo: "\<forall>n. closedin mtopology (C n)"
+      and ne: "\<forall>n. C n \<noteq> ({}::'a set)"
+      and dec: "decseq C"
+      and cover [rule_format]: "\<forall>\<epsilon>>0. \<exists>n a. C n \<subseteq> mcball a \<epsilon>"
+    obtain \<sigma> where \<sigma>: "\<And>n. \<sigma> n \<in> C n"
+      by (meson ne empty_iff set_eq_iff)
+    have "MCauchy \<sigma>"
+      unfolding MCauchy_def
+    proof (intro conjI strip)
+      show "range \<sigma> \<subseteq> M"
+        using \<sigma> clo metric_closedin_iff_sequentially_closed by auto 
+      fix \<epsilon> :: real
+      assume "\<epsilon> > 0"
+      then obtain N a where N: "C N \<subseteq> mcball a (\<epsilon>/3)"
+        using cover by fastforce
+      have "d (\<sigma> m) (\<sigma> n) < \<epsilon>" if "N \<le> m" "N \<le> n" for m n
+      proof -
+        have "d a (\<sigma> m) \<le> \<epsilon>/3" "d a (\<sigma> n) \<le> \<epsilon>/3"
+          using dec N \<sigma> that by (fastforce simp: decseq_def)+
+        then have "d (\<sigma> m) (\<sigma> n) \<le> \<epsilon>/3 + \<epsilon>/3"
+          using triangle \<sigma> commute dec decseq_def subsetD that N
+          by (smt (verit, ccfv_threshold) in_mcball)
+        also have "... < \<epsilon>"
+          using \<open>\<epsilon> > 0\<close> by auto
+        finally show ?thesis .
+      qed
+      then show "\<exists>N. \<forall>m n. N \<le> m \<longrightarrow> N \<le> n \<longrightarrow> d (\<sigma> m) (\<sigma> n) < \<epsilon>"
+        by blast
+    qed
+    then obtain x where x: "limitin mtopology \<sigma> x sequentially"
+      using L mcomplete_def by blast
+    have "x \<in> C n" for n
+    proof (rule limitin_closedin [OF x])
+      show "closedin mtopology (C n)"
+        by (simp add: clo)
+      show "\<forall>\<^sub>F x in sequentially. \<sigma> x \<in> C n"
+        by (metis \<sigma> dec decseq_def eventually_sequentiallyI subsetD)
+    qed auto
+    then show "\<Inter> (range C) \<noteq> {}"
+      by blast
+qed
+next
+  assume R: ?rhs  
+  show ?lhs
+    unfolding mcomplete_def
+  proof (intro strip)
+    fix \<sigma>
+    assume "MCauchy \<sigma>"
+    then have "range \<sigma> \<subseteq> M"
+      using MCauchy_def by blast
+    define C where "C \<equiv> \<lambda>n. mtopology closure_of (\<sigma> ` {n..})"
+    have "\<forall>n. closedin mtopology (C n)" 
+      by (auto simp: C_def)
+    moreover
+    have ne: "\<And>n. C n \<noteq> {}"
+      using \<open>MCauchy \<sigma>\<close> by (auto simp: C_def MCauchy_def disjnt_iff closure_of_eq_empty_gen)
+    moreover
+    have dec: "decseq C"
+      unfolding monotone_on_def
+    proof (intro strip)
+      fix m n::nat
+      assume "m \<le> n"
+      then have "{n..} \<subseteq> {m..}"
+        by auto
+      then show "C n \<subseteq> C m"
+        unfolding C_def by (meson closure_of_mono image_mono)
+    qed
+    moreover
+    have C: "\<exists>N u. C N \<subseteq> mcball u \<epsilon>" if "\<epsilon>>0" for \<epsilon>
+    proof -
+      obtain N where "\<And>m n. N \<le> m \<and> N \<le> n \<Longrightarrow> d (\<sigma> m) (\<sigma> n) < \<epsilon>"
+        by (meson MCauchy_def \<open>0 < \<epsilon>\<close> \<open>MCauchy \<sigma>\<close>)
+      then have "\<sigma> ` {N..} \<subseteq> mcball (\<sigma> N) \<epsilon>"
+        using MCauchy_def \<open>MCauchy \<sigma>\<close> by (force simp: less_eq_real_def)
+      then have "C N \<subseteq> mcball (\<sigma> N) \<epsilon>"
+        by (simp add: C_def closure_of_minimal)
+      then show ?thesis
+        by blast
+    qed
+    ultimately obtain l where x: "l \<in> \<Inter> (range C)"
+      by (metis R ex_in_conv)
+    then have *: "\<And>\<epsilon> N. 0 < \<epsilon> \<Longrightarrow> \<exists>n'. N \<le> n' \<and> l \<in> M \<and> \<sigma> n' \<in> M \<and> d l (\<sigma> n') < \<epsilon>"
+      by (force simp: C_def metric_closure_of)
+    then have "l \<in> M"
+      using gt_ex by blast
+    show "\<exists>l. limitin mtopology \<sigma> l sequentially"
+      unfolding limitin_metric
+    proof (intro conjI strip exI)
+      show "l \<in> M"
+        using \<open>\<forall>n. closedin mtopology (C n)\<close> closedin_subset x by fastforce
+      fix \<epsilon>::real
+      assume "\<epsilon> > 0"
+      obtain N where N: "\<And>m n. N \<le> m \<and> N \<le> n \<Longrightarrow> d (\<sigma> m) (\<sigma> n) < \<epsilon>/2"
+        by (meson MCauchy_def \<open>0 < \<epsilon>\<close> \<open>MCauchy \<sigma>\<close> half_gt_zero)
+      with * [of "\<epsilon>/2" N]
+      have "\<forall>n\<ge>N. \<sigma> n \<in> M \<and> d (\<sigma> n) l < \<epsilon>"
+        by (smt (verit) \<open>range \<sigma> \<subseteq> M\<close> commute field_sum_of_halves range_subsetD triangle)
+      then show "\<forall>\<^sub>F n in sequentially. \<sigma> n \<in> M \<and> d (\<sigma> n) l < \<epsilon>"
+        using eventually_sequentially by blast
+    qed
+  qed
+qed
+
+
+lemma mcomplete_nest_sing:
+   "mcomplete \<longleftrightarrow>
+    (\<forall>C. (\<forall>n. closedin mtopology (C n)) \<and>
+          (\<forall>n. C n \<noteq> {}) \<and> decseq C \<and> (\<forall>e>0. \<exists>n a. C n \<subseteq> mcball a e)
+         \<longrightarrow> (\<exists>l. l \<in> M \<and> \<Inter> (range C) = {l}))"
+proof -
+  have *: False
+    if clo: "\<forall>n. closedin mtopology (C n)"
+      and cover: "\<forall>\<epsilon>>0. \<exists>n a. C n \<subseteq> mcball a \<epsilon>"
+      and no_sing: "\<And>y. y \<in> M \<Longrightarrow> \<Inter> (range C) \<noteq> {y}"
+      and l: "\<forall>n. l \<in> C n"
+    for C :: "nat \<Rightarrow> 'a set" and l
+  proof -
+    have inM: "\<And>x. x \<in> \<Inter> (range C) \<Longrightarrow> x \<in> M"
+      using closedin_metric clo by fastforce
+    then have "l \<in> M"
+      by (simp add: l)
+    have False if l': "l' \<in> \<Inter> (range C)" and "l' \<noteq> l" for l'
+    proof -
+      have "l' \<in> M"
+        using inM l' by blast
+      obtain n a where na: "C n \<subseteq> mcball a (d l l' / 3)"
+        using inM \<open>l \<in> M\<close> l' \<open>l' \<noteq> l\<close> cover by force
+      then have "d a l \<le> (d l l' / 3)" "d a l' \<le> (d l l' / 3)" "a \<in> M"
+        using l l' na in_mcball by auto
+      then have "d l l' \<le> (d l l' / 3) + (d l l' / 3)"
+        using \<open>l \<in> M\<close> \<open>l' \<in> M\<close> mdist_reverse_triangle by fastforce
+      then show False
+        using nonneg [of l l'] \<open>l' \<noteq> l\<close> \<open>l \<in> M\<close> \<open>l' \<in> M\<close> zero by force
+    qed
+    then show False
+      by (metis l \<open>l \<in> M\<close> no_sing INT_I empty_iff insertI1 is_singletonE is_singletonI')
+  qed
+  show ?thesis
+    unfolding mcomplete_nest imp_conjL 
+    apply (intro all_cong1 imp_cong refl)
+    using * 
+    by (smt (verit) Inter_iff ex_in_conv range_constant range_eqI)
+qed
+
+lemma mcomplete_fip:
+   "mcomplete \<longleftrightarrow>
+    (\<forall>\<C>. (\<forall>C \<in> \<C>. closedin mtopology C) \<and>
+         (\<forall>e>0. \<exists>C a. C \<in> \<C> \<and> C \<subseteq> mcball a e) \<and> (\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<C> \<longrightarrow> \<Inter> \<F> \<noteq> {})
+        \<longrightarrow> \<Inter> \<C> \<noteq> {})" 
+   (is "?lhs = ?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+    unfolding mcomplete_nest_sing imp_conjL
+  proof (intro strip)
+    fix \<C> :: "'a set set"
+    assume clo: "\<forall>C\<in>\<C>. closedin mtopology C"
+      and cover: "\<forall>e>0. \<exists>C a. C \<in> \<C> \<and> C \<subseteq> mcball a e"
+      and fip: "\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> \<C> \<longrightarrow> \<Inter> \<F> \<noteq> {}"
+    then have "\<forall>n. \<exists>C. C \<in> \<C> \<and> (\<exists>a. C \<subseteq> mcball a (inverse (Suc n)))"
+      by simp
+    then obtain C where C: "\<And>n. C n \<in> \<C>" 
+          and coverC: "\<And>n. \<exists>a. C n \<subseteq> mcball a (inverse (Suc n))"
+      by metis
+    define D where "D \<equiv> \<lambda>n. \<Inter> (C ` {..n})"
+    have cloD: "closedin mtopology (D n)" for n
+      unfolding D_def using clo C by blast
+    have neD: "D n \<noteq> {}" for n
+      using fip C by (simp add: D_def image_subset_iff)
+    have decD: "decseq D"
+      by (force simp: D_def decseq_def)
+    have coverD: "\<exists>n a. D n \<subseteq> mcball a \<epsilon>" if " \<epsilon> >0" for \<epsilon>
+    proof -
+      obtain n where "inverse (Suc n) < \<epsilon>"
+        using \<open>0 < \<epsilon>\<close> reals_Archimedean by blast
+      then obtain a where "C n \<subseteq> mcball a \<epsilon>"
+        by (meson coverC less_eq_real_def mcball_subset_concentric order_trans)
+      then show ?thesis
+        unfolding D_def by blast
+    qed
+    have *: "a \<in> \<Inter>\<C>" if a: "\<Inter> (range D) = {a}" and "a \<in> M" for a
+    proof -
+      have aC: "a \<in> C n" for n
+        using that by (auto simp: D_def)
+      have eqa: "\<And>u. (\<forall>n. u \<in> C n) \<Longrightarrow> a = u"
+        using that by (auto simp: D_def)
+      have "a \<in> T" if "T \<in> \<C>" for T
+      proof -
+        have cloT: "closedin mtopology (T \<inter> D n)" for n
+          using clo cloD that by blast
+        have "\<Inter> (insert T (C ` {..n})) \<noteq> {}" for n
+          using that C by (intro fip [rule_format]) auto
+        then have neT: "T \<inter> D n \<noteq> {}" for n
+          by (simp add: D_def)
+        have decT: "decseq (\<lambda>n. T \<inter> D n)"
+          by (force simp: D_def decseq_def)
+        have coverT: "\<exists>n a. T \<inter> D n \<subseteq> mcball a \<epsilon>" if " \<epsilon> >0" for \<epsilon>
+          by (meson coverD le_infI2 that)
+        show ?thesis
+          using L [unfolded mcomplete_nest_sing, rule_format, of "\<lambda>n. T \<inter> D n"] a
+          by (force simp: cloT neT decT coverT)
+      qed
+      then show ?thesis by auto
+    qed
+    show "\<Inter> \<C> \<noteq> {}"
+      by (metis L cloD neD decD coverD * empty_iff mcomplete_nest_sing)
+  qed
+next
+  assume R [rule_format]: ?rhs
+  show ?lhs
+    unfolding mcomplete_nest imp_conjL
+  proof (intro strip)
+    fix C :: "nat \<Rightarrow> 'a set"
+    assume clo: "\<forall>n. closedin mtopology (C n)"
+      and ne: "\<forall>n. C n \<noteq> {}"
+      and dec: "decseq C"
+      and cover: "\<forall>\<epsilon>>0. \<exists>n a. C n \<subseteq> mcball a \<epsilon>"
+    have "\<Inter>(C ` N) \<noteq> {}" if "finite N" for N
+    proof -
+      obtain k where "N \<subseteq> {..k}"
+        using \<open>finite N\<close> finite_nat_iff_bounded_le by auto
+      with dec have "C k \<subseteq> \<Inter>(C ` N)" by (auto simp: decseq_def)
+      then show ?thesis
+        using ne by force
+    qed
+    with clo cover R [of "range C"] show "\<Inter> (range C) \<noteq> {}"
+      by (metis (no_types, opaque_lifting) finite_subset_image image_iff UNIV_I)
+  qed
+qed
+
+
+lemma mcomplete_fip_sing:
+   "mcomplete \<longleftrightarrow>
+    (\<forall>\<C>. (\<forall>C\<in>\<C>. closedin mtopology C) \<and>
+     (\<forall>e>0. \<exists>c a. c \<in> \<C> \<and> c \<subseteq> mcball a e) \<and>
+     (\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<C> \<longrightarrow> \<Inter> \<F> \<noteq> {}) \<longrightarrow>
+     (\<exists>l. l \<in> M \<and> \<Inter> \<C> = {l}))"
+   (is "?lhs = ?rhs")
+proof
+  have *: "l \<in> M" "\<Inter> \<C> = {l}"
+    if clo: "Ball \<C> (closedin mtopology)"
+      and cover: "\<forall>e>0. \<exists>C a. C \<in> \<C> \<and> C \<subseteq> mcball a e"
+      and fin: "\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> \<C> \<longrightarrow> \<Inter> \<F> \<noteq> {}"
+      and l: "l \<in> \<Inter> \<C>"
+    for \<C> :: "'a set set" and l
+  proof -
+    show "l \<in> M"
+      by (meson Inf_lower2 clo cover gt_ex metric_closedin_iff_sequentially_closed subsetD that(4))
+    show  "\<Inter> \<C> = {l}"
+    proof (cases "\<C> = {}")
+      case True
+      then show ?thesis
+        using cover mbounded_pos by auto
+    next
+      case False
+      have CM: "\<And>a. a \<in> \<Inter>\<C> \<Longrightarrow> a \<in> M"
+        using False clo closedin_subset by fastforce
+      have "l' \<notin> \<Inter> \<C>" if "l' \<noteq> l" for l'
+      proof 
+        assume l': "l' \<in> \<Inter> \<C>"
+        with CM have "l' \<in> M" by blast
+        with that \<open>l \<in> M\<close> have gt0: "0 < d l l'"
+          by simp
+        then obtain C a where "C \<in> \<C>" and C: "C \<subseteq> mcball a (d l l' / 3)"
+          using cover [rule_format, of "d l l' / 3"] by auto
+        then have "d a l \<le> (d l l' / 3)" "d a l' \<le> (d l l' / 3)" "a \<in> M"
+          using l l' in_mcball by auto
+        then have "d l l' \<le> (d l l' / 3) + (d l l' / 3)"
+          using \<open>l \<in> M\<close> \<open>l' \<in> M\<close> mdist_reverse_triangle by fastforce
+        with gt0 show False by auto
+      qed
+      then show ?thesis
+        using l by fastforce
+    qed
+  qed
+  assume L: ?lhs
+  with * show ?rhs
+    unfolding mcomplete_fip imp_conjL ex_in_conv [symmetric]
+    by (elim all_forward imp_forward2 asm_rl) (blast intro:  elim: )
+next
+  assume ?rhs then show ?lhs
+    unfolding mcomplete_fip by (force elim!: all_forward)
+qed
+
+end
+
+lemma MCauchy_iff_Cauchy [iff]: "Met_TC.MCauchy = Cauchy"
+  by (force simp: Cauchy_def Met_TC.MCauchy_def)
+
+lemma mcomplete_iff_complete [iff]:
+  "Met_TC.mcomplete (Pure.type ::'a::metric_space itself) \<longleftrightarrow> complete (UNIV::'a set)"
+  by (auto simp: Met_TC.mcomplete_def complete_def)
+
+lemma euclidean_metric: "Met_TC.mcomplete (Pure.type ::'a::euclidean_space itself)"
+  using complete_UNIV mcomplete_iff_complete by blast
+
+context submetric
+begin 
+
+lemma MCauchy_submetric:
+   "sub.MCauchy \<sigma> \<longleftrightarrow> range \<sigma> \<subseteq> A \<and> MCauchy \<sigma>"
+  using MCauchy_def sub.MCauchy_def subset by force
+
+lemma closedin_mcomplete_imp_mcomplete:
+  assumes clo: "closedin mtopology A" and "mcomplete"
+  shows "sub.mcomplete"
+  unfolding sub.mcomplete_def
+proof (intro strip)
+  fix \<sigma>
+  assume "sub.MCauchy \<sigma>"
+  then have \<sigma>: "MCauchy \<sigma>" "range \<sigma> \<subseteq> A"
+    using MCauchy_submetric by blast+
+  then obtain x where x: "limitin mtopology \<sigma> x sequentially"
+    using \<open>mcomplete\<close> unfolding mcomplete_def by blast
+  then have "x \<in> A"
+    using \<sigma> clo metric_closedin_iff_sequentially_closed by force
+  with \<sigma> x show "\<exists>x. limitin sub.mtopology \<sigma> x sequentially"
+    using limitin_submetric_iff range_subsetD by fastforce
+qed
+
+
+lemma sequentially_closedin_mcomplete_imp_mcomplete:
+  assumes "mcomplete" and "\<And>\<sigma> l. range \<sigma> \<subseteq> A \<and> limitin mtopology \<sigma> l sequentially \<Longrightarrow> l \<in> A"
+  shows "sub.mcomplete"
+  using assms closedin_mcomplete_imp_mcomplete metric_closedin_iff_sequentially_closed subset by blast
+
+end
+
+
+context Metric_space
+begin
+
+lemma mcomplete_Un:
+  assumes A: "submetric M d A" "Metric_space.mcomplete A d" 
+      and B: "submetric M d B" "Metric_space.mcomplete B d"
+  shows   "submetric M d (A \<union> B)" "Metric_space.mcomplete (A \<union> B) d" 
+proof -
+  show "submetric M d (A \<union> B)"
+    by (meson assms le_sup_iff submetric_axioms_def submetric_def) 
+  then interpret MAB: Metric_space "A \<union> B" d
+    by (meson submetric.subset subspace)
+  interpret MA: Metric_space A d
+    by (meson A submetric.subset subspace)
+  interpret MB: Metric_space B d
+    by (meson B submetric.subset subspace)
+  show "Metric_space.mcomplete (A \<union> B) d"
+    unfolding MAB.mcomplete_def
+  proof (intro strip)
+    fix \<sigma>
+    assume "MAB.MCauchy \<sigma>"
+    then have "range \<sigma> \<subseteq> A \<union> B"
+      using MAB.MCauchy_def by blast
+    then have "UNIV \<subseteq> \<sigma> -` A \<union> \<sigma> -` B"
+      by blast
+    then consider "infinite (\<sigma> -` A)" | "infinite (\<sigma> -` B)"
+      using finite_subset by auto
+    then show "\<exists>x. limitin MAB.mtopology \<sigma> x sequentially"
+    proof cases
+      case 1
+      then obtain r where "strict_mono r" and r: "\<And>n::nat. r n \<in> \<sigma> -` A"
+        using infinite_enumerate by blast 
+      then have "MA.MCauchy (\<sigma> \<circ> r)"
+        using MA.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence \<open>MAB.MCauchy \<sigma>\<close> by auto
+      with A obtain x where "limitin MA.mtopology (\<sigma> \<circ> r) x sequentially"
+        using MA.mcomplete_def by blast
+      then have "limitin MAB.mtopology (\<sigma> \<circ> r) x sequentially"
+        by (metis MA.limit_metric_sequentially MAB.limit_metric_sequentially UnCI)
+      then show ?thesis
+        using MAB.MCauchy_convergent_subsequence \<open>MAB.MCauchy \<sigma>\<close> \<open>strict_mono r\<close> by blast
+    next
+      case 2
+      then obtain r where "strict_mono r" and r: "\<And>n::nat. r n \<in> \<sigma> -` B"
+        using infinite_enumerate by blast 
+      then have "MB.MCauchy (\<sigma> \<circ> r)"
+        using MB.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence \<open>MAB.MCauchy \<sigma>\<close> by auto
+      with B obtain x where "limitin MB.mtopology (\<sigma> \<circ> r) x sequentially"
+        using MB.mcomplete_def by blast
+      then have "limitin MAB.mtopology (\<sigma> \<circ> r) x sequentially"
+        by (metis MB.limit_metric_sequentially MAB.limit_metric_sequentially UnCI)
+      then show ?thesis
+        using MAB.MCauchy_convergent_subsequence \<open>MAB.MCauchy \<sigma>\<close> \<open>strict_mono r\<close> by blast
+    qed
+  qed
+qed
+
+lemma mcomplete_Union:
+  assumes "finite \<S>"
+    and "\<And>A. A \<in> \<S> \<Longrightarrow> submetric M d A" "\<And>A. A \<in> \<S> \<Longrightarrow> Metric_space.mcomplete A d"
+  shows   "submetric M d (\<Union>\<S>)" "Metric_space.mcomplete (\<Union>\<S>) d" 
+  using assms
+  by (induction rule: finite_induct) (auto simp: mcomplete_Un)
+
+lemma mcomplete_Inter:
+  assumes "finite \<S>" "\<S> \<noteq> {}"
+    and sub: "\<And>A. A \<in> \<S> \<Longrightarrow> submetric M d A" 
+    and comp: "\<And>A. A \<in> \<S> \<Longrightarrow> Metric_space.mcomplete A d"
+  shows "submetric M d (\<Inter>\<S>)" "Metric_space.mcomplete (\<Inter>\<S>) d"
+proof -
+  show "submetric M d (\<Inter>\<S>)"
+    using assms unfolding submetric_def submetric_axioms_def
+    by (metis Inter_lower equals0I inf.orderE le_inf_iff) 
+  then interpret MS: submetric M d "\<Inter>\<S>" 
+    by (meson submetric.subset subspace)
+  show "Metric_space.mcomplete (\<Inter>\<S>) d"
+    unfolding MS.sub.mcomplete_def
+  proof (intro strip)
+    fix \<sigma>
+    assume "MS.sub.MCauchy \<sigma>"
+    then have "range \<sigma> \<subseteq> \<Inter>\<S>"
+      using MS.MCauchy_submetric by blast
+    obtain A where "A \<in> \<S>" and A: "Metric_space.mcomplete A d"
+      using assms by blast
+    then have "range \<sigma> \<subseteq> A"
+      using \<open>range \<sigma> \<subseteq> \<Inter>\<S>\<close> by blast
+    interpret SA: submetric M d A
+      by (meson \<open>A \<in> \<S>\<close> sub submetric.subset subspace)
+    have "MCauchy \<sigma>"
+      using MS.MCauchy_submetric \<open>MS.sub.MCauchy \<sigma>\<close> by blast
+    then obtain x where x: "limitin SA.sub.mtopology \<sigma> x sequentially"
+      by (metis A SA.sub.MCauchy_def SA.sub.mcomplete_alt MCauchy_def \<open>range \<sigma> \<subseteq> A\<close>)
+    show "\<exists>x. limitin MS.sub.mtopology \<sigma> x sequentially"
+      apply (rule_tac x="x" in exI)
+      unfolding MS.limitin_submetric_iff
+    proof (intro conjI)
+      show "x \<in> \<Inter> \<S>"
+      proof clarsimp
+        fix U
+        assume "U \<in> \<S>"
+        interpret SU: submetric M d U 
+          by (meson \<open>U \<in> \<S>\<close> sub submetric.subset subspace)
+        have "range \<sigma> \<subseteq> U"
+          using \<open>U \<in> \<S>\<close> \<open>range \<sigma> \<subseteq> \<Inter> \<S>\<close> by blast
+        moreover have "Metric_space.mcomplete U d"
+          by (simp add: \<open>U \<in> \<S>\<close> comp)
+        ultimately obtain x' where x': "limitin SU.sub.mtopology \<sigma> x' sequentially"
+          using MCauchy_def SU.sub.MCauchy_def SU.sub.mcomplete_alt \<open>MCauchy \<sigma>\<close> by meson 
+        have "x' = x"
+        proof (intro limitin_metric_unique)
+          show "limitin mtopology \<sigma> x' sequentially"
+            by (meson SU.submetric_axioms submetric.limitin_submetric_iff x')
+          show "limitin mtopology \<sigma> x sequentially"
+            by (meson SA.submetric_axioms submetric.limitin_submetric_iff x)
+        qed auto
+        then show "x \<in> U"
+          using SU.sub.limitin_mspace x' by blast
+      qed
+      show "\<forall>\<^sub>F n in sequentially. \<sigma> n \<in> \<Inter>\<S>"
+        by (meson \<open>range \<sigma> \<subseteq> \<Inter> \<S>\<close> always_eventually range_subsetD)
+      show "limitin mtopology \<sigma> x sequentially"
+        by (meson SA.submetric_axioms submetric.limitin_submetric_iff x)
+    qed
+  qed
+qed
+
+
+lemma mcomplete_Int:
+  assumes A: "submetric M d A" "Metric_space.mcomplete A d" 
+      and B: "submetric M d B" "Metric_space.mcomplete B d"
+    shows   "submetric M d (A \<inter> B)" "Metric_space.mcomplete (A \<inter> B) d"
+  using mcomplete_Inter [of "{A,B}"] assms by force+
+
+subsection\<open>Totally bounded subsets of metric spaces\<close>
+
+definition mtotally_bounded 
+  where "mtotally_bounded S \<equiv> \<forall>\<epsilon>>0. \<exists>K. finite K \<and> K \<subseteq> S \<and> S \<subseteq> (\<Union>x\<in>K. mball x \<epsilon>)"
+
+lemma mtotally_bounded_empty [iff]: "mtotally_bounded {}"
+by (simp add: mtotally_bounded_def)
+
+lemma finite_imp_mtotally_bounded:
+   "\<lbrakk>finite S; S \<subseteq> M\<rbrakk> \<Longrightarrow> mtotally_bounded S"
+  by (auto simp: mtotally_bounded_def)
+
+lemma mtotally_bounded_imp_subset: "mtotally_bounded S \<Longrightarrow> S \<subseteq> M"
+  by (force simp: mtotally_bounded_def intro!: zero_less_one)
+
+lemma mtotally_bounded_sing [simp]:
+   "mtotally_bounded {x} \<longleftrightarrow> x \<in> M"
+  by (meson empty_subsetI finite.simps finite_imp_mtotally_bounded insert_subset mtotally_bounded_imp_subset)
+
+lemma mtotally_bounded_Un:
+  assumes  "mtotally_bounded S" "mtotally_bounded T"
+  shows "mtotally_bounded (S \<union> T)"
+proof -
+  have "\<exists>K. finite K \<and> K \<subseteq> S \<union> T \<and> S \<union> T \<subseteq> (\<Union>x\<in>K. mball x e)"
+    if  "e>0" and K: "finite K \<and> K \<subseteq> S \<and> S \<subseteq> (\<Union>x\<in>K. mball x e)"
+      and L: "finite L \<and> L \<subseteq> T \<and> T \<subseteq> (\<Union>x\<in>L. mball x e)" for K L e
+    using that by (rule_tac x="K \<union> L" in exI) auto
+  with assms show ?thesis
+    unfolding mtotally_bounded_def by presburger
+qed
+ 
+lemma mtotally_bounded_Union:
+  assumes "finite f" "\<And>S. S \<in> f \<Longrightarrow> mtotally_bounded S"
+  shows "mtotally_bounded (\<Union>f)"
+  using assms by (induction f) (auto simp: mtotally_bounded_Un)
+
+lemma mtotally_bounded_imp_mbounded:
+  assumes "mtotally_bounded S"
+  shows "mbounded S"
+proof -
+  obtain K where "finite K \<and> K \<subseteq> S \<and> S \<subseteq> (\<Union>x\<in>K. mball x 1)" 
+    using assms by (force simp: mtotally_bounded_def)
+  then show ?thesis
+    by (smt (verit) finite_imageI image_iff mbounded_Union mbounded_mball mbounded_subset)
+qed
+
+
+lemma mtotally_bounded_sequentially:
+  "mtotally_bounded S \<longleftrightarrow>
+    S \<subseteq> M \<and> (\<forall>\<sigma>::nat \<Rightarrow> 'a. range \<sigma> \<subseteq> S \<longrightarrow> (\<exists>r. strict_mono r \<and> MCauchy (\<sigma> \<circ> r)))"
+  (is "_ \<longleftrightarrow> _ \<and> ?rhs")
+proof (cases "S \<subseteq> M")
+  case True
+  show ?thesis
+  proof -
+    { fix \<sigma> :: "nat \<Rightarrow> 'a"                                                            
+      assume L: "mtotally_bounded S" and \<sigma>: "range \<sigma> \<subseteq> S"
+      have "\<exists>j > i. d (\<sigma> i) (\<sigma> j) < 3*\<epsilon>/2 \<and> infinite (\<sigma> -` mball (\<sigma> j) (\<epsilon>/2))"
+        if inf: "infinite (\<sigma> -` mball (\<sigma> i) \<epsilon>)" and "\<epsilon> > 0" for i \<epsilon>
+      proof -
+        obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> (\<Union>x\<in>K. mball x (\<epsilon>/4))"
+          by (metis L mtotally_bounded_def \<open>\<epsilon> > 0\<close> zero_less_divide_iff zero_less_numeral)
+        then have K_imp_ex: "\<And>y. y \<in> S \<Longrightarrow> \<exists>x\<in>K. d x y < \<epsilon>/4"
+          by fastforce
+        have False if "\<forall>x\<in>K. d x (\<sigma> i) < \<epsilon> + \<epsilon>/4 \<longrightarrow> finite (\<sigma> -` mball x (\<epsilon>/4))"
+        proof -
+          have "\<exists>w. w \<in> K \<and> d w (\<sigma> i) < 5 * \<epsilon>/4 \<and> d w (\<sigma> j) < \<epsilon>/4"
+            if "d (\<sigma> i) (\<sigma> j) < \<epsilon>" for j
+          proof -
+            obtain w where w: "d w (\<sigma> j) < \<epsilon>/4" "w \<in> K"
+              using K_imp_ex \<sigma> by blast
+            then have "d w (\<sigma> i) < \<epsilon> + \<epsilon>/4"
+              by (smt (verit, ccfv_SIG) True \<open>K \<subseteq> S\<close> \<sigma> rangeI subset_eq that triangle')
+            with w show ?thesis
+              using in_mball by auto
+          qed
+          then have "(\<sigma> -` mball (\<sigma> i) \<epsilon>) \<subseteq> (\<Union>x\<in>K. if d x (\<sigma> i) < \<epsilon> + \<epsilon>/4 then \<sigma> -` mball x (\<epsilon>/4) else {})"
+            using True \<open>K \<subseteq> S\<close> by force
+          then show False
+            using finite_subset inf \<open>finite K\<close> that by fastforce
+        qed
+        then obtain x where "x \<in> K" and dxi: "d x (\<sigma> i) < \<epsilon> + \<epsilon>/4" and infx: "infinite (\<sigma> -` mball x (\<epsilon>/4))"
+          by blast
+        then obtain j where "j \<in> (\<sigma> -` mball x (\<epsilon>/4)) - {..i}"
+          using bounded_nat_set_is_finite by (meson Diff_infinite_finite finite_atMost)
+        then have "j > i" and dxj: "d x (\<sigma> j) < \<epsilon>/4" 
+          by auto
+        have "(\<sigma> -` mball x (\<epsilon>/4)) \<subseteq> (\<sigma> -` mball y (\<epsilon>/2))" if "d x y < \<epsilon>/4" "y \<in> M" for y
+          using that by (simp add: mball_subset Metric_space_axioms vimage_mono)
+        then have infj: "infinite (\<sigma> -` mball (\<sigma> j) (\<epsilon>/2))"
+          by (meson True \<open>d x (\<sigma> j) < \<epsilon>/4\<close> \<sigma> in_mono infx rangeI finite_subset)
+        have "\<sigma> i \<in> M" "\<sigma> j \<in> M" "x \<in> M"  
+          using True \<open>K \<subseteq> S\<close> \<open>x \<in> K\<close> \<sigma> by force+
+        then have "d (\<sigma> i) (\<sigma> j) \<le> d x (\<sigma> i) + d x (\<sigma> j)"
+          using triangle'' by blast
+        also have "\<dots> < 3*\<epsilon>/2"
+          using dxi dxj by auto
+        finally have "d (\<sigma> i) (\<sigma> j) < 3*\<epsilon>/2" .
+        with \<open>i < j\<close> infj show ?thesis by blast
+      qed
+      then obtain nxt where nxt: "\<And>i \<epsilon>. \<lbrakk>\<epsilon> > 0; infinite (\<sigma> -` mball (\<sigma> i) \<epsilon>)\<rbrakk> \<Longrightarrow> 
+                 nxt i \<epsilon> > i \<and> d (\<sigma> i) (\<sigma> (nxt i \<epsilon>)) < 3*\<epsilon>/2 \<and> infinite (\<sigma> -` mball (\<sigma> (nxt i \<epsilon>)) (\<epsilon>/2))"
+        by metis
+      have "mbounded S"
+        using L by (simp add: mtotally_bounded_imp_mbounded)
+      then obtain B where B: "\<forall>y \<in> S. d (\<sigma> 0) y \<le> B" and "B > 0"
+        by (meson \<sigma> mbounded_alt_pos range_subsetD)
+      define eps where "eps \<equiv> \<lambda>n. (B+1) / 2^n"
+      have [simp]: "eps (Suc n) = eps n / 2" "eps n > 0" for n
+        using \<open>B > 0\<close> by (auto simp: eps_def)
+      have "UNIV \<subseteq> \<sigma> -` mball (\<sigma> 0) (B+1)"
+        using B True \<sigma> unfolding image_iff subset_iff
+        by (smt (verit, best) UNIV_I in_mball vimageI)
+      then have inf0: "infinite (\<sigma> -` mball (\<sigma> 0) (eps 0))"
+        using finite_subset by (auto simp: eps_def)
+      define r where "r \<equiv> rec_nat 0 (\<lambda>n rec. nxt rec (eps n))"
+      have [simp]: "r 0 = 0" "r (Suc n) = nxt (r n) (eps n)" for n
+        by (auto simp: r_def)
+      have \<sigma>rM[simp]: "\<sigma> (r n) \<in> M" for n
+        using True \<sigma> by blast
+      have inf: "infinite (\<sigma> -` mball (\<sigma> (r n)) (eps n))" for n
+      proof (induction n)
+        case 0 then show ?case  
+          by (simp add: inf0)
+      next
+        case (Suc n) then show ?case
+          using nxt [of "eps n" "r n"] by simp
+      qed
+      then have "r (Suc n) > r n" for n
+        by (simp add: nxt)
+      then have "strict_mono r"
+        by (simp add: strict_mono_Suc_iff)
+      have d_less: "d (\<sigma> (r n)) (\<sigma> (r (Suc n))) < 3 * eps n / 2" for n
+        using nxt [OF _ inf] by simp
+      have eps_plus: "eps (k + n) = eps n * (1/2)^k" for k n
+        by (simp add: eps_def power_add field_simps)
+      have *: "d (\<sigma> (r n)) (\<sigma> (r (k + n))) < 3 * eps n" for n k
+      proof -
+        have "d (\<sigma> (r n)) (\<sigma> (r (k+n))) \<le> 3/2 * eps n * (\<Sum>i<k. (1/2)^i)"
+        proof (induction k)
+          case 0 then show ?case 
+            by simp
+        next
+          case (Suc k)
+          have "d (\<sigma> (r n)) (\<sigma> (r (Suc k + n))) \<le> d (\<sigma> (r n)) (\<sigma> (r (k + n))) + d (\<sigma> (r (k + n))) (\<sigma> (r (Suc (k + n))))"
+            by (metis \<sigma>rM add.commute add_Suc_right triangle)
+          with d_less[of "k+n"] Suc show ?case
+            by (simp add: algebra_simps eps_plus)
+        qed
+        also have "\<dots> < 3/2 * eps n * 2"
+          using geometric_sum [of "1/2::real" k] by simp
+        finally show ?thesis by simp
+      qed
+      have "\<exists>N. \<forall>n\<ge>N. \<forall>n'\<ge>N. d (\<sigma> (r n)) (\<sigma> (r n')) < \<epsilon>" if "\<epsilon> > 0" for \<epsilon>
+      proof -
+        define N where "N \<equiv> nat \<lceil>(log 2 (6*(B+1) / \<epsilon>))\<rceil>"
+        have \<section>: "b \<le> 2 ^ nat \<lceil>log 2 b\<rceil>" for b
+          by (smt (verit) less_log_of_power real_nat_ceiling_ge)
+        have N: "6 * eps N \<le> \<epsilon>"
+          using \<section> [of "(6*(B+1) / \<epsilon>)"] that by (auto simp: N_def eps_def field_simps)
+        have "d (\<sigma> (r N)) (\<sigma> (r n)) < 3 * eps N" if "n \<ge> N" for n
+          by (metis * add.commute nat_le_iff_add that)
+        then have "\<forall>n\<ge>N. \<forall>n'\<ge>N. d (\<sigma> (r n)) (\<sigma> (r n')) < 3 * eps N + 3 * eps N"
+          by (smt (verit, best) \<sigma>rM triangle'')
+        with N show ?thesis
+          by fastforce
+      qed
+      then have "MCauchy (\<sigma> \<circ> r)"
+        unfolding MCauchy_def using True \<sigma> by auto
+      then have "\<exists>r. strict_mono r \<and> MCauchy (\<sigma> \<circ> r)"
+        using \<open>strict_mono r\<close> by blast      
+    }
+    moreover
+    { assume R: ?rhs
+      have "mtotally_bounded S"
+        unfolding mtotally_bounded_def
+      proof (intro strip)
+        fix \<epsilon> :: real
+        assume "\<epsilon> > 0"
+        have False if \<section>: "\<And>K. \<lbrakk>finite K; K \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>s\<in>S. s \<notin> (\<Union>x\<in>K. mball x \<epsilon>)"
+        proof -
+          obtain f where f: "\<And>K. \<lbrakk>finite K; K \<subseteq> S\<rbrakk> \<Longrightarrow> f K \<in> S \<and> f K \<notin> (\<Union>x\<in>K. mball x \<epsilon>)"
+            using \<section> by metis
+          define \<sigma> where "\<sigma> \<equiv> wfrec less_than (\<lambda>seq n. f (seq ` {..<n}))"
+          have \<sigma>_eq: "\<sigma> n = f (\<sigma> ` {..<n})" for n
+            by (simp add: cut_apply def_wfrec [OF \<sigma>_def])
+          have [simp]: "\<sigma> n \<in> S" for n
+            using wf_less_than
+          proof (induction n rule: wf_induct_rule)
+            case (less n) with f show ?case
+              by (auto simp: \<sigma>_eq [of n])
+          qed
+          then have "range \<sigma> \<subseteq> S" by blast
+          have \<sigma>: "p < n \<Longrightarrow> \<epsilon> \<le> d (\<sigma> p) (\<sigma> n)" for n p
+            using f[of "\<sigma> ` {..<n}"] True by (fastforce simp: \<sigma>_eq [of n] Ball_def)
+          then obtain r where "strict_mono r" "MCauchy (\<sigma> \<circ> r)"
+            by (meson R \<open>range \<sigma> \<subseteq> S\<close>)
+          with \<open>0 < \<epsilon>\<close> obtain N 
+            where N: "\<And>n n'. \<lbrakk>n\<ge>N; n'\<ge>N\<rbrakk> \<Longrightarrow> d (\<sigma> (r n)) (\<sigma> (r n')) < \<epsilon>"
+            by (force simp: MCauchy_def)
+          show ?thesis
+            using N [of N "Suc (r N)"] \<open>strict_mono r\<close>
+            by (smt (verit) Suc_le_eq \<sigma> le_SucI order_refl strict_mono_imp_increasing)
+        qed
+        then show "\<exists>K. finite K \<and> K \<subseteq> S \<and> S \<subseteq> (\<Union>x\<in>K. mball x \<epsilon>)"
+          by blast
+      qed
+    }
+    ultimately show ?thesis 
+      using True by blast
+  qed
+qed (use mtotally_bounded_imp_subset in auto)
+
+
+lemma mtotally_bounded_subset:
+   "\<lbrakk>mtotally_bounded S; T \<subseteq> S\<rbrakk> \<Longrightarrow> mtotally_bounded T"
+  by (meson mtotally_bounded_sequentially order_trans) 
+
+lemma mtotally_bounded_submetric:
+  assumes "mtotally_bounded S" "S \<subseteq> T" "T \<subseteq> M"
+  shows "Metric_space.mtotally_bounded T d S"
+proof -
+  interpret submetric M d T
+    by (simp add: Metric_space_axioms assms submetric.intro submetric_axioms.intro)
+  show ?thesis
+    using assms
+    unfolding sub.mtotally_bounded_def mtotally_bounded_def
+    by (force simp: subset_iff elim!: all_forward ex_forward)
+qed
+
+lemma mtotally_bounded_absolute:
+   "mtotally_bounded S \<longleftrightarrow> S \<subseteq> M \<and> Metric_space.mtotally_bounded S d S "
+proof -
+  have "mtotally_bounded S" if "S \<subseteq> M" "Metric_space.mtotally_bounded S d S"
+  proof -
+    interpret submetric M d S
+      by (simp add: Metric_space_axioms submetric_axioms.intro submetric_def \<open>S \<subseteq> M\<close>)
+    show ?thesis
+      using that
+      by (metis MCauchy_submetric Metric_space.mtotally_bounded_sequentially Metric_space_axioms subspace)
+  qed
+  moreover have "mtotally_bounded S \<Longrightarrow> Metric_space.mtotally_bounded S d S"
+    by (simp add: mtotally_bounded_imp_subset mtotally_bounded_submetric)
+  ultimately show ?thesis
+    using mtotally_bounded_imp_subset by blast
+qed
+
+lemma mtotally_bounded_closure_of:
+  assumes "mtotally_bounded S"
+  shows "mtotally_bounded (mtopology closure_of S)"
+proof -
+  have "S \<subseteq> M"
+    by (simp add: assms mtotally_bounded_imp_subset)
+  have "mtotally_bounded(mtopology closure_of S)"
+    unfolding mtotally_bounded_def
+  proof (intro strip)
+    fix \<epsilon>::real
+    assume "\<epsilon> > 0"
+    then obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> (\<Union>x\<in>K. mball x (\<epsilon>/2))"
+      by (metis assms mtotally_bounded_def half_gt_zero)
+    have "mtopology closure_of S \<subseteq> (\<Union>x\<in>K. mball x \<epsilon>)"
+      unfolding metric_closure_of
+    proof clarsimp
+      fix x
+      assume "x \<in> M" and x: "\<forall>r>0. \<exists>y\<in>S. y \<in> M \<and> d x y < r"
+      then obtain y where "y \<in> S" and y: "d x y < \<epsilon>/2"
+        using \<open>0 < \<epsilon>\<close> half_gt_zero by blast
+      then obtain x' where "x' \<in> K" "y \<in> mball x' (\<epsilon>/2)"
+        using K by auto
+      then have "d x' x < \<epsilon>/2 + \<epsilon>/2"
+        using triangle y \<open>x \<in> M\<close> commute by fastforce
+      then show "\<exists>x'\<in>K. x' \<in> M \<and> d x' x < \<epsilon>"
+        using \<open>K \<subseteq> S\<close> \<open>S \<subseteq> M\<close> \<open>x' \<in> K\<close> by force
+    qed
+    then show "\<exists>K. finite K \<and> K \<subseteq> mtopology closure_of S \<and> mtopology closure_of S \<subseteq> (\<Union>x\<in>K. mball x \<epsilon>)"
+      using closure_of_subset_Int  \<open>K \<subseteq> S\<close> \<open>finite K\<close> K by fastforce
+  qed
+  then show ?thesis
+    by (simp add: assms inf.absorb2 mtotally_bounded_imp_subset)
+qed
+
+lemma mtotally_bounded_closure_of_eq:
+   "S \<subseteq> M \<Longrightarrow> mtotally_bounded (mtopology closure_of S) \<longleftrightarrow> mtotally_bounded S"
+  by (metis closure_of_subset mtotally_bounded_closure_of mtotally_bounded_subset topspace_mtopology)
+
+lemma mtotally_bounded_cauchy_sequence:
+  assumes "MCauchy \<sigma>"
+  shows "mtotally_bounded (range \<sigma>)"
+  unfolding MCauchy_def mtotally_bounded_def
+proof (intro strip)
+  fix \<epsilon>::real
+  assume "\<epsilon> > 0"
+  then obtain N where "\<And>n. N \<le> n \<Longrightarrow> d (\<sigma> N) (\<sigma> n) < \<epsilon>"
+    using assms by (force simp: MCauchy_def)
+  then have "\<And>m. \<exists>n\<le>N. \<sigma> n \<in> M \<and> \<sigma> m \<in> M \<and> d (\<sigma> n) (\<sigma> m) < \<epsilon>"
+    by (metis MCauchy_def assms mdist_zero nle_le range_subsetD)
+  then
+  show "\<exists>K. finite K \<and> K \<subseteq> range \<sigma> \<and> range \<sigma> \<subseteq> (\<Union>x\<in>K. mball x \<epsilon>)"
+    by (rule_tac x="\<sigma> ` {0..N}" in exI) force
+qed
+
+lemma MCauchy_imp_mbounded:
+   "MCauchy \<sigma> \<Longrightarrow> mbounded (range \<sigma>)"
+  by (simp add: mtotally_bounded_cauchy_sequence mtotally_bounded_imp_mbounded)
+
+subsection\<open>Compactness in metric spaces\<close>
+
+lemma Bolzano_Weierstrass_property:
+  assumes "S \<subseteq> U" "S \<subseteq> M"
+  shows
+   "(\<forall>\<sigma>::nat\<Rightarrow>'a. range \<sigma> \<subseteq> S
+         \<longrightarrow> (\<exists>l r. l \<in> U \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially)) \<longleftrightarrow>
+    (\<forall>T. T \<subseteq> S \<and> infinite T \<longrightarrow> U \<inter> mtopology derived_set_of T \<noteq> {})"  (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+  proof clarify
+    fix T
+    assume "T \<subseteq> S" and "infinite T"
+      and T: "U \<inter> mtopology derived_set_of T = {}"
+    then obtain \<sigma> :: "nat\<Rightarrow>'a" where "inj \<sigma>" "range \<sigma> \<subseteq> T"
+      by (meson infinite_countable_subset)
+    with L obtain l r where "l \<in> U" "strict_mono r" 
+           and lr: "limitin mtopology (\<sigma> \<circ> r) l sequentially"
+      by (meson \<open>T \<subseteq> S\<close> subset_trans)
+    then obtain \<epsilon> where "\<epsilon> > 0" and \<epsilon>: "\<And>y. y \<in> T \<Longrightarrow> y = l \<or> \<not> d l y < \<epsilon>"
+      using T \<open>T \<subseteq> S\<close> \<open>S \<subseteq> M\<close> 
+      by (force simp: metric_derived_set_of limitin_metric disjoint_iff)
+    with lr have "\<forall>\<^sub>F n in sequentially. \<sigma> (r n) \<in> M \<and> d (\<sigma> (r n)) l < \<epsilon>"
+      by (auto simp: limitin_metric)
+    then obtain N where N: "d (\<sigma> (r N)) l < \<epsilon>" "d (\<sigma> (r (Suc N))) l < \<epsilon>"
+      using less_le_not_le by (auto simp: eventually_sequentially)
+    moreover have "\<sigma> (r N) \<noteq> l \<or> \<sigma> (r (Suc N)) \<noteq> l"
+      by (meson \<open>inj \<sigma>\<close> \<open>strict_mono r\<close> injD n_not_Suc_n strict_mono_eq)
+    ultimately
+    show False
+       using \<epsilon> \<open>range \<sigma> \<subseteq> T\<close> commute by fastforce
+  qed
+next
+  assume R: ?rhs 
+  show ?lhs
+  proof (intro strip)
+    fix \<sigma> :: "nat \<Rightarrow> 'a"
+    assume "range \<sigma> \<subseteq> S"
+    show "\<exists>l r. l \<in> U \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially"
+    proof (cases "finite (range \<sigma>)")
+      case True
+      then obtain m where "infinite (\<sigma> -` {\<sigma> m})"
+        by (metis image_iff inf_img_fin_dom nat_not_finite)
+      then obtain r where [iff]: "strict_mono r" and r: "\<And>n::nat. r n \<in> \<sigma> -` {\<sigma> m}"
+        using infinite_enumerate by blast 
+      have [iff]: "\<sigma> m \<in> U" "\<sigma> m \<in> M"
+        using \<open>range \<sigma> \<subseteq> S\<close> assms by blast+
+      show ?thesis
+      proof (intro conjI exI)
+        show "limitin mtopology (\<sigma> \<circ> r) (\<sigma> m) sequentially"
+          using r by (simp add: limitin_metric)
+      qed auto
+    next
+      case False
+      then obtain l where "l \<in> U" and l: "l \<in> mtopology derived_set_of (range \<sigma>)"
+        by (meson R \<open>range \<sigma> \<subseteq> S\<close> disjoint_iff)
+      then obtain g where g: "\<And>\<epsilon>. \<epsilon>>0 \<Longrightarrow> \<sigma> (g \<epsilon>) \<noteq> l \<and> d l (\<sigma> (g \<epsilon>)) < \<epsilon>"
+        by (simp add: metric_derived_set_of) metis
+      have "range \<sigma> \<subseteq> M"
+        using \<open>range \<sigma> \<subseteq> S\<close> assms by auto
+      have "l \<in> M"
+        using l metric_derived_set_of by auto
+      define E where  \<comment>\<open>a construction to ensure monotonicity\<close>
+        "E \<equiv> \<lambda>rec n. insert (inverse (Suc n)) ((\<lambda>i. d l (\<sigma> i)) ` (\<Union>k<n. {0..rec k})) - {0}"
+      define r where "r \<equiv> wfrec less_than (\<lambda>rec n. g (Min (E rec n)))"
+      have "(\<Union>k<n. {0..cut r less_than n k}) = (\<Union>k<n. {0..r k})" for n
+        by (auto simp: cut_apply)
+      then have r_eq: "r n = g (Min (E r n))" for n
+        by (metis E_def def_wfrec [OF r_def] wf_less_than)
+      have dl_pos[simp]: "d l (\<sigma> (r n)) > 0" for n
+        using wf_less_than
+      proof (induction n rule: wf_induct_rule)
+        case (less n) 
+        then have *: "Min (E r n) > 0"
+          using \<open>l \<in> M\<close> \<open>range \<sigma> \<subseteq> M\<close> by (auto simp: E_def image_subset_iff)
+        show ?case
+          using g [OF *] r_eq [of n]
+          by (metis \<open>l \<in> M\<close> \<open>range \<sigma> \<subseteq> M\<close> mdist_pos_less range_subsetD)
+      qed
+      then have non_l: "\<sigma> (r n) \<noteq> l" for n
+        using \<open>range \<sigma> \<subseteq> M\<close> mdist_pos_eq by blast
+      have Min_pos: "Min (E r n) > 0" for n
+        using dl_pos \<open>l \<in> M\<close> \<open>range \<sigma> \<subseteq> M\<close> by (auto simp: E_def image_subset_iff)
+      have d_small: "d (\<sigma>(r n)) l < inverse(Suc n)" for n
+      proof -
+        have "d (\<sigma>(r n)) l < Min (E r n)"
+          by (simp add: \<open>0 < Min (E r n)\<close> commute g r_eq) 
+        also have "... \<le> inverse(Suc n)"
+          by (simp add: E_def)
+        finally show ?thesis .
+      qed
+      have d_lt_d: "d l (\<sigma> (r n)) < d l (\<sigma> i)" if \<section>: "p < n" "i \<le> r p" "\<sigma> i \<noteq> l" for i p n
+      proof -
+        have 1: "d l (\<sigma> i) \<in> E r n"
+          using \<section> \<open>l \<in> M\<close> \<open>range \<sigma> \<subseteq> M\<close> 
+          by (force simp: E_def image_subset_iff image_iff)
+        have "d l (\<sigma> (g (Min (E r n)))) < Min (E r n)"
+          by (rule conjunct2 [OF g [OF Min_pos]])
+        also have "Min (E r n) \<le> d l (\<sigma> i)"
+          using 1 unfolding E_def by (force intro!: Min.coboundedI)
+        finally show ?thesis
+          by (simp add: r_eq) 
+      qed
+      have r: "r p < r n" if "p < n" for p n
+        using d_lt_d [OF that] non_l by (meson linorder_not_le order_less_irrefl) 
+      show ?thesis
+      proof (intro exI conjI)
+        show "strict_mono r"
+          by (simp add: r strict_monoI)
+        show "limitin mtopology (\<sigma> \<circ> r) l sequentially"
+          unfolding limitin_metric
+        proof (intro conjI strip \<open>l \<in> M\<close>)
+          fix \<epsilon> :: real
+          assume "\<epsilon> > 0"
+          then have "\<forall>\<^sub>F n in sequentially. inverse(Suc n) < \<epsilon>"
+            using Archimedean_eventually_inverse by auto
+          then show "\<forall>\<^sub>F n in sequentially. (\<sigma> \<circ> r) n \<in> M \<and> d ((\<sigma> \<circ> r) n) l < \<epsilon>"
+            by (smt (verit) \<open>range \<sigma> \<subseteq> M\<close> commute comp_apply d_small eventually_mono range_subsetD)
+        qed
+      qed (use \<open>l \<in> U\<close> in auto)
+    qed
+  qed
+qed
+
+subsubsection \<open>More on Bolzano Weierstrass\<close>
+
+lemma Bolzano_Weierstrass_A: 
+  assumes "compactin mtopology S" "T \<subseteq> S" "infinite T"
+  shows "S \<inter> mtopology derived_set_of T \<noteq> {}"
+  by (simp add: assms compactin_imp_Bolzano_Weierstrass)
+
+lemma Bolzano_Weierstrass_B:
+  fixes \<sigma> :: "nat \<Rightarrow> 'a"
+  assumes "S \<subseteq> M" "range \<sigma> \<subseteq> S"
+    and "\<And>T. \<lbrakk>T \<subseteq> S \<and> infinite T\<rbrakk> \<Longrightarrow> S \<inter> mtopology derived_set_of T \<noteq> {}"
+  shows "\<exists>l r. l \<in> S \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially"
+  using Bolzano_Weierstrass_property assms by blast
+
+lemma Bolzano_Weierstrass_C:
+  assumes "S \<subseteq> M"
+  assumes "\<And>\<sigma>:: nat \<Rightarrow> 'a. range \<sigma> \<subseteq> S \<Longrightarrow>
+                (\<exists>l r. l \<in> S \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially)"
+  shows "mtotally_bounded S"
+  unfolding mtotally_bounded_sequentially
+  by (metis convergent_imp_MCauchy assms image_comp image_mono subset_UNIV subset_trans)
+
+lemma Bolzano_Weierstrass_D:
+  assumes "S \<subseteq> M" "S \<subseteq> \<Union>\<C>" and opeU: "\<And>U. U \<in> \<C> \<Longrightarrow> openin mtopology U"
+  assumes \<section>: "(\<forall>\<sigma>::nat\<Rightarrow>'a. range \<sigma> \<subseteq> S
+         \<longrightarrow> (\<exists>l r. l \<in> S \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially))"
+  shows "\<exists>\<epsilon>>0. \<forall>x \<in> S. \<exists>U \<in> \<C>. mball x \<epsilon> \<subseteq> U"
+proof (rule ccontr)
+  assume "\<not> (\<exists>\<epsilon>>0. \<forall>x \<in> S. \<exists>U \<in> \<C>. mball x \<epsilon> \<subseteq> U)"
+  then have "\<forall>n. \<exists>x\<in>S. \<forall>U\<in>\<C>. \<not> mball x (inverse (Suc n)) \<subseteq> U"
+    by simp
+  then obtain \<sigma> where "\<And>n. \<sigma> n \<in> S" 
+       and \<sigma>: "\<And>n U. U \<in> \<C> \<Longrightarrow> \<not> mball (\<sigma> n) (inverse (Suc n)) \<subseteq> U"
+    by metis
+  then obtain l r where "l \<in> S" "strict_mono r" 
+         and lr: "limitin mtopology (\<sigma> \<circ> r) l sequentially"
+    by (meson \<section> image_subsetI)
+  with \<open>S \<subseteq> \<Union>\<C>\<close> obtain B where "l \<in> B" "B \<in> \<C>"
+    by auto
+  then obtain \<epsilon> where "\<epsilon> > 0" and \<epsilon>: "\<And>z. \<lbrakk>z \<in> M; d z l < \<epsilon>\<rbrakk> \<Longrightarrow> z \<in> B"
+    by (metis opeU [OF \<open>B \<in> \<C>\<close>] commute in_mball openin_mtopology subset_iff)
+  then have "\<forall>\<^sub>F n in sequentially. \<sigma> (r n) \<in> M \<and> d (\<sigma> (r n)) l < \<epsilon>/2"
+    using lr half_gt_zero unfolding limitin_metric o_def by blast
+  moreover have "\<forall>\<^sub>F n in sequentially. inverse (real (Suc n)) < \<epsilon>/2"
+    using Archimedean_eventually_inverse \<open>0 < \<epsilon>\<close> half_gt_zero by blast
+  ultimately obtain n where n: "d (\<sigma> (r n)) l < \<epsilon>/2" "inverse (real (Suc n)) < \<epsilon>/2"
+    by (smt (verit, del_insts) eventually_sequentially le_add1 le_add2)
+  have "x \<in> B" if "d (\<sigma> (r n)) x < inverse (Suc(r n))" "x \<in> M" for x
+  proof -
+    have rle: "inverse (real (Suc (r n))) \<le> inverse (real (Suc n))"
+      using \<open>strict_mono r\<close> strict_mono_imp_increasing by auto
+    have "d x l \<le> d (\<sigma> (r n)) x + d (\<sigma> (r n)) l"
+      using that by (metis triangle \<open>\<And>n. \<sigma> n \<in> S\<close> \<open>l \<in> S\<close> \<open>S \<subseteq> M\<close> commute subsetD)
+    also have "... < \<epsilon>"
+      using that n rle by linarith
+    finally show ?thesis
+      by (simp add: \<epsilon> that)
+  qed
+  then show False
+    using \<sigma> [of B "r n"] by (simp add: \<open>B \<in> \<C>\<close> subset_iff)
+qed
+
+
+lemma Bolzano_Weierstrass_E:
+  assumes "mtotally_bounded S" "S \<subseteq> M"
+  and S: "\<And>\<C>. \<lbrakk>\<And>U. U \<in> \<C> \<Longrightarrow> openin mtopology U; S \<subseteq> \<Union>\<C>\<rbrakk> \<Longrightarrow> \<exists>\<epsilon>>0. \<forall>x \<in> S. \<exists>U \<in> \<C>. mball x \<epsilon> \<subseteq> U"
+  shows "compactin mtopology S"
+proof (clarsimp simp: compactin_def assms)
+  fix \<U> :: "'a set set"
+  assume \<U>: "\<forall>x\<in>\<U>. openin mtopology x" and "S \<subseteq> \<Union>\<U>"
+  then obtain \<epsilon> where "\<epsilon>>0" and \<epsilon>: "\<And>x. x \<in> S \<Longrightarrow> \<exists>U \<in> \<U>. mball x \<epsilon> \<subseteq> U"
+    by (metis S)
+  then obtain f where f: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> \<U> \<and> mball x \<epsilon> \<subseteq> f x"
+    by metis
+  then obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> (\<Union>x\<in>K. mball x \<epsilon>)"
+    by (metis \<open>0 < \<epsilon>\<close> \<open>mtotally_bounded S\<close> mtotally_bounded_def)
+  show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> S \<subseteq> \<Union>\<F>"
+  proof (intro conjI exI)
+    show "finite (f ` K)"
+      by (simp add: \<open>finite K\<close>)
+    show "f ` K \<subseteq> \<U>"
+      using \<open>K \<subseteq> S\<close> f by blast
+    show "S \<subseteq> \<Union>(f ` K)"
+      using K \<open>K \<subseteq> S\<close> by (force dest: f)
+  qed
+qed
+
+
+lemma compactin_eq_Bolzano_Weierstrass:
+  "compactin mtopology S \<longleftrightarrow>
+   S \<subseteq> M \<and> (\<forall>T. T \<subseteq> S \<and> infinite T \<longrightarrow> S \<inter> mtopology derived_set_of T \<noteq> {})"
+  using Bolzano_Weierstrass_C Bolzano_Weierstrass_D Bolzano_Weierstrass_E
+  by (smt (verit, del_insts) Bolzano_Weierstrass_property compactin_imp_Bolzano_Weierstrass compactin_subspace subset_refl topspace_mtopology)
+
+lemma compactin_sequentially:
+  shows "compactin mtopology S \<longleftrightarrow>
+        S \<subseteq> M \<and>
+        ((\<forall>\<sigma>::nat\<Rightarrow>'a. range \<sigma> \<subseteq> S
+         \<longrightarrow> (\<exists>l r. l \<in> S \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially)))"
+  by (metis Bolzano_Weierstrass_property compactin_eq_Bolzano_Weierstrass subset_refl)
+
+lemma compactin_imp_mtotally_bounded: 
+  "compactin mtopology S \<Longrightarrow> mtotally_bounded S"
+  by (simp add: Bolzano_Weierstrass_C compactin_sequentially)
+
+lemma lebesgue_number:
+    "\<lbrakk>compactin mtopology S; S \<subseteq> \<Union>\<C>; \<And>U. U \<in> \<C> \<Longrightarrow> openin mtopology U\<rbrakk>
+    \<Longrightarrow> \<exists>\<epsilon>>0. \<forall>x \<in> S. \<exists>U \<in> \<C>. mball x \<epsilon> \<subseteq> U"
+  by (simp add: Bolzano_Weierstrass_D compactin_sequentially)
+
+lemma compact_space_sequentially:
+   "compact_space mtopology \<longleftrightarrow>
+    (\<forall>\<sigma>::nat\<Rightarrow>'a. range \<sigma> \<subseteq> M
+         \<longrightarrow> (\<exists>l r. l \<in> M \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially))"
+  by (simp add: compact_space_def compactin_sequentially)
+
+lemma compact_space_eq_Bolzano_Weierstrass:
+   "compact_space mtopology \<longleftrightarrow>
+    (\<forall>S. S \<subseteq> M \<and> infinite S \<longrightarrow> mtopology derived_set_of S \<noteq> {})"
+  using Int_absorb1 [OF derived_set_of_subset_topspace [of mtopology]]
+  by (force simp: compact_space_def compactin_eq_Bolzano_Weierstrass)
+
+lemma compact_space_nest:
+   "compact_space mtopology \<longleftrightarrow>
+    (\<forall>C. (\<forall>n::nat. closedin mtopology (C n)) \<and> (\<forall>n. C n \<noteq> {}) \<and> decseq C \<longrightarrow> \<Inter>(range C) \<noteq> {})"
+   (is "?lhs=?rhs")
+proof
+  assume L: ?lhs
+  show ?rhs
+  proof clarify
+    fix C :: "nat \<Rightarrow> 'a set"
+    assume "\<forall>n. closedin mtopology (C n)"
+      and "\<forall>n. C n \<noteq> {}"
+      and "decseq C"
+      and "\<Inter> (range C) = {}"
+    then obtain K where K: "finite K" "\<Inter>(C ` K) = {}"
+      by (metis L compact_space_imp_nest)
+    then obtain k where "K \<subseteq> {..k}"
+      using finite_nat_iff_bounded_le by auto
+    then have "C k \<subseteq> \<Inter>(C ` K)"
+      using \<open>decseq C\<close> by (auto simp:decseq_def)
+    then show False
+      by (simp add: K \<open>\<forall>n. C n \<noteq> {}\<close>)
+  qed
+next
+  assume R [rule_format]: ?rhs
+  show ?lhs
+    unfolding compact_space_sequentially
+  proof (intro strip)
+    fix \<sigma> :: "nat \<Rightarrow> 'a"
+    assume \<sigma>: "range \<sigma> \<subseteq> M"
+    have "mtopology closure_of \<sigma> ` {n..} \<noteq> {}" for n
+      using \<open>range \<sigma> \<subseteq> M\<close> by (auto simp: closure_of_eq_empty image_subset_iff)
+    moreover have "decseq (\<lambda>n. mtopology closure_of \<sigma> ` {n..})"
+      using closure_of_mono image_mono by (smt (verit) atLeast_subset_iff decseq_def) 
+    ultimately obtain l where l: "\<And>n. l \<in> mtopology closure_of \<sigma> ` {n..}"
+      using R [of "\<lambda>n. mtopology closure_of (\<sigma> ` {n..})"] by auto
+    then have "l \<in> M" and "\<And>n. \<forall>r>0. \<exists>k\<ge>n. \<sigma> k \<in> M \<and> d l (\<sigma> k) < r"
+      using metric_closure_of by fastforce+
+    then obtain f where f: "\<And>n r. r>0 \<Longrightarrow> f n r \<ge> n \<and> \<sigma> (f n r) \<in> M \<and> d l (\<sigma> (f n r)) < r"
+      by metis
+    define r where "r = rec_nat (f 0 1) (\<lambda>n rec. (f (Suc rec) (inverse (Suc (Suc n)))))"
+    have r: "d l (\<sigma>(r n)) < inverse(Suc n)" for n
+      by (induction n) (auto simp: rec_nat_0_imp [OF r_def] rec_nat_Suc_imp [OF r_def] f)
+    have "r n < r(Suc n)" for n
+      by (simp add: Suc_le_lessD f r_def)
+    then have "strict_mono r"
+      by (simp add: strict_mono_Suc_iff)
+    moreover have "limitin mtopology (\<sigma> \<circ> r) l sequentially"
+      proof (clarsimp simp: limitin_metric \<open>l \<in> M\<close>)
+        fix \<epsilon> :: real
+        assume "\<epsilon> > 0"
+        then have "(\<forall>\<^sub>F n in sequentially. inverse (real (Suc n)) < \<epsilon>)"
+          using Archimedean_eventually_inverse by blast
+        then show "\<forall>\<^sub>F n in sequentially. \<sigma> (r n) \<in> M \<and> d (\<sigma> (r n)) l < \<epsilon>"
+          by eventually_elim (metis commute \<open>range \<sigma> \<subseteq> M\<close>  order_less_trans r range_subsetD)
+      qed
+    ultimately show "\<exists>l r. l \<in> M \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially"
+      using \<open>l \<in> M\<close> by blast
+  qed
+qed
+
+
+lemma (in discrete_metric) mcomplete_discrete_metric: "disc.mcomplete"
+proof (clarsimp simp: disc.mcomplete_def)
+  fix \<sigma> :: "nat \<Rightarrow> 'a"
+  assume "disc.MCauchy \<sigma>"
+  then obtain N where "\<And>n. N \<le> n \<Longrightarrow> \<sigma> N = \<sigma> n"
+    unfolding disc.MCauchy_def by (metis dd_def dual_order.refl order_less_irrefl zero_less_one)
+  moreover have "range \<sigma> \<subseteq> M"
+    using \<open>disc.MCauchy \<sigma>\<close> disc.MCauchy_def by blast
+  ultimately have "limitin disc.mtopology \<sigma> (\<sigma> N) sequentially"
+    by (metis disc.limit_metric_sequentially disc.zero range_subsetD)
+  then show "\<exists>x. limitin disc.mtopology \<sigma> x sequentially" ..
+qed
+
+lemma compact_space_imp_mcomplete: "compact_space mtopology \<Longrightarrow> mcomplete"
+  by (simp add: compact_space_nest mcomplete_nest)
+
+lemma (in submetric) compactin_imp_mcomplete:
+   "compactin mtopology A \<Longrightarrow> sub.mcomplete"
+  by (simp add: compactin_subspace mtopology_submetric sub.compact_space_imp_mcomplete)
+
+lemma (in submetric) mcomplete_imp_closedin:
+  assumes "sub.mcomplete"
+  shows "closedin mtopology A"
+proof -
+  have "l \<in> A"
+    if "range \<sigma> \<subseteq> A" and l: "limitin mtopology \<sigma> l sequentially"
+    for \<sigma> :: "nat \<Rightarrow> 'a" and l
+  proof -
+    have "sub.MCauchy \<sigma>"
+      using convergent_imp_MCauchy subset that by (force simp: MCauchy_submetric)
+    then have "limitin sub.mtopology \<sigma> l sequentially"
+      using assms unfolding sub.mcomplete_def
+      using l limitin_metric_unique limitin_submetric_iff trivial_limit_sequentially by blast
+    then show ?thesis
+      using limitin_submetric_iff by blast
+  qed
+  then show ?thesis
+    using metric_closedin_iff_sequentially_closed subset by auto
+qed
+
+lemma (in submetric) closedin_eq_mcomplete:
+   "mcomplete \<Longrightarrow> (closedin mtopology A \<longleftrightarrow> sub.mcomplete)"
+  using closedin_mcomplete_imp_mcomplete mcomplete_imp_closedin by blast
+
+lemma compact_space_eq_mcomplete_mtotally_bounded:
+   "compact_space mtopology \<longleftrightarrow> mcomplete \<and> mtotally_bounded M"
+  by (meson Bolzano_Weierstrass_C compact_space_imp_mcomplete compact_space_sequentially limitin_mspace 
+            mcomplete_alt mtotally_bounded_sequentially subset_refl)
+
+
+lemma compact_closure_of_imp_mtotally_bounded:
+   "\<lbrakk>compactin mtopology (mtopology closure_of S); S \<subseteq> M\<rbrakk>
+      \<Longrightarrow> mtotally_bounded S"
+  using compactin_imp_mtotally_bounded mtotally_bounded_closure_of_eq by blast
+
+lemma mtotally_bounded_eq_compact_closure_of:
+  assumes "mcomplete"
+  shows "mtotally_bounded S \<longleftrightarrow> S \<subseteq> M \<and> compactin mtopology (mtopology closure_of S)"
+  (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+    unfolding compactin_subspace
+  proof (intro conjI)
+    show "S \<subseteq> M"
+      using L by (simp add: mtotally_bounded_imp_subset)
+    show "mtopology closure_of S \<subseteq> topspace mtopology"
+      by (simp add: \<open>S \<subseteq> M\<close> closure_of_minimal)
+    then have MSM: "mtopology closure_of S \<subseteq> M"
+      by auto
+    interpret S: submetric M d "mtopology closure_of S"
+    proof qed (use MSM in auto)
+    have "S.sub.mtotally_bounded (mtopology closure_of S)"
+      using L mtotally_bounded_absolute mtotally_bounded_closure_of by blast
+    then
+    show "compact_space (subtopology mtopology (mtopology closure_of S))"
+      using S.closedin_mcomplete_imp_mcomplete S.mtopology_submetric S.sub.compact_space_eq_mcomplete_mtotally_bounded assms by force
+  qed
+qed (auto simp: compact_closure_of_imp_mtotally_bounded)
+
+
+
+lemma compact_closure_of_eq_Bolzano_Weierstrass:
+   "compactin mtopology (mtopology closure_of S) \<longleftrightarrow>
+    (\<forall>T. infinite T \<and> T \<subseteq> S \<and> T \<subseteq> M \<longrightarrow> mtopology derived_set_of T \<noteq> {})"  (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+  proof (intro strip)
+    fix T
+    assume T: "infinite T \<and> T \<subseteq> S \<and> T \<subseteq> M"
+    show "mtopology derived_set_of T \<noteq> {}"
+    proof (intro compact_closure_of_imp_Bolzano_Weierstrass)
+      show "compactin mtopology (mtopology closure_of S)"
+        by (simp add: L)
+    qed (use T in auto)
+  qed
+next
+  have "compactin mtopology (mtopology closure_of S)"
+    if \<section>: "\<And>T. \<lbrakk>infinite T; T \<subseteq> S\<rbrakk> \<Longrightarrow> mtopology derived_set_of T \<noteq> {}" and "S \<subseteq> M" for S
+    unfolding compactin_sequentially
+  proof (intro conjI strip)
+    show MSM: "mtopology closure_of S \<subseteq> M"
+      using closure_of_subset_topspace by fastforce
+    fix \<sigma> :: "nat \<Rightarrow> 'a"
+    assume \<sigma>: "range \<sigma> \<subseteq> mtopology closure_of S"
+    then have "\<exists>y \<in> S. d (\<sigma> n) y < inverse(Suc n)" for n
+      by (simp add: metric_closure_of image_subset_iff) (metis inverse_Suc of_nat_Suc)
+    then obtain \<tau> where \<tau>: "\<And>n. \<tau> n \<in> S \<and> d (\<sigma> n) (\<tau> n) < inverse(Suc n)"
+      by metis
+    then have "range \<tau> \<subseteq> S"
+      by blast
+    moreover
+    have *: "\<forall>T. T \<subseteq> S \<and> infinite T \<longrightarrow> mtopology closure_of S \<inter> mtopology derived_set_of T \<noteq> {}"
+      using "\<section>"(1) derived_set_of_mono derived_set_of_subset_closure_of by fastforce
+    moreover have "S \<subseteq> mtopology closure_of S"
+      by (simp add: \<open>S \<subseteq> M\<close> closure_of_subset)
+    ultimately obtain l r where lr:
+      "l \<in> mtopology closure_of S" "strict_mono r" "limitin mtopology (\<tau> \<circ> r) l sequentially"
+      using Bolzano_Weierstrass_property \<open>S \<subseteq> M\<close> by metis
+    then have "l \<in> M"
+      using limitin_mspace by blast
+    have dr_less: "d ((\<sigma> \<circ> r) n) ((\<tau> \<circ> r) n) < inverse(Suc n)" for n
+    proof -
+      have "d ((\<sigma> \<circ> r) n) ((\<tau> \<circ> r) n) < inverse(Suc (r n))"
+        using \<tau> by auto
+      also have "... \<le> inverse(Suc n)"
+        using lr strict_mono_imp_increasing by auto
+      finally show ?thesis .
+    qed
+    have "limitin mtopology (\<sigma> \<circ> r) l sequentially"
+      unfolding limitin_metric
+    proof (intro conjI strip)
+      show "l \<in> M"
+        using limitin_mspace lr by blast
+      fix \<epsilon> :: real
+      assume "\<epsilon> > 0"
+      then have "\<forall>\<^sub>F n in sequentially. (\<tau> \<circ> r) n \<in> M \<and> d ((\<tau> \<circ> r) n) l < \<epsilon>/2"
+        using lr half_gt_zero limitin_metric by blast 
+      moreover have "\<forall>\<^sub>F n in sequentially. inverse (real (Suc n)) < \<epsilon>/2"
+        using Archimedean_eventually_inverse \<open>0 < \<epsilon>\<close> half_gt_zero by blast
+      then have "\<forall>\<^sub>F n in sequentially. d ((\<sigma> \<circ> r) n) ((\<tau> \<circ> r) n) < \<epsilon>/2"
+        by eventually_elim (smt (verit, del_insts) dr_less)
+      ultimately have "\<forall>\<^sub>F n in sequentially. d ((\<sigma> \<circ> r) n) l < \<epsilon>/2 + \<epsilon>/2"
+        by eventually_elim (smt (verit) triangle \<open>l \<in> M\<close> MSM \<sigma> comp_apply order_trans range_subsetD)      
+      then show "\<forall>\<^sub>F n in sequentially. (\<sigma> \<circ> r) n \<in> M \<and> d ((\<sigma> \<circ> r) n) l < \<epsilon>"
+        apply eventually_elim
+        using \<open>mtopology closure_of S \<subseteq> M\<close> \<sigma> by auto
+    qed
+    with lr show "\<exists>l r. l \<in> mtopology closure_of S \<and> strict_mono r \<and> limitin mtopology (\<sigma> \<circ> r) l sequentially"
+      by blast
+  qed
+  then show "?rhs \<Longrightarrow> ?lhs"
+    by (metis Int_subset_iff closure_of_restrict inf_le1 topspace_mtopology)
+qed
+
+end
+
+lemma (in discrete_metric) mtotally_bounded_discrete_metric:
+   "disc.mtotally_bounded S \<longleftrightarrow> finite S \<and> S \<subseteq> M" (is "?lhs=?rhs")
+proof
+  assume L: ?lhs 
+  show ?rhs
+  proof
+    show "finite S"
+      by (metis (no_types) L closure_of_subset_Int compactin_discrete_topology disc.mtotally_bounded_eq_compact_closure_of
+          disc.topspace_mtopology discrete_metric.mcomplete_discrete_metric inf.absorb_iff2 mtopology_discrete_metric finite_subset)
+    show "S \<subseteq> M"
+      by (simp add: L disc.mtotally_bounded_imp_subset)
+  qed
+qed (simp add: disc.finite_imp_mtotally_bounded)
+
+
+context Metric_space
+begin
+
+lemma derived_set_of_infinite_openin_metric:
+   "mtopology derived_set_of S =
+    {x \<in> M. \<forall>U. x \<in> U \<and> openin mtopology U \<longrightarrow> infinite(S \<inter> U)}"
+  by (simp add: derived_set_of_infinite_openin Hausdorff_space_mtopology)
+
+lemma derived_set_of_infinite_1: 
+  assumes "infinite (S \<inter> mball x \<epsilon>)" 
+  shows "infinite (S \<inter> mcball x \<epsilon>)"
+  by (meson Int_mono assms finite_subset mball_subset_mcball subset_refl)
+
+lemma derived_set_of_infinite_2:
+  assumes "openin mtopology U" "\<And>\<epsilon>. 0 < \<epsilon> \<Longrightarrow> infinite (S \<inter> mcball x \<epsilon>)" and "x \<in> U"
+  shows "infinite (S \<inter> U)"
+  by (metis assms openin_mtopology_mcball finite_Int inf.absorb_iff2 inf_assoc)
+
+lemma derived_set_of_infinite_mball:
+  "mtopology derived_set_of S = {x \<in> M. \<forall>e>0. infinite(S \<inter> mball x e)}"
+  unfolding derived_set_of_infinite_openin_metric
+  by (meson centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2)
+
+lemma derived_set_of_infinite_mcball:
+  "mtopology derived_set_of S = {x \<in> M. \<forall>e>0. infinite(S \<inter> mcball x e)}"
+  unfolding derived_set_of_infinite_openin_metric
+  by (meson centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2)
+
+end
+
+subsection\<open>Continuous functions on metric spaces\<close>
+
+context Metric_space
+begin
+
+lemma continuous_map_to_metric:
+   "continuous_map X mtopology f \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. \<forall>\<epsilon>>0. \<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y\<in>U. f y \<in> mball (f x) \<epsilon>))"
+   (is "?lhs=?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    unfolding continuous_map_eq_topcontinuous_at topcontinuous_at_def
+    by (metis centre_in_mball_iff openin_mball topspace_mtopology)
+next
+  assume R: ?rhs
+  then have "\<forall>x\<in>topspace X. f x \<in> M"
+    by (meson gt_ex in_mball)
+  moreover 
+  have "\<And>x V. \<lbrakk>x \<in> topspace X; openin mtopology V; f x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y\<in>U. f y \<in> V)"
+    unfolding openin_mtopology by (metis Int_iff R inf.orderE)
+  ultimately
+  show ?lhs
+    by (simp add: continuous_map_eq_topcontinuous_at topcontinuous_at_def)
+qed 
+
+lemma continuous_map_from_metric:
+   "continuous_map mtopology X f \<longleftrightarrow>
+    f ` M \<subseteq> topspace X \<and>
+    (\<forall>a \<in> M. \<forall>U. openin X U \<and> f a \<in> U \<longrightarrow> (\<exists>r>0. \<forall>x. x \<in> M \<and> d a x < r \<longrightarrow> f x \<in> U))"
+proof (cases "f ` M \<subseteq> topspace X")
+  case True
+  then show ?thesis
+    by (fastforce simp: continuous_map openin_mtopology subset_eq)
+next
+  case False
+  then show ?thesis
+    by (simp add: continuous_map_eq_image_closure_subset_gen)
+qed
+
+text \<open>An abstract formulation, since the limits do not have to be sequential\<close>
+lemma continuous_map_uniform_limit:
+  assumes contf: "\<forall>\<^sub>F \<xi> in F. continuous_map X mtopology (f \<xi>)"
+    and dfg: "\<And>\<epsilon>. 0 < \<epsilon> \<Longrightarrow> \<forall>\<^sub>F \<xi> in F. \<forall>x \<in> topspace X. g x \<in> M \<and> d (f \<xi> x) (g x) < \<epsilon>"
+    and nontriv: "\<not> trivial_limit F"
+  shows "continuous_map X mtopology g"
+  unfolding continuous_map_to_metric
+proof (intro strip)
+  fix x and \<epsilon>::real
+  assume "x \<in> topspace X" and "\<epsilon> > 0"
+  then obtain \<xi> where k: "continuous_map X mtopology (f \<xi>)" 
+    and gM: "\<forall>x \<in> topspace X. g x \<in> M" 
+    and third: "\<forall>x \<in> topspace X. d (f \<xi> x) (g x) < \<epsilon>/3"
+    using eventually_conj [OF contf] contf dfg [of "\<epsilon>/3"] eventually_happens' [OF nontriv]
+    by (smt (verit, ccfv_SIG) zero_less_divide_iff)
+  then obtain U where U: "openin X U" "x \<in> U" and Uthird: "\<forall>y\<in>U. d (f \<xi> y) (f \<xi> x) < \<epsilon>/3"
+    unfolding continuous_map_to_metric
+    by (metis \<open>0 < \<epsilon>\<close> \<open>x \<in> topspace X\<close> commute divide_pos_pos in_mball zero_less_numeral)
+  have f_inM: "f \<xi> y \<in> M" if "y\<in>U" for y
+    using U k openin_subset that by (fastforce simp: continuous_map_def)
+  have "d (g y) (g x) < \<epsilon>" if "y\<in>U" for y
+  proof -
+    have "g y \<in> M"
+      using U gM openin_subset that by blast
+    have "d (g y) (g x) \<le> d (g y) (f \<xi> x) + d (f \<xi> x) (g x)"
+      by (simp add: U \<open>g y \<in> M\<close> \<open>x \<in> topspace X\<close> f_inM gM triangle)
+    also have "\<dots> \<le> d (g y) (f \<xi> y) + d (f \<xi> y) (f \<xi> x) + d (f \<xi> x) (g x)"
+      by (simp add: U \<open>g y \<in> M\<close> commute f_inM that triangle')
+    also have "\<dots> < \<epsilon>/3 + \<epsilon>/3 + \<epsilon>/3"
+      by (smt (verit) U(1) Uthird \<open>x \<in> topspace X\<close> commute openin_subset subsetD that third)
+    finally show ?thesis by simp
+  qed
+  with U gM show "\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y\<in>U. g y \<in> mball (g x) \<epsilon>)"
+    by (metis commute in_mball in_mono openin_subset)
+qed
+
+
+lemma continuous_map_uniform_limit_alt:
+  assumes contf: "\<forall>\<^sub>F \<xi> in F. continuous_map X mtopology (f \<xi>)"
+    and gim: "g ` (topspace X) \<subseteq> M"
+    and dfg: "\<And>\<epsilon>. 0 < \<epsilon> \<Longrightarrow> \<forall>\<^sub>F \<xi> in F. \<forall>x \<in> topspace X. d (f \<xi> x) (g x) < \<epsilon>"
+    and nontriv: "\<not> trivial_limit F"
+  shows "continuous_map X mtopology g"
+proof (rule continuous_map_uniform_limit [OF contf])
+  fix \<epsilon> :: real
+  assume "\<epsilon> > 0"
+  with gim dfg show "\<forall>\<^sub>F \<xi> in F. \<forall>x\<in>topspace X. g x \<in> M \<and> d (f \<xi> x) (g x) < \<epsilon>"
+    by (simp add: image_subset_iff)
+qed (use nontriv in auto)
+
+
+lemma continuous_map_uniformly_Cauchy_limit:
+  assumes "mcomplete"
+  assumes contf: "\<forall>\<^sub>F n in sequentially. continuous_map X mtopology (f n)"
+    and Cauchy': "\<And>\<epsilon>. \<epsilon> > 0 \<Longrightarrow> \<exists>N. \<forall>m n x. N \<le> m \<longrightarrow> N \<le> n \<longrightarrow> x \<in> topspace X \<longrightarrow> d (f m x) (f n x) < \<epsilon>"
+  obtains g where
+    "continuous_map X mtopology g"
+    "\<And>\<epsilon>. 0 < \<epsilon> \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>topspace X. d (f n x) (g x) < \<epsilon>"
+proof -
+  have "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>l. limitin mtopology (\<lambda>n. f n x) l sequentially"
+    using \<open>mcomplete\<close> [unfolded mcomplete, rule_format] assms
+    by (smt (verit) contf continuous_map_def eventually_mono topspace_mtopology)
+  then obtain g where g: "\<And>x. x \<in> topspace X \<Longrightarrow> limitin mtopology (\<lambda>n. f n x) (g x) sequentially"
+    by metis
+  show thesis
+  proof
+    show "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>topspace X. d (f n x) (g x) < \<epsilon>"
+      if "\<epsilon> > 0" for \<epsilon> :: real
+    proof -
+      obtain N where N: "\<And>m n x. \<lbrakk>N \<le> m; N \<le> n; x \<in> topspace X\<rbrakk> \<Longrightarrow> d (f m x) (f n x) < \<epsilon>/2"
+        by (meson Cauchy' \<open>0 < \<epsilon>\<close> half_gt_zero)
+      obtain P where P: "\<And>n x. \<lbrakk>n \<ge> P; x \<in> topspace X\<rbrakk> \<Longrightarrow> f n x \<in> M"
+        using contf by (auto simp: eventually_sequentially continuous_map_def)
+      show ?thesis
+      proof (intro eventually_sequentiallyI strip)
+        fix n x
+        assume "max N P \<le> n" and x: "x \<in> topspace X"
+        obtain L where "g x \<in> M" and L: "\<forall>n\<ge>L. f n x \<in> M \<and> d (f n x) (g x) < \<epsilon>/2"
+          using g [OF x] \<open>\<epsilon> > 0\<close> unfolding limitin_metric
+          by (metis (no_types, lifting) eventually_sequentially half_gt_zero)
+        define n' where "n' \<equiv> Max{L,N,P}"
+        have L': "\<forall>m \<ge> n'. f m x \<in> M \<and> d (f m x) (g x) < \<epsilon>/2"
+          using L by (simp add: n'_def)
+        moreover
+        have "d (f n x) (f n' x) < \<epsilon>/2"
+          using N [of n n' x] \<open>max N P \<le> n\<close> n'_def x by fastforce
+        ultimately have "d (f n x) (g x) < \<epsilon>/2 + \<epsilon>/2"
+          by (smt (verit, ccfv_SIG) P \<open>g x \<in> M\<close> \<open>max N P \<le> n\<close> le_refl max.bounded_iff mdist_zero triangle' x)
+        then show "d (f n x) (g x) < \<epsilon>" by simp
+      qed
+    qed
+    then show "continuous_map X mtopology g"
+      by (smt (verit, del_insts) eventually_mono g limitin_mspace trivial_limit_sequentially continuous_map_uniform_limit [OF contf])
+  qed
+qed
+
+lemma metric_continuous_map:
+  assumes "Metric_space M' d'"
+  shows
+   "continuous_map mtopology (Metric_space.mtopology M' d') f \<longleftrightarrow>
+    f ` M \<subseteq> M' \<and> (\<forall>a \<in> M. \<forall>\<epsilon>>0. \<exists>\<delta>>0.  (\<forall>x. x \<in> M \<and> d a x < \<delta> \<longrightarrow> d' (f a) (f x) < \<epsilon>))"
+   (is "?lhs = ?rhs")
+proof -
+  interpret M': Metric_space M' d'
+    by (simp add: assms)
+  show ?thesis
+  proof
+    assume L: ?lhs
+    show ?rhs
+    proof (intro conjI strip)
+      show "f ` M \<subseteq> M'"
+        using L by (auto simp: continuous_map_def)
+      fix a and \<epsilon> :: real
+      assume "a \<in> M" and "\<epsilon> > 0"
+      then have "openin mtopology {x \<in> M. f x \<in> M'.mball (f a) \<epsilon>}" "f a \<in> M'"
+        using L unfolding continuous_map_def by fastforce+
+      then obtain \<delta> where "\<delta> > 0" "mball a \<delta> \<subseteq> {x \<in> M. f x \<in> M' \<and> d' (f a) (f x) < \<epsilon>}"
+        using \<open>0 < \<epsilon>\<close> \<open>a \<in> M\<close> openin_mtopology by auto
+      then show "\<exists>\<delta>>0. \<forall>x. x \<in> M \<and> d a x < \<delta> \<longrightarrow> d' (f a) (f x) < \<epsilon>"
+        using \<open>a \<in> M\<close> in_mball by blast
+    qed
+  next
+    assume R: ?rhs    
+    show ?lhs
+      unfolding continuous_map_def
+    proof (intro conjI strip)
+      fix U
+      assume "openin M'.mtopology U"
+      then show "openin mtopology {x \<in> topspace mtopology. f x \<in> U}"
+        apply (simp add: continuous_map_def openin_mtopology M'.openin_mtopology subset_iff)
+        by (metis R image_subset_iff)
+    qed (use R in auto)
+  qed
+qed
+
+end (*Metric_space*)
+
+end
+
--- a/src/HOL/Analysis/Abstract_Topological_Spaces.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy	Fri May 26 09:56:20 2023 +0200
@@ -1276,6 +1276,10 @@
     using Hausdorff_space_closedin_diagonal embedding_imp_closed_map_eq by blast
 qed
 
+lemma proper_map_diag_eq [simp]:
+   "proper_map X (prod_topology X X) (\<lambda>x. (x,x)) \<longleftrightarrow> Hausdorff_space X"
+  by (simp add: closed_map_diag_eq inj_on_convol_ident injective_imp_proper_eq_closed_map)
+  
 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}"
@@ -1606,6 +1610,340 @@
    "\<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>Technical results about proper maps, perfect maps, etc\<close>
+
+lemma compact_imp_proper_map_gen:
+  assumes Y: "\<And>S. \<lbrakk>S \<subseteq> topspace Y; \<And>K. compactin Y K \<Longrightarrow> compactin Y (S \<inter> K)\<rbrakk>
+               \<Longrightarrow> closedin Y S"
+    and fim: "f ` (topspace X) \<subseteq> topspace Y"
+    and f: "continuous_map X Y f \<or> kc_space X"
+    and YX: "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
+  shows "proper_map X Y f"
+  unfolding proper_map_alt closed_map_def
+proof (intro conjI strip)
+  fix C
+  assume C: "closedin X C"
+  show "closedin Y (f ` C)"
+  proof (intro Y)
+    show "f ` C \<subseteq> topspace Y"
+      using C closedin_subset fim by blast
+    fix K
+    assume K: "compactin Y K"
+    define A where "A \<equiv> {x \<in> topspace X. f x \<in> K}"
+    have eq: "f ` C \<inter> K = f ` ({x \<in> topspace X. f x \<in> K} \<inter> C)"
+      using C closedin_subset by auto
+    show "compactin Y (f ` C \<inter> K)"
+      unfolding eq
+    proof (rule image_compactin)
+      show "compactin (subtopology X A) ({x \<in> topspace X. f x \<in> K} \<inter> C)"
+      proof (rule closedin_compact_space)
+        show "compact_space (subtopology X A)"
+          by (simp add: A_def K YX compact_space_subtopology)
+        show "closedin (subtopology X A) ({x \<in> topspace X. f x \<in> K} \<inter> C)"
+          using A_def C closedin_subtopology by blast
+      qed
+      have "continuous_map (subtopology X A) (subtopology Y K) f" if "kc_space X"
+        unfolding continuous_map_closedin
+      proof (intro conjI strip)
+        show "f x \<in> topspace (subtopology Y K)"
+          if "x \<in> topspace (subtopology X A)" for x
+          using that A_def K compactin_subset_topspace by auto
+      next
+        fix C
+        assume C: "closedin (subtopology Y K) C"
+        show "closedin (subtopology X A) {x \<in> topspace (subtopology X A). f x \<in> C}"
+        proof (rule compactin_imp_closedin_gen)
+          show "kc_space (subtopology X A)"
+            by (simp add: kc_space_subtopology that)
+          have [simp]: "{x \<in> topspace X. f x \<in> K \<and> f x \<in> C} = {x \<in> topspace X. f x \<in> C}"
+            using C closedin_imp_subset by auto
+          have "compactin (subtopology Y K) C"
+            by (simp add: C K closedin_compact_space compact_space_subtopology)
+          then have "compactin X {x \<in> topspace X. x \<in> A \<and> f x \<in> C}"
+            by (auto simp: A_def compactin_subtopology dest: YX)
+          then show "compactin (subtopology X A) {x \<in> topspace (subtopology X A). f x \<in> C}"
+            by (auto simp add: compactin_subtopology)
+        qed
+      qed
+      with f show "continuous_map (subtopology X A) Y f"
+        using continuous_map_from_subtopology continuous_map_in_subtopology by blast
+    qed
+  qed
+qed (simp add: YX)
+
+lemma tube_lemma_left:
+  assumes W: "openin (prod_topology X Y) W" and C: "compactin X C" 
+    and y: "y \<in> topspace Y" and subW: "C \<times> {y} \<subseteq> W"
+  shows "\<exists>U V. openin X U \<and> openin Y V \<and> C \<subseteq> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W"
+proof (cases "C = {}")
+  case True
+  with y 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 "x \<in> C" for x
+    using W openin_prod_topology_alt subW subsetD that by fastforce
+  then obtain U V where UV: "\<And>x. x \<in> C \<Longrightarrow> openin X (U x) \<and> openin Y (V x) \<and> x \<in> U x \<and> y \<in> V x \<and> U x \<times> V x \<subseteq> W" 
+    by metis
+  then obtain D where D: "finite D" "D \<subseteq> C" "C \<subseteq> \<Union> (U ` D)"
+    using compactinD [OF C, of "U`C"]
+    by (smt (verit) UN_I finite_subset_image imageE subsetI)
+  show ?thesis
+  proof (intro exI conjI)
+    show "openin X (\<Union> (U ` D))" "openin Y (\<Inter> (V ` D))"
+      using D False UV by blast+
+    show "y \<in> \<Inter> (V ` D)" "C \<subseteq> \<Union> (U ` D)" "\<Union>(U ` D) \<times> \<Inter>(V ` D) \<subseteq> W"
+      using D UV by force+
+  qed
+qed
+
+lemma Wallace_theorem_prod_topology:
+  assumes "compactin X K" "compactin Y L" 
+    and W: "openin (prod_topology X Y) W" and subW: "K \<times> L \<subseteq> W"
+  obtains U V where "openin X U" "openin Y V" "K \<subseteq> U" "L \<subseteq> V" "U \<times> V \<subseteq> W"
+proof -
+  have "\<And>y. y \<in> L \<Longrightarrow> \<exists>U V. openin X U \<and> openin Y V \<and> K \<subseteq> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W"
+  proof (intro tube_lemma_left assms)
+    fix y assume "y \<in> L"
+    show "y \<in> topspace Y"
+      using assms \<open>y \<in> L\<close> compactin_subset_topspace by blast 
+    show "K \<times> {y} \<subseteq> W"
+      using \<open>y \<in> L\<close> subW by force
+  qed
+  then obtain U V where UV: 
+         "\<And>y. y \<in> L \<Longrightarrow> openin X (U y) \<and> openin Y (V y) \<and> K \<subseteq> U y \<and> y \<in> V y \<and> U y \<times> V y \<subseteq> W"
+    by metis
+  then obtain M where "finite M" "M \<subseteq> L" and M: "L \<subseteq> \<Union> (V ` M)"
+    using \<open>compactin Y L\<close> unfolding compactin_def
+    by (smt (verit) UN_iff finite_subset_image imageE subset_iff)
+  show thesis
+  proof (cases "M={}")
+    case True
+    with M have "L={}"
+      by blast
+    then show ?thesis
+      using \<open>compactin X K\<close> compactin_subset_topspace that by fastforce
+  next
+    case False
+    show ?thesis
+    proof
+      show "openin X (\<Inter>(U`M))"
+        using False UV \<open>M \<subseteq> L\<close> \<open>finite M\<close> by blast
+      show "openin Y (\<Union>(V`M))"
+        using UV \<open>M \<subseteq> L\<close> by blast
+      show "K \<subseteq> \<Inter>(U`M)"
+        by (meson INF_greatest UV \<open>M \<subseteq> L\<close> subsetD)
+      show "L \<subseteq> \<Union>(V`M)"
+        by (simp add: M)
+      show "\<Inter>(U`M) \<times> \<Union>(V`M) \<subseteq> W"
+        using UV \<open>M \<subseteq> L\<close> by fastforce
+    qed   
+  qed
+qed
+
+lemma proper_map_prod:
+   "proper_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) \<longleftrightarrow>
+    topspace(prod_topology X Y) = {} \<or> proper_map X X' f \<and> proper_map Y Y' g"
+   (is "?lhs \<longleftrightarrow> _ \<or> ?rhs")
+proof (cases "topspace(prod_topology X Y) = {}")
+  case True
+  then show ?thesis
+    by (simp add: proper_map_on_empty)
+next
+  case False
+  then have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
+    by auto
+  define h where "h \<equiv> \<lambda>(x,y). (f x, g y)"
+  have "proper_map X X' f" "proper_map Y Y' g" if ?lhs
+  proof -
+    have cm: "closed_map X X' f" "closed_map Y Y' g"
+      using that False closed_map_prod proper_imp_closed_map by blast+
+    show "proper_map X X' f"
+    proof (clarsimp simp add: proper_map_def cm)
+      fix y
+      assume y: "y \<in> topspace X'"
+      obtain z where z: "z \<in> topspace Y"
+        using ne by blast
+      then have eq: "{x \<in> topspace X. f x = y} =
+                     fst ` {u \<in> topspace X \<times> topspace Y. h u = (y,g z)}"
+        by (force simp: h_def)
+      show "compactin X {x \<in> topspace X. f x = y}"
+        unfolding eq
+      proof (intro image_compactin)
+        have "g z \<in> topspace Y'"
+          by (meson closed_map_def closedin_subset closedin_topspace cm image_subset_iff z)
+        with y show "compactin (prod_topology X Y) {u \<in> topspace X \<times> topspace Y. (h u) = (y, g z)}"
+          using that by (simp add: h_def proper_map_def)
+        show "continuous_map (prod_topology X Y) X fst"
+          by (simp add: continuous_map_fst)
+      qed
+    qed
+    show "proper_map Y Y' g"
+    proof (clarsimp simp add: proper_map_def cm)
+      fix y
+      assume y: "y \<in> topspace Y'"
+      obtain z where z: "z \<in> topspace X"
+        using ne by blast
+      then have eq: "{x \<in> topspace Y. g x = y} =
+                     snd ` {u \<in> topspace X \<times> topspace Y. h u = (f z,y)}"
+        by (force simp: h_def)
+      show "compactin Y {x \<in> topspace Y. g x = y}"
+        unfolding eq
+      proof (intro image_compactin)
+        have "f z \<in> topspace X'"
+          by (meson closed_map_def closedin_subset closedin_topspace cm image_subset_iff z)
+        with y show "compactin (prod_topology X Y) {u \<in> topspace X \<times> topspace Y. (h u) = (f z, y)}"
+          using that by (simp add: proper_map_def h_def)
+        show "continuous_map (prod_topology X Y) Y snd"
+          by (simp add: continuous_map_snd)
+      qed
+    qed
+  qed
+  moreover
+  { assume R: ?rhs
+    then have fgim: "f ` topspace X \<subseteq> topspace X'" "g ` topspace Y \<subseteq> topspace Y'" 
+          and cm: "closed_map X X' f" "closed_map Y Y' g"
+      by (auto simp: proper_map_def closed_map_imp_subset_topspace)
+    have "closed_map (prod_topology X Y) (prod_topology X' Y') h"
+      unfolding closed_map_fibre_neighbourhood imp_conjL
+    proof (intro conjI strip)
+      show "h ` topspace (prod_topology X Y) \<subseteq> topspace (prod_topology X' Y')"
+        unfolding h_def using fgim by auto
+      fix W w
+      assume W: "openin (prod_topology X Y) W"
+        and w: "w \<in> topspace (prod_topology X' Y')"
+        and subW: "{x \<in> topspace (prod_topology X Y). h x = w} \<subseteq> W"
+      then obtain x' y' where weq: "w = (x',y')" "x' \<in> topspace X'" "y' \<in> topspace Y'"
+        by auto
+      have eq: "{u \<in> topspace X \<times> topspace Y. h u = (x',y')} = {x \<in> topspace X. f x = x'} \<times> {y \<in> topspace Y. g y = y'}"
+        by (auto simp: h_def)
+      obtain U V where "openin X U" "openin Y V" "U \<times> V \<subseteq> W"
+        and U: "{x \<in> topspace X. f x = x'} \<subseteq> U" 
+        and V: "{x \<in> topspace Y. g x = y'} \<subseteq> V" 
+      proof (rule Wallace_theorem_prod_topology)
+        show "compactin X {x \<in> topspace X. f x = x'}" "compactin Y {x \<in> topspace Y. g x = y'}"
+          using R weq unfolding proper_map_def closed_map_fibre_neighbourhood by fastforce+
+        show "{x \<in> topspace X. f x = x'} \<times> {x \<in> topspace Y. g x = y'} \<subseteq> W"
+          using weq subW by (auto simp: h_def)
+      qed (use W in auto)
+      obtain U' where "openin X' U'" "x' \<in> U'" and U': "{x \<in> topspace X. f x \<in> U'} \<subseteq> U"
+        using cm U \<open>openin X U\<close> weq unfolding closed_map_fibre_neighbourhood by meson
+      obtain V' where "openin Y' V'" "y' \<in> V'" and V': "{x \<in> topspace Y. g x \<in> V'} \<subseteq> V"
+        using cm V \<open>openin Y V\<close> weq unfolding closed_map_fibre_neighbourhood by meson
+      show "\<exists>V. openin (prod_topology X' Y') V \<and> w \<in> V \<and> {x \<in> topspace (prod_topology X Y). h x \<in> V} \<subseteq> W"
+      proof (intro conjI exI)
+        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 "w \<in> U' \<times> V'"
+          using \<open>x' \<in> U'\<close> \<open>y' \<in> V'\<close> weq by blast
+        show "{x \<in> topspace (prod_topology X Y). h x \<in> U' \<times> V'} \<subseteq> W"
+          using \<open>U \<times> V \<subseteq> W\<close> U' V' h_def by auto
+      qed
+    qed
+    moreover
+    have "compactin (prod_topology X Y) {u \<in> topspace X \<times> topspace Y. h u = (w, z)}"
+      if "w \<in> topspace X'" and "z \<in> topspace Y'" for w z
+    proof -
+      have eq: "{u \<in> topspace X \<times> topspace Y. h u = (w,z)} =
+                {u \<in> topspace X. f u = w} \<times> {y. y \<in> topspace Y \<and> g y = z}"
+        by (auto simp: h_def)
+      show ?thesis
+        using R that by (simp add: eq compactin_Times proper_map_def)
+    qed
+    ultimately have ?lhs
+      by (auto simp: h_def proper_map_def) 
+  }
+  ultimately show ?thesis using False by metis
+qed
+
+lemma proper_map_paired:
+  assumes "Hausdorff_space X \<and> proper_map X Y f \<and> proper_map X Z g \<or>
+        Hausdorff_space Y \<and> continuous_map X Y f \<and> proper_map X Z g \<or>
+        Hausdorff_space Z \<and> proper_map X Y f \<and> continuous_map X Z g"
+  shows "proper_map X (prod_topology Y Z) (\<lambda>x. (f x,g x))"
+  using assms
+proof (elim disjE conjE)
+  assume \<section>: "Hausdorff_space X" "proper_map X Y f" "proper_map X Z g"
+  have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x, y). (f x, g y)) \<circ> (\<lambda>x. (x, x))"
+    by auto
+  show "proper_map X (prod_topology Y Z) (\<lambda>x. (f x, g x))"
+    unfolding eq
+  proof (rule proper_map_compose)
+    show "proper_map X (prod_topology X X) (\<lambda>x. (x,x))"
+      by (simp add: \<section>)
+    show "proper_map (prod_topology X X) (prod_topology Y Z) (\<lambda>(x,y). (f x, g y))"
+      by (simp add: \<section> proper_map_prod)
+  qed
+next
+  assume \<section>: "Hausdorff_space Y" "continuous_map X Y f" "proper_map X Z g"
+  have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x,y). (x,g y)) \<circ> (\<lambda>x. (f x,x))"
+    by auto
+  show "proper_map X (prod_topology Y Z) (\<lambda>x. (f x, g x))"
+    unfolding eq
+  proof (rule proper_map_compose)
+    show "proper_map X (prod_topology Y X) (\<lambda>x. (f x,x))"
+      by (simp add: \<section> proper_map_paired_continuous_map_left)
+    show "proper_map (prod_topology Y X) (prod_topology Y Z) (\<lambda>(x,y). (x,g y))"
+      by (simp add: \<section> proper_map_prod proper_map_id [unfolded id_def])
+  qed
+next
+  assume \<section>: "Hausdorff_space Z" "proper_map X Y f" "continuous_map X Z g"
+  have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x,y). (f x,y)) \<circ> (\<lambda>x. (x,g x))"
+    by auto
+  show "proper_map X (prod_topology Y Z) (\<lambda>x. (f x, g x))"
+    unfolding eq
+  proof (rule proper_map_compose)
+    show "proper_map X (prod_topology X Z) (\<lambda>x. (x, g x))"
+      using \<section> proper_map_paired_continuous_map_right by auto
+    show "proper_map (prod_topology X Z) (prod_topology Y Z) (\<lambda>(x,y). (f x,y))"
+      by (simp add: \<section> proper_map_prod proper_map_id [unfolded id_def])
+  qed
+qed
+
+lemma proper_map_pairwise:
+  assumes
+    "Hausdorff_space X \<and> proper_map X Y (fst \<circ> f) \<and> proper_map X Z (snd \<circ> f) \<or>
+     Hausdorff_space Y \<and> continuous_map X Y (fst \<circ> f) \<and> proper_map X Z (snd \<circ> f) \<or>
+     Hausdorff_space Z \<and> proper_map X Y (fst \<circ> f) \<and> continuous_map X Z (snd \<circ> f)"
+  shows "proper_map X (prod_topology Y Z) f"
+  using proper_map_paired [OF assms] by (simp add: o_def)
+
+lemma proper_map_from_composition_right:
+  assumes "Hausdorff_space Y" "proper_map X Z (g \<circ> f)" and "continuous_map X Y f"
+    and contg: "continuous_map Y Z g"
+  shows "proper_map X Y f"
+proof -
+  define YZ where "YZ \<equiv> subtopology (prod_topology Y Z) ((\<lambda>x. (x, g x)) ` topspace Y)"
+  have "proper_map X Y (fst \<circ> (\<lambda>x. (f x, (g \<circ> f) x)))"
+  proof (rule proper_map_compose)
+    have [simp]: "x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y" for x
+      by (meson assms(3) continuous_map_def)
+    show "proper_map X YZ (\<lambda>x. (f x, (g \<circ> f) x))"
+      unfolding YZ_def
+      using assms
+      by (force intro!: proper_map_into_subtopology proper_map_paired simp: o_def image_iff)+
+    show "proper_map YZ Y fst"
+      using contg 
+      by (simp flip: homeomorphic_maps_graph add: YZ_def homeomorphic_maps_map homeomorphic_imp_proper_map)
+  qed
+  moreover have "fst \<circ> (\<lambda>x. (f x, (g \<circ> f) x)) = f"
+    by auto
+  ultimately show ?thesis
+    by auto
+qed
+
+lemma perfect_map_from_composition_right:
+   "\<lbrakk>Hausdorff_space Y; perfect_map X Z (g \<circ> f);
+     continuous_map X Y f; continuous_map Y Z g; f ` topspace X = topspace Y\<rbrakk>
+    \<Longrightarrow> perfect_map X Y f"
+  by (meson perfect_map_def proper_map_from_composition_right)
+
+lemma perfect_map_from_composition_right_inj:
+   "\<lbrakk>perfect_map X Z (g \<circ> f); f ` topspace X = topspace Y;
+     continuous_map X Y f; continuous_map Y Z g; inj_on g (topspace Y)\<rbrakk>
+    \<Longrightarrow> perfect_map X Y f"
+  by (meson continuous_map_image_subset_topspace perfect_map_def proper_map_from_composition_right_inj)
+
 
 subsection \<open>Regular spaces\<close>
 
@@ -2022,6 +2360,10 @@
     by auto
 qed
 
+lemma continuous_imp_proper_map:
+   "\<lbrakk>compact_space X; kc_space Y; continuous_map X Y f\<rbrakk> \<Longrightarrow> proper_map X Y f"
+  by (simp add: continuous_closed_imp_proper_map continuous_imp_closed_map_gen kc_imp_t1_space)
+
 
 lemma tube_lemma_right:
   assumes W: "openin (prod_topology X Y) W" and C: "compactin Y C" 
@@ -2060,8 +2402,7 @@
   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
+    using tube_lemma_right[of X Y _ "topspace Y"] assms by (fastforce simp: compact_space_def)
   show ?thesis
     unfolding closed_map_fibre_neighbourhood
     by (force simp: * openin_subset cong: conj_cong intro: **)
@@ -2492,17 +2833,19 @@
 qed
 
 lemma locally_compact_space_open_subset:
-  assumes reg: "regular_space X" and loc: "locally_compact_space X" and "openin X S"
+  assumes X: "Hausdorff_space X \<or> 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)"
+  moreover have reg: "regular_space X"
+    using X loc locally_compact_Hausdorff_imp_regular_space by blast
+  moreover have "openin X (U \<inter> S)"
     by (simp add: UK \<open>openin X S\<close> openin_Int)
-  with UK reg x obtain V C 
+  ultimately 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)
+    by (metis \<open>x \<in> S\<close> 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)
@@ -2575,19 +2918,17 @@
   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)
+  proof (rule locally_compact_space_open_subset)
+    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 (simp add: Hausdorff_space_subtopology \<open>Hausdorff_space X\<close>)
+  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
+    using assms locally_compact_subspace_closed_Int_openin by blast
 qed
 
 lemma dense_locally_compact_openin_Hausdorff_space:
@@ -3464,5 +3805,1505 @@
     by force
 qed
 
+subsection\<open>Quasi-components\<close>
+
+definition quasi_component_of :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+  where
+  "quasi_component_of X x y \<equiv>
+        x \<in> topspace X \<and> y \<in> topspace X \<and>
+        (\<forall>T. closedin X T \<and> openin X T \<longrightarrow> (x \<in> T \<longleftrightarrow> y \<in> T))"
+
+abbreviation "quasi_component_of_set S x \<equiv> Collect (quasi_component_of S x)"
+
+definition quasi_components_of :: "'a topology \<Rightarrow> ('a set) set"
+  where
+  "quasi_components_of X = quasi_component_of_set X ` topspace X"
+
+lemma quasi_component_in_topspace:
+   "quasi_component_of X x y \<Longrightarrow> x \<in> topspace X \<and> y \<in> topspace X"
+  by (simp add: quasi_component_of_def)
+
+lemma quasi_component_of_refl [simp]:
+   "quasi_component_of X x x \<longleftrightarrow> x \<in> topspace X"
+  by (simp add: quasi_component_of_def)
+
+lemma quasi_component_of_sym:
+   "quasi_component_of X x y \<longleftrightarrow> quasi_component_of X y x"
+  by (meson quasi_component_of_def)
+
+lemma quasi_component_of_trans:
+   "\<lbrakk>quasi_component_of X x y; quasi_component_of X y z\<rbrakk> \<Longrightarrow> quasi_component_of X x z"
+  by (simp add: quasi_component_of_def)
+
+lemma quasi_component_of_subset_topspace:
+   "quasi_component_of_set X x \<subseteq> topspace X"
+  using quasi_component_of_def by fastforce
+
+lemma quasi_component_of_eq_empty:
+   "quasi_component_of_set X x = {} \<longleftrightarrow> (x \<notin> topspace X)"
+  using quasi_component_of_def by fastforce
+
+lemma quasi_component_of:
+   "quasi_component_of X x y \<longleftrightarrow>
+    x \<in> topspace X \<and> y \<in> topspace X \<and> (\<forall>T. x \<in> T \<and> closedin X T \<and> openin X T \<longrightarrow> y \<in> T)"
+  unfolding quasi_component_of_def by (metis Diff_iff closedin_def openin_closedin_eq) 
+
+lemma quasi_component_of_alt:
+  "quasi_component_of X x y \<longleftrightarrow>
+      x \<in> topspace X \<and> y \<in> topspace X \<and>
+      \<not> (\<exists>U V. openin X U \<and> openin X V \<and> U \<union> V = topspace X \<and> disjnt U V \<and> x \<in> U \<and> y \<in> V)" 
+  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    unfolding quasi_component_of_def
+    by (metis disjnt_iff separatedin_full separatedin_open_sets)
+  show "?rhs \<Longrightarrow> ?lhs"
+    unfolding quasi_component_of_def
+    by (metis Diff_disjoint Diff_iff Un_Diff_cancel closedin_def disjnt_def inf_commute sup.orderE sup_commute)
+qed
+
+lemma quasi_component_of_separated:
+   "quasi_component_of X x y \<longleftrightarrow>
+     x \<in> topspace X \<and> y \<in> topspace X \<and>
+     \<not> (\<exists>U V. separatedin X U V \<and> U \<union> V = topspace X \<and> x \<in> U \<and> y \<in> V)"
+  by (meson quasi_component_of_alt separatedin_full separatedin_open_sets)
+
+lemma quasi_component_of_subtopology:
+  "quasi_component_of (subtopology X s) x y \<Longrightarrow> quasi_component_of X x y"
+  unfolding quasi_component_of_def
+  by (simp add: closedin_subtopology) (metis Int_iff inf_commute openin_subtopology_Int2)
+
+lemma quasi_component_of_mono:
+   "quasi_component_of (subtopology X S) x y \<and> S \<subseteq> T
+        \<Longrightarrow> quasi_component_of (subtopology X T) x y"
+  by (metis inf.absorb_iff2 quasi_component_of_subtopology subtopology_subtopology)
+
+lemma quasi_component_of_equiv:
+   "quasi_component_of X x y \<longleftrightarrow>
+    x \<in> topspace X \<and> y \<in> topspace X \<and> quasi_component_of X x = quasi_component_of X y"
+  using quasi_component_of_def by fastforce
+
+lemma quasi_component_of_disjoint [simp]:
+   "disjnt (quasi_component_of_set X x) (quasi_component_of_set X y) \<longleftrightarrow> \<not> (quasi_component_of X x y)"
+  by (metis disjnt_iff quasi_component_of_equiv mem_Collect_eq)
+
+lemma quasi_component_of_eq:
+   "quasi_component_of X x = quasi_component_of X y \<longleftrightarrow>
+    (x \<notin> topspace X \<and> y \<notin> topspace X) 
+  \<or> x \<in> topspace X \<and> y \<in> topspace X \<and> quasi_component_of X x y"
+  by (metis Collect_empty_eq_bot quasi_component_of_eq_empty quasi_component_of_equiv)
+
+lemma topspace_imp_quasi_components_of:
+  assumes "x \<in> topspace X"
+  obtains C where "C \<in> quasi_components_of X" "x \<in> C"
+  by (metis assms imageI mem_Collect_eq quasi_component_of_refl quasi_components_of_def)
+
+lemma Union_quasi_components_of: "\<Union> (quasi_components_of X) = topspace X"
+  by (auto simp: quasi_components_of_def quasi_component_of_def)
+
+lemma pairwise_disjoint_quasi_components_of:
+   "pairwise disjnt (quasi_components_of X)"
+  by (auto simp: quasi_components_of_def quasi_component_of_def disjoint_def)
+
+lemma complement_quasi_components_of_Union:
+  assumes "C \<in> quasi_components_of X"
+  shows "topspace X - C = \<Union> (quasi_components_of X - {C})"  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    using Union_quasi_components_of by fastforce
+  show "?rhs \<subseteq> ?lhs"
+    using assms
+    using quasi_component_of_equiv by (fastforce simp add: quasi_components_of_def image_iff subset_iff)
+qed
+
+lemma nonempty_quasi_components_of:
+   "C \<in> quasi_components_of X \<Longrightarrow> C \<noteq> {}"
+  by (metis imageE quasi_component_of_eq_empty quasi_components_of_def)
+
+lemma quasi_components_of_subset:
+   "C \<in> quasi_components_of X \<Longrightarrow> C \<subseteq> topspace X"
+  using Union_quasi_components_of by force
+
+lemma quasi_component_in_quasi_components_of:
+   "quasi_component_of_set X a \<in> quasi_components_of X \<longleftrightarrow> a \<in> topspace X"
+  by (metis (no_types, lifting) image_iff quasi_component_of_eq_empty quasi_components_of_def)
+
+lemma quasi_components_of_eq_empty [simp]:
+   "quasi_components_of X = {} \<longleftrightarrow> topspace X = {}"
+  by (simp add: quasi_components_of_def)
+
+lemma quasi_components_of_empty_space:
+   "topspace X = {} \<Longrightarrow> quasi_components_of X = {}"
+  by simp
+
+lemma quasi_component_of_set:
+   "quasi_component_of_set X x =
+        (if x \<in> topspace X
+        then \<Inter> {t. closedin X t \<and> openin X t \<and> x \<in> t}
+        else {})"
+  by (auto simp: quasi_component_of)
+
+lemma closedin_quasi_component_of: "closedin X (quasi_component_of_set X x)"
+  by (auto simp: quasi_component_of_set)
+
+lemma closedin_quasi_components_of:
+   "C \<in> quasi_components_of X \<Longrightarrow> closedin X C"
+  by (auto simp: quasi_components_of_def closedin_quasi_component_of)
+
+lemma openin_finite_quasi_components:
+  "\<lbrakk>finite(quasi_components_of X); C \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> openin X C"
+  apply (simp add:openin_closedin_eq quasi_components_of_subset complement_quasi_components_of_Union)
+  by (meson DiffD1 closedin_Union closedin_quasi_components_of finite_Diff)
+
+lemma quasi_component_of_eq_overlap:
+   "quasi_component_of X x = quasi_component_of X y \<longleftrightarrow>
+      (x \<notin> topspace X \<and> y \<notin> topspace X) \<or>
+      \<not> (quasi_component_of_set X x \<inter> quasi_component_of_set X y = {})"
+  using quasi_component_of_equiv by fastforce
+
+lemma quasi_component_of_nonoverlap:
+   "quasi_component_of_set X x \<inter> quasi_component_of_set X y = {} \<longleftrightarrow>
+     (x \<notin> topspace X) \<or> (y \<notin> topspace X) \<or>
+     \<not> (quasi_component_of X x = quasi_component_of X y)"
+  by (metis inf.idem quasi_component_of_eq_empty quasi_component_of_eq_overlap)
+
+lemma quasi_component_of_overlap:
+   "\<not> (quasi_component_of_set X x \<inter> quasi_component_of_set X y = {}) \<longleftrightarrow>
+    x \<in> topspace X \<and> y \<in> topspace X \<and> quasi_component_of X x = quasi_component_of X y"
+  by (meson quasi_component_of_nonoverlap)
+
+lemma quasi_components_of_disjoint:
+   "\<lbrakk>C \<in> quasi_components_of X; D \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> disjnt C D \<longleftrightarrow> C \<noteq> D"
+  by (metis disjnt_self_iff_empty nonempty_quasi_components_of pairwiseD pairwise_disjoint_quasi_components_of)
+
+lemma quasi_components_of_overlap:
+   "\<lbrakk>C \<in> quasi_components_of X; D \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> \<not> (C \<inter> D = {}) \<longleftrightarrow> C = D"
+  by (metis disjnt_def quasi_components_of_disjoint)
+
+lemma pairwise_separated_quasi_components_of:
+   "pairwise (separatedin X) (quasi_components_of X)"
+  by (metis closedin_quasi_components_of pairwise_def pairwise_disjoint_quasi_components_of separatedin_closed_sets)
+
+lemma finite_quasi_components_of_finite:
+   "finite(topspace X) \<Longrightarrow> finite(quasi_components_of X)"
+  by (simp add: Union_quasi_components_of finite_UnionD)
+
+lemma connected_imp_quasi_component_of:
+  assumes "connected_component_of X x y"
+  shows "quasi_component_of X x y"
+proof -
+  have "x \<in> topspace X" "y \<in> topspace X"
+    by (meson assms connected_component_of_equiv)+
+  with assms show ?thesis
+    apply (clarsimp simp add: quasi_component_of connected_component_of_def)
+    by (meson connectedin_clopen_cases disjnt_iff subsetD)
+qed
+
+lemma connected_component_subset_quasi_component_of:
+   "connected_component_of_set X x \<subseteq> quasi_component_of_set X x"
+  using connected_imp_quasi_component_of by force
+
+lemma quasi_component_as_connected_component_Union:
+   "quasi_component_of_set X x =
+    \<Union> (connected_component_of_set X ` quasi_component_of_set X x)" 
+    (is "?lhs = ?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs"
+    using connected_component_of_refl quasi_component_of by fastforce
+  show "?rhs \<subseteq> ?lhs"
+    apply (rule SUP_least)
+    by (simp add: connected_component_subset_quasi_component_of quasi_component_of_equiv)
+qed
+
+lemma quasi_components_as_connected_components_Union:
+  assumes "C \<in> quasi_components_of X"
+  obtains \<T> where "\<T> \<subseteq> connected_components_of X" "\<Union>\<T> = C"
+proof -
+  obtain x where "x \<in> topspace X" and Ceq: "C = quasi_component_of_set X x"
+    by (metis assms imageE quasi_components_of_def)
+  define \<T> where "\<T> \<equiv> connected_component_of_set X ` quasi_component_of_set X x"
+  show thesis
+  proof
+    show "\<T> \<subseteq> connected_components_of X"
+      by (simp add: \<T>_def connected_components_of_def image_mono quasi_component_of_subset_topspace)
+    show "\<Union>\<T> = C"
+      by (metis \<T>_def Ceq quasi_component_as_connected_component_Union)
+  qed
+qed
+
+lemma path_imp_quasi_component_of:
+   "path_component_of X x y \<Longrightarrow> quasi_component_of X x y"
+  by (simp add: connected_imp_quasi_component_of path_imp_connected_component_of)
+
+lemma path_component_subset_quasi_component_of:
+   "path_component_of_set X x \<subseteq> quasi_component_of_set X x"
+  by (simp add: Collect_mono path_imp_quasi_component_of)
+
+lemma connected_space_iff_quasi_component:
+   "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. quasi_component_of X x y)"
+  unfolding connected_space_clopen_in closedin_def quasi_component_of
+  by blast
+
+lemma connected_space_imp_quasi_component_of:
+   " \<lbrakk>connected_space X; a \<in> topspace X; b \<in> topspace X\<rbrakk> \<Longrightarrow> quasi_component_of X a b"
+  by (simp add: connected_space_iff_quasi_component)
+
+lemma connected_space_quasi_component_set:
+   "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. quasi_component_of_set X x = topspace X)"
+  by (metis Ball_Collect connected_space_iff_quasi_component quasi_component_of_subset_topspace subset_antisym)
+
+lemma connected_space_iff_quasi_components_eq:
+  "connected_space X \<longleftrightarrow>
+    (\<forall>C \<in> quasi_components_of X. \<forall>D \<in> quasi_components_of X. C = D)"
+  apply (simp add: quasi_components_of_def)
+  by (metis connected_space_iff_quasi_component mem_Collect_eq quasi_component_of_equiv)
+
+lemma quasi_components_of_subset_sing:
+   "quasi_components_of X \<subseteq> {S} \<longleftrightarrow> connected_space X \<and> (topspace X = {} \<or> topspace X = S)"
+proof (cases "quasi_components_of X = {}")
+  case True
+  then show ?thesis
+    by (simp add: connected_space_topspace_empty subset_singleton_iff)
+next
+  case False
+  then show ?thesis
+    apply (simp add: connected_space_iff_quasi_components_eq subset_iff Ball_def)
+    by (metis quasi_components_of_subset subsetI subset_antisym subset_empty topspace_imp_quasi_components_of)
+qed
+
+lemma connected_space_iff_quasi_components_subset_sing:
+   "connected_space X \<longleftrightarrow> (\<exists>a. quasi_components_of X \<subseteq> {a})"
+  by (simp add: quasi_components_of_subset_sing)
+
+lemma quasi_components_of_eq_singleton:
+   "quasi_components_of X = {S} \<longleftrightarrow>
+        connected_space X \<and> \<not> (topspace X = {}) \<and> S = topspace X"
+  by (metis ccpo_Sup_singleton insert_not_empty quasi_components_of_subset_sing subset_singleton_iff)
+
+lemma quasi_components_of_connected_space:
+   "connected_space X
+        \<Longrightarrow> quasi_components_of X = (if topspace X = {} then {} else {topspace X})"
+  by (simp add: quasi_components_of_eq_singleton)
+
+lemma separated_between_singletons:
+   "separated_between X {x} {y} \<longleftrightarrow>
+    x \<in> topspace X \<and> y \<in> topspace X \<and> \<not> (quasi_component_of X x y)"
+proof (cases "x \<in> topspace X \<and> y \<in> topspace X")
+  case True
+  then show ?thesis
+    by (auto simp add: separated_between_def quasi_component_of_alt)
+qed (use separated_between_imp_subset in blast)
+
+lemma quasi_component_nonseparated:
+   "quasi_component_of X x y \<longleftrightarrow> x \<in> topspace X \<and> y \<in> topspace X \<and> \<not> (separated_between X {x} {y})"
+  by (metis quasi_component_of_equiv separated_between_singletons)
+
+lemma separated_between_quasi_component_pointwise_left:
+  assumes "C \<in> quasi_components_of X"
+  shows "separated_between X C S \<longleftrightarrow> (\<exists>x \<in> C. separated_between X {x} S)"  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    using assms quasi_components_of_disjoint separated_between_mono by fastforce
+next
+  assume ?rhs
+  then obtain y where "separated_between X {y} S" and "y \<in> C"
+    by metis
+  with assms show ?lhs
+    by (force simp add: separated_between quasi_components_of_def quasi_component_of_def)
+qed
+
+lemma separated_between_quasi_component_pointwise_right:
+   "C \<in> quasi_components_of X \<Longrightarrow> separated_between X S C \<longleftrightarrow> (\<exists>x \<in> C. separated_between X S {x})"
+  by (simp add: separated_between_quasi_component_pointwise_left separated_between_sym)
+
+lemma separated_between_quasi_component_point:
+  assumes "C \<in> quasi_components_of X"
+  shows "separated_between X C {x} \<longleftrightarrow> x \<in> topspace X - C" (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (meson DiffI disjnt_insert2 insert_subset separated_between_imp_disjoint separated_between_imp_subset)
+next
+  assume ?rhs
+  with assms show ?lhs
+    unfolding quasi_components_of_def image_iff Diff_iff separated_between_quasi_component_pointwise_left [OF assms]
+    by (metis mem_Collect_eq quasi_component_of_refl separated_between_singletons)
+qed
+
+lemma separated_between_point_quasi_component:
+   "C \<in> quasi_components_of X \<Longrightarrow> separated_between X {x} C \<longleftrightarrow> x \<in> topspace X - C"
+  by (simp add: separated_between_quasi_component_point separated_between_sym)
+
+lemma separated_between_quasi_component_compact:
+   "\<lbrakk>C \<in> quasi_components_of X; compactin X K\<rbrakk> \<Longrightarrow> (separated_between X C K \<longleftrightarrow> disjnt C K)"
+  unfolding disjnt_iff
+  using compactin_subset_topspace quasi_components_of_subset separated_between_pointwise_right separated_between_quasi_component_point by fastforce
+
+lemma separated_between_compact_quasi_component:
+   "\<lbrakk>compactin X K; C \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> separated_between X K C \<longleftrightarrow> disjnt K C"
+  using disjnt_sym separated_between_quasi_component_compact separated_between_sym by blast
+
+lemma separated_between_quasi_components:
+  assumes C: "C \<in> quasi_components_of X" and D: "D \<in> quasi_components_of X"
+  shows "separated_between X C D \<longleftrightarrow> disjnt C D"   (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (simp add: separated_between_imp_disjoint)
+next
+  assume ?rhs
+  obtain x y where x: "C = quasi_component_of_set X x" and "x \<in> C"
+               and y: "D = quasi_component_of_set X y" and "y \<in> D"
+    using assms by (auto simp: quasi_components_of_def)
+  then have "separated_between X {x} {y}"
+    using \<open>disjnt C D\<close> separated_between_singletons by fastforce
+  with \<open>x \<in> C\<close> \<open>y \<in> D\<close> show ?lhs
+    by (auto simp: assms separated_between_quasi_component_pointwise_left separated_between_quasi_component_pointwise_right)
+qed
+
+lemma quasi_eq_connected_component_of_eq:
+   "quasi_component_of X x = connected_component_of X x \<longleftrightarrow>
+    connectedin X (quasi_component_of_set X x)"  (is "?lhs = ?rhs")
+proof (cases "x \<in> topspace X")
+  case True
+  show ?thesis
+  proof
+    show "?lhs \<Longrightarrow> ?rhs"
+      by (simp add: connectedin_connected_component_of)
+  next
+    assume ?rhs
+    then have "\<And>y. quasi_component_of X x y = connected_component_of X x y"
+      by (metis connected_component_of_def connected_imp_quasi_component_of mem_Collect_eq quasi_component_of_equiv)
+    then show ?lhs
+      by force
+  qed
+next
+  case False
+  then show ?thesis
+    by (metis Collect_empty_eq_bot connected_component_of_eq_empty connectedin_empty quasi_component_of_eq_empty)
+qed
+
+lemma connected_quasi_component_of:
+  assumes "C \<in> quasi_components_of X"
+  shows "C \<in> connected_components_of X \<longleftrightarrow> connectedin X C"  (is "?lhs = ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    using assms
+    by (simp add: connectedin_connected_components_of)
+next
+  assume ?rhs
+  with assms show ?lhs
+    unfolding quasi_components_of_def connected_components_of_def image_iff
+    by (metis quasi_eq_connected_component_of_eq)
+qed
+
+lemma quasi_component_of_clopen_cases:
+   "\<lbrakk>C \<in> quasi_components_of X; closedin X T; openin X T\<rbrakk> \<Longrightarrow> C \<subseteq> T \<or> disjnt C T"
+  by (smt (verit) disjnt_iff image_iff mem_Collect_eq quasi_component_of_def quasi_components_of_def subset_iff)
+
+lemma quasi_components_of_set:
+  assumes "C \<in> quasi_components_of X"
+  shows "\<Inter> {T. closedin X T \<and> openin X T \<and> C \<subseteq> T} = C"  (is "?lhs = ?rhs")
+proof
+  have "x \<in> C" if "x \<in> \<Inter> {T. closedin X T \<and> openin X T \<and> C \<subseteq> T}" for x
+  proof (rule ccontr)
+    assume "x \<notin> C"
+    have "x \<in> topspace X"
+      using assms quasi_components_of_subset that by force
+    then have "separated_between X C {x}"
+      by (simp add: \<open>x \<notin> C\<close> assms separated_between_quasi_component_point)
+    with that show False
+      by (auto simp: separated_between)
+  qed
+  then show "?lhs \<subseteq> ?rhs"
+    by auto
+qed blast
+
+lemma open_quasi_eq_connected_components_of:
+  assumes "openin X C"
+  shows "C \<in> quasi_components_of X \<longleftrightarrow> C \<in> connected_components_of X"  (is "?lhs = ?rhs")
+proof (cases "closedin X C")
+  case True
+  show ?thesis
+  proof
+    assume L: ?lhs
+    have "T = {} \<or> T = topspace X \<inter> C"
+      if "openin (subtopology X C) T" "closedin (subtopology X C) T" for T
+    proof -
+      have "C \<subseteq> T \<or> disjnt C T"
+        by (meson L True assms closedin_trans_full openin_trans_full quasi_component_of_clopen_cases that)
+      with that show ?thesis
+        by (metis Int_absorb2 True closedin_imp_subset closure_of_subset_eq disjnt_def inf_absorb2)
+    qed
+    with L assms show "?rhs"
+      by (simp add: connected_quasi_component_of connected_space_clopen_in connectedin_def openin_subset)
+  next
+    assume ?rhs
+    then obtain x where "x \<in> topspace X" and x: "C = connected_component_of_set X x"
+      by (metis connected_components_of_def imageE)
+    have "C = quasi_component_of_set X x"
+      using True assms connected_component_of_refl connected_imp_quasi_component_of quasi_component_of_def x by fastforce
+    then show ?lhs
+      using \<open>x \<in> topspace X\<close> quasi_components_of_def by fastforce
+  qed
+next
+  case False
+  then show ?thesis
+    using closedin_connected_components_of closedin_quasi_components_of by blast
+qed
+
+lemma quasi_component_of_continuous_image:
+  assumes f:  "continuous_map X Y f" and qc: "quasi_component_of X x y"
+  shows "quasi_component_of Y (f x) (f y)"
+  unfolding quasi_component_of_def
+proof (intro strip conjI)
+  show "f x \<in> topspace Y" "f y \<in> topspace Y"
+    using assms by (simp_all add: continuous_map_def quasi_component_of_def)
+  fix T
+  assume "closedin Y T \<and> openin Y T"
+  with assms show "(f x \<in> T) = (f y \<in> T)"
+    by (smt (verit) continuous_map_closedin continuous_map_def mem_Collect_eq quasi_component_of_def)
+qed
+
+lemma quasi_component_of_discrete_topology:
+   "quasi_component_of_set (discrete_topology U) x = (if x \<in> U then {x} else {})"
+proof -
+  have "quasi_component_of_set (discrete_topology U) y = {y}" if "y \<in> U" for y
+    using that
+    apply (simp add: set_eq_iff quasi_component_of_def)
+    by (metis Set.set_insert insertE subset_insertI)
+  then show ?thesis
+    by (simp add: quasi_component_of)
+qed
+
+lemma quasi_components_of_discrete_topology:
+   "quasi_components_of (discrete_topology U) = (\<lambda>x. {x}) ` U"
+  by (auto simp add: quasi_components_of_def quasi_component_of_discrete_topology)
+
+lemma homeomorphic_map_quasi_component_of:
+  assumes hmf: "homeomorphic_map X Y f" and "x \<in> topspace X"
+  shows "quasi_component_of_set Y (f x) = f ` (quasi_component_of_set X x)"
+proof -
+  obtain g where hmg: "homeomorphic_map Y X g"
+    and contf: "continuous_map X Y f" and contg: "continuous_map Y X g"
+    and fg: "(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> topspace Y. f(g y) = y)"
+    by (smt (verit, best) hmf homeomorphic_map_maps homeomorphic_maps_def)
+  show ?thesis
+  proof
+    show "quasi_component_of_set Y (f x) \<subseteq> f ` quasi_component_of_set X x"
+      using quasi_component_of_continuous_image [OF contg]
+         \<open>x \<in> topspace X\<close> fg image_iff quasi_component_of_subset_topspace by fastforce
+    show "f ` quasi_component_of_set X x \<subseteq> quasi_component_of_set Y (f x)"
+      using quasi_component_of_continuous_image [OF contf] by blast
+  qed
+qed
+
+
+lemma homeomorphic_map_quasi_components_of:
+  assumes "homeomorphic_map X Y f"
+  shows "quasi_components_of Y = image (image f) (quasi_components_of X)"
+  using assms
+proof -
+  have "\<exists>x\<in>topspace X. quasi_component_of_set Y y = f ` quasi_component_of_set X x"
+    if "y \<in> topspace Y" for y 
+    by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of image_iff)
+  moreover have "\<exists>x\<in>topspace Y. f ` quasi_component_of_set X u = quasi_component_of_set Y x"
+    if  "u \<in> topspace X" for u
+    by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of imageI)
+  ultimately show ?thesis
+    by (auto simp: quasi_components_of_def image_iff)
+qed
+
+lemma openin_quasi_component_of_locally_connected_space:
+  assumes "locally_connected_space X"
+  shows "openin X (quasi_component_of_set X x)"
+proof -
+  have *: "openin X (connected_component_of_set X x)"
+    by (simp add: assms openin_connected_component_of_locally_connected_space)
+  moreover have "connected_component_of_set X x = quasi_component_of_set X x"
+    using * closedin_connected_component_of connected_component_of_refl connected_imp_quasi_component_of
+            quasi_component_of_def by fastforce
+  ultimately show ?thesis
+    by simp
+qed
+
+lemma openin_quasi_components_of_locally_connected_space:
+   "locally_connected_space X \<and> c \<in> quasi_components_of X
+        \<Longrightarrow> openin X c"
+  by (smt (verit, best) image_iff openin_quasi_component_of_locally_connected_space quasi_components_of_def)
+
+lemma quasi_eq_connected_components_of_alt:
+  "quasi_components_of X = connected_components_of X \<longleftrightarrow> (\<forall>C \<in> quasi_components_of X. connectedin X C)"
+  (is "?lhs = ?rhs")
+proof
+  assume R: ?rhs
+  moreover have "connected_components_of X \<subseteq> quasi_components_of X"
+    using R unfolding quasi_components_of_def connected_components_of_def
+    by (force simp flip: quasi_eq_connected_component_of_eq)
+  ultimately show ?lhs
+    using connected_quasi_component_of by blast
+qed (use connected_quasi_component_of in blast)
+  
+lemma connected_subset_quasi_components_of_pointwise:
+   "connected_components_of X \<subseteq> quasi_components_of X \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. quasi_component_of X x = connected_component_of X x)"
+  (is "?lhs = ?rhs")
+proof
+  assume L: ?lhs
+  have "connectedin X (quasi_component_of_set X x)" if "x \<in> topspace X" for x
+  proof -
+    have "\<exists>y\<in>topspace X. connected_component_of_set X x = quasi_component_of_set X y"
+      using L that by (force simp: quasi_components_of_def connected_components_of_def image_subset_iff)
+    then show ?thesis
+      by (metis connected_component_of_equiv connectedin_connected_component_of mem_Collect_eq quasi_component_of_eq)
+  qed
+  then show ?rhs
+    by (simp add: quasi_eq_connected_component_of_eq)
+qed (simp add: connected_components_of_def quasi_components_of_def)
+
+lemma quasi_subset_connected_components_of_pointwise:
+   "quasi_components_of X \<subseteq> connected_components_of X \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. quasi_component_of X x = connected_component_of X x)"
+  by (simp add: connected_quasi_component_of image_subset_iff quasi_components_of_def quasi_eq_connected_component_of_eq)
+
+lemma quasi_eq_connected_components_of_pointwise:
+   "quasi_components_of X = connected_components_of X \<longleftrightarrow>
+    (\<forall>x \<in> topspace X. quasi_component_of X x = connected_component_of X x)"
+  using connected_subset_quasi_components_of_pointwise quasi_subset_connected_components_of_pointwise by fastforce
+
+lemma quasi_eq_connected_components_of_pointwise_alt:
+   "quasi_components_of X = connected_components_of X \<longleftrightarrow>
+    (\<forall>x. quasi_component_of X x = connected_component_of X x)"
+  unfolding quasi_eq_connected_components_of_pointwise
+  by (metis connectedin_empty quasi_component_of_eq_empty quasi_eq_connected_component_of_eq)
+
+lemma quasi_eq_connected_components_of_inclusion:
+   "quasi_components_of X = connected_components_of X \<longleftrightarrow>
+        connected_components_of X \<subseteq> quasi_components_of X \<or>
+        quasi_components_of X \<subseteq> connected_components_of X"
+  by (simp add: connected_subset_quasi_components_of_pointwise dual_order.eq_iff quasi_subset_connected_components_of_pointwise)
+
+
+lemma quasi_eq_connected_components_of:
+  "finite(connected_components_of X) \<or>
+      finite(quasi_components_of X) \<or>
+      locally_connected_space X \<or>
+      compact_space X \<and> (Hausdorff_space X \<or> regular_space X \<or> normal_space X)
+      \<Longrightarrow> quasi_components_of X = connected_components_of X"
+proof (elim disjE)
+  show "quasi_components_of X = connected_components_of X"
+    if "finite (connected_components_of X)"
+    unfolding quasi_eq_connected_components_of_inclusion
+    using that open_in_finite_connected_components open_quasi_eq_connected_components_of by blast
+  show "quasi_components_of X = connected_components_of X"
+    if "finite (quasi_components_of X)"
+    unfolding quasi_eq_connected_components_of_inclusion
+    using that open_quasi_eq_connected_components_of openin_finite_quasi_components by blast 
+  show "quasi_components_of X = connected_components_of X"
+    if "locally_connected_space X"
+    unfolding quasi_eq_connected_components_of_inclusion
+    using that open_quasi_eq_connected_components_of openin_quasi_components_of_locally_connected_space by auto 
+  show "quasi_components_of X = connected_components_of X"
+    if "compact_space X \<and> (Hausdorff_space X \<or> regular_space X \<or> normal_space X)"
+  proof -
+    show ?thesis
+      unfolding quasi_eq_connected_components_of_alt
+    proof (intro strip)
+      fix C
+      assume C: "C \<in> quasi_components_of X"
+      then have cloC: "closedin X C"
+        by (simp add: closedin_quasi_components_of)
+      have "normal_space X"
+        using that compact_Hausdorff_or_regular_imp_normal_space by blast
+      show "connectedin X C"
+      proof (clarsimp simp add: connectedin_def connected_space_closedin_eq closedin_closed_subtopology cloC closedin_subset [OF cloC])
+        fix S T
+        assume "S \<subseteq> C" and "closedin X S" and "S \<inter> T = {}" and SUT: "S \<union> T = topspace X \<inter> C"
+          and T: "T \<subseteq> C" "T \<noteq> {}" and "closedin X T" 
+        with \<open>normal_space X\<close> obtain U V where UV: "openin X U" "openin X V" "S \<subseteq> U" "T \<subseteq> V" "disjnt U V"
+          by (meson disjnt_def normal_space_def)
+        moreover have "compactin X (topspace X - (U \<union> V))"
+          using UV that by (intro closedin_compact_space closedin_diff openin_Un) auto
+        ultimately have "separated_between X C (topspace X - (U \<union> V)) \<longleftrightarrow> disjnt C (topspace X - (U \<union> V))"
+          by (simp add: \<open>C \<in> quasi_components_of X\<close> separated_between_quasi_component_compact)
+        moreover have "disjnt C (topspace X - (U \<union> V))"
+          using UV SUT disjnt_def by fastforce
+        ultimately have "separated_between X C (topspace X - (U \<union> V))"
+          by simp
+        then obtain A B where "openin X A" "openin X B" "A \<union> B = topspace X" "disjnt A B" "C \<subseteq> A" 
+                        and subB: "topspace X - (U \<union> V) \<subseteq> B"
+          by (meson separated_between_def)
+        have "B \<union> U = topspace X - (A \<inter> V)"
+        proof
+          show "B \<union> U \<subseteq> topspace X - A \<inter> V"
+            using \<open>openin X U\<close> \<open>disjnt U V\<close> \<open>disjnt A B\<close> \<open>openin X B\<close> disjnt_iff openin_closedin_eq by fastforce
+          show "topspace X - A \<inter> V \<subseteq> B \<union> U"
+            using \<open>A \<union> B = topspace X\<close> subB by fastforce
+        qed
+        then have "closedin X (B \<union> U)"
+          using \<open>openin X V\<close> \<open>openin X A\<close> by auto
+        then have "C \<subseteq> B \<union> U \<or> disjnt C (B \<union> U)"
+          using quasi_component_of_clopen_cases [OF C] \<open>openin X U\<close> \<open>openin X B\<close> by blast
+        with UV show "S = {}"
+          by (metis UnE \<open>C \<subseteq> A\<close> \<open>S \<subseteq> C\<close> T \<open>disjnt A B\<close> all_not_in_conv disjnt_Un2 disjnt_iff subset_eq)
+      qed
+    qed
+  qed
+qed
+
+
+lemma quasi_eq_connected_component_of:
+   "finite(connected_components_of X) \<or>
+      finite(quasi_components_of X) \<or>
+      locally_connected_space X \<or>
+      compact_space X \<and> (Hausdorff_space X \<or> regular_space X \<or> normal_space X)
+      \<Longrightarrow> quasi_component_of X x = connected_component_of X x"
+  by (metis quasi_eq_connected_components_of quasi_eq_connected_components_of_pointwise_alt)
+
+
+subsection\<open>Additional quasicomponent and continuum properties like Boundary Bumping\<close>
+
+lemma cut_wire_fence_theorem_gen:
+  assumes "compact_space X" and X: "Hausdorff_space X \<or> regular_space X \<or> normal_space X"
+    and S: "compactin X S" and T: "closedin X T"
+    and dis: "\<And>C. connectedin X C \<Longrightarrow> disjnt C S \<or> disjnt C T"
+  shows "separated_between X S T"
+  proof -
+  have "x \<in> topspace X" if "x \<in> S" and "T = {}" for x
+    using that S compactin_subset_topspace by auto
+  moreover have "separated_between X {x} {y}" if "x \<in> S" and "y \<in> T" for x y
+  proof (cases "x \<in> topspace X \<and> y \<in> topspace X")
+    case True
+    then have "\<not> connected_component_of X x y"
+      by (meson dis connected_component_of_def disjnt_iff that)
+    with True X \<open>compact_space X\<close> show ?thesis
+      by (metis quasi_component_nonseparated quasi_eq_connected_component_of)
+  next
+    case False
+    then show ?thesis
+      using S T compactin_subset_topspace closedin_subset that by blast
+  qed
+  ultimately show ?thesis
+    using assms
+    by (simp add: separated_between_pointwise_left separated_between_pointwise_right 
+              closedin_compact_space closedin_subset)
+qed
+
+lemma cut_wire_fence_theorem:
+   "\<lbrakk>compact_space X; Hausdorff_space X; closedin X S; closedin X T;
+     \<And>C. connectedin X C \<Longrightarrow> disjnt C S \<or> disjnt C T\<rbrakk>
+        \<Longrightarrow> separated_between X S T"
+  by (simp add: closedin_compact_space cut_wire_fence_theorem_gen)
+
+lemma separated_between_from_closed_subtopology:
+  assumes XC: "separated_between (subtopology X C) S (X frontier_of C)" 
+    and ST: "separated_between (subtopology X C) S T"
+  shows "separated_between X S T"
+proof -
+  obtain U where clo: "closedin (subtopology X C) U" and ope: "openin (subtopology X C) U" 
+             and "S \<subseteq> U" and sub: "X frontier_of C \<union> T \<subseteq> topspace (subtopology X C) - U"
+    by (meson assms separated_between separated_between_Un)
+  then have "X frontier_of C \<union> T \<subseteq> topspace X \<inter> C - U"
+    by auto
+  have "closedin X (topspace X \<inter> C)"
+    by (metis XC frontier_of_restrict frontier_of_subset_eq inf_le1 separated_between_imp_subset topspace_subtopology)
+  then have "closedin X U"
+    by (metis clo closedin_closed_subtopology subtopology_restrict)
+  moreover have "openin (subtopology X C) U \<longleftrightarrow> openin X U \<and> U \<subseteq> C"
+    using disjnt_iff sub by (force intro!: openin_subset_topspace_eq)
+  with ope have "openin X U"
+    by blast
+  moreover have "T \<subseteq> topspace X - U"
+    using ope openin_closedin_eq sub by auto
+  ultimately show ?thesis
+    using \<open>S \<subseteq> U\<close> separated_between by blast
+qed
+
+lemma separated_between_from_closed_subtopology_frontier:
+   "separated_between (subtopology X T) S (X frontier_of T)
+        \<Longrightarrow> separated_between X S (X frontier_of T)"
+  using separated_between_from_closed_subtopology by blast
+
+lemma separated_between_from_frontier_of_closed_subtopology:
+  assumes "separated_between (subtopology X T) S (X frontier_of T)"
+  shows "separated_between X S (topspace X - T)"
+proof -
+  have "disjnt S (topspace X - T)"
+    using assms disjnt_iff separated_between_imp_subset by fastforce
+  then show ?thesis
+    by (metis Diff_subset assms frontier_of_complement separated_between_from_closed_subtopology separated_between_frontier_of_eq')
+qed
+
+lemma separated_between_compact_connected_component:
+  assumes "locally_compact_space X" "Hausdorff_space X" 
+    and C: "C \<in> connected_components_of X" 
+    and "compactin X C" "closedin X T" "disjnt C T"
+  shows "separated_between X C T"
+proof -
+  have Csub: "C \<subseteq> topspace X"
+    by (simp add: assms(4) compactin_subset_topspace)
+  have "Hausdorff_space (subtopology X (topspace X - T))"
+    using Hausdorff_space_subtopology assms(2) by blast
+  moreover have "compactin (subtopology X (topspace X - T)) C"
+    using assms Csub by (metis Diff_Int_distrib Diff_empty compact_imp_compactin_subtopology disjnt_def le_iff_inf)
+  moreover have "locally_compact_space (subtopology X (topspace X - T))"
+    by (meson assms closedin_def locally_compact_Hausdorff_imp_regular_space locally_compact_space_open_subset)
+  ultimately
+  obtain N L where "openin X N" "compactin X L" "closedin X L" "C \<subseteq> N" "N \<subseteq> L" 
+    and Lsub: "L \<subseteq> topspace X - T"
+    using \<open>Hausdorff_space X\<close> \<open>closedin X T\<close>
+    apply (simp add: locally_compact_space_compact_closed_compact compactin_subtopology)
+    by (meson closedin_def compactin_imp_closedin  openin_trans_full)
+  then have disC: "disjnt C (topspace X - L)"
+    by (meson DiffD2 disjnt_iff subset_iff)
+  have "separated_between (subtopology X L) C (X frontier_of L)"
+  proof (rule cut_wire_fence_theorem)
+    show "compact_space (subtopology X L)"
+      by (simp add: \<open>compactin X L\<close> compact_space_subtopology)
+    show "Hausdorff_space (subtopology X L)"
+      by (simp add: Hausdorff_space_subtopology \<open>Hausdorff_space X\<close>)
+    show "closedin (subtopology X L) C"
+      by (meson \<open>C \<subseteq> N\<close> \<open>N \<subseteq> L\<close> \<open>Hausdorff_space X\<close> \<open>compactin X C\<close> closedin_subset_topspace compactin_imp_closedin subset_trans)
+    show "closedin (subtopology X L) (X frontier_of L)"
+      by (simp add: \<open>closedin X L\<close> closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin)
+    show "disjnt D C \<or> disjnt D (X frontier_of L)"
+      if "connectedin (subtopology X L) D" for D 
+    proof (rule ccontr)
+      assume "\<not> (disjnt D C \<or> disjnt D (X frontier_of L))"
+      moreover have "connectedin X D"
+        using connectedin_subtopology that by blast
+      ultimately show False
+        using that connected_components_of_maximal [of C X D] C
+        apply (simp add: disjnt_iff)
+        by (metis Diff_eq_empty_iff \<open>C \<subseteq> N\<close> \<open>N \<subseteq> L\<close> \<open>openin X N\<close> disjoint_iff frontier_of_openin_straddle_Int(2) subsetD)
+    qed
+  qed
+  then have "separated_between X (X frontier_of C) (topspace X - L)"
+    using separated_between_from_frontier_of_closed_subtopology separated_between_frontier_of_eq by blast
+  with \<open>closedin X T\<close>  
+    separated_between_frontier_of [OF Csub disC] 
+  show ?thesis
+    unfolding separated_between by (smt (verit) Diff_iff Lsub closedin_subset subset_iff)
+qed
+
+lemma wilder_locally_compact_component_thm:
+  assumes "locally_compact_space X" "Hausdorff_space X" 
+    and "C \<in> connected_components_of X" "compactin X C" "openin X W" "C \<subseteq> W"
+  obtains U V where "openin X U" "openin X V" "disjnt U V" "U \<union> V = topspace X" "C \<subseteq> U" "U \<subseteq> W"
+proof -
+  have "closedin X (topspace X - W)"
+    using \<open>openin X W\<close> by blast
+  moreover have "disjnt C (topspace X - W)"
+    using \<open>C \<subseteq> W\<close> disjnt_def by fastforce
+  ultimately have "separated_between X C (topspace X - W)"
+    using separated_between_compact_connected_component assms by blast
+  then show thesis
+    by (smt (verit, del_insts) DiffI disjnt_iff openin_subset separated_between_def subset_iff that)
+qed
+
+lemma compact_quasi_eq_connected_components_of:
+  assumes "locally_compact_space X" "Hausdorff_space X" "compactin X C"
+  shows "C \<in> quasi_components_of X \<longleftrightarrow> C \<in> connected_components_of X"
+proof -
+  have "compactin X (connected_component_of_set X x)" 
+    if "x \<in> topspace X" "compactin X (quasi_component_of_set X x)" for x
+  proof (rule closed_compactin)
+    show "compactin X (quasi_component_of_set X x)"
+      by (simp add: that)
+    show "connected_component_of_set X x \<subseteq> quasi_component_of_set X x"
+      by (simp add: connected_component_subset_quasi_component_of)
+    show "closedin X (connected_component_of_set X x)"
+      by (simp add: closedin_connected_component_of)
+  qed
+  moreover have "connected_component_of X x = quasi_component_of X x"
+    if \<section>: "x \<in> topspace X" "compactin X (connected_component_of_set X x)" for x
+  proof -
+    have "\<And>y. connected_component_of X x y \<Longrightarrow> quasi_component_of X x y"
+      by (simp add: connected_imp_quasi_component_of)
+    moreover have False if non: "\<not> connected_component_of X x y" and quasi: "quasi_component_of X x y" for y
+    proof -
+      have "y \<in> topspace X"
+        by (meson quasi_component_of_equiv that)
+      then have "closedin X {y}"
+        by (simp add: \<open>Hausdorff_space X\<close> compactin_imp_closedin)
+      moreover have "disjnt (connected_component_of_set X x) {y}"
+        by (simp add: non)
+      moreover have "\<not> separated_between X (connected_component_of_set X x) {y}"
+        using \<section> quasi separated_between_pointwise_left 
+        by (fastforce simp: quasi_component_nonseparated connected_component_of_refl)
+      ultimately show False
+        using assms by (metis \<section> connected_component_in_connected_components_of separated_between_compact_connected_component)
+    qed
+    ultimately show ?thesis
+      by blast
+  qed
+  ultimately show ?thesis
+    using \<open>compactin X C\<close> unfolding connected_components_of_def image_iff quasi_components_of_def by metis
+qed
+
+
+lemma boundary_bumping_theorem_closed_gen:
+  assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" "closedin X S" 
+    "S \<noteq> topspace X" and C: "compactin X C" "C \<in> connected_components_of (subtopology X S)"
+  shows "C \<inter> X frontier_of S \<noteq> {}"
+proof 
+  assume \<section>: "C \<inter> X frontier_of S = {}"
+  consider "C \<noteq> {}" "X frontier_of S \<subseteq> topspace X" | "C \<subseteq> topspace X" "S = {}"
+    using C by (metis frontier_of_subset_topspace nonempty_connected_components_of)
+  then show False
+  proof cases
+    case 1
+    have "separated_between (subtopology X S) C (X frontier_of S)"
+    proof (rule separated_between_compact_connected_component)
+      show "compactin (subtopology X S) C"
+        using C compact_imp_compactin_subtopology connected_components_of_subset by fastforce
+      show "closedin (subtopology X S) (X frontier_of S)"
+        by (simp add: \<open>closedin X S\<close> closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin)
+      show "disjnt C (X frontier_of S)"
+        using \<section> by (simp add: disjnt_def)
+    qed (use assms Hausdorff_space_subtopology locally_compact_space_closed_subset in auto)
+    then have "separated_between X C (X frontier_of S)"
+      using separated_between_from_closed_subtopology by auto
+    then have "X frontier_of S = {}"
+      using \<open>C \<noteq> {}\<close> \<open>connected_space X\<close> connected_space_separated_between by blast
+    moreover have "C \<subseteq> S"
+      using C connected_components_of_subset by fastforce
+    ultimately show False
+      using 1 assms by (metis closedin_subset connected_space_eq_frontier_eq_empty subset_empty)
+  next
+    case 2
+    then show False
+      using C connected_components_of_eq_empty by fastforce
+  qed
+qed
+
+lemma boundary_bumping_theorem_closed:
+  assumes "connected_space X" "compact_space X" "Hausdorff_space X" "closedin X S" 
+          "S \<noteq> topspace X" "C \<in> connected_components_of(subtopology X S)"
+  shows "C \<inter> X frontier_of S \<noteq> {}"
+  by (meson assms boundary_bumping_theorem_closed_gen closedin_compact_space closedin_connected_components_of
+            closedin_trans_full compact_imp_locally_compact_space)
+
+
+lemma intermediate_continuum_exists:
+  assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" 
+    and C: "compactin X C" "connectedin X C" "C \<noteq> {}" "C \<noteq> topspace X"
+    and U: "openin X U" "C \<subseteq> U"
+  obtains D where "compactin X D" "connectedin X D" "C \<subset> D" "D \<subset> U"
+proof -
+  have "C \<subseteq> topspace X"
+    by (simp add: C compactin_subset_topspace)
+  with C obtain a where a: "a \<in> topspace X" "a \<notin> C"
+    by blast
+  moreover have "compactin (subtopology X (U - {a})) C"
+    by (simp add: C U a compact_imp_compactin_subtopology subset_Diff_insert)
+  moreover have "Hausdorff_space (subtopology X (U - {a}))"
+    using Hausdorff_space_subtopology assms(3) by blast
+  moreover
+  have "locally_compact_space (subtopology X (U - {a}))"
+    by (rule locally_compact_space_open_subset)
+       (auto simp: locally_compact_Hausdorff_imp_regular_space open_in_Hausdorff_delete assms)
+  ultimately obtain V K where V: "openin X V" "a \<notin> V" "V \<subseteq> U" and K: "compactin X K" "a \<notin> K" "K \<subseteq> U" 
+    and cloK: "closedin (subtopology X (U - {a})) K" and "C \<subseteq> V" "V \<subseteq> K"
+    using locally_compact_space_compact_closed_compact [of "subtopology X (U - {a})"] assms
+    by (smt (verit, del_insts) Diff_empty compactin_subtopology open_in_Hausdorff_delete openin_open_subtopology subset_Diff_insert)
+  then obtain D where D: "D \<in> connected_components_of (subtopology X K)" and "C \<subseteq> D"
+    using C by (metis bot.extremum_unique connectedin_subtopology order.trans exists_connected_component_of_superset subtopology_topspace)
+  show thesis
+  proof
+    have cloD: "closedin (subtopology X K) D"
+      by (simp add: D closedin_connected_components_of)
+    then have XKD: "compactin (subtopology X K) D"
+      by (simp add: K closedin_compact_space compact_space_subtopology)
+    then show "compactin X D"
+      using compactin_subtopology_imp_compact by blast
+    show "connectedin X D"
+      using D connectedin_connected_components_of connectedin_subtopology by blast
+    have "K \<noteq> topspace X"
+      using K a by blast
+    moreover have "V \<subseteq> X interior_of K"
+      by (simp add: \<open>openin X V\<close> \<open>V \<subseteq> K\<close> interior_of_maximal)
+    ultimately have "C \<noteq> D"
+      using boundary_bumping_theorem_closed_gen [of X K C] D \<open>C \<subseteq> V\<close> 
+      by (auto simp add: assms K compactin_imp_closedin frontier_of_def)
+    then show "C \<subset> D"
+      using \<open>C \<subseteq> D\<close> by blast
+    have "D \<subseteq> U"
+      using K(3) \<open>closedin (subtopology X K) D\<close> closedin_imp_subset by blast
+    moreover have "D \<noteq> U"
+      using K XKD \<open>C \<subset> D\<close> assms
+      by (metis \<open>K \<noteq> topspace X\<close> cloD closedin_imp_subset compactin_imp_closedin connected_space_clopen_in
+                inf_bot_left inf_le2 subset_antisym)
+    ultimately
+    show "D \<subset> U" by blast
+  qed
+qed
+
+lemma boundary_bumping_theorem_gen:
+  assumes X: "connected_space X" "locally_compact_space X" "Hausdorff_space X" 
+   and "S \<subset> topspace X" and C: "C \<in> connected_components_of(subtopology X S)" 
+   and compC: "compactin X (X closure_of C)"
+ shows "X frontier_of C \<inter> X frontier_of S \<noteq> {}"
+proof -
+  have Csub: "C \<subseteq> topspace X" "C \<subseteq> S" and "connectedin X C"
+    using C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology
+    by fastforce+
+  have "C \<noteq> {}"
+    using C nonempty_connected_components_of by blast
+  obtain "X interior_of C \<subseteq> X interior_of S" "X closure_of C \<subseteq> X closure_of S"
+    by (simp add: Csub closure_of_mono interior_of_mono)
+  moreover have False if "X closure_of C \<subseteq> X interior_of S"
+  proof -
+    have "X closure_of C = C"
+      by (meson C closedin_connected_component_of_subtopology closure_of_eq interior_of_subset order_trans that)
+    with that have "C \<subseteq> X interior_of S"
+      by simp
+    then obtain D where  "compactin X D" and "connectedin X D" and "C \<subset> D" and "D \<subset> X interior_of S"
+      using intermediate_continuum_exists assms  \<open>X closure_of C = C\<close> compC Csub
+      by (metis \<open>C \<noteq> {}\<close> \<open>connectedin X C\<close> openin_interior_of psubsetE)
+    then have "D \<subseteq> C"
+      by (metis C \<open>C \<noteq> {}\<close> connected_components_of_maximal connectedin_subtopology disjnt_def inf.orderE interior_of_subset order_trans psubsetE)
+    then show False
+      using \<open>C \<subset> D\<close> by blast
+  qed
+  ultimately show ?thesis
+    by (smt (verit, ccfv_SIG) DiffI disjoint_iff_not_equal frontier_of_def subset_eq)
+qed
+
+lemma boundary_bumping_theorem:
+   "\<lbrakk>connected_space X; compact_space X; Hausdorff_space X; S \<subset> topspace X; 
+     C \<in> connected_components_of(subtopology X S)\<rbrakk>
+    \<Longrightarrow> X frontier_of C \<inter> X frontier_of S \<noteq> {}"
+  by (simp add: boundary_bumping_theorem_gen closedin_compact_space compact_imp_locally_compact_space)
+
+subsection \<open>Compactly generated spaces (k-spaces)\<close>
+
+text \<open>These don't have to be Hausdorff\<close>
+
+definition k_space where
+  "k_space X \<equiv>
+    \<forall>S. S \<subseteq> topspace X \<longrightarrow> 
+        (closedin X S \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> S)))"
+
+lemma k_space:
+   "k_space X \<longleftrightarrow>
+    (\<forall>S. S \<subseteq> topspace X \<and>
+         (\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> S)) \<longrightarrow> closedin X S)"
+  by (metis closedin_subtopology inf_commute k_space_def)
+
+lemma k_space_open:
+   "k_space X \<longleftrightarrow>
+    (\<forall>S. S \<subseteq> topspace X \<and>
+         (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S)) \<longrightarrow> openin X S)"
+proof -
+  have "openin X S"
+    if "k_space X" "S \<subseteq> topspace X"
+      and "\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S)" for S
+    using that unfolding k_space openin_closedin_eq
+    by (metis Diff_Int_distrib2 Diff_subset inf_commute topspace_subtopology)
+  moreover have "k_space X"
+    if "\<forall>S. S \<subseteq> topspace X \<and> (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S)) \<longrightarrow> openin X S"
+    unfolding k_space openin_closedin_eq
+    by (simp add: Diff_Int_distrib closedin_def inf_commute that)
+  ultimately show ?thesis
+    by blast
+qed
+
+lemma k_space_alt:
+   "k_space X \<longleftrightarrow>
+    (\<forall>S. S \<subseteq> topspace X
+        \<longrightarrow> (openin X S \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S))))"
+  by (meson k_space_open openin_subtopology_Int2)
+
+lemma k_space_quotient_map_image:
+  assumes q: "quotient_map X Y q" and X: "k_space X"
+  shows "k_space Y"
+  unfolding k_space
+proof clarify
+  fix S
+  assume "S \<subseteq> topspace Y" and S: "\<forall>K. compactin Y K \<longrightarrow> closedin (subtopology Y K) (K \<inter> S)"
+  then have iff: "closedin X {x \<in> topspace X. q x \<in> S} \<longleftrightarrow> closedin Y S"
+    using q quotient_map_closedin by fastforce
+  have "closedin (subtopology X K) (K \<inter> {x \<in> topspace X. q x \<in> S})" if "compactin X K" for K
+  proof -
+    have "{x \<in> topspace X. q x \<in> q ` K} \<inter> K = K"
+      using compactin_subset_topspace that by blast
+    then have *: "subtopology X K = subtopology (subtopology X {x \<in> topspace X. q x \<in> q ` K}) K"
+      by (simp add: subtopology_subtopology)
+    have **: "K \<inter> {x \<in> topspace X. q x \<in> S} =
+              K \<inter> {x \<in> topspace (subtopology X {x \<in> topspace X. q x \<in> q ` K}). q x \<in> q ` K \<inter> S}"
+      by auto
+    have "K \<subseteq> topspace X"
+      by (simp add: compactin_subset_topspace that)
+    show ?thesis
+      unfolding * **
+    proof (intro closedin_continuous_map_preimage closedin_subtopology_Int_closed)
+      show "continuous_map (subtopology X {x \<in> topspace X. q x \<in> q ` K}) (subtopology Y (q ` K)) q"
+        by (auto simp add: continuous_map_in_subtopology continuous_map_from_subtopology q quotient_imp_continuous_map)
+      show "closedin (subtopology Y (q ` K)) (q ` K \<inter> S)"
+        by (meson S image_compactin q quotient_imp_continuous_map that)
+    qed
+  qed
+  then have "closedin X {x \<in> topspace X. q x \<in> S}"
+    by (metis (no_types, lifting) X k_space mem_Collect_eq subsetI)
+  with iff show "closedin Y S" by simp
+qed
+
+lemma k_space_retraction_map_image:
+   "\<lbrakk>retraction_map X Y r; k_space X\<rbrakk> \<Longrightarrow> k_space Y"
+  using k_space_quotient_map_image retraction_imp_quotient_map by blast
+
+lemma homeomorphic_k_space:
+   "X homeomorphic_space Y \<Longrightarrow> k_space X \<longleftrightarrow> k_space Y"
+  by (meson homeomorphic_map_def homeomorphic_space homeomorphic_space_sym k_space_quotient_map_image)
+
+lemma k_space_perfect_map_image:
+   "\<lbrakk>k_space X; perfect_map X Y f\<rbrakk> \<Longrightarrow> k_space Y"
+  using k_space_quotient_map_image perfect_imp_quotient_map by blast
+
+lemma locally_compact_imp_k_space:
+  assumes "locally_compact_space X"
+  shows "k_space X"
+  unfolding k_space
+proof clarify
+  fix S
+  assume "S \<subseteq> topspace X" and S: "\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> S)"
+  have False if non: "\<not> (X closure_of S \<subseteq> S)"
+  proof -
+    obtain x where "x \<in> X closure_of S" "x \<notin> S"
+      using non by blast
+    then have "x \<in> topspace X"
+      by (simp add: in_closure_of)
+    then obtain K U where "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K"
+      by (meson assms locally_compact_space_def)
+    then show False
+      using \<open>x \<in> X closure_of S\<close>  openin_Int_closure_of_eq [OF \<open>openin X U\<close>]
+      by (smt (verit, ccfv_threshold) Int_iff S \<open>x \<notin> S\<close> closedin_Int_closure_of inf.orderE inf_assoc)
+  qed
+  then show "closedin X S"
+    using S \<open>S \<subseteq> topspace X\<close> closure_of_subset_eq by blast
+qed
+
+lemma compact_imp_k_space:
+   "compact_space X \<Longrightarrow> k_space X"
+  by (simp add: compact_imp_locally_compact_space locally_compact_imp_k_space)
+
+lemma k_space_discrete_topology: "k_space(discrete_topology U)"
+  by (simp add: k_space_open)
+
+lemma k_space_closed_subtopology:
+  assumes "k_space X" "closedin X C"
+  shows "k_space (subtopology X C)"
+unfolding k_space compactin_subtopology
+proof clarsimp
+  fix S
+  assume Ssub: "S \<subseteq> topspace X" "S \<subseteq> C"
+      and S: "\<forall>K. compactin X K \<and> K \<subseteq> C \<longrightarrow> closedin (subtopology (subtopology X C) K) (K \<inter> S)"
+  have "closedin (subtopology X K) (K \<inter> S)" if "compactin X K" for K
+  proof -
+    have "closedin (subtopology (subtopology X C) (K \<inter> C)) ((K \<inter> C) \<inter> S)"
+      by (simp add: S \<open>closedin X C\<close> compact_Int_closedin that)
+    then show ?thesis
+      using \<open>closedin X C\<close> Ssub by (auto simp add: closedin_subtopology)
+  qed
+  then show "closedin (subtopology X C) S"
+    by (metis Ssub \<open>k_space X\<close> closedin_subset_topspace k_space_def)
+qed
+
+lemma k_space_subtopology:
+  assumes 1: "\<And>T. \<lbrakk>T \<subseteq> topspace X; T \<subseteq> S;
+                   \<And>K. compactin X K \<Longrightarrow> closedin (subtopology X (K \<inter> S)) (K \<inter> T)\<rbrakk> \<Longrightarrow> closedin (subtopology X S) T"
+  assumes 2: "\<And>K. compactin X K \<Longrightarrow> k_space(subtopology X (K \<inter> S))"
+  shows "k_space (subtopology X S)"
+  unfolding k_space
+proof (intro conjI strip)
+  fix U
+  assume \<section>: "U \<subseteq> topspace (subtopology X S) \<and> (\<forall>K. compactin (subtopology X S) K \<longrightarrow> closedin (subtopology (subtopology X S) K) (K \<inter> U))"
+  have "closedin (subtopology X (K \<inter> S)) (K \<inter> U)" if "compactin X K" for K
+  proof -
+    have "K \<inter> U \<subseteq> topspace (subtopology X (K \<inter> S))"
+      using "\<section>" by auto
+    moreover
+    have "\<And>K'. compactin (subtopology X (K \<inter> S)) K' \<Longrightarrow> closedin (subtopology (subtopology X (K \<inter> S)) K') (K' \<inter> K \<inter> U)"
+      by (metis "\<section>" compactin_subtopology inf.orderE inf_commute subtopology_subtopology)
+    ultimately show ?thesis
+      by (metis (no_types, opaque_lifting) "2" inf.assoc k_space_def that)
+  qed
+  then show "closedin (subtopology X S) U"
+    using "1" \<section> by auto
+qed
+
+lemma k_space_subtopology_open:
+  assumes 1: "\<And>T. \<lbrakk>T \<subseteq> topspace X; T \<subseteq> S;
+                    \<And>K. compactin X K \<Longrightarrow> openin (subtopology X (K \<inter> S)) (K \<inter> T)\<rbrakk> \<Longrightarrow> openin (subtopology X S) T"
+  assumes 2: "\<And>K. compactin X K \<Longrightarrow> k_space(subtopology X (K \<inter> S))"
+  shows "k_space (subtopology X S)"
+  unfolding k_space_open
+proof (intro conjI strip)
+  fix U
+  assume \<section>: "U \<subseteq> topspace (subtopology X S) \<and> (\<forall>K. compactin (subtopology X S) K \<longrightarrow> openin (subtopology (subtopology X S) K) (K \<inter> U))"
+  have "openin (subtopology X (K \<inter> S)) (K \<inter> U)" if "compactin X K" for K
+  proof -
+    have "K \<inter> U \<subseteq> topspace (subtopology X (K \<inter> S))"
+      using "\<section>" by auto
+    moreover
+    have "\<And>K'. compactin (subtopology X (K \<inter> S)) K' \<Longrightarrow> openin (subtopology (subtopology X (K \<inter> S)) K') (K' \<inter> K \<inter> U)"
+      by (metis "\<section>" compactin_subtopology inf.orderE inf_commute subtopology_subtopology)
+    ultimately show ?thesis
+      by (metis (no_types, opaque_lifting) "2" inf.assoc k_space_open that)
+  qed
+  then show "openin (subtopology X S) U"
+    using "1" \<section> by auto
+qed
+
+
+lemma k_space_open_subtopology_aux:
+  assumes "kc_space X" "compact_space X" "openin X V"
+  shows "k_space (subtopology X V)"
+proof (clarsimp simp: k_space subtopology_subtopology compactin_subtopology Int_absorb1)
+  fix S
+  assume "S \<subseteq> topspace X"
+    and "S \<subseteq> V"
+    and S: "\<forall>K. compactin X K \<and> K \<subseteq> V \<longrightarrow> closedin (subtopology X K) (K \<inter> S)"
+  then have "V \<subseteq> topspace X"
+    using assms openin_subset by blast
+  have "S = V \<inter> ((topspace X - V) \<union> S)"
+    using \<open>S \<subseteq> V\<close> by auto
+  moreover have "closedin (subtopology X V) (V \<inter> ((topspace X - V) \<union> S))"
+  proof (intro closedin_subtopology_Int_closed compactin_imp_closedin_gen \<open>kc_space X\<close>)
+    show "compactin X (topspace X - V \<union> S)"
+      unfolding compactin_def
+    proof (intro conjI strip)
+      show "topspace X - V \<union> S \<subseteq> topspace X"
+        by (simp add: \<open>S \<subseteq> topspace X\<close>)
+      fix \<U>
+      assume \<U>: "Ball \<U> (openin X) \<and> topspace X - V \<union> S \<subseteq> \<Union>\<U>"
+      moreover
+      have "compactin X (topspace X - V)"
+        using assms closedin_compact_space by blast
+      ultimately obtain \<G> where "finite \<G>" "\<G> \<subseteq> \<U>" and \<G>: "topspace X - V \<subseteq> \<Union>\<G>"
+        unfolding compactin_def using \<open>V \<subseteq> topspace X\<close> by (metis le_sup_iff)
+      then have "topspace X - \<Union>\<G> \<subseteq> V"
+        by blast
+      then have "closedin (subtopology X (topspace X - \<Union>\<G>)) ((topspace X - \<Union>\<G>) \<inter> S)"
+        by (meson S \<U> \<open>\<G> \<subseteq> \<U>\<close> \<open>compact_space X\<close> closedin_compact_space openin_Union openin_closedin_eq subset_iff)
+      then have "compactin X ((topspace X - \<Union>\<G>) \<inter> S)"
+        by (meson \<U> \<open>\<G> \<subseteq> \<U>\<close>\<open>compact_space X\<close> closedin_compact_space closedin_trans_full openin_Union openin_closedin_eq subset_iff)
+      then obtain \<H> where "finite \<H>" "\<H> \<subseteq> \<U>" "(topspace X - \<Union>\<G>) \<inter> S \<subseteq> \<Union>\<H>"
+        unfolding compactin_def by (smt (verit, best) \<U> inf_le2 subset_trans sup.boundedE)
+      with \<G> have "topspace X - V \<union> S \<subseteq> \<Union>(\<G> \<union> \<H>)"
+        using \<open>S \<subseteq> topspace X\<close> by auto
+      then show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> topspace X - V \<union> S \<subseteq> \<Union>\<F>"
+        by (metis \<open>\<G> \<subseteq> \<U>\<close> \<open>\<H> \<subseteq> \<U>\<close> \<open>finite \<G>\<close> \<open>finite \<H>\<close> finite_Un le_sup_iff)
+    qed
+  qed
+  ultimately show "closedin (subtopology X V) S"
+    by metis
+qed
+
+
+lemma k_space_open_subtopology:
+  assumes X: "kc_space X \<or> Hausdorff_space X \<or> regular_space X" and "k_space X" "openin X S"
+  shows "k_space(subtopology X S)"
+proof (rule k_space_subtopology_open)
+  fix T
+  assume "T \<subseteq> topspace X"
+    and "T \<subseteq> S"
+    and T: "\<And>K. compactin X K \<Longrightarrow> openin (subtopology X (K \<inter> S)) (K \<inter> T)"
+  have "openin (subtopology X K) (K \<inter> T)" if "compactin X K" for K
+    by (smt (verit, ccfv_threshold) T assms(3) inf_assoc inf_commute openin_Int openin_subtopology that)
+  then show "openin (subtopology X S) T"
+    by (metis \<open>T \<subseteq> S\<close> \<open>T \<subseteq> topspace X\<close> assms(2) k_space_alt subset_openin_subtopology)
+next
+  fix K
+  assume "compactin X K"
+  then have KS: "openin (subtopology X K) (K \<inter> S)"
+    by (simp add: \<open>openin X S\<close> openin_subtopology_Int2)
+  have XK: "compact_space (subtopology X K)"
+    by (simp add: \<open>compactin X K\<close> compact_space_subtopology)
+  show "k_space (subtopology X (K \<inter> S))"
+    using X
+  proof (rule disjE)
+    assume "kc_space X"
+    then show "k_space (subtopology X (K \<inter> S))"
+      using k_space_open_subtopology_aux [of "subtopology X K" "K \<inter> S"]
+      by (simp add: KS XK kc_space_subtopology subtopology_subtopology)
+  next
+    assume "Hausdorff_space X \<or> regular_space X"
+    then have "locally_compact_space (subtopology (subtopology X K) (K \<inter> S))"
+      using locally_compact_space_open_subset Hausdorff_space_subtopology KS XK 
+        compact_imp_locally_compact_space regular_space_subtopology by blast
+    then show "k_space (subtopology X (K \<inter> S))"
+      by (simp add: locally_compact_imp_k_space subtopology_subtopology)
+  qed
+qed
+
+lemma k_kc_space_subtopology:
+   "\<lbrakk>k_space X; kc_space X; openin X S \<or> closedin X S\<rbrakk> \<Longrightarrow> k_space(subtopology X S) \<and> kc_space(subtopology X S)"
+  by (metis k_space_closed_subtopology k_space_open_subtopology kc_space_subtopology)
+
+
+lemma k_space_as_quotient_explicit:
+  "k_space X \<longleftrightarrow> quotient_map (sum_topology (subtopology X) {K. compactin X K}) X snd"
+proof -
+  have [simp]: "{x \<in> topspace X. x \<in> K \<and> x \<in> U} = K \<inter> U" if "U \<subseteq> topspace X" for K U
+    using that by blast
+  show "?thesis"
+    apply (simp add: quotient_map_def openin_sum_topology snd_image_Sigma k_space_alt)
+    by (smt (verit, del_insts) Union_iff compactin_sing inf.orderE mem_Collect_eq singletonI subsetI)
+qed
+
+lemma k_space_as_quotient:
+  fixes X :: "'a topology"
+  shows "k_space X \<longleftrightarrow> (\<exists>q. \<exists>Y:: ('a set * 'a) topology. locally_compact_space Y \<and> quotient_map Y X q)" 
+         (is "?lhs=?rhs")
+proof
+  show "k_space X" if ?rhs
+    using that k_space_quotient_map_image locally_compact_imp_k_space by blast 
+next
+  assume "k_space X"
+  show ?rhs
+  proof (intro exI conjI)
+    show "locally_compact_space (sum_topology (subtopology X) {K. compactin X K})"
+      by (simp add: compact_imp_locally_compact_space compact_space_subtopology locally_compact_space_sum_topology)
+    show "quotient_map (sum_topology (subtopology X) {K. compactin X K}) X snd"
+      using \<open>k_space X\<close> k_space_as_quotient_explicit by blast
+  qed
+qed
+
+lemma k_space_prod_topology_left:
+  assumes X: "locally_compact_space X" "Hausdorff_space X \<or> regular_space X" and "k_space Y"
+  shows "k_space (prod_topology X Y)"
+proof -
+  obtain q and Z :: "('b set * 'b) topology" where "locally_compact_space Z" and q: "quotient_map Z Y q"
+    using \<open>k_space Y\<close> k_space_as_quotient by blast
+  then show ?thesis
+    using quotient_map_prod_right [OF X q] X k_space_quotient_map_image locally_compact_imp_k_space 
+          locally_compact_space_prod_topology by blast
+qed
+
+text \<open>Essentially the same proof\<close>
+lemma k_space_prod_topology_right:
+  assumes "k_space X" and Y: "locally_compact_space Y" "Hausdorff_space Y \<or> regular_space Y"
+  shows "k_space (prod_topology X Y)"
+proof -
+  obtain q and Z :: "('a set * 'a) topology" where "locally_compact_space Z" and q: "quotient_map Z X q"
+    using \<open>k_space X\<close> k_space_as_quotient by blast
+  then show ?thesis
+    using quotient_map_prod_left [OF Y q] using Y k_space_quotient_map_image locally_compact_imp_k_space 
+          locally_compact_space_prod_topology by blast
+qed
+
+
+lemma continuous_map_from_k_space:
+  assumes "k_space X" and f: "\<And>K. compactin X K \<Longrightarrow> continuous_map(subtopology X K) Y f"
+  shows "continuous_map X Y f"
+proof -
+  have "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y"
+    by (metis compactin_absolute compactin_sing f image_compactin image_empty image_insert)
+  moreover have "closedin X {x \<in> topspace X. f x \<in> C}" if "closedin Y C" for C
+  proof -
+    have "{x \<in> topspace X. f x \<in> C} \<subseteq> topspace X"
+      by fastforce
+    moreover 
+    have eq: "K \<inter> {x \<in> topspace X. f x \<in> C} = {x. x \<in> topspace(subtopology X K) \<and> f x \<in> (f ` K \<inter> C)}" for K
+      by auto
+    have "closedin (subtopology X K) (K \<inter> {x \<in> topspace X. f x \<in> C})" if "compactin X K" for K
+      unfolding eq
+    proof (rule closedin_continuous_map_preimage)
+      show "continuous_map (subtopology X K) (subtopology Y (f`K)) f"
+        by (simp add: continuous_map_in_subtopology f image_mono that)
+      show "closedin (subtopology Y (f`K)) (f ` K \<inter> C)"
+        using \<open>closedin Y C\<close> closedin_subtopology by blast
+    qed
+    ultimately show ?thesis
+      using \<open>k_space X\<close> k_space by blast
+  qed
+  ultimately show ?thesis
+    by (simp add: continuous_map_closedin)
+qed
+
+lemma closed_map_into_k_space:
+  assumes "k_space Y" and fim: "f ` (topspace X) \<subseteq> topspace Y"
+    and f: "\<And>K. compactin Y K
+                \<Longrightarrow> closed_map(subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f"
+  shows "closed_map X Y f"
+  unfolding closed_map_def
+proof (intro strip)
+  fix C
+  assume "closedin X C"
+  have "closedin (subtopology Y K) (K \<inter> f ` C)"
+    if "compactin Y K" for K
+  proof -
+    have eq: "K \<inter> f ` C = f ` ({x \<in> topspace X. f x \<in> K} \<inter> C)"
+      using \<open>closedin X C\<close> closedin_subset by auto
+    show ?thesis
+      unfolding eq
+      by (metis (no_types, lifting) \<open>closedin X C\<close> closed_map_def closedin_subtopology f inf_commute that)
+  qed
+  then show "closedin Y (f ` C)"
+    using \<open>k_space Y\<close> unfolding k_space
+    by (meson \<open>closedin X C\<close> closedin_subset dual_order.trans fim image_mono)
+qed
+
+
+text \<open>Essentially the same proof\<close>
+lemma open_map_into_k_space:
+  assumes "k_space Y" and fim: "f ` (topspace X) \<subseteq> topspace Y"
+    and f: "\<And>K. compactin Y K
+                 \<Longrightarrow> open_map (subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f"
+  shows "open_map X Y f"
+  unfolding open_map_def
+proof (intro strip)
+  fix C
+  assume "openin X C"
+  have "openin (subtopology Y K) (K \<inter> f ` C)"
+    if "compactin Y K" for K
+  proof -
+    have eq: "K \<inter> f ` C = f ` ({x \<in> topspace X. f x \<in> K} \<inter> C)"
+      using \<open>openin X C\<close> openin_subset by auto
+    show ?thesis
+      unfolding eq
+      by (metis (no_types, lifting) \<open>openin X C\<close> open_map_def openin_subtopology f inf_commute that)
+  qed
+  then show "openin Y (f ` C)"
+    using \<open>k_space Y\<close> unfolding k_space_open
+    by (meson \<open>openin X C\<close> openin_subset dual_order.trans fim image_mono)
+qed
+
+lemma quotient_map_into_k_space:
+  fixes f :: "'a \<Rightarrow> 'b"
+  assumes "k_space Y" and cmf: "continuous_map X Y f" 
+    and fim: "f ` (topspace X) = topspace Y"
+    and f: "\<And>k. compactin Y k
+                 \<Longrightarrow> quotient_map (subtopology X {x \<in> topspace X. f x \<in> k})
+                                  (subtopology Y k) f"
+  shows "quotient_map X Y f"
+proof -
+  have "closedin Y C"
+    if "C \<subseteq> topspace Y" and K: "closedin X {x \<in> topspace X. f x \<in> C}" for C
+  proof -
+    have "closedin (subtopology Y K) (K \<inter> C)" if "compactin Y K" for K
+    proof -
+      define Kf where "Kf \<equiv> {x \<in> topspace X. f x \<in> K}"
+      have *: "K \<inter> C \<subseteq> topspace Y \<and> K \<inter> C \<subseteq> K"
+        using \<open>C \<subseteq> topspace Y\<close> by blast
+      then have eq: "closedin (subtopology X Kf) (Kf \<inter> {x \<in> topspace X. f x \<in> C}) =
+                 closedin (subtopology Y K) (K \<inter> C)"
+        using f [OF that] * unfolding quotient_map_closedin Kf_def
+        by (smt (verit, ccfv_SIG) Collect_cong Int_def compactin_subset_topspace mem_Collect_eq that topspace_subtopology topspace_subtopology_subset)
+      have dd: "{x \<in> topspace X \<inter> Kf. f x \<in> K \<inter> C} = Kf \<inter> {x \<in> topspace X. f x \<in> C}"
+        by (auto simp add: Kf_def)
+      have "closedin (subtopology X Kf) {x \<in> topspace X. x \<in> Kf \<and> f x \<in> K \<and> f x \<in> C}"
+        using K closedin_subtopology by (fastforce simp add: Kf_def)
+      with K closedin_subtopology_Int_closed eq show ?thesis
+        by blast
+    qed
+    then show ?thesis 
+      using \<open>k_space Y\<close> that unfolding k_space by blast
+  qed
+  moreover have "closedin X {x \<in> topspace X. f x \<in> K}"
+    if "K \<subseteq> topspace Y" "closedin Y K" for K
+    using that cmf continuous_map_closedin by fastforce
+  ultimately show ?thesis
+    unfolding quotient_map_closedin using fim by blast
+qed
+
+lemma quotient_map_into_k_space_eq:
+  assumes "k_space Y" "kc_space Y"
+  shows "quotient_map X Y f \<longleftrightarrow>
+         continuous_map X Y f \<and> f ` (topspace X) = topspace Y \<and>
+         (\<forall>K. compactin Y K
+              \<longrightarrow> quotient_map (subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f)"
+  using assms
+  by (auto simp: kc_space_def intro: quotient_map_into_k_space quotient_map_restriction
+       dest: quotient_imp_continuous_map quotient_imp_surjective_map)
+
+lemma open_map_into_k_space_eq:
+  assumes "k_space Y"
+  shows "open_map X Y f \<longleftrightarrow>
+         f ` (topspace X) \<subseteq> topspace Y \<and>
+         (\<forall>k. compactin Y k
+              \<longrightarrow> open_map (subtopology X {x \<in> topspace X. f x \<in> k}) (subtopology Y k) f)"
+       (is "?lhs=?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (simp add: open_map_imp_subset_topspace open_map_restriction)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (simp add: assms open_map_into_k_space)
+qed
+
+lemma closed_map_into_k_space_eq:
+  assumes "k_space Y"
+  shows "closed_map X Y f \<longleftrightarrow>
+         f ` (topspace X) \<subseteq> topspace Y \<and>
+         (\<forall>k. compactin Y k
+              \<longrightarrow> closed_map (subtopology X {x \<in> topspace X. f x \<in> k}) (subtopology Y k) f)"
+       (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (simp add: closed_map_imp_subset_topspace closed_map_restriction)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (simp add: assms closed_map_into_k_space)
+qed
+
+lemma proper_map_into_k_space:
+  assumes "k_space Y" and fim: "f ` (topspace X) \<subseteq> topspace Y"
+    and f: "\<And>K. compactin Y K
+                 \<Longrightarrow> proper_map (subtopology X {x \<in> topspace X. f x \<in> K})
+                                (subtopology Y K) f"
+  shows "proper_map X Y f"
+proof -
+  have "closed_map X Y f"
+    by (meson assms closed_map_into_k_space fim proper_map_def)
+  with f topspace_subtopology_subset show ?thesis
+    apply (simp add: proper_map_alt)
+    by (smt (verit, best) Collect_cong compactin_absolute)
+qed
+
+lemma proper_map_into_k_space_eq:
+  assumes "k_space Y"
+  shows "proper_map X Y f \<longleftrightarrow>
+         f ` (topspace X) \<subseteq> topspace Y \<and>
+         (\<forall>K. compactin Y K
+              \<longrightarrow> proper_map (subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f)"
+         (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (simp add: proper_map_imp_subset_topspace proper_map_restriction)
+  show "?rhs \<Longrightarrow> ?lhs"
+    by (simp add: assms proper_map_into_k_space)
+qed
+
+lemma compact_imp_proper_map:
+  assumes "k_space Y" "kc_space Y" and fim: "f ` (topspace X) \<subseteq> topspace Y" 
+    and f: "continuous_map X Y f \<or> kc_space X" 
+    and comp: "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
+  shows "proper_map X Y f"
+proof (rule compact_imp_proper_map_gen)
+  fix S
+  assume "S \<subseteq> topspace Y"
+      and "\<And>K. compactin Y K \<Longrightarrow> compactin Y (S \<inter> K)"
+  with assms show "closedin Y S"
+    by (simp add: closedin_subset_topspace inf_commute k_space kc_space_def)
+qed (use assms in auto)
+
+lemma proper_eq_compact_map:
+  assumes "k_space Y" "kc_space Y" 
+    and f: "continuous_map X Y f \<or> kc_space X" 
+  shows  "proper_map X Y f \<longleftrightarrow>
+             f ` (topspace X) \<subseteq> topspace Y \<and>
+             (\<forall>K. compactin Y K \<longrightarrow> compactin X {x \<in> topspace X. f x \<in> K})"
+         (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  show "?lhs \<Longrightarrow> ?rhs"
+    by (simp add: proper_map_alt proper_map_imp_subset_topspace)
+qed (use assms compact_imp_proper_map in auto)
+
+lemma compact_imp_perfect_map:
+  assumes "k_space Y" "kc_space Y" and "f ` (topspace X) = topspace Y" 
+    and "continuous_map X Y f" 
+    and "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
+  shows "perfect_map X Y f"
+  by (simp add: assms compact_imp_proper_map perfect_map_def)
+
 end
 
--- a/src/HOL/Analysis/Abstract_Topology.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Analysis/Abstract_Topology.thy	Fri May 26 09:56:20 2023 +0200
@@ -269,6 +269,10 @@
   unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
   by auto
 
+lemma subset_openin_subtopology:
+   "\<lbrakk>openin X S; S \<subseteq> T\<rbrakk> \<Longrightarrow> openin (subtopology X T) S"
+  by (metis inf.orderE openin_subtopology)
+
 lemma openin_subtopology_Int:
    "openin X S \<Longrightarrow> openin (subtopology X T) (S \<inter> T)"
   using openin_subtopology by auto
@@ -296,6 +300,10 @@
   unfolding closedin_def topspace_subtopology
   by (auto simp: openin_subtopology)
 
+lemma closedin_subtopology_Int_closed:
+   "closedin X T \<Longrightarrow> closedin (subtopology X S) (S \<inter> T)"
+  using closedin_subtopology inf_commute by blast
+
 lemma closedin_subset_topspace:
    "\<lbrakk>closedin X S; S \<subseteq> T\<rbrakk> \<Longrightarrow> closedin (subtopology X T) S"
   using closedin_subtopology by fastforce
@@ -1098,7 +1106,6 @@
      "X frontier_of S = X closure_of S \<inter> X closure_of (topspace X - S)"
   by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of)
 
-
 lemma interior_of_union_frontier_of [simp]:
      "X interior_of S \<union> X frontier_of S = X closure_of S"
   by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym)
@@ -1265,6 +1272,37 @@
      "(discrete_topology U) frontier_of S = {}"
   by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures)
 
+lemma openin_subset_topspace_eq:
+  assumes "disjnt S (X frontier_of U)"
+  shows "openin (subtopology X U) S \<longleftrightarrow> openin X S \<and> S \<subseteq> U"
+proof
+  assume S: "openin (subtopology X U) S"
+  show "openin X S \<and> S \<subseteq> U"
+  proof
+    show "S \<subseteq> U"
+      using S openin_imp_subset by blast
+    have "disjnt S (X frontier_of (topspace X \<inter> U))"
+      by (metis assms frontier_of_restrict)
+    moreover
+    have "openin (subtopology X (topspace X \<inter> U)) S"
+      by (simp add: S subtopology_restrict)
+    moreover
+    have "openin X S" 
+      if dis: "disjnt S (X frontier_of U)" and ope: "openin (subtopology X U) S" and "U \<subseteq> topspace X" 
+      for S U
+    proof -
+      obtain T where T: "openin X T" "S = T \<inter> U"
+        using ope by (auto simp: openin_subtopology)
+      have "T \<inter> U = T \<inter> X interior_of U"
+        using that T interior_of_subset in_closure_of by (fastforce simp: disjnt_iff frontier_of_def)
+      then show ?thesis
+        using \<open>S = T \<inter> U\<close> \<open>openin X T\<close> by auto
+    qed
+  ultimately show "openin X S"
+      by blast
+  qed 
+qed (metis inf.absorb_iff1 openin_subtopology_Int)
+
 
 subsection\<open>Locally finite collections\<close>
 
@@ -4778,6 +4816,18 @@
    "proper_map X Y f \<Longrightarrow> f ` (topspace X) \<subseteq> topspace Y"
   by (simp add: closed_map_imp_subset_topspace proper_map_def)
 
+lemma proper_map_restriction:
+  assumes "proper_map X Y f" "{x \<in> topspace X. f x \<in> V} = U"
+  shows "proper_map (subtopology X U) (subtopology Y V) f"
+proof -
+  have [simp]: "{x \<in> topspace X. f x \<in> V \<and> f x = y} = {x \<in> topspace X. f x = y}" 
+    if "y \<in> V" for y
+    using that by auto
+  show ?thesis
+    using assms unfolding proper_map_def using closed_map_restriction
+    by (force simp: compactin_subtopology)
+qed
+
 lemma closed_injective_imp_proper_map:
   assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)"
   shows "proper_map X Y f"
@@ -4946,6 +4996,45 @@
    "S \<subseteq> topspace X \<Longrightarrow> proper_map (subtopology X S) X id \<longleftrightarrow> closedin X S \<and> (\<forall>k. compactin X k \<longrightarrow> compactin X (S \<inter> k))"
   by (metis closed_Int_compactin closed_map_inclusion_eq inf.absorb_iff2 inj_on_id injective_imp_proper_eq_closed_map)
 
+lemma proper_map_into_subtopology:
+   "\<lbrakk>proper_map X Y f; f ` topspace X \<subseteq> C\<rbrakk> \<Longrightarrow> proper_map X (subtopology Y C) f"
+  by (simp add: closed_map_into_subtopology compactin_subtopology proper_map_alt)
+
+lemma proper_map_from_composition_left:
+  assumes gf: "proper_map X Z (g \<circ> f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y"
+  shows "proper_map Y Z g"
+  unfolding proper_map_def
+proof (intro strip conjI)
+  show "closed_map Y Z g"
+    by (meson closed_map_from_composition_left gf contf fim proper_imp_closed_map)
+  fix z assume "z \<in> topspace Z"
+  have eq: "{y \<in> topspace Y. g y = z} = f ` {x \<in> topspace X. (g \<circ> f) x = z}"
+    using fim by force
+  show "compactin Y {x \<in> topspace Y. g x = z}"
+    unfolding eq
+  proof (rule image_compactin [OF _ contf])
+    show "compactin X {x \<in> topspace X. (g \<circ> f) x = z}"
+      using \<open>z \<in> topspace Z\<close> gf proper_map_def by fastforce
+  qed
+qed
+
+lemma proper_map_from_composition_right_inj:
+  assumes gf: "proper_map X Z (g \<circ> f)" and fim: "f ` topspace X \<subseteq> topspace Y" 
+    and contf: "continuous_map Y Z g" and inj: "inj_on g (topspace Y)"
+  shows "proper_map X Y f"
+  unfolding proper_map_def
+proof (intro strip conjI)
+  show "closed_map X Y f"
+    by (meson closed_map_from_composition_right assms proper_imp_closed_map)
+  fix y
+  assume "y \<in> topspace Y"
+  with fim inj have eq: "{x \<in> topspace X. f x = y} = {x \<in> topspace X. (g \<circ> f) x = g y}"
+    by (auto simp: image_subset_iff inj_onD)
+  show "compactin X {x \<in> topspace X. f x = y}"
+    unfolding eq
+    by (smt (verit) Collect_cong \<open>y \<in> topspace Y\<close> contf continuous_map_closedin gf proper_map_def)
+qed
+
 
 subsection\<open>Perfect maps (proper, continuous and surjective)\<close>
 
@@ -4991,4 +5080,11 @@
    "perfect_map X Y f \<Longrightarrow> f ` (topspace X) = topspace Y"
   by (simp add: perfect_map_def)
 
+lemma perfect_map_from_composition_left:
+  assumes "perfect_map X Z (g \<circ> f)" and "continuous_map X Y f"
+    and "continuous_map Y Z g" and "f ` topspace X = topspace Y"
+  shows "perfect_map Y Z g"
+  using assms unfolding perfect_map_def
+  by (metis image_comp proper_map_from_composition_left)
+
 end
--- a/src/HOL/Analysis/Analysis.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Analysis/Analysis.thy	Fri May 26 09:56:20 2023 +0200
@@ -7,6 +7,7 @@
   FSigma
   Sum_Topology
   Abstract_Topological_Spaces
+  Abstract_Metric_Spaces
   Connected
   Abstract_Limits
   Isolated
--- a/src/HOL/Analysis/Measure_Space.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy	Fri May 26 09:56:20 2023 +0200
@@ -2288,12 +2288,8 @@
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close>
 
 lemma strict_monoI_Suc:
-  assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
-  unfolding strict_mono_def
-proof safe
-  fix n m :: nat assume "n < m" then show "f n < f m"
-    by (induct m) (auto simp: less_Suc_eq intro: less_trans ord)
-qed
+  assumes "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
+  by (simp add: assms strict_mono_Suc_iff)
 
 lemma emeasure_count_space:
   assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \<infinity>)"
--- a/src/HOL/Analysis/Product_Topology.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Analysis/Product_Topology.thy	Fri May 26 09:56:20 2023 +0200
@@ -139,6 +139,47 @@
   qed
 qed
 
+text \<open>Missing the opposite direction. Does it hold? A converse is proved for proper maps, a stronger condition\<close>
+lemma closed_map_prod:
+  assumes "closed_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y))"
+  shows "topspace(prod_topology X Y) = {} \<or> closed_map X X' f \<and> closed_map Y Y' g"
+proof (cases "topspace(prod_topology X Y) = {}")
+  case False
+  then have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
+    by auto
+  have "closed_map X X' f"
+    unfolding closed_map_def
+  proof (intro strip)
+    fix C
+    assume "closedin X C"
+    show "closedin X' (f ` C)"
+    proof (cases "C={}")
+      case False
+      with assms have "closedin (prod_topology X' Y') ((\<lambda>(x,y). (f x, g y)) ` (C \<times> topspace Y))"
+        by (simp add: \<open>closedin X C\<close> closed_map_def closedin_prod_Times_iff)
+      with False ne show ?thesis
+        by (simp add: image_paired_Times closedin_Times closedin_prod_Times_iff)
+    qed auto
+  qed
+  moreover
+  have "closed_map Y Y' g"
+    unfolding closed_map_def
+  proof (intro strip)
+    fix C
+    assume "closedin Y C"
+    show "closedin Y' (g ` C)"
+    proof (cases "C={}")
+      case False
+      with assms have "closedin (prod_topology X' Y') ((\<lambda>(x,y). (f x, g y)) ` (topspace X \<times> C))"
+        by (simp add: \<open>closedin Y C\<close> closed_map_def closedin_prod_Times_iff)
+      with False ne show ?thesis
+        by (simp add: image_paired_Times closedin_Times closedin_prod_Times_iff)
+    qed auto
+  qed
+  ultimately show ?thesis
+    by (auto simp: False)
+qed auto
+
 subsection \<open>Continuity\<close>
 
 lemma continuous_map_pairwise:
--- a/src/HOL/Analysis/T1_Spaces.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Analysis/T1_Spaces.thy	Fri May 26 09:56:20 2023 +0200
@@ -49,6 +49,12 @@
   apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton)
   using t1_space_alt by auto
 
+lemma continuous_closed_imp_proper_map:
+   "\<lbrakk>compact_space X; t1_space Y; continuous_map X Y f; closed_map X Y f\<rbrakk> \<Longrightarrow> proper_map X Y f"
+  unfolding proper_map_def
+  by (smt (verit) closedin_compact_space closedin_continuous_map_preimage 
+      Collect_cong singleton_iff t1_space_closedin_singleton)
+
 lemma t1_space_euclidean: "t1_space (euclidean :: 'a::metric_space topology)"
   by (simp add: t1_space_closedin_singleton)
 
--- a/src/HOL/Bali/Eval.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Bali/Eval.thy	Fri May 26 09:56:20 2023 +0200
@@ -851,28 +851,28 @@
   by (auto, frule eval_Inj_elim, auto, frule eval_Inj_elim, auto)
 
 simproc_setup eval_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm eval_expr_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm eval_expr_eq}))))\<close>
 
 simproc_setup eval_var ("G\<turnstile>s \<midarrow>In2 t\<succ>\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm eval_var_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm eval_var_eq}))))\<close>
 
 simproc_setup eval_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm eval_exprs_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm eval_exprs_eq}))))\<close>
 
 simproc_setup eval_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm eval_stmt_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm eval_stmt_eq}))))\<close>
 
 ML \<open>
 ML_Thms.bind_thms ("AbruptIs", sum3_instantiate \<^context> @{thm eval.Abrupt})
@@ -947,10 +947,10 @@
 done
 
 simproc_setup eval_no_abrupt ("G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ (Const (\<^const_name>\<open>Pair\<close>, _) $ (Const (\<^const_name>\<open>None\<close>, _)) $ _) $ _  $ _ $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm eval_no_abrupt}))
+    | _ => SOME (mk_meta_eq @{thm eval_no_abrupt}))))
 \<close>
 
 
@@ -967,10 +967,10 @@
 done
 
 simproc_setup eval_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<rightarrow> (w,s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ _ $ (Const (\<^const_name>\<open>Pair\<close>, _) $ (Const (\<^const_name>\<open>Some\<close>, _) $ _)$ _)) => NONE
-    | _ => SOME (mk_meta_eq @{thm eval_abrupt}))
+    | _ => SOME (mk_meta_eq @{thm eval_abrupt}))))
 \<close>
 
 lemma LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else undefined)\<rightarrow> s"
--- a/src/HOL/Bali/Evaln.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Bali/Evaln.thy	Fri May 26 09:56:20 2023 +0200
@@ -269,28 +269,28 @@
   by (auto, frule evaln_Inj_elim, auto, frule evaln_Inj_elim, auto)
 
 simproc_setup evaln_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm evaln_expr_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm evaln_expr_eq}))))\<close>
 
 simproc_setup evaln_var ("G\<turnstile>s \<midarrow>In2 t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm evaln_var_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm evaln_var_eq}))))\<close>
 
 simproc_setup evaln_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm evaln_exprs_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm evaln_exprs_eq}))))\<close>
 
 simproc_setup evaln_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
-    | _ => SOME (mk_meta_eq @{thm evaln_stmt_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm evaln_stmt_eq}))))\<close>
 
 ML \<open>ML_Thms.bind_thms ("evaln_AbruptIs", sum3_instantiate \<^context> @{thm evaln.Abrupt})\<close>
 declare evaln_AbruptIs [intro!]
@@ -356,11 +356,11 @@
 done
 
 simproc_setup evaln_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s')") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ _ $ _ $ (Const (\<^const_name>\<open>Pair\<close>, _) $ (Const (\<^const_name>\<open>Some\<close>,_) $ _)$ _))
         => NONE
-    | _ => SOME (mk_meta_eq @{thm evaln_abrupt}))
+    | _ => SOME (mk_meta_eq @{thm evaln_abrupt}))))
 \<close>
 
 lemma evaln_LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else undefined)\<midarrow>n\<rightarrow> s"
--- a/src/HOL/Bali/WellType.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Bali/WellType.thy	Fri May 26 09:56:20 2023 +0200
@@ -605,28 +605,28 @@
   by (auto, frule wt_Inj_elim, auto, frule wt_Inj_elim, auto)
 
 simproc_setup wt_expr ("E,dt\<Turnstile>In1l t\<Colon>U") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
-    | _ => SOME (mk_meta_eq @{thm wt_expr_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm wt_expr_eq}))))\<close>
 
 simproc_setup wt_var ("E,dt\<Turnstile>In2 t\<Colon>U") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
-    | _ => SOME (mk_meta_eq @{thm wt_var_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm wt_var_eq}))))\<close>
 
 simproc_setup wt_exprs ("E,dt\<Turnstile>In3 t\<Colon>U") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
-    | _ => SOME (mk_meta_eq @{thm wt_exprs_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm wt_exprs_eq}))))\<close>
 
 simproc_setup wt_stmt ("E,dt\<Turnstile>In1r t\<Colon>U") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
-    | _ => SOME (mk_meta_eq @{thm wt_stmt_eq}))\<close>
+    | _ => SOME (mk_meta_eq @{thm wt_stmt_eq}))))\<close>
 
 lemma wt_elim_BinOp:
   "\<lbrakk>E,dt\<Turnstile>In1l (BinOp binop e1 e2)\<Colon>T;
--- a/src/HOL/Boolean_Algebras.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Boolean_Algebras.thy	Fri May 26 09:56:20 2023 +0200
@@ -513,10 +513,10 @@
 ML_file \<open>Tools/boolean_algebra_cancel.ML\<close>
 
 simproc_setup boolean_algebra_cancel_sup ("sup a b::'a::boolean_algebra") =
-  \<open>fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_sup_conv\<close>
+  \<open>K (K (try Boolean_Algebra_Cancel.cancel_sup_conv))\<close>
 
 simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") =
-  \<open>fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_inf_conv\<close>
+  \<open>K (K (try Boolean_Algebra_Cancel.cancel_inf_conv))\<close>
 
 
 context boolean_algebra
--- a/src/HOL/Data_Structures/Less_False.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Data_Structures/Less_False.thy	Fri May 26 09:56:20 2023 +0200
@@ -6,7 +6,7 @@
 imports Main
 begin
 
-simproc_setup less_False ("(x::'a::order) < y") = \<open>fn _ => fn ctxt => fn ct =>
+simproc_setup less_False ("(x::'a::order) < y") = \<open>K (fn ctxt => fn ct =>
   let
     fun prp t thm = Thm.full_prop_of thm aconv t;
 
@@ -25,7 +25,7 @@
              end
          | SOME thm => NONE
       end;
-  in prove_less_False (Thm.term_of ct) end
+  in prove_less_False (Thm.term_of ct) end)
 \<close>
 
 end
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy	Fri May 26 09:56:20 2023 +0200
@@ -942,16 +942,15 @@
 context cring
 begin
 
-local_setup \<open>
-Local_Theory.declaration {syntax = false, pervasive = false}
-  (fn phi => Ring_Tac.Ring_Simps.map (Ring_Tac.insert_rules Ring_Tac.eq_ring_simps
+declaration \<open>fn phi =>
+  Ring_Tac.Ring_Simps.map (Ring_Tac.insert_rules Ring_Tac.eq_ring_simps
     (Morphism.term phi \<^term>\<open>R\<close>,
      (Morphism.fact phi @{thms Ipol.simps [meta] Ipol_Pc [meta]},
       Morphism.fact phi @{thms Ipolex.simps [meta] Ipolex_Const [meta]},
       Morphism.fact phi @{thms Ipolex_polex_list.simps [meta]},
       Morphism.fact phi @{thms in_carrier_Nil in_carrier_Cons},
       singleton (Morphism.fact phi) @{thm head.simps(2) [meta]},
-      singleton (Morphism.fact phi) @{thm norm_subst_correct [meta]}))))
+      singleton (Morphism.fact phi) @{thm norm_subst_correct [meta]})))
 \<close>
 
 end
--- a/src/HOL/Decision_Procs/Conversions.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Decision_Procs/Conversions.thy	Fri May 26 09:56:20 2023 +0200
@@ -20,7 +20,7 @@
 \<close>
 
 attribute_setup meta =
-  \<open>Scan.succeed (fn (ctxt, th) => (NONE, SOME (mk_meta_eq th)))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (K mk_meta_eq))\<close>
   \<open>convert equality to meta equality\<close>
 
 ML \<open>
--- a/src/HOL/Decision_Procs/Reflective_Field.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Decision_Procs/Reflective_Field.thy	Fri May 26 09:56:20 2023 +0200
@@ -909,15 +909,14 @@
 context field
 begin
 
-local_setup \<open>
-Local_Theory.declaration {syntax = false, pervasive = false}
-  (fn phi => Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps
+declaration \<open>fn phi =>
+  Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps
     (Morphism.term phi \<^term>\<open>R\<close>,
      (Morphism.fact phi @{thms feval.simps [meta] feval_Cnst [meta]},
       Morphism.fact phi @{thms peval.simps [meta] peval_Cnst [meta]},
       Morphism.fact phi @{thms nonzero.simps [meta] nonzero_singleton [meta]},
       singleton (Morphism.fact phi) @{thm nth_el_Cons [meta]},
-      singleton (Morphism.fact phi) @{thm feval_eq}))))
+      singleton (Morphism.fact phi) @{thm feval_eq})))
 \<close>
 
 end
--- a/src/HOL/Decision_Procs/ferrante_rackoff_data.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Decision_Procs/ferrante_rackoff_data.ML	Fri May 26 09:56:20 2023 +0200
@@ -15,7 +15,7 @@
   val funs: thm -> 
     {isolate_conv: morphism -> Proof.context -> cterm list -> cterm -> thm,
      whatis: morphism -> cterm -> cterm -> ord,
-     simpset: morphism -> Proof.context -> simpset} -> declaration
+     simpset: morphism -> Proof.context -> simpset} -> Morphism.declaration
   val match: Proof.context -> cterm -> entry option
 end;
 
--- a/src/HOL/Eisbach/Tests.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Eisbach/Tests.thy	Fri May 26 09:56:20 2023 +0200
@@ -432,9 +432,7 @@
   assumes A : "A"
 begin
 
-local_setup
-  \<open>Local_Theory.declaration {pervasive = false, syntax = false}
-    (fn phi => Data.put (Morphism.fact phi @{thms A}))\<close>
+declaration \<open>fn phi => Data.put (Morphism.fact phi @{thms A})\<close>
 
 lemma A by dynamic_thms_test
 
--- a/src/HOL/Eisbach/method_closure.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Eisbach/method_closure.ML	Fri May 26 09:56:20 2023 +0200
@@ -66,13 +66,14 @@
   let
     val name = Local_Theory.full_name lthy binding;
   in
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
-      Data.map
-        (Symtab.update (name,
-          {vars = map (Morphism.term phi) (#vars closure),
-           named_thms = #named_thms closure,
-           methods = #methods closure,
-           body = (Method.map_source o map) (Token.transform phi) (#body closure)})))
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding}
+      (fn phi =>
+        Data.map
+          (Symtab.update (name,
+            {vars = map (Morphism.term phi) (#vars closure),
+             named_thms = #named_thms closure,
+             methods = #methods closure,
+             body = (Method.map_source o map) (Token.transform phi) (#body closure)})))
   end;
 
 
--- a/src/HOL/Enum.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Enum.thy	Fri May 26 09:56:20 2023 +0200
@@ -564,10 +564,10 @@
 by(cases x) simp
 
 simproc_setup finite_1_eq ("x::finite_1") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     (case Thm.term_of ct of
       Const (\<^const_name>\<open>a\<^sub>1\<close>, _) => NONE
-    | _ => SOME (mk_meta_eq @{thm finite_1_eq}))
+    | _ => SOME (mk_meta_eq @{thm finite_1_eq}))))
 \<close>
 
 instantiation finite_1 :: complete_boolean_algebra
--- a/src/HOL/Finite_Set.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Finite_Set.thy	Fri May 26 09:56:20 2023 +0200
@@ -193,7 +193,7 @@
 lemma finite_subset: "A \<subseteq> B \<Longrightarrow> finite B \<Longrightarrow> finite A"
   by (rule rev_finite_subset)
 
-simproc_setup finite ("finite A") = \<open>fn _ =>
+simproc_setup finite ("finite A") = \<open>
 let
   val finite_subset = @{thm finite_subset}
   val Eq_TrueI = @{thm Eq_TrueI}
@@ -209,17 +209,17 @@
 
   fun comb (A,sub_th) (A',fin_th) ths = if A aconv A' then (sub_th,fin_th) :: ths else ths
 
-  fun proc ss ct =
+  fun proc ctxt ct =
     (let
        val _ $ A = Thm.term_of ct
-       val prems = Simplifier.prems_of ss
+       val prems = Simplifier.prems_of ctxt
        val fins = map_filter is_finite prems
        val subsets = map_filter (is_subset A) prems
      in case fold_product comb subsets fins [] of
           (sub_th,fin_th) :: _ => SOME((fin_th RS (sub_th RS finite_subset)) RS Eq_TrueI)
         | _ => NONE
      end)
-in proc end
+in K proc end
 \<close>
 
 (* Needs to be used with care *)
--- a/src/HOL/Fun.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Fun.thy	Fri May 26 09:56:20 2023 +0200
@@ -1353,7 +1353,7 @@
 
 text \<open>Simplify terms of the form \<open>f(\<dots>,x:=y,\<dots>,x:=z,\<dots>)\<close> to \<open>f(\<dots>,x:=z,\<dots>)\<close>\<close>
 
-simproc_setup fun_upd2 ("f(v := w, x := y)") = \<open>fn _ =>
+simproc_setup fun_upd2 ("f(v := w, x := y)") = \<open>
   let
     fun gen_fun_upd NONE T _ _ = NONE
       | gen_fun_upd (SOME f) T x y = SOME (Const (\<^const_name>\<open>fun_upd\<close>, T) $ f $ x $ y)
@@ -1380,7 +1380,7 @@
                 resolve_tac ctxt @{thms ext} 1 THEN
                 simp_tac (put_simpset ss ctxt) 1)))
       end
-  in proc end
+  in K proc end
 \<close>
 
 
--- a/src/HOL/Groups.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Groups.thy	Fri May 26 09:56:20 2023 +0200
@@ -178,8 +178,8 @@
       | _ => false)
 \<close>
 
-simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc
-simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc
+simproc_setup reorient_zero ("0 = x") = \<open>K Reorient_Proc.proc\<close>
+simproc_setup reorient_one ("1 = x") = \<open>K Reorient_Proc.proc\<close>
 
 typed_print_translation \<open>
   let
--- a/src/HOL/HOL.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/HOL.thy	Fri May 26 09:56:20 2023 +0200
@@ -1236,7 +1236,7 @@
 
 text \<open>Simproc for proving \<open>(y = x) \<equiv> False\<close> from premise \<open>\<not> (x = y)\<close>:\<close>
 
-simproc_setup neq ("x = y") = \<open>fn _ =>
+simproc_setup neq ("x = y") = \<open>
   let
     val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI};
     fun is_neq eq lhs rhs thm =
@@ -1252,7 +1252,7 @@
             SOME thm => SOME (thm RS neq_to_EQ_False)
           | NONE => NONE)
        | _ => NONE);
-  in proc end
+  in K proc end
 \<close>
 
 simproc_setup let_simp ("Let x f") = \<open>
@@ -1266,7 +1266,7 @@
         Abs (_, _, t') => count_loose t' 0 <= 1
       | _ => true);
   in
-    fn _ => fn ctxt => fn ct =>
+    K (fn ctxt => fn ct =>
       if is_trivial_let (Thm.term_of ct)
       then SOME @{thm Let_def} (*no or one ocurrence of bound variable*)
       else
@@ -1308,7 +1308,7 @@
                     in SOME (rl OF [Thm.transitive fx_g g_g'x]) end
                 end
             | _ => NONE)
-        end
+        end)
   end
 \<close>
 
@@ -1649,7 +1649,7 @@
 signature REORIENT_PROC =
 sig
   val add : (term -> bool) -> theory -> theory
-  val proc : morphism -> Proof.context -> cterm -> thm option
+  val proc : Proof.context -> cterm -> thm option
 end;
 
 structure Reorient_Proc : REORIENT_PROC =
@@ -1664,7 +1664,7 @@
   fun matches thy t = exists (fn (m, _) => m t) (Data.get thy);
 
   val meta_reorient = @{thm eq_commute [THEN eq_reflection]};
-  fun proc phi ctxt ct =
+  fun proc ctxt ct =
     let
       val thy = Proof_Context.theory_of ctxt;
     in
@@ -1851,12 +1851,12 @@
 
 declare [[coercion_args NO_MATCH - -]]
 
-simproc_setup NO_MATCH ("NO_MATCH pat val") = \<open>fn _ => fn ctxt => fn ct =>
+simproc_setup NO_MATCH ("NO_MATCH pat val") = \<open>K (fn ctxt => fn ct =>
   let
     val thy = Proof_Context.theory_of ctxt
     val dest_binop = Term.dest_comb #> apfst (Term.dest_comb #> snd)
     val m = Pattern.matches thy (dest_binop (Thm.term_of ct))
-  in if m then NONE else SOME @{thm NO_MATCH_def} end
+  in if m then NONE else SOME @{thm NO_MATCH_def} end)
 \<close>
 
 text \<open>
--- a/src/HOL/HOLCF/Cfun.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/HOLCF/Cfun.thy	Fri May 26 09:56:20 2023 +0200
@@ -142,14 +142,14 @@
 \<close>
 
 simproc_setup beta_cfun_proc ("Rep_cfun (Abs_cfun f)") = \<open>
-  fn phi => fn ctxt => fn ct =>
+  K (fn ctxt => fn ct =>
     let
       val f = #2 (Thm.dest_comb (#2 (Thm.dest_comb ct)));
       val [T, U] = Thm.dest_ctyp (Thm.ctyp_of_cterm f);
       val tr = Thm.instantiate' [SOME T, SOME U] [SOME f] (mk_meta_eq @{thm Abs_cfun_inverse2});
       val rules = Named_Theorems.get ctxt \<^named_theorems>\<open>cont2cont\<close>;
       val tac = SOLVED' (REPEAT_ALL_NEW (match_tac ctxt (rev rules)));
-    in SOME (perhaps (SINGLE (tac 1)) tr) end
+    in SOME (perhaps (SINGLE (tac 1)) tr) end)
 \<close>
 
 text \<open>Eta-equality for continuous functions\<close>
--- a/src/HOL/HOLCF/Pcpo.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/HOLCF/Pcpo.thy	Fri May 26 09:56:20 2023 +0200
@@ -157,7 +157,7 @@
 
 text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close>
 setup \<open>Reorient_Proc.add (fn \<^Const_>\<open>bottom _\<close> => true | _ => false)\<close>
-simproc_setup reorient_bottom ("\<bottom> = x") = Reorient_Proc.proc
+simproc_setup reorient_bottom ("\<bottom> = x") = \<open>K Reorient_Proc.proc\<close>
 
 text \<open>useful lemmas about \<^term>\<open>\<bottom>\<close>\<close>
 
--- a/src/HOL/HOLCF/Tools/fixrec.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/HOLCF/Tools/fixrec.ML	Fri May 26 09:56:20 2023 +0200
@@ -171,7 +171,7 @@
     fun unfold_note (name, thm) : Attrib.binding * Thm.thm list =
       let
         val thm_name = Binding.qualify true name (Binding.name "unfold")
-        val src = Attrib.internal (K add_unfold)
+        val src = Attrib.internal \<^here> (K add_unfold)
       in
         ((thm_name, [src]), [thm])
       end
--- a/src/HOL/Library/Extended_Nat.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Library/Extended_Nat.thy	Fri May 26 09:56:20 2023 +0200
@@ -588,15 +588,15 @@
 
 simproc_setup enat_eq_cancel
   ("(l::enat) + m = n" | "(l::enat) = m + n") =
-  \<open>fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct))\<close>
 
 simproc_setup enat_le_cancel
   ("(l::enat) + m \<le> n" | "(l::enat) \<le> m + n") =
-  \<open>fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct))\<close>
 
 simproc_setup enat_less_cancel
   ("(l::enat) + m < n" | "(l::enat) < m + n") =
-  \<open>fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct))\<close>
 
 text \<open>TODO: add regression tests for these simprocs\<close>
 
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy	Fri May 26 09:56:20 2023 +0200
@@ -433,15 +433,15 @@
 
 simproc_setup ennreal_eq_cancel
   ("(l::ennreal) + m = n" | "(l::ennreal) = m + n") =
-  \<open>fn phi => fn ctxt => fn ct => Eq_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => Eq_Ennreal_Cancel.proc ctxt (Thm.term_of ct))\<close>
 
 simproc_setup ennreal_le_cancel
   ("(l::ennreal) + m \<le> n" | "(l::ennreal) \<le> m + n") =
-  \<open>fn phi => fn ctxt => fn ct => Le_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => Le_Ennreal_Cancel.proc ctxt (Thm.term_of ct))\<close>
 
 simproc_setup ennreal_less_cancel
   ("(l::ennreal) + m < n" | "(l::ennreal) < m + n") =
-  \<open>fn phi => fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct))\<close>
 
 
 subsection \<open>Order with top\<close>
--- a/src/HOL/Library/Multiset.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Library/Multiset.thy	Fri May 26 09:56:20 2023 +0200
@@ -1003,28 +1003,28 @@
    "add_mset a m = n" | "m = add_mset a n" |
    "replicate_mset p a = n" | "m = replicate_mset p a" |
    "repeat_mset p m = n" | "m = repeat_mset p m") =
-  \<open>fn phi => Cancel_Simprocs.eq_cancel\<close>
+  \<open>K Cancel_Simprocs.eq_cancel\<close>
 
 simproc_setup msetsubset_cancel
   ("(l::'a multiset) + m \<subset># n" | "(l::'a multiset) \<subset># m + n" |
    "add_mset a m \<subset># n" | "m \<subset># add_mset a n" |
    "replicate_mset p r \<subset># n" | "m \<subset># replicate_mset p r" |
    "repeat_mset p m \<subset># n" | "m \<subset># repeat_mset p m") =
-  \<open>fn phi => Multiset_Simprocs.subset_cancel_msets\<close>
+  \<open>K Multiset_Simprocs.subset_cancel_msets\<close>
 
 simproc_setup msetsubset_eq_cancel
   ("(l::'a multiset) + m \<subseteq># n" | "(l::'a multiset) \<subseteq># m + n" |
    "add_mset a m \<subseteq># n" | "m \<subseteq># add_mset a n" |
    "replicate_mset p r \<subseteq># n" | "m \<subseteq># replicate_mset p r" |
    "repeat_mset p m \<subseteq># n" | "m \<subseteq># repeat_mset p m") =
-  \<open>fn phi => Multiset_Simprocs.subseteq_cancel_msets\<close>
+  \<open>K Multiset_Simprocs.subseteq_cancel_msets\<close>
 
 simproc_setup msetdiff_cancel
   ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" |
    "add_mset a m - n" | "m - add_mset a n" |
    "replicate_mset p r - n" | "m - replicate_mset p r" |
    "repeat_mset p m - n" | "m - repeat_mset p m") =
-  \<open>fn phi => Cancel_Simprocs.diff_cancel\<close>
+  \<open>K Cancel_Simprocs.diff_cancel\<close>
 
 
 subsubsection \<open>Conditionally complete lattice\<close>
--- a/src/HOL/Library/Multiset_Order.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Library/Multiset_Order.thy	Fri May 26 09:56:20 2023 +0200
@@ -776,14 +776,14 @@
    "add_mset a m < n" | "m < add_mset a n" |
    "replicate_mset p a < n" | "m < replicate_mset p a" |
    "repeat_mset p m < n" | "m < repeat_mset p n") =
-  \<open>fn phi => Cancel_Simprocs.less_cancel\<close>
+  \<open>K Cancel_Simprocs.less_cancel\<close>
 
 simproc_setup msetle_cancel
   ("(l::'a::preorder multiset) + m \<le> n" | "(l::'a multiset) \<le> m + n" |
    "add_mset a m \<le> n" | "m \<le> add_mset a n" |
    "replicate_mset p a \<le> n" | "m \<le> replicate_mset p a" |
    "repeat_mset p m \<le> n" | "m \<le> repeat_mset p n") =
-  \<open>fn phi => Cancel_Simprocs.less_eq_cancel\<close>
+  \<open>K Cancel_Simprocs.less_eq_cancel\<close>
 
 
 subsection \<open>Additional facts and instantiations\<close>
--- a/src/HOL/Library/adhoc_overloading.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Library/adhoc_overloading.ML	Fri May 26 09:56:20 2023 +0200
@@ -227,7 +227,7 @@
       |> map (apfst (const_name lthy))
       |> map (apsnd (map (read_term lthy)));
   in
-    Local_Theory.declaration {syntax = true, pervasive = false}
+    Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()}
       (adhoc_overloading_cmd' add args) lthy
   end;
 
--- a/src/HOL/List.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/List.thy	Fri May 26 09:56:20 2023 +0200
@@ -1045,7 +1045,7 @@
         else if lastl aconv lastr then rearr @{thm append_same_eq}
         else NONE
       end;
-  in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end
+  in K (fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct)) end
 \<close>
 
 
--- a/src/HOL/Nat.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Nat.thy	Fri May 26 09:56:20 2023 +0200
@@ -1968,19 +1968,19 @@
 
 simproc_setup nateq_cancel_sums
   ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") =
-  \<open>fn phi => try o Nat_Arith.cancel_eq_conv\<close>
+  \<open>K (try o Nat_Arith.cancel_eq_conv)\<close>
 
 simproc_setup natless_cancel_sums
   ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") =
-  \<open>fn phi => try o Nat_Arith.cancel_less_conv\<close>
+  \<open>K (try o Nat_Arith.cancel_less_conv)\<close>
 
 simproc_setup natle_cancel_sums
   ("(l::nat) + m \<le> n" | "(l::nat) \<le> m + n" | "Suc m \<le> n" | "m \<le> Suc n") =
-  \<open>fn phi => try o Nat_Arith.cancel_le_conv\<close>
+  \<open>K (try o Nat_Arith.cancel_le_conv)\<close>
 
 simproc_setup natdiff_cancel_sums
   ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") =
-  \<open>fn phi => try o Nat_Arith.cancel_diff_conv\<close>
+  \<open>K (try o Nat_Arith.cancel_diff_conv)\<close>
 
 context order
 begin
@@ -2029,6 +2029,17 @@
 lemma antimono_iff_le_Suc: "antimono f \<longleftrightarrow> (\<forall>n. f (Suc n) \<le> f n)"
   unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f])
 
+lemma strict_mono_Suc_iff: "strict_mono f \<longleftrightarrow> (\<forall>n. f n < f (Suc n))"
+proof (intro iffI strict_monoI)
+  assume *: "\<forall>n. f n < f (Suc n)"
+  fix m n :: nat assume "m < n"
+  thus "f m < f n"
+    by (induction rule: less_Suc_induct) (use * in auto)
+qed (auto simp: strict_mono_def)
+
+lemma strict_mono_add: "strict_mono (\<lambda>n::'a::linordered_semidom. n + k)"
+  by (auto simp: strict_mono_def)
+
 lemma mono_nat_linear_lb:
   fixes f :: "nat \<Rightarrow> nat"
   assumes "\<And>m n. m < n \<Longrightarrow> f m < f n"
--- a/src/HOL/Nominal/nominal_inductive.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Fri May 26 09:56:20 2023 +0200
@@ -566,7 +566,7 @@
         val strong_cases = map (mk_cases_proof ##> Inductive.rulify lthy1)
           (thsss ~~ elims ~~ cases_prems ~~ cases_prems');
         val strong_induct_atts =
-          map (Attrib.internal o K)
+          map (Attrib.internal \<^here> o K)
             [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))];
         val strong_induct =
           if length names > 1 then strong_raw_induct
@@ -580,12 +580,12 @@
         Local_Theory.notes
           [((rec_qualified (Binding.name "strong_inducts"), []),
             strong_inducts |> map (fn th => ([th],
-              [Attrib.internal (K ind_case_names),
-               Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |>
+              [Attrib.internal \<^here> (K ind_case_names),
+               Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |>
         Local_Theory.notes (map (fn ((name, elim), (_, cases)) =>
             ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"),
-              [Attrib.internal (K (Rule_Cases.case_names (map snd cases))),
-               Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])]))
+              [Attrib.internal \<^here> (K (Rule_Cases.case_names (map snd cases))),
+               Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])]))
           (strong_cases ~~ induct_cases')) |> snd
       end)
       (map (map (rulify_term thy #> rpair [])) vc_compat)
@@ -675,7 +675,7 @@
     lthy |>
     Local_Theory.notes (map (fn (name, ths) =>
         ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"),
-          [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])]))
+          @{attributes [eqvt]}), [(ths, [])]))
       (names ~~ transp thss)) |> snd
   end;
 
--- a/src/HOL/Nominal/nominal_inductive2.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Fri May 26 09:56:20 2023 +0200
@@ -466,7 +466,7 @@
         val strong_raw_induct =
           mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1;
         val strong_induct_atts =
-          map (Attrib.internal o K)
+          map (Attrib.internal \<^here> o K)
             [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))];
         val strong_induct =
           if length names > 1 then strong_raw_induct
@@ -484,8 +484,8 @@
         lthy2 |>
         Local_Theory.notes [((inducts_name, []),
           strong_inducts |> map (fn th => ([th],
-            [Attrib.internal (K ind_case_names),
-             Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd
+            [Attrib.internal \<^here> (K ind_case_names),
+             Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd
       end)
       (map (map (rulify_term thy #> rpair [])) vc_compat)
   end;
--- a/src/HOL/Nonstandard_Analysis/NSA.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Nonstandard_Analysis/NSA.thy	Fri May 26 09:56:20 2023 +0200
@@ -549,12 +549,12 @@
   ("0 \<approx> x" | "1 \<approx> y" | "numeral w \<approx> z" | "- 1 \<approx> y" | "- numeral w \<approx> r") =
 \<open>
   let val rule = @{thm approx_reorient} RS eq_reflection
-      fun proc phi ss ct =
+      fun proc ct =
         case Thm.term_of ct of
           _ $ t $ u => if can HOLogic.dest_number u then NONE
             else if can HOLogic.dest_number t then SOME rule else NONE
         | _ => NONE
-  in proc end
+  in K (K proc) end
 \<close>
 
 lemma Infinitesimal_approx_minus: "x - y \<in> Infinitesimal \<longleftrightarrow> x \<approx> y"
--- a/src/HOL/Num.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Num.thy	Fri May 26 09:56:20 2023 +0200
@@ -1394,7 +1394,7 @@
 \<close>
 
 simproc_setup reorient_numeral ("numeral w = x" | "- numeral w = y") =
-  Reorient_Proc.proc
+  \<open>K Reorient_Proc.proc\<close>
 
 
 subsubsection \<open>Simplification of arithmetic operations on integer constants\<close>
--- a/src/HOL/Numeral_Simprocs.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Numeral_Simprocs.thy	Fri May 26 09:56:20 2023 +0200
@@ -104,17 +104,17 @@
 
 simproc_setup semiring_assoc_fold
   ("(a::'a::comm_semiring_1_cancel) * b") =
-  \<open>fn phi => Numeral_Simprocs.assoc_fold\<close>
+  \<open>K Numeral_Simprocs.assoc_fold\<close>
 
 (* TODO: see whether the type class can be generalized further *)
 simproc_setup int_combine_numerals
   ("(i::'a::comm_ring_1) + j" | "(i::'a::comm_ring_1) - j") =
-  \<open>fn phi => Numeral_Simprocs.combine_numerals\<close>
+  \<open>K Numeral_Simprocs.combine_numerals\<close>
 
 simproc_setup field_combine_numerals
   ("(i::'a::{field,ring_char_0}) + j"
   |"(i::'a::{field,ring_char_0}) - j") =
-  \<open>fn phi => Numeral_Simprocs.field_combine_numerals\<close>
+  \<open>K Numeral_Simprocs.field_combine_numerals\<close>
 
 simproc_setup inteq_cancel_numerals
   ("(l::'a::comm_ring_1) + m = n"
@@ -125,7 +125,7 @@
   |"(l::'a::comm_ring_1) = m * n"
   |"- (l::'a::comm_ring_1) = m"
   |"(l::'a::comm_ring_1) = - m") =
-  \<open>fn phi => Numeral_Simprocs.eq_cancel_numerals\<close>
+  \<open>K Numeral_Simprocs.eq_cancel_numerals\<close>
 
 simproc_setup intless_cancel_numerals
   ("(l::'a::linordered_idom) + m < n"
@@ -136,7 +136,7 @@
   |"(l::'a::linordered_idom) < m * n"
   |"- (l::'a::linordered_idom) < m"
   |"(l::'a::linordered_idom) < - m") =
-  \<open>fn phi => Numeral_Simprocs.less_cancel_numerals\<close>
+  \<open>K Numeral_Simprocs.less_cancel_numerals\<close>
 
 simproc_setup intle_cancel_numerals
   ("(l::'a::linordered_idom) + m \<le> n"
@@ -147,138 +147,138 @@
   |"(l::'a::linordered_idom) \<le> m * n"
   |"- (l::'a::linordered_idom) \<le> m"
   |"(l::'a::linordered_idom) \<le> - m") =
-  \<open>fn phi => Numeral_Simprocs.le_cancel_numerals\<close>
+  \<open>K Numeral_Simprocs.le_cancel_numerals\<close>
 
 simproc_setup ring_eq_cancel_numeral_factor
   ("(l::'a::{idom,ring_char_0}) * m = n"
   |"(l::'a::{idom,ring_char_0}) = m * n") =
-  \<open>fn phi => Numeral_Simprocs.eq_cancel_numeral_factor\<close>
+  \<open>K Numeral_Simprocs.eq_cancel_numeral_factor\<close>
 
 simproc_setup ring_less_cancel_numeral_factor
   ("(l::'a::linordered_idom) * m < n"
   |"(l::'a::linordered_idom) < m * n") =
-  \<open>fn phi => Numeral_Simprocs.less_cancel_numeral_factor\<close>
+  \<open>K Numeral_Simprocs.less_cancel_numeral_factor\<close>
 
 simproc_setup ring_le_cancel_numeral_factor
   ("(l::'a::linordered_idom) * m <= n"
   |"(l::'a::linordered_idom) <= m * n") =
-  \<open>fn phi => Numeral_Simprocs.le_cancel_numeral_factor\<close>
+  \<open>K Numeral_Simprocs.le_cancel_numeral_factor\<close>
 
 (* TODO: remove comm_ring_1 constraint if possible *)
 simproc_setup int_div_cancel_numeral_factors
   ("((l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) * m) div n"
   |"(l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) div (m * n)") =
-  \<open>fn phi => Numeral_Simprocs.div_cancel_numeral_factor\<close>
+  \<open>K Numeral_Simprocs.div_cancel_numeral_factor\<close>
 
 simproc_setup divide_cancel_numeral_factor
   ("((l::'a::{field,ring_char_0}) * m) / n"
   |"(l::'a::{field,ring_char_0}) / (m * n)"
   |"((numeral v)::'a::{field,ring_char_0}) / (numeral w)") =
-  \<open>fn phi => Numeral_Simprocs.divide_cancel_numeral_factor\<close>
+  \<open>K Numeral_Simprocs.divide_cancel_numeral_factor\<close>
 
 simproc_setup ring_eq_cancel_factor
   ("(l::'a::idom) * m = n" | "(l::'a::idom) = m * n") =
-  \<open>fn phi => Numeral_Simprocs.eq_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.eq_cancel_factor\<close>
 
 simproc_setup linordered_ring_le_cancel_factor
   ("(l::'a::linordered_idom) * m <= n"
   |"(l::'a::linordered_idom) <= m * n") =
-  \<open>fn phi => Numeral_Simprocs.le_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.le_cancel_factor\<close>
 
 simproc_setup linordered_ring_less_cancel_factor
   ("(l::'a::linordered_idom) * m < n"
   |"(l::'a::linordered_idom) < m * n") =
-  \<open>fn phi => Numeral_Simprocs.less_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.less_cancel_factor\<close>
 
 simproc_setup int_div_cancel_factor
   ("((l::'a::euclidean_semiring_cancel) * m) div n"
   |"(l::'a::euclidean_semiring_cancel) div (m * n)") =
-  \<open>fn phi => Numeral_Simprocs.div_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.div_cancel_factor\<close>
 
 simproc_setup int_mod_cancel_factor
   ("((l::'a::euclidean_semiring_cancel) * m) mod n"
   |"(l::'a::euclidean_semiring_cancel) mod (m * n)") =
-  \<open>fn phi => Numeral_Simprocs.mod_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.mod_cancel_factor\<close>
 
 simproc_setup dvd_cancel_factor
   ("((l::'a::idom) * m) dvd n"
   |"(l::'a::idom) dvd (m * n)") =
-  \<open>fn phi => Numeral_Simprocs.dvd_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.dvd_cancel_factor\<close>
 
 simproc_setup divide_cancel_factor
   ("((l::'a::field) * m) / n"
   |"(l::'a::field) / (m * n)") =
-  \<open>fn phi => Numeral_Simprocs.divide_cancel_factor\<close>
+  \<open>K Numeral_Simprocs.divide_cancel_factor\<close>
 
 ML_file \<open>Tools/nat_numeral_simprocs.ML\<close>
 
 simproc_setup nat_combine_numerals
   ("(i::nat) + j" | "Suc (i + j)") =
-  \<open>fn phi => Nat_Numeral_Simprocs.combine_numerals\<close>
+  \<open>K Nat_Numeral_Simprocs.combine_numerals\<close>
 
 simproc_setup nateq_cancel_numerals
   ("(l::nat) + m = n" | "(l::nat) = m + n" |
    "(l::nat) * m = n" | "(l::nat) = m * n" |
    "Suc m = n" | "m = Suc n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_numerals\<close>
+  \<open>K Nat_Numeral_Simprocs.eq_cancel_numerals\<close>
 
 simproc_setup natless_cancel_numerals
   ("(l::nat) + m < n" | "(l::nat) < m + n" |
    "(l::nat) * m < n" | "(l::nat) < m * n" |
    "Suc m < n" | "m < Suc n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_numerals\<close>
+  \<open>K Nat_Numeral_Simprocs.less_cancel_numerals\<close>
 
 simproc_setup natle_cancel_numerals
   ("(l::nat) + m \<le> n" | "(l::nat) \<le> m + n" |
    "(l::nat) * m \<le> n" | "(l::nat) \<le> m * n" |
    "Suc m \<le> n" | "m \<le> Suc n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_numerals\<close>
+  \<open>K Nat_Numeral_Simprocs.le_cancel_numerals\<close>
 
 simproc_setup natdiff_cancel_numerals
   ("((l::nat) + m) - n" | "(l::nat) - (m + n)" |
    "(l::nat) * m - n" | "(l::nat) - m * n" |
    "Suc m - n" | "m - Suc n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.diff_cancel_numerals\<close>
+  \<open>K Nat_Numeral_Simprocs.diff_cancel_numerals\<close>
 
 simproc_setup nat_eq_cancel_numeral_factor
   ("(l::nat) * m = n" | "(l::nat) = m * n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_numeral_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.eq_cancel_numeral_factor\<close>
 
 simproc_setup nat_less_cancel_numeral_factor
   ("(l::nat) * m < n" | "(l::nat) < m * n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_numeral_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.less_cancel_numeral_factor\<close>
 
 simproc_setup nat_le_cancel_numeral_factor
   ("(l::nat) * m <= n" | "(l::nat) <= m * n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_numeral_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.le_cancel_numeral_factor\<close>
 
 simproc_setup nat_div_cancel_numeral_factor
   ("((l::nat) * m) div n" | "(l::nat) div (m * n)") =
-  \<open>fn phi => Nat_Numeral_Simprocs.div_cancel_numeral_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.div_cancel_numeral_factor\<close>
 
 simproc_setup nat_dvd_cancel_numeral_factor
   ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
-  \<open>fn phi => Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\<close>
 
 simproc_setup nat_eq_cancel_factor
   ("(l::nat) * m = n" | "(l::nat) = m * n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.eq_cancel_factor\<close>
 
 simproc_setup nat_less_cancel_factor
   ("(l::nat) * m < n" | "(l::nat) < m * n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.less_cancel_factor\<close>
 
 simproc_setup nat_le_cancel_factor
   ("(l::nat) * m <= n" | "(l::nat) <= m * n") =
-  \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.le_cancel_factor\<close>
 
 simproc_setup nat_div_cancel_factor
   ("((l::nat) * m) div n" | "(l::nat) div (m * n)") =
-  \<open>fn phi => Nat_Numeral_Simprocs.div_cancel_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.div_cancel_factor\<close>
 
 simproc_setup nat_dvd_cancel_factor
   ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
-  \<open>fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor\<close>
+  \<open>K Nat_Numeral_Simprocs.dvd_cancel_factor\<close>
 
 declaration \<open>
   K (Lin_Arith.add_simprocs
--- a/src/HOL/Parity.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Parity.thy	Fri May 26 09:56:20 2023 +0200
@@ -1311,24 +1311,22 @@
       let
         val thm = Simplifier.rewrite ctxt ct
       in if Thm.is_reflexive thm then NONE else SOME thm end;
-  in fn phi =>
-    let
-      val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
-        one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral
-        one_div_minus_numeral one_mod_minus_numeral
-        numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral
-        numeral_div_minus_numeral numeral_mod_minus_numeral
-        div_minus_minus mod_minus_minus Parity.adjust_div_eq of_bool_eq one_neq_zero
-        numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial
-        divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One
-        case_prod_beta rel_simps Parity.adjust_mod_def div_minus1_right mod_minus1_right
-        minus_minus numeral_times_numeral mult_zero_right mult_1_right
-        euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral}
-        @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
-      fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt
-        (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps)
-    in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end
-  end
+    val simps = @{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
+      one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral
+      one_div_minus_numeral one_mod_minus_numeral
+      numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral
+      numeral_div_minus_numeral numeral_mod_minus_numeral
+      div_minus_minus mod_minus_minus Parity.adjust_div_eq of_bool_eq one_neq_zero
+      numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial
+      divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One
+      case_prod_beta rel_simps Parity.adjust_mod_def div_minus1_right mod_minus1_right
+      minus_minus numeral_times_numeral mult_zero_right mult_1_right
+      euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral}
+      @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}];
+    val simpset =
+      HOL_ss |> Simplifier.simpset_map \<^context>
+        (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps);
+  in K (fn ctxt => successful_rewrite (Simplifier.put_simpset simpset ctxt)) end
 \<close> \<comment> \<open>
   There is space for improvement here: the calculation itself
   could be carried out outside the logic, and a generic simproc
--- a/src/HOL/Product_Type.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Product_Type.thy	Fri May 26 09:56:20 2023 +0200
@@ -73,9 +73,9 @@
 \<close>
 
 simproc_setup unit_eq ("x::unit") = \<open>
-  fn _ => fn _ => fn ct =>
+  K (K (fn ct =>
     if HOLogic.is_unit (Thm.term_of ct) then NONE
-    else SOME (mk_meta_eq @{thm unit_eq})
+    else SOME (mk_meta_eq @{thm unit_eq})))
 \<close>
 
 free_constructors case_unit for "()"
@@ -570,9 +570,9 @@
 end;
 \<close>
 simproc_setup case_prod_beta ("case_prod f z") =
-  \<open>fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct))\<close>
 simproc_setup case_prod_eta ("case_prod f") =
-  \<open>fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct)\<close>
+  \<open>K (fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct))\<close>
 
 lemma case_prod_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
   by (auto simp: fun_eq_iff)
@@ -1326,7 +1326,7 @@
 
 (* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
 simproc_setup Collect_mem ("Collect t") = \<open>
-  fn _ => fn ctxt => fn ct =>
+  K (fn ctxt => fn ct =>
     (case Thm.term_of ct of
       S as Const (\<^const_name>\<open>Collect\<close>, Type (\<^type_name>\<open>fun\<close>, [_, T])) $ t =>
         let val (u, _, ps) = HOLogic.strip_ptupleabs t in
@@ -1355,7 +1355,7 @@
                   else NONE)
           | _ => NONE)
         end
-    | _ => NONE)
+    | _ => NONE))
 \<close>
 
 ML_file \<open>Tools/inductive_set.ML\<close>
--- a/src/HOL/Real_Asymp/exp_log_expression.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Real_Asymp/exp_log_expression.ML	Fri May 26 09:56:20 2023 +0200
@@ -229,7 +229,7 @@
     fun decl phi =
       register_custom' binding (Morphism.term phi pat) expand
   in
-    Local_Theory.declaration {syntax = false, pervasive = false} decl
+    Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} decl
   end
 
 fun register_custom_from_thm binding thm expand = 
--- a/src/HOL/Set.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Set.thy	Fri May 26 09:56:20 2023 +0200
@@ -65,13 +65,13 @@
 \<close>
 
 simproc_setup defined_Collect ("{x. P x \<and> Q x}") = \<open>
-  fn _ => Quantifier1.rearrange_Collect
+  K (Quantifier1.rearrange_Collect
     (fn ctxt =>
       resolve_tac ctxt @{thms Collect_cong} 1 THEN
       resolve_tac ctxt @{thms iffI} 1 THEN
       ALLGOALS
         (EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE},
-          DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})]))
+          DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})])))
 \<close>
 
 lemmas CollectE = CollectD [elim_format]
@@ -330,13 +330,11 @@
 \<close>
 
 simproc_setup defined_Bex ("\<exists>x\<in>A. P x \<and> Q x") = \<open>
-  fn _ => Quantifier1.rearrange_Bex
-    (fn ctxt => unfold_tac ctxt @{thms Bex_def})
+  K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}))
 \<close>
 
 simproc_setup defined_All ("\<forall>x\<in>A. P x \<longrightarrow> Q x") = \<open>
-  fn _ => Quantifier1.rearrange_Ball
-    (fn ctxt => unfold_tac ctxt @{thms Ball_def})
+  K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def}))
 \<close>
 
 lemma ballI [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> P x) \<Longrightarrow> \<forall>x\<in>A. P x"
--- a/src/HOL/Statespace/state_space.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Statespace/state_space.ML	Fri May 26 09:56:20 2023 +0200
@@ -228,12 +228,14 @@
   else (tracing ("variable not fixed or declared: " ^ name); NONE)
 
 
-val get_dist_thm = Symtab.lookup o get_distinctthm;
+fun get_dist_thm context name =
+  Symtab.lookup_list (get_distinctthm context) name
+  |> map (Thm.transfer'' context)
 
 fun get_dist_thm2 ctxt x y =
  (let
     val dist_thms = [x, y] |> map (#1 o dest_Free)
-      |> map (these o get_dist_thm (Context.Proof ctxt)) |> flat;
+      |> maps (get_dist_thm (Context.Proof ctxt));
 
     fun get_paths dist_thm =
       let
@@ -331,7 +333,8 @@
           val declinfo = get_declinfo context
           val tt = get_distinctthm context;
           val all_names = comps_of_distinct_thm thm;
-          fun upd name = Symtab.map_default (name, [thm]) (insert_distinct_thm thm)
+          val thm0 = Thm.trim_context thm;
+          fun upd name = Symtab.map_default (name, [thm0]) (insert_distinct_thm thm0)
 
           val tt' = tt |> fold upd all_names;
           val context' =
@@ -343,7 +346,7 @@
               |> map_distinctthm (K tt');
         in context' end));
 
-    val attr = Attrib.internal type_attr;
+    val attr = Attrib.internal \<^here> type_attr;
 
     val assume =
       ((Binding.name dist_thm_name, [attr]),
@@ -375,7 +378,9 @@
 fun add_declaration name decl thy =
   thy
   |> Named_Target.init [] name
-  |> (fn lthy => Local_Theory.declaration {syntax = true, pervasive = false} (decl lthy) lthy)
+  |> (fn lthy =>
+    Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()}
+      (decl lthy) lthy)
   |> Local_Theory.exit_global;
 
 fun parent_components thy (Ts, pname, renaming) =
--- a/src/HOL/Tools/BNF/bnf_def.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/BNF/bnf_def.ML	Fri May 26 09:56:20 2023 +0200
@@ -2081,7 +2081,7 @@
 val interpret_bnf = BNF_Plugin.data;
 
 fun register_bnf_raw key bnf =
-  Local_Theory.declaration {syntax = false, pervasive = true}
+  Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
     (fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf)));
 
 fun register_bnf plugins key bnf =
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Fri May 26 09:56:20 2023 +0200
@@ -418,7 +418,7 @@
 
 val register_fp_sugars_raw =
   fold (fn fp_sugar as {T = Type (s, _), ...} =>
-    Local_Theory.declaration {syntax = false, pervasive = true}
+    Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
       (fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar))));
 
 fun register_fp_sugars plugins fp_sugars =
@@ -1069,7 +1069,7 @@
                   |> rotate_prems ~1;
 
                 val cases_set_attr =
-                  Attrib.internal (K (Induct.cases_pred (name_of_set set)));
+                  Attrib.internal \<^here> (K (Induct.cases_pred (name_of_set set)));
 
                 val ctr_names = quasi_unambiguous_case_names (flat
                   (map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs)));
@@ -2031,7 +2031,7 @@
           |> Thm.close_derivation \<^here>;
 
         val case_names_attr = Attrib.case_names (quasi_unambiguous_case_names case_names);
-        val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets;
+        val induct_set_attrs = map (Attrib.internal \<^here> o K o Induct.induct_pred o name_of_set) sets;
       in
         (thm, case_names_attr :: induct_set_attrs)
       end
@@ -2647,8 +2647,8 @@
         val rec_transfer_thmss =
           map single (derive_rec_transfer_thms lthy recs rec_defs recs_args_types);
 
-        val induct_type_attr = Attrib.internal o K o Induct.induct_type;
-        val induct_pred_attr = Attrib.internal o K o Induct.induct_pred;
+        val induct_type_attr = Attrib.internal \<^here> o K o Induct.induct_type;
+        val induct_pred_attr = Attrib.internal \<^here> o K o Induct.induct_pred;
 
         val ((rel_induct_thmss, common_rel_induct_thms),
              (rel_induct_attrs, common_rel_induct_attrs)) =
@@ -2808,8 +2808,8 @@
         val corec_code_thms = map (eq_ifIN lthy) corec_thmss;
         val corec_sel_thmss = map flat corec_sel_thmsss;
 
-        val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type;
-        val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred;
+        val coinduct_type_attr = Attrib.internal \<^here> o K o Induct.coinduct_type;
+        val coinduct_pred_attr = Attrib.internal \<^here> o K o Induct.coinduct_pred;
 
         val flat_corec_thms = append oo append;
 
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Fri May 26 09:56:20 2023 +0200
@@ -59,7 +59,7 @@
   #> Option.map (transfer_n2m_sugar (Proof_Context.theory_of ctxt));
 
 fun register_n2m_sugar key n2m_sugar =
-  Local_Theory.declaration {syntax = false, pervasive = false}
+  Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
     (fn phi => Data.map (Typtab.update (key, morph_n2m_sugar phi n2m_sugar)));
 
 fun unfold_lets_splits (Const (\<^const_name>\<open>Let\<close>, _) $ t $ u) =
--- a/src/HOL/Tools/BNF/bnf_gfp_grec.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_grec.ML	Fri May 26 09:56:20 2023 +0200
@@ -2909,7 +2909,7 @@
   end;
 
 fun set_corec_info_exprs fpT_name f =
-  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
+  Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} (fn phi =>
     let val exprs = f phi in
       Data.map (apsnd (fn [info_tab] => [Symtab.map_default (fpT_name, exprs) (K exprs) info_tab]))
     end);
@@ -3214,7 +3214,7 @@
     let
       val (info_tab', lthy) = fold_rev (uncurry o merge_corec_info_tabs) info_tabs (info_tab, lthy);
     in
-      Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
+      Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} (fn phi =>
           Data.map (apsnd (fn _ => [Symtab.map (K (map (morph_corec_info_expr phi))) info_tab'])))
         lthy
     end);
--- a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML	Fri May 26 09:56:20 2023 +0200
@@ -143,8 +143,9 @@
 );
 
 fun register_codatatype_extra fpT_name extra =
-  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
-    Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra))));
+  Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()}
+    (fn phi =>
+      Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra))));
 
 fun codatatype_extra_of ctxt =
   Symtab.lookup (#1 (Data.get (Context.Proof ctxt)))
@@ -154,19 +155,21 @@
   Symtab.dest (#1 (Data.get (Context.Proof ctxt)));
 
 fun register_coinduct_extra fpT_name extra =
-  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
-    Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra))));
+  Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()}
+    (fn phi =>
+      Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra))));
 
 fun coinduct_extra_of ctxt =
   Symtab.lookup (#2 (Data.get (Context.Proof ctxt)))
   #> Option.map (transfer_coinduct_extra (Proof_Context.theory_of ctxt));
 
 fun register_friend_extra fun_name eq_algrho algrho_eq =
-  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
-    Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra)
-      (fn {eq_algrhos, algrho_eqs} =>
-        {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos,
-         algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs}))));
+  Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()}
+    (fn phi =>
+      Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra)
+        (fn {eq_algrhos, algrho_eqs} =>
+          {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos,
+           algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs}))));
 
 fun all_friend_extras_of ctxt =
   Symtab.dest (#3 (Data.get (Context.Proof ctxt)));
--- a/src/HOL/Tools/BNF/bnf_lfp_size.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML	Fri May 26 09:56:20 2023 +0200
@@ -382,7 +382,7 @@
       val phi0 = substitute_noted_thm noted;
     in
       lthy3
-      |> Local_Theory.declaration {syntax = false, pervasive = true}
+      |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
         (fn phi => Data.map (@{fold 3} (fn T_name => fn Const (size_name, _) =>
             fn overloaded_size_def =>
                let val morph = Morphism.thm (phi0 $> phi) in
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Fri May 26 09:56:20 2023 +0200
@@ -210,7 +210,7 @@
 val interpret_ctr_sugar = Ctr_Sugar_Plugin.data;
 
 fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) =
-  Local_Theory.declaration {syntax = false, pervasive = true}
+  Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
     (fn phi => fn context =>
       let val pos = Position.thread_data ()
       in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end);
@@ -1072,8 +1072,8 @@
             |> Thm.close_derivation \<^here>
           end;
 
-        val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
-        val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name));
+        val exhaust_case_names_attr = Attrib.internal \<^here> (K (Rule_Cases.case_names exhaust_cases));
+        val cases_type_attr = Attrib.internal \<^here> (K (Induct.cases_type fcT_name));
 
         val nontriv_disc_eq_thmss =
           map (map (fn th => th RS @{thm eq_False[THEN iffD2]}
@@ -1119,12 +1119,12 @@
             (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms))
           |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational))
             (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss))
-          |> Local_Theory.declaration {syntax = false, pervasive = true}
+          |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
             (fn phi => Case_Translation.register
                (Morphism.term phi casex) (map (Morphism.term phi) ctrs))
           |> plugins code_plugin ?
              (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms))
-             #> Local_Theory.declaration {syntax = false, pervasive = false}
+             #> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
                (fn phi => Context.mapping
                   (add_ctr_code fcT_name (map (Morphism.typ phi) As)
                      (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms)
--- a/src/HOL/Tools/Function/function.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Function/function.ML	Fri May 26 09:56:20 2023 +0200
@@ -129,7 +129,7 @@
             (K false) (map fst fixes)
       in
         (info,
-         lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false}
+         lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
           (fn phi => add_function_data (transform_function_data phi info)))
       end
   in
@@ -209,7 +209,7 @@
           in
             (info',
              lthy2
-             |> Local_Theory.declaration {syntax = false, pervasive = false}
+             |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
                (fn phi => add_function_data (transform_function_data phi info'))
              |> Spec_Rules.add Binding.empty Spec_Rules.equational_recdef fs tsimps)
           end)
--- a/src/HOL/Tools/Function/partial_function.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Function/partial_function.ML	Fri May 26 09:56:20 2023 +0200
@@ -6,7 +6,7 @@
 
 signature PARTIAL_FUNCTION =
 sig
-  val init: string -> term -> term -> thm -> thm -> thm option -> declaration
+  val init: string -> term -> term -> thm -> thm -> thm option -> Morphism.declaration
   val mono_tac: Proof.context -> int -> tactic
   val add_partial_function: string -> (binding * typ option * mixfix) list ->
     Attrib.binding * term -> local_theory -> (term * thm) * local_theory
--- a/src/HOL/Tools/Lifting/lifting_def.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Lifting/lifting_def.ML	Fri May 26 09:56:20 2023 +0200
@@ -528,7 +528,7 @@
 
   val register_code_eq_attribute = Thm.declaration_attribute
     (fn thm => Context.mapping (register_encoded_code_eq thm) I)
-  val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute)
+  val register_code_eq_attrib = Attrib.internal \<^here> (K register_code_eq_attribute)
 
 in
 
@@ -547,7 +547,7 @@
             which code equation is going to be used. This is going to be resolved at the
             point when an interpretation of the locale is executed. *)
         val lthy'' = lthy'
-          |> Local_Theory.declaration {syntax = false, pervasive = true} (K (Data.put NONE))
+          |> Local_Theory.declaration {syntax = false, pervasive = true, pos =  \<^here>} (K (Data.put NONE))
       in (code_eq, lthy'') end
     else
       (NONE_EQ, lthy)
--- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML	Fri May 26 09:56:20 2023 +0200
@@ -146,7 +146,7 @@
 
 fun update_code_dt code_dt =
   (snd o Local_Theory.begin_nested)
-  #> Local_Theory.declaration {syntax = false, pervasive = true}
+  #> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
     (fn phi => Data.map (Item_Net.update (morph_code_dt phi code_dt)))
   #> Local_Theory.end_nested
 
@@ -346,7 +346,7 @@
       |> update_rep_isom rep_isom (transfer_rules_of_lift_def rep_isom_lift_def |> hd)
           (bundle_name_of_bundle_binding qty_isom_bundle phi context) pointer;
     val lthy3 = lthy2  
-      |> Local_Theory.declaration {syntax = false, pervasive = true}
+      |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
         (fn phi => fn context => Data.map (Item_Net.update (morph_code_dt phi (code_dt phi context))) context)
       |> Local_Theory.end_nested
 
--- a/src/HOL/Tools/Lifting/lifting_setup.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Lifting/lifting_setup.ML	Fri May 26 09:56:20 2023 +0200
@@ -245,12 +245,10 @@
 
     val dummy_thm = Thm.transfer thy Drule.dummy_thm
     val restore_lifting_att = 
-      ([dummy_thm],
-        [map (Token.make_string o rpair Position.none)
-          ["Lifting.lifting_restore_internal", bundle_name]])
+      ([dummy_thm], [map Token.make_string0 ["Lifting.lifting_restore_internal", bundle_name]])
   in
     lthy
-    |> Local_Theory.declaration {syntax = false, pervasive = true}
+    |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
         (fn phi => Lifting_Info.init_restore_data bundle_name (phi_qinfo phi))
     |> Bundle.bundle ((binding, [restore_lifting_att])) []
     |> pair binding
@@ -274,7 +272,7 @@
     val quotients = { quot_thm = quot_thm, pcr_info = pcr_info }
     val qty_full_name = (fst o dest_Type) qty
     fun quot_info phi = Lifting_Info.transform_quotient phi quotients
-    val reflexivity_rule_attr = Attrib.internal (K Lifting_Info.add_reflexivity_rule_attribute)
+    val reflexivity_rule_attr = Attrib.internal \<^here> (K Lifting_Info.add_reflexivity_rule_attribute)
     val lthy3 =
       case opt_reflp_thm of
         SOME reflp_thm =>
@@ -285,7 +283,7 @@
       | NONE => lthy2 |> define_abs_type quot_thm
   in
     lthy3
-    |> Local_Theory.declaration {syntax = false, pervasive = true}
+    |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
         (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi))
     |> lifting_bundle qty_full_name quotients
   end
@@ -522,7 +520,7 @@
     val quot_thm = Morphism.thm (Local_Theory.target_morphism lthy0) quot_thm
     (**)
     val (rty, qty) = quot_thm_rty_qty quot_thm
-    val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty))))
+    val induct_attr = Attrib.internal \<^here> (K (Induct.induct_type (fst (dest_Type qty))))
     val qty_full_name = (fst o dest_Type) qty
     val qty_name = (Binding.name o Long_Name.base_name) qty_full_name
     val qualify = Binding.qualify_name true qty_name
@@ -999,7 +997,7 @@
           val quot_thm = #quot_thm qinfo
           val transfer_rules = get_transfer_rules_to_delete qinfo lthy
         in
-          Local_Theory.declaration {syntax = false, pervasive = true}
+          Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
             (K (fold (Transfer.transfer_raw_del) transfer_rules #> Lifting_Info.delete_quotients quot_thm))
             lthy
         end
@@ -1031,7 +1029,7 @@
   in
     case Lifting_Info.lookup_restore_data lthy pointer of
       SOME refresh_data => 
-        Local_Theory.declaration {syntax = false, pervasive = true}
+        Local_Theory.declaration {syntax = false, pervasive = true, pos =  \<^here>}
           (fn phi => Lifting_Info.add_transfer_rules_in_restore_data pointer 
             (new_transfer_rules refresh_data lthy phi)) lthy
     | NONE => error "The lifting bundle refers to non-existent restore data."
--- a/src/HOL/Tools/Quotient/quotient_def.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Quotient/quotient_def.ML	Fri May 26 09:56:20 2023 +0200
@@ -74,7 +74,7 @@
     val rsp_thm_name = qualify lhs_name "rsp"
 
     val lthy'' = lthy'
-      |> Local_Theory.declaration {syntax = false, pervasive = true}
+      |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
         (fn phi =>
           (case qconst_data phi of
             qcinfo as {qconst = Const (c, _), ...} =>
--- a/src/HOL/Tools/Quotient/quotient_type.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Quotient/quotient_type.ML	Fri May 26 09:56:20 2023 +0200
@@ -149,7 +149,7 @@
       Quotient_Info.transform_abs_rep phi {abs = abs, rep = rep}
   in
     lthy
-    |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
+    |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi =>
         Quotient_Info.update_quotients (qty_full_name, quotients phi) #>
         Quotient_Info.update_abs_rep (qty_full_name, abs_rep phi))
     |> setup_lifting_package quot_thm equiv_thm opt_par_thm
--- a/src/HOL/Tools/Transfer/transfer_bnf.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/Transfer/transfer_bnf.ML	Fri May 26 09:56:20 2023 +0200
@@ -208,7 +208,7 @@
     lthy
     |> Local_Theory.notes notes
     |> snd
-    |> Local_Theory.declaration {syntax = false, pervasive = true}
+    |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
       (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data))
   end
 
@@ -260,7 +260,7 @@
       |> Transfer.update_pred_simps pred_injects
   in
     lthy
-    |> Local_Theory.declaration {syntax = false, pervasive = true}
+    |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
       (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data))
   end
 
--- a/src/HOL/Tools/functor.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/functor.ML	Fri May 26 09:56:20 2023 +0200
@@ -263,7 +263,7 @@
         |> Local_Theory.note ((qualify identityN, []),
             [prove_identity lthy id_thm])
         |> snd
-        |> Local_Theory.declaration {syntax = false, pervasive = false}
+        |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
           (mapper_declaration comp_thm id_thm))
   in
     lthy
--- a/src/HOL/Tools/groebner.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/groebner.ML	Fri May 26 09:56:20 2023 +0200
@@ -781,7 +781,7 @@
   in
     Simplifier.cert_simproc (Thm.theory_of_thm idl_sub) "poly_eq_simproc"
      {lhss = [Thm.term_of (Thm.lhs_of idl_sub)],
-      proc = fn _ => fn _ => proc}
+      proc = Morphism.entity (fn _ => fn _ => proc)}
   end;
 
 val poly_eq_ss =
--- a/src/HOL/Tools/inductive.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/inductive.ML	Fri May 26 09:56:20 2023 +0200
@@ -1032,7 +1032,7 @@
           [Attrib.case_names [rec_name],
            Attrib.case_conclusion (rec_name, intr_names),
            Attrib.consumes (1 - Thm.nprems_of raw_induct),
-           Attrib.internal (K (Induct.coinduct_pred (hd cnames)))])
+           Attrib.internal \<^here> (K (Induct.coinduct_pred (hd cnames)))])
       else if no_ind orelse length cnames > 1 then
         (raw_induct, ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))])
       else
@@ -1055,7 +1055,7 @@
           ((Binding.qualify true (Long_Name.base_name name) (Binding.name "cases"),
             ((if forall (equal "") cases then [] else [Attrib.case_names cases]) @
               [Attrib.consumes (1 - Thm.nprems_of elim), Attrib.constraints k,
-               Attrib.internal (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})),
+               Attrib.internal \<^here> (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})),
             [elim]) #>
         apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>>
       Local_Theory.note
@@ -1065,7 +1065,7 @@
     val (eqs', lthy3) = lthy2 |>
       fold_map (fn (name, eq) => Local_Theory.note
           ((Binding.qualify true (Long_Name.base_name name) (Binding.name "simps"),
-            [Attrib.internal (K equation_add_permissive)]), [eq])
+            [Attrib.internal \<^here> (K equation_add_permissive)]), [eq])
           #> apfst (hd o snd))
         (if null eqs then [] else (cnames ~~ eqs))
     val (inducts, lthy4) =
@@ -1077,7 +1077,7 @@
             inducts |> map (fn (name, th) => ([th],
               ind_case_names @
                 [Attrib.consumes (1 - Thm.nprems_of th),
-                 Attrib.internal (K (Induct.induct_pred name))])))] |>> snd o hd
+                 Attrib.internal \<^here> (K (Induct.induct_pred name))])))] |>> snd o hd
         end;
   in (intrs', elims', eqs', induct', inducts, lthy4) end;
 
@@ -1144,7 +1144,7 @@
        eqs = eqs'};
 
     val lthy4 = lthy3
-      |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi =>
+      |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi =>
         let val result' = transform_result phi result;
         in put_inductives ({names = cnames, coind = coind}, result') end);
   in (result, lthy4) end;
--- a/src/HOL/Tools/inductive_set.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/inductive_set.ML	Fri May 26 09:56:20 2023 +0200
@@ -494,7 +494,7 @@
                 [def, @{thm mem_Collect_eq}, @{thm case_prod_conv}]) 1))
         in
           lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"),
-            [Attrib.internal (K pred_set_conv_att)]),
+            [Attrib.internal \<^here> (K pred_set_conv_att)]),
               [conv_thm]) |> snd
         end) (preds ~~ cs ~~ cs_info ~~ defs) lthy2;
 
--- a/src/HOL/Tools/semiring_normalizer.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/semiring_normalizer.ML	Fri May 26 09:56:20 2023 +0200
@@ -173,7 +173,7 @@
     val raw_ring = prepare_ops raw_ring0;
     val raw_field = prepare_ops raw_field0;
   in
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context =>
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => fn context =>
       let
         val ctxt = Context.proof_of context;
         val key = Morphism.thm phi raw_key;
--- a/src/HOL/Tools/typedef.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Tools/typedef.ML	Fri May 26 09:56:20 2023 +0200
@@ -259,19 +259,19 @@
               make @{thm type_definition.Abs_inject})
           ||>> note ((Binding.suffix_name "_cases" Rep_name,
                 [Attrib.case_names [Binding.name_of Rep_name],
-                 Attrib.internal (K (Induct.cases_pred full_name))]),
+                 Attrib.internal \<^here> (K (Induct.cases_pred full_name))]),
               make @{thm type_definition.Rep_cases})
           ||>> note ((Binding.suffix_name "_cases" Abs_name,
                 [Attrib.case_names [Binding.name_of Abs_name],
-                 Attrib.internal (K (Induct.cases_type full_name))]),
+                 Attrib.internal \<^here> (K (Induct.cases_type full_name))]),
               make @{thm type_definition.Abs_cases})
           ||>> note ((Binding.suffix_name "_induct" Rep_name,
                 [Attrib.case_names [Binding.name_of Rep_name],
-                 Attrib.internal (K (Induct.induct_pred full_name))]),
+                 Attrib.internal \<^here> (K (Induct.induct_pred full_name))]),
               make @{thm type_definition.Rep_induct})
           ||>> note ((Binding.suffix_name "_induct" Abs_name,
                 [Attrib.case_names [Binding.name_of Abs_name],
-                 Attrib.internal (K (Induct.induct_type full_name))]),
+                 Attrib.internal \<^here> (K (Induct.induct_type full_name))]),
               make @{thm type_definition.Abs_induct});
 
         val info =
@@ -283,7 +283,7 @@
           Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct});
       in
         lthy3
-        |> Local_Theory.declaration {syntax = false, pervasive = true}
+        |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
           (fn phi => put_info full_name (transform_info phi info))
         |> Typedef_Plugin.data Plugin_Name.default_filter full_name
         |> pair (full_name, info)
--- a/src/HOL/Topological_Spaces.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/HOL/Topological_Spaces.thy	Fri May 26 09:56:20 2023 +0200
@@ -1319,17 +1319,6 @@
 
 subsubsection \<open>Subsequence (alternative definition, (e.g. Hoskins)\<close>
 
-lemma strict_mono_Suc_iff: "strict_mono f \<longleftrightarrow> (\<forall>n. f n < f (Suc n))"
-proof (intro iffI strict_monoI)
-  assume *: "\<forall>n. f n < f (Suc n)"
-  fix m n :: nat assume "m < n"
-  thus "f m < f n"
-    by (induction rule: less_Suc_induct) (use * in auto)
-qed (auto simp: strict_mono_def)
-
-lemma strict_mono_add: "strict_mono (\<lambda>n::'a::linordered_semidom. n + k)"
-  by (auto simp: strict_mono_def)
-
 text \<open>For any sequence, there is a monotonic subsequence.\<close>
 lemma seq_monosub:
   fixes s :: "nat \<Rightarrow> 'a::linorder"
--- a/src/Provers/order_tac.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Provers/order_tac.ML	Fri May 26 09:56:20 2023 +0200
@@ -368,15 +368,16 @@
   )
 
   fun declare (octxt as {kind = kind, raw_proc = raw_proc, ...}) lthy =
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context =>
-      let
-        val ops = map (Morphism.term phi) (#ops octxt)
-        val thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#thms octxt)
-        val conv_thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#conv_thms octxt)
-        val octxt' = {kind = kind, ops = ops, thms = thms, conv_thms = conv_thms}
-      in
-        context |> Data.map (Library.insert order_data_eq (octxt', raw_proc))
-      end)
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
+      (fn phi => fn context =>
+        let
+          val ops = map (Morphism.term phi) (#ops octxt)
+          val thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#thms octxt)
+          val conv_thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#conv_thms octxt)
+          val octxt' = {kind = kind, ops = ops, thms = thms, conv_thms = conv_thms}
+        in
+          context |> Data.map (Library.insert order_data_eq (octxt', raw_proc))
+        end)
 
   fun declare_order {
       ops = {eq = eq, le = le, lt = lt},
--- a/src/Pure/Admin/build_release.scala	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Admin/build_release.scala	Fri May 26 09:56:20 2023 +0200
@@ -242,7 +242,7 @@
           ssh.write_file(remote_tmp_tar, local_tmp_tar)
 
           val build_command =
-            "bin/isabelle build -o system_heaps -b -- " + Bash.strings(build_sessions)
+            "bin/isabelle build -o parallel_proofs=0 -o system_heaps -b -- " + Bash.strings(build_sessions)
           def system_apple(b: Boolean): String =
             """{ echo "ML_system_apple = """ + b + """" > "$(bin/isabelle getenv -b ISABELLE_HOME_USER)/etc/preferences"; }"""
 
--- a/src/Pure/Admin/isabelle_cronjob.scala	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Admin/isabelle_cronjob.scala	Fri May 26 09:56:20 2023 +0200
@@ -57,8 +57,9 @@
             """ -a --include="*/" --include="plain_identify*" --exclude="*" """ +
             Bash.string(backup + "/log/.") + " " + File.bash_path(main_dir) + "/log/.").check
 
-        if (!Isabelle_Devel.cronjob_log.is_file)
+        if (!Isabelle_Devel.cronjob_log.is_file) {
           Files.createSymbolicLink(Isabelle_Devel.cronjob_log.java_path, current_log.java_path)
+        }
       })
 
   val exit: Logger_Task =
@@ -292,7 +293,7 @@
       Remote_Build("macOS 10.8 Mountain Lion", "macbroy30", options = "-m32 -M2", args = "-a",
         detect = Build_Log.Prop.build_start.toString + " < date '2017-03-03'")) :::
       {
-        for { (n, hosts) <- List(1 -> List("lxbroy6"), 2 -> List("lxbroy8", "lxbroy5")) }
+        for ((n, hosts) <- List(1 -> List("lxbroy6"), 2 -> List("lxbroy8", "lxbroy5")))
         yield {
           Remote_Build("AFP old", host = hosts.head, more_hosts = hosts.tail,
             options = "-m32 -M1x2 -t AFP -P" + n +
--- a/src/Pure/General/basics.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/General/basics.ML	Fri May 26 09:56:20 2023 +0200
@@ -32,6 +32,7 @@
   val the_default: 'a -> 'a option -> 'a
   val perhaps: ('a -> 'a option) -> 'a -> 'a
   val merge_options: 'a option * 'a option -> 'a option
+  val join_options: ('a * 'a -> 'a) -> 'a option * 'a option -> 'a option
   val eq_option: ('a * 'b -> bool) -> 'a option * 'b option -> bool
 
   (*partiality*)
@@ -94,6 +95,9 @@
 
 fun merge_options (x, y) = if is_some x then x else y;
 
+fun join_options f (SOME x, SOME y) = SOME (f (x, y))
+  | join_options _ args = merge_options args;
+
 fun eq_option eq (SOME x, SOME y) = eq (x, y)
   | eq_option _ (NONE, NONE) = true
   | eq_option _ _ = false;
--- a/src/Pure/Isar/args.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/args.ML	Fri May 26 09:56:20 2023 +0200
@@ -36,15 +36,16 @@
   val internal_typ: typ parser
   val internal_term: term parser
   val internal_fact: thm list parser
-  val internal_attribute: (morphism -> attribute) parser
-  val internal_declaration: declaration parser
+  val internal_attribute: attribute Morphism.entity parser
+  val internal_declaration: Morphism.declaration_entity parser
   val named_source: (Token.T -> Token.src) -> Token.src parser
   val named_typ: (string -> typ) -> typ parser
   val named_term: (string -> term) -> term parser
   val named_fact: (string -> string option * thm list) -> thm list parser
-  val named_attribute: (string * Position.T -> morphism -> attribute) ->
-    (morphism -> attribute) parser
-  val embedded_declaration: (Input.source -> declaration) -> declaration parser
+  val named_attribute: (string * Position.T -> attribute Morphism.entity) ->
+    attribute Morphism.entity parser
+  val embedded_declaration: (Input.source -> Morphism.declaration_entity) ->
+    Morphism.declaration_entity parser
   val typ_abbrev: typ context_parser
   val typ: typ context_parser
   val term: term context_parser
--- a/src/Pure/Isar/attrib.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/attrib.ML	Fri May 26 09:56:20 2023 +0200
@@ -48,8 +48,8 @@
     local_theory -> local_theory
   val attribute_setup: bstring * Position.T -> Input.source -> string ->
     local_theory -> local_theory
-  val internal: (morphism -> attribute) -> Token.src
-  val internal_declaration: declaration -> thms
+  val internal: Position.T -> (morphism -> attribute) -> Token.src
+  val internal_declaration: Position.T -> Morphism.declaration_entity -> thms
   val add_del: attribute -> attribute -> attribute context_parser
   val thm: thm context_parser
   val thms: thm list context_parser
@@ -169,7 +169,7 @@
         val name = #1 (Token.name_of_src src);
         val label = Long_Name.qualify Markup.attributeN name;
         val att = #1 (Name_Space.get table name) src;
-      in Position.setmp_thread_data_label label att end
+      in Position.setmp_thread_data_label label att : attribute end
   end;
 
 val attribute = attribute_generic o Context.Proof;
@@ -187,12 +187,19 @@
 fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g)));
 
 
-(* fact expressions *)
+(* implicit context *)
+
+val trim_context_binding: Attrib.binding -> Attrib.binding =
+  apsnd ((map o map) Token.trim_context);
 
-val trim_context_binding: Attrib.binding -> Attrib.binding = apsnd (map Token.trim_context_src);
-val trim_context_thms: thms -> thms = (map o apfst o map) Thm.trim_context;
+val trim_context_thms: thms -> thms =
+  map (fn (thms, atts) => (map Thm.trim_context thms, (map o map) Token.trim_context atts));
+
 fun trim_context_fact (binding, thms) = (trim_context_binding binding, trim_context_thms thms);
 
+
+(* fact expressions *)
+
 fun global_notes kind facts thy = thy |>
   Global_Theory.note_thmss kind (map_facts (map (attribute_global thy)) facts);
 
@@ -230,23 +237,31 @@
 
 (* internal attribute *)
 
-val _ = Theory.setup
-  (setup (Binding.make ("attribute", \<^here>))
-    (Scan.lift Args.internal_attribute >> Morphism.form)
-    "internal attribute");
-
-fun internal_name ctxt name =
+fun make_name ctxt name =
   Token.make_src (name, Position.none) [] |> check_src ctxt |> hd;
 
-val internal_attribute_name =
-  internal_name (Context.the_local_context ()) "attribute";
+local
+
+val internal_binding = Binding.make ("attribute", \<^here>);
+
+val _ = Theory.setup
+  (setup internal_binding
+    (Scan.lift Args.internal_attribute >> Morphism.form ||
+      Scan.lift Args.internal_declaration >> (Thm.declaration_attribute o K o Morphism.form))
+    "internal attribute");
 
-fun internal att =
-  internal_attribute_name ::
-    [Token.make_string ("<attribute>", Position.none) |> Token.assign (SOME (Token.Attribute att))];
+val internal_name = make_name (Context.the_local_context ()) (Binding.name_of internal_binding);
+fun internal_source name value = [internal_name, Token.assign (SOME value) (Token.make_string name)];
+
+in
 
-fun internal_declaration decl =
-  [([Drule.dummy_thm], [internal (fn phi => Thm.declaration_attribute (K (decl phi)))])];
+fun internal pos arg =
+  internal_source ("internal", pos) (Token.Attribute (Morphism.entity arg));
+
+fun internal_declaration pos arg =
+  [([Drule.dummy_thm], [internal_source ("declaration", pos) (Token.Declaration arg)])];
+
+end;
 
 
 (* add/del syntax *)
@@ -426,7 +441,8 @@
 fun register binding config thy =
   let val name = Sign.full_name thy binding in
     thy
-    |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form) "configuration option"
+    |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form_entity)
+        "configuration option"
     |> Configs.map (Symtab.update (name, config))
   end;
 
@@ -623,23 +639,21 @@
 
 local
 
-val internal = internal_name (Context.the_local_context ());
+val make_name0 = make_name (Context.the_local_context ());
 
-val consumes_name = internal "consumes";
-val constraints_name = internal "constraints";
-val cases_open_name = internal "cases_open";
-val case_names_name = internal "case_names";
-val case_conclusion_name = internal "case_conclusion";
-
-fun make_string s = Token.make_string (s, Position.none);
+val consumes_name = make_name0 "consumes";
+val constraints_name = make_name0 "constraints";
+val cases_open_name = make_name0 "cases_open";
+val case_names_name = make_name0 "case_names";
+val case_conclusion_name = make_name0 "case_conclusion";
 
 in
 
 fun consumes i = consumes_name :: Token.make_int i;
 fun constraints i = constraints_name :: Token.make_int i;
 val cases_open = [cases_open_name];
-fun case_names names = case_names_name :: map make_string names;
-fun case_conclusion (name, names) = case_conclusion_name :: map make_string (name :: names);
+fun case_names names = case_names_name :: map Token.make_string0 names;
+fun case_conclusion (name, names) = case_conclusion_name :: map Token.make_string0 (name :: names);
 
 end;
 
--- a/src/Pure/Isar/bundle.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/bundle.ML	Fri May 26 09:56:20 2023 +0200
@@ -59,13 +59,6 @@
 
 fun extern ctxt = Name_Space.extern ctxt (Name_Space.space_of_table (get_all ctxt));
 
-fun define_bundle (b, bundle) context =
-  let
-    val bundle' = Attrib.trim_context_thms bundle;
-    val (name, bundles') = Name_Space.define context true (b, bundle') (get_all_generic context);
-    val context' = (Data.map o apfst o K) bundles' context;
-  in (name, context') end;
-
 
 (* target -- bundle under construction *)
 
@@ -102,9 +95,24 @@
 
 (** define bundle **)
 
+(* context and morphisms *)
+
+val trim_context_bundle =
+  map (fn (fact, atts) => (map Thm.trim_context fact, (map o map) Token.trim_context atts));
+
+fun transfer_bundle thy =
+  map (fn (fact, atts) => (map (Thm.transfer thy) fact, (map o map) (Token.transfer thy) atts));
+
 fun transform_bundle phi =
   map (fn (fact, atts) => (Morphism.fact phi fact, (map o map) (Token.transform phi) atts));
 
+fun define_bundle (b, bundle) context =
+  let
+    val (name, bundles') = get_all_generic context
+      |> Name_Space.define context true (b, trim_context_bundle bundle);
+    val context' = (Data.map o apfst o K) bundles' context;
+  in (name, context') end;
+
 
 (* command *)
 
@@ -116,11 +124,15 @@
     val bundle0 = raw_bundle
       |> map (fn (fact, atts) => (prep_fact ctxt' fact, map (prep_att ctxt') atts));
     val bundle =
-      Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)] |> map snd |> flat
-      |> transform_bundle (Proof_Context.export_morphism ctxt' lthy);
+      Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)]
+      |> maps #2
+      |> transform_bundle (Proof_Context.export_morphism ctxt' lthy)
+      |> trim_context_bundle;
   in
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = true}
-      (fn phi => #2 o define_bundle (Morphism.binding phi binding, transform_bundle phi bundle))
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding}
+      (fn phi => fn context =>
+        let val psi = Morphism.set_trim_context'' context phi
+        in #2 (define_bundle (Morphism.binding psi binding, transform_bundle psi bundle) context) end)
   end;
 
 in
@@ -169,9 +181,9 @@
     |> Attrib.local_notes kind facts
   end;
 
-fun bundle_declaration decl lthy =
+fun bundle_declaration pos decl lthy =
   lthy
-  |> (augment_target o Attrib.internal_declaration)
+  |> (augment_target o Attrib.internal_declaration pos)
     (Morphism.transform (Local_Theory.standard_morphism_theory lthy) decl)
   |> Generic_Target.standard_declaration (K true) decl;
 
@@ -186,10 +198,10 @@
      {define = bad_operation,
       notes = bundle_notes,
       abbrev = bad_operation,
-      declaration = K bundle_declaration,
+      declaration = fn {pos, ...} => bundle_declaration pos,
       theory_registration = bad_operation,
       locale_dependency = bad_operation,
-      pretty = pretty binding}
+      pretty = pretty binding};
 
 end;
 
@@ -200,7 +212,10 @@
 local
 
 fun gen_activate notes prep_bundle args ctxt =
-  let val decls = maps (prep_bundle ctxt) args in
+  let
+    val thy = Proof_Context.theory_of ctxt;
+    val decls = maps (prep_bundle ctxt) args |> transfer_bundle thy;
+  in
     ctxt
     |> Context_Position.set_visible false
     |> notes [(Binding.empty_atts, decls)] |> #2
--- a/src/Pure/Isar/class.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/class.ML	Fri May 26 09:56:20 2023 +0200
@@ -157,9 +157,9 @@
 val base_morphism = #base_morph oo the_class_data;
 
 fun morphism thy class =
-  (case Element.eq_morphism thy (these_defs thy [class]) of
-    SOME eq_morph => base_morphism thy class $> eq_morph
-  | NONE => base_morphism thy class);
+  Morphism.set_context thy
+   (base_morphism thy class $>
+    Morphism.default (Element.eq_morphism (these_defs thy [class])));
 
 val export_morphism = #export_morph oo the_class_data;
 
@@ -221,21 +221,23 @@
       (c, (class, ((ty, Free v_ty), false)))) params;
     val add_class = Graph.new_node (class,
         make_class_data (((map o apply2) fst params, base_sort,
-          base_morph, export_morph, Option.map Thm.trim_context some_assm_intro,
-          Thm.trim_context of_class, Option.map Thm.trim_context some_axiom),
+          Morphism.reset_context base_morph,
+          Morphism.reset_context export_morph,
+          Option.map Thm.trim_context some_assm_intro,
+          Thm.trim_context of_class,
+          Option.map Thm.trim_context some_axiom),
           (Thm.item_net, operations)))
       #> fold (curry Graph.add_edge class) sups;
   in Class_Data.map add_class thy end;
 
 fun activate_defs class thms thy =
-  (case Element.eq_morphism thy thms of
+  (case Element.eq_morphism thms of
     SOME eq_morph =>
-      fold (fn cls => fn thy =>
-        Context.theory_map
-          (Locale.amend_registration
-            {inst = (cls, base_morphism thy cls),
-              mixin = SOME (eq_morph, true),
-              export = export_morphism thy cls}) thy) (heritage thy [class]) thy
+      fold (fn cls => fn thy' =>
+        (Context.theory_map o Locale.amend_registration)
+         {inst = (cls, base_morphism thy' cls),
+          mixin = SOME (eq_morph, true),
+          export = export_morphism thy' cls} thy') (heritage thy [class]) thy
   | NONE => thy);
 
 fun register_operation class (c, t) input_only thy =
@@ -358,7 +360,8 @@
 
 fun target_const class phi0 prmode (b, rhs) lthy =
   let
-    val export = Variable.export_morphism lthy (Local_Theory.target_of lthy);
+    val export =
+      Morphism.set_context' lthy (Variable.export_morphism lthy (Local_Theory.target_of lthy));
     val guess_identity = guess_morphism_identity (b, rhs) export;
     val guess_canonical = guess_morphism_identity (b, rhs) (export $> phi0);
   in
@@ -478,9 +481,11 @@
 fun register_subclass (sub, sup) some_dep_morph some_witn export lthy =
   let
     val thy = Proof_Context.theory_of lthy;
-    val intros = (snd o rules thy) sup :: map_filter I
-      [Option.map (Drule.export_without_context_open o Element.conclude_witness lthy) some_witn,
-        (fst o rules thy) sub];
+    val conclude_witness =
+      Thm.trim_context o Drule.export_without_context_open o Element.conclude_witness lthy;
+    val intros =
+      (snd o rules thy) sup ::
+        map_filter I [Option.map conclude_witness some_witn, (fst o rules thy) sub];
     val classrel =
       Goal.prove_sorry_global thy [] [] (Logic.mk_classrel (sub, sup))
         (fn {context = ctxt, ...} => EVERY (map (TRYALL o resolve_tac ctxt o single) intros));
--- a/src/Pure/Isar/class_declaration.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/class_declaration.ML	Fri May 26 09:56:20 2023 +0200
@@ -31,6 +31,7 @@
 
     val certT = Thm.trim_context_ctyp o Thm.global_ctyp_of thy;
     val cert = Thm.trim_context_cterm o Thm.global_cterm_of thy;
+    val conclude_witness = Thm.trim_context o Element.conclude_witness thy_ctxt;
 
     (* instantiation of canonical interpretation *)
     val a_tfree = (Name.aT, base_sort);
@@ -67,13 +68,15 @@
         val tac = loc_intro_tac
           THEN ALLGOALS (Proof_Context.fact_tac thy_ctxt (sup_axioms @ the_list assm_axiom));
       in Element.prove_witness thy_ctxt prop tac end) some_prop;
-    val some_axiom = Option.map (Element.conclude_witness thy_ctxt) some_witn;
+    val some_axiom = Option.map conclude_witness some_witn;
 
     (* canonical interpretation *)
     val base_morph = inst_morph
       $> Morphism.binding_morphism "class_binding" (Binding.prefix false (Class.class_prefix class))
       $> Element.satisfy_morphism (the_list some_witn);
-    val eq_morph = Element.eq_morphism thy (Class.these_defs thy sups);
+    val eq_morph =
+      Element.eq_morphism (Class.these_defs thy sups)
+      |> Option.map (Morphism.set_context thy);
 
     (* assm_intro *)
     fun prove_assm_intro thm =
--- a/src/Pure/Isar/code.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/code.ML	Fri May 26 09:56:20 2023 +0200
@@ -393,21 +393,23 @@
 
 type data = Any.T Datatab.table;
 
-fun make_dataref thy =
-  (Context.theory_long_name thy,
-    Synchronized.var "code data" (NONE : (data * Context.theory_id) option));
+fun make_dataref () =
+  Synchronized.var "code data" (NONE : (data * Context.theory_id) option);
 
 structure Code_Data = Theory_Data
 (
   type T = specs * (string * (data * Context.theory_id) option Synchronized.var);
-  val empty = (empty_specs, make_dataref (Context.the_global_context ()));
+  val empty =
+    (empty_specs, (Context.theory_long_name (Context.the_global_context ()), make_dataref ()));
   fun merge ((specs1, dataref), (specs2, _)) =
     (merge_specs (specs1, specs2), dataref);
 );
 
 fun init_dataref thy =
-  if #1 (#2 (Code_Data.get thy)) = Context.theory_long_name thy then NONE
-  else SOME ((Code_Data.map o apsnd) (fn _ => make_dataref thy) thy)
+  let val thy_name = Context.theory_long_name thy in
+    if #1 (#2 (Code_Data.get thy)) = thy_name then NONE
+    else SOME ((Code_Data.map o apsnd) (K (thy_name, make_dataref ())) thy)
+  end;
 
 in
 
@@ -419,7 +421,8 @@
 val specs_of : theory -> specs = fst o Code_Data.get;
 
 fun modify_specs f thy =
-  Code_Data.map (fn (specs, _) => (f specs, make_dataref thy)) thy;
+  let val thy_name = Context.theory_long_name thy
+  in Code_Data.map (fn (specs, _) => (f specs, (thy_name, make_dataref ()))) thy end;
 
 
 (* access to data dependent on executable specifications *)
@@ -569,7 +572,7 @@
 
 fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of
     SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) =>
-      (vs, {abs_rep = abs_rep, abstractor = abstractor, projection = projection})
+      (vs, {abs_rep = Thm.transfer thy abs_rep, abstractor = abstractor, projection = projection})
   | _ => error ("Not an abstract type: " ^ tyco);
 
 fun get_type_of_constr_or_abstr thy c =
@@ -1014,13 +1017,13 @@
       Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco);
 
 fun conclude_cert (Nothing cert_thm) =
-      Nothing (Thm.close_derivation \<^here> cert_thm)
+      Nothing (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context)
   | conclude_cert (Equations (cert_thm, propers)) =
-      Equations (Thm.close_derivation \<^here> cert_thm, propers)
+      Equations (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context, propers)
   | conclude_cert (cert as Projection _) =
       cert
   | conclude_cert (Abstract (abs_thm, tyco)) =
-      Abstract (Thm.close_derivation \<^here> abs_thm, tyco);
+      Abstract (Thm.close_derivation \<^here> abs_thm |> Thm.trim_context, tyco);
 
 fun typscheme_of_cert thy (Nothing cert_thm) =
       fst (get_head thy cert_thm)
@@ -1057,6 +1060,7 @@
         val tyscm = typscheme_of_cert thy cert;
         val thms = if null propers then [] else
           cert_thm
+          |> Thm.transfer thy
           |> Local_Defs.expand [snd (get_head thy cert_thm)]
           |> Thm.varifyT_global
           |> Conjunction.elim_balanced (length propers);
@@ -1072,7 +1076,7 @@
   | equations_of_cert thy (Abstract (abs_thm, tyco)) =
       let
         val tyscm = typscheme_abs thy abs_thm;
-        val (abs, concrete_thm) = concretify_abs thy tyco abs_thm;
+        val (abs, concrete_thm) = concretify_abs thy tyco (Thm.transfer thy abs_thm);
         fun abstractions (args, rhs) = (map (pair NONE) args, (SOME abs, rhs));
       in
         (tyscm, SOME [((abstractions o dest_eqn o Thm.prop_of) concrete_thm,
@@ -1305,20 +1309,10 @@
 
 (* abstract code declarations *)
 
-local
-
-fun generic_code_declaration strictness lift_phi f x =
-  Local_Theory.declaration
-    {syntax = false, pervasive = false}
+fun code_declaration strictness lift_phi f x =
+  Local_Theory.declaration {syntax = false, pervasive = false, pos = Position.thread_data ()}
     (fn phi => Context.mapping (f strictness (lift_phi phi x)) I);
 
-in
-
-fun silent_code_declaration lift_phi = generic_code_declaration Silent lift_phi;
-fun code_declaration lift_phi = generic_code_declaration Liberal lift_phi;
-
-end;
-
 
 (* types *)
 
@@ -1363,16 +1357,15 @@
     SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) =>
       thy
       |> modify_specs (register_type
-            (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj, abs_rep = abs_rep, more_abstract_functions = []}))
+            (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj,
+                abs_rep = Thm.trim_context abs_rep, more_abstract_functions = []}))
           #> register_fun_spec proj
             (Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs))))
       |> register_tyco_for_plugin tyco
   | NONE => thy;
 
 val declare_abstype_global = generic_declare_abstype Strict;
-
-val declare_abstype =
-  code_declaration Morphism.thm generic_declare_abstype;
+val declare_abstype = code_declaration Liberal Morphism.thm generic_declare_abstype;
 
 
 (* functions *)
@@ -1422,7 +1415,7 @@
     else specs);
 
 fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) =
-  modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm, tyco_abs))
+  modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm |> Thm.trim_context, tyco_abs))
     #> map_types (History.modify_entry tyco (add_abstract_function c)))
 
 in
@@ -1445,14 +1438,10 @@
 end;
 
 val declare_default_eqns_global = generic_declare_eqns true Silent;
-
-val declare_default_eqns =
-  silent_code_declaration (map o apfst o Morphism.thm) (generic_declare_eqns true);
+val declare_default_eqns = code_declaration Silent (map o apfst o Morphism.thm) (generic_declare_eqns true);
 
 val declare_eqns_global = generic_declare_eqns false Strict;
-
-val declare_eqns =
-  code_declaration (map o apfst o Morphism.thm) (generic_declare_eqns false);
+val declare_eqns = code_declaration Liberal (map o apfst o Morphism.thm) (generic_declare_eqns false);
 
 val add_eqn_global = generic_add_eqn Strict;
 
@@ -1463,9 +1452,7 @@
   | NONE => thy;
 
 val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict;
-
-val declare_abstract_eqn =
-  code_declaration Morphism.thm generic_declare_abstract_eqn;
+val declare_abstract_eqn = code_declaration Liberal Morphism.thm generic_declare_abstract_eqn;
 
 fun declare_aborting_global c =
   modify_specs (register_fun_spec c aborting);
--- a/src/Pure/Isar/element.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/element.ML	Fri May 26 09:56:20 2023 +0200
@@ -29,6 +29,8 @@
     ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
   val map_ctxt_attrib: (Token.src -> Token.src) ->
     ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt
+  val trim_context_ctxt: context_i -> context_i
+  val transfer_ctxt: theory -> context_i -> context_i
   val transform_ctxt: morphism -> context_i -> context_i
   val pretty_stmt: Proof.context -> statement_i -> Pretty.T list
   val pretty_ctxt: Proof.context -> context_i -> Pretty.T list
@@ -49,8 +51,8 @@
   val pretty_witness: Proof.context -> witness -> Pretty.T
   val instantiate_normalize_morphism: ctyp TFrees.table * cterm Frees.table -> morphism
   val satisfy_morphism: witness list -> morphism
-  val eq_term_morphism: theory -> term list -> morphism option
-  val eq_morphism: theory -> thm list -> morphism option
+  val eq_term_morphism: term list -> morphism option
+  val eq_morphism: thm list -> morphism option
   val init: context_i -> Context.generic -> Context.generic
   val activate_i: context_i -> Proof.context -> context_i * Proof.context
   val activate: (typ, term, Facts.ref) ctxt -> Proof.context -> context_i * Proof.context
@@ -103,6 +105,16 @@
 fun map_ctxt_attrib attrib =
   map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = I, attrib = attrib};
 
+val trim_context_ctxt: context_i -> context_i = map_ctxt
+ {binding = I, typ = I, term = I, pattern = I,
+  fact = map Thm.trim_context,
+  attrib = map Token.trim_context};
+
+fun transfer_ctxt thy: context_i -> context_i = map_ctxt
+ {binding = I, typ = I, term = I, pattern = I,
+  fact = map (Thm.transfer thy),
+  attrib = map (Token.transfer thy)};
+
 fun transform_ctxt phi = map_ctxt
  {binding = Morphism.binding phi,
   typ = Morphism.typ phi,
@@ -275,7 +287,8 @@
   Witness (t,
     Goal.prove ctxt [] [] (mark_witness t)
       (fn _ => resolve_tac ctxt [Drule.protectI] 1 THEN tac)
-    |> Thm.close_derivation \<^here>);
+    |> Thm.close_derivation \<^here>
+    |> Thm.trim_context);
 
 
 local
@@ -290,7 +303,9 @@
       (map o map) (fn prop => (mark_witness prop, [])) wit_propss @
         [map (rpair []) eq_props];
     fun after_qed' thmss =
-      let val (wits, eqs) = split_last ((map o map) (Thm.close_derivation \<^here>) thmss);
+      let
+        val (wits, eqs) =
+          split_last ((map o map) (Thm.close_derivation \<^here> #> Thm.trim_context) thmss);
       in after_qed ((map2 o map2) (curry Witness) wit_propss wits) eqs end;
   in proof after_qed' propss #> refine_witness end;
 
@@ -322,7 +337,7 @@
 end;
 
 fun conclude_witness ctxt (Witness (_, th)) =
-  Goal.conclude th
+  Goal.conclude (Thm.transfer' ctxt th)
   |> Raw_Simplifier.norm_hhf_protect ctxt
   |> Thm.close_derivation \<^here>;
 
@@ -359,7 +374,7 @@
 
 fun compose_witness (Witness (_, th)) r =
   let
-    val th' = Goal.conclude th;
+    val th' = Goal.conclude (Thm.transfer (Thm.theory_of_thm r) th);
     val A = Thm.cprem_of r 1;
   in
     Thm.implies_elim
@@ -384,36 +399,42 @@
 (* rewriting with equalities *)
 
 (* for activating declarations only *)
-fun eq_term_morphism _ [] = NONE
-  | eq_term_morphism thy props =
+fun eq_term_morphism [] = NONE
+  | eq_term_morphism props =
       let
-        fun decomp_simp prop =
+        fun decomp_simp ctxt prop =
           let
-            val ctxt = Proof_Context.init_global thy;
             val _ = Logic.no_prems prop orelse
               error ("Bad conditional rewrite rule " ^ Syntax.string_of_term ctxt prop);
-            val lhsrhs = Logic.dest_equals prop
-              handle TERM _ => error ("Rewrite rule not a meta-equality " ^ Syntax.string_of_term ctxt prop);
-          in lhsrhs end;
+          in
+            Logic.dest_equals prop handle TERM _ =>
+              error ("Rewrite rule not a meta-equality " ^ Syntax.string_of_term ctxt prop)
+          end;
+        fun rewrite_term thy =
+          let val ctxt = Proof_Context.init_global thy
+          in Pattern.rewrite_term thy (map (decomp_simp ctxt) props) [] end;
         val phi =
           Morphism.morphism "Element.eq_term_morphism"
            {binding = [],
             typ = [],
-            term = [Pattern.rewrite_term thy (map decomp_simp props) []],
-            fact = [fn _ => error "Illegal application of Element.eq_term_morphism"]};
+            term = [rewrite_term o Morphism.the_theory],
+            fact = [fn _ => fn _ => error "Illegal application of Element.eq_term_morphism"]};
       in SOME phi end;
 
-fun eq_morphism _ [] = NONE
-  | eq_morphism thy thms =
+fun eq_morphism [] = NONE
+  | eq_morphism thms =
       let
-        (* FIXME proper context!? *)
-        fun rewrite th = rewrite_rule (Proof_Context.init_global (Thm.theory_of_thm th)) thms th;
+        val thms0 = map Thm.trim_context thms;
+        fun rewrite_term thy =
+          Raw_Simplifier.rewrite_term thy (map (Thm.transfer thy) thms0) [];
+        fun rewrite thy =
+          Raw_Simplifier.rewrite_rule (Proof_Context.init_global thy) (map (Thm.transfer thy) thms0);
         val phi =
           Morphism.morphism "Element.eq_morphism"
            {binding = [],
             typ = [],
-            term = [Raw_Simplifier.rewrite_term thy thms []],
-            fact = [map rewrite]};
+            term = [rewrite_term o Morphism.the_theory],
+            fact = [map o rewrite o Morphism.the_theory]};
       in SOME phi end;
 
 
--- a/src/Pure/Isar/entity.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/entity.ML	Fri May 26 09:56:20 2023 +0200
@@ -35,12 +35,13 @@
 (* local definition *)
 
 fun alias {get_data, put_data} binding name =
-  Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context =>
-    let
-      val naming = Name_Space.naming_of context;
-      val binding' = Morphism.binding phi binding;
-      val data' = Name_Space.alias_table naming binding' name (get_data context);
-    in put_data data' context end);
+  Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding}
+    (fn phi => fn context =>
+      let
+        val naming = Name_Space.naming_of context;
+        val binding' = Morphism.binding phi binding;
+        val data' = Name_Space.alias_table naming binding' name (get_data context);
+      in put_data data' context end);
 
 fun transfer {get_data, put_data} ctxt =
   let
--- a/src/Pure/Isar/expression.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/expression.ML	Fri May 26 09:56:20 2023 +0200
@@ -406,8 +406,8 @@
         val rewrite_morph = eqns'
           |> map (abs_def ctxt')
           |> Variable.export_terms ctxt' ctxt
-          |> Element.eq_term_morphism (Proof_Context.theory_of ctxt)
-          |> the_default Morphism.identity;
+          |> Element.eq_term_morphism
+          |> Morphism.default;
        val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt;
        val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns'];
       in (i + 1, insts', eqnss', ctxt'') end;
@@ -556,7 +556,7 @@
     val exp_typ = Logic.type_map exp_term;
     val export' =
       Morphism.morphism "Expression.prep_goal"
-        {binding = [], typ = [exp_typ], term = [exp_term], fact = [exp_fact]};
+        {binding = [], typ = [K exp_typ], term = [K exp_term], fact = [K exp_fact]};
   in ((propss, eq_propss, deps, eqnss, export'), goal_ctxt) end;
 
 in
@@ -743,41 +743,41 @@
     val ctxt = Proof_Context.init_global thy;
     val defs' = map (Thm.cterm_of ctxt #> Assumption.assume ctxt #> Drule.abs_def) defs;
 
-    val (a_pred, a_intro, a_axioms, thy'') =
+    val (a_pred, a_intro, a_axioms, thy2) =
       if null exts then (NONE, NONE, [], thy)
       else
         let
           val abinding =
             if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding;
-          val ((statement, intro, axioms), thy') =
+          val ((statement, intro, axioms), thy1) =
             thy
             |> def_pred abinding parms defs' exts exts';
-          val ((_, [intro']), thy'') =
-            thy'
+          val ((_, [intro']), thy2) =
+            thy1
             |> Sign.qualified_path true abinding
             |> Global_Theory.note_thms ""
               ((Binding.name introN, []), [([intro], [Locale.unfold_add])])
-            ||> Sign.restore_naming thy';
-          in (SOME statement, SOME intro', axioms, thy'') end;
-    val (b_pred, b_intro, b_axioms, thy'''') =
-      if null ints then (NONE, NONE, [], thy'')
+            ||> Sign.restore_naming thy1;
+          in (SOME statement, SOME intro', axioms, thy2) end;
+    val (b_pred, b_intro, b_axioms, thy4) =
+      if null ints then (NONE, NONE, [], thy2)
       else
         let
-          val ((statement, intro, axioms), thy''') =
-            thy''
+          val ((statement, intro, axioms), thy3) =
+            thy2
             |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred);
-          val ctxt''' = Proof_Context.init_global thy''';
-          val ([(_, [intro']), _], thy'''') =
-            thy'''
+          val conclude_witness =
+            Drule.export_without_context o Element.conclude_witness (Proof_Context.init_global thy3);
+          val ([(_, [intro']), _], thy4) =
+            thy3
             |> Sign.qualified_path true binding
             |> Global_Theory.note_thmss ""
                  [((Binding.name introN, []), [([intro], [Locale.intro_add])]),
                   ((Binding.name axiomsN, []),
-                    [(map (Drule.export_without_context o Element.conclude_witness ctxt''') axioms,
-                      [])])]
-            ||> Sign.restore_naming thy''';
-        in (SOME statement, SOME intro', axioms, thy'''') end;
-  in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
+                    [(map conclude_witness axioms, [])])]
+            ||> Sign.restore_naming thy3;
+        in (SOME statement, SOME intro', axioms, thy4) end;
+  in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy4) end;
 
 end;
 
@@ -794,7 +794,7 @@
 fun defines_to_notes ctxt (Defines defs) =
       Notes ("", map (fn (a, (def, _)) =>
         (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)],
-          [(Attrib.internal o K) Locale.witness_add])])) defs)
+          [Attrib.internal @{here} (K Locale.witness_add)])])) defs)
   | defines_to_notes _ e = e;
 
 val is_hyp = fn Assumes _ => true | Defines _ => true | _ => false;
@@ -846,11 +846,12 @@
       if is_some asm then
         [("", [((Binding.suffix_name ("_" ^ axiomsN) binding, []),
           [([Assumption.assume pred_ctxt (Thm.cterm_of pred_ctxt (the asm))],
-            [(Attrib.internal o K) Locale.witness_add])])])]
+            [Attrib.internal @{here} (K Locale.witness_add)])])])]
       else [];
 
     val notes' =
       body_elems
+      |> map (Element.transfer_ctxt thy')
       |> map (defines_to_notes pred_ctxt)
       |> map (Element.transform_ctxt a_satisfy)
       |> (fn elems =>
--- a/src/Pure/Isar/generic_target.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/generic_target.ML	Fri May 26 09:56:20 2023 +0200
@@ -16,7 +16,7 @@
   (*background primitives*)
   val background_foundation: ((binding * typ) * mixfix) * (binding * term) ->
     term list * term list -> local_theory -> (term * thm) * local_theory
-  val background_declaration: declaration -> local_theory -> local_theory
+  val background_declaration: Morphism.declaration_entity -> local_theory -> local_theory
   val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory
   val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) ->
     theory -> theory
@@ -25,8 +25,8 @@
   val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list
   val standard_notes: (int * int -> bool) -> string -> Attrib.fact list ->
     local_theory -> local_theory
-  val standard_declaration: (int * int -> bool) ->
-    (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory
+  val standard_declaration: (int * int -> bool) -> Morphism.declaration_entity ->
+    local_theory -> local_theory
   val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term ->
     local_theory -> local_theory
   val local_interpretation: Locale.registration ->
@@ -55,7 +55,7 @@
   (*theory target operations*)
   val theory_abbrev: Syntax.mode -> (binding * mixfix) * term ->
     local_theory -> (term * term) * local_theory
-  val theory_declaration: declaration -> local_theory -> local_theory
+  val theory_declaration: Morphism.declaration_entity -> local_theory -> local_theory
   val theory_registration: Locale.registration -> local_theory -> local_theory
 
   (*locale target primitives*)
@@ -63,15 +63,16 @@
     local_theory -> local_theory
   val locale_target_abbrev: string -> Syntax.mode ->
     (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory
-  val locale_target_declaration: string -> bool -> declaration -> local_theory -> local_theory
+  val locale_target_declaration: string -> {syntax: bool, pos: Position.T} ->
+    Morphism.declaration_entity -> local_theory -> local_theory
   val locale_target_const: string -> (morphism -> bool) -> Syntax.mode ->
     (binding * mixfix) * term -> local_theory -> local_theory
 
   (*locale operations*)
   val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term ->
     local_theory -> (term * term) * local_theory
-  val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration ->
-    local_theory -> local_theory
+  val locale_declaration: string -> {syntax: bool, pervasive: bool, pos: Position.T} ->
+    Morphism.declaration_entity -> local_theory -> local_theory
   val locale_const: string -> Syntax.mode -> (binding * mixfix) * term ->
     local_theory -> local_theory
   val locale_dependency: string -> Locale.registration ->
@@ -123,7 +124,7 @@
   | same_const (t $ _, t' $ _) = same_const (t, t')
   | same_const (_, _) = false;
 
-fun const_decl phi_pred prmode ((b, mx), rhs) phi context =
+fun const_decl phi_pred prmode ((b, mx), rhs) = Morphism.entity (fn phi => fn context =>
   if phi_pred phi then
     let
       val b' = Morphism.binding phi b;
@@ -149,7 +150,7 @@
         SOME c =>
           context
           |> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c)
-          |> Morphism.form (Proof_Context.generic_notation true prmode [(rhs', mx)])
+          |> Morphism.form_entity (Proof_Context.generic_notation true prmode [(rhs', mx)])
       | NONE =>
           context
           |> Proof_Context.generic_add_abbrev Print_Mode.internal
@@ -157,9 +158,9 @@
           |-> (fn (const as Const (c, _), _) => same_stem ?
                 (Proof_Context.generic_revert_abbrev (#1 prmode) c #>
                  same_shape ?
-                  Morphism.form (Proof_Context.generic_notation true prmode [(const, mx)]))))
+                  Morphism.form_entity (Proof_Context.generic_notation true prmode [(const, mx)]))))
     end
-  else context;
+  else context);
 
 
 
@@ -167,18 +168,18 @@
 
 structure Foundation_Interpretations = Theory_Data
 (
-  type T = (binding * (term * term list) -> Context.generic -> Context.generic) Inttab.table;
-  val empty = Inttab.empty;
-  val merge = Inttab.merge (K true);
+  type T = ((binding * (term * term list) -> Context.generic -> Context.generic) * stamp) list
+  val empty = [];
+  val merge = Library.merge (eq_snd (op =));
 );
 
 fun add_foundation_interpretation f =
-  Foundation_Interpretations.map (Inttab.update_new (serial (), f));
+  Foundation_Interpretations.map (cons (f, stamp ()));
 
 fun foundation_interpretation binding_const_params lthy =
   let
     val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy);
-    val interp = Inttab.fold (fn (_, f) => f binding_const_params) interps;
+    val interp = fold (fn (f, _) => f binding_const_params) interps;
   in
     lthy
     |> Local_Theory.background_theory (Context.theory_map interp)
@@ -442,20 +443,21 @@
       Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #>
   standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts;
 
-fun locale_target_declaration locale syntax decl lthy = lthy
+fun locale_target_declaration locale params decl lthy = lthy
   |> Local_Theory.target (fn ctxt => ctxt |>
-    Locale.add_declaration locale syntax
+    Locale.add_declaration locale params
       (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl));
 
 fun locale_target_const locale phi_pred prmode ((b, mx), rhs) =
-  locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs))
+  locale_target_declaration locale {syntax = true, pos = Binding.pos_of b}
+    (const_decl phi_pred prmode ((b, mx), rhs));
 
 
 (** locale operations **)
 
-fun locale_declaration locale {syntax, pervasive} decl =
+fun locale_declaration locale {syntax, pervasive, pos} decl =
   pervasive ? background_declaration decl
-  #> locale_target_declaration locale syntax decl
+  #> locale_target_declaration locale {syntax = syntax, pos = pos} decl
   #> standard_declaration (fn (_, other) => other <> 0) decl;
 
 fun locale_const locale prmode ((b, mx), rhs) =
--- a/src/Pure/Isar/interpretation.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/interpretation.ML	Fri May 26 09:56:20 2023 +0200
@@ -98,20 +98,20 @@
   let
     val factss = thms
       |> unflat ((map o map) #1 eqnss)
-      |> map2 (map2 (fn b => fn eq => (b, [([Morphism.thm export eq], [])]))) ((map o map) #1 eqnss);
+      |> map2 (map2 (fn b => fn eq =>
+          (b, [([Morphism.thm export (Thm.transfer' ctxt eq)], [])]))) ((map o map) #1 eqnss);
     val (eqnss', ctxt') =
       fold_map (fn facts => note Thm.theoremK facts #-> abs_def_rule) factss ctxt;
     val defs = (Binding.empty_atts, [(map (Morphism.thm (export' $> export)) def_eqns, [])]);
     val (eqns', ctxt'') = ctxt' |> note Thm.theoremK [defs] |-> abs_def_rule;
+    val transform_witness = Element.transform_witness (Morphism.set_trim_context' ctxt' export');
     val deps' =
       (deps ~~ witss) |> map (fn ((dep, morph), wits) =>
-        (dep, morph $> Element.satisfy_morphism (map (Element.transform_witness export') wits)));
+        (dep, morph $> Element.satisfy_morphism (map transform_witness wits)));
     fun register (dep, eqns) ctxt =
       ctxt |> add_registration
         {inst = dep,
-          mixin =
-            Option.map (rpair true)
-              (Element.eq_morphism (Proof_Context.theory_of ctxt) (eqns @ eqns')),
+          mixin = Option.map (rpair true) (Element.eq_morphism (eqns @ eqns')),
           export = export};
   in ctxt'' |> fold register (deps' ~~ eqnss') end;
 
--- a/src/Pure/Isar/isar_cmd.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/isar_cmd.ML	Fri May 26 09:56:20 2023 +0200
@@ -130,7 +130,8 @@
   ML_Context.expression (Input.pos_of source)
     (ML_Lex.read
       ("Theory.local_setup (Local_Theory.declaration {syntax = " ^
-        Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^ "} (") @
+        Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^
+        ", pos = " ^ ML_Syntax.print_position (Position.thread_data ()) ^ "} (") @
      ML_Lex.read_source source @ ML_Lex.read "))")
   |> Context.proof_map;
 
--- a/src/Pure/Isar/local_defs.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/local_defs.ML	Fri May 26 09:56:20 2023 +0200
@@ -155,7 +155,7 @@
                   if t aconv u then (asm, false)
                   else (Drule.abs_def (Variable.gen_all outer asm), true))
             end)));
-  in (apply2 (map #1) (List.partition #2 defs_asms), Assumption.export false inner outer th) end;
+  in (apply2 (map #1) (List.partition #2 defs_asms), Assumption.export inner outer th) end;
 
 (*
   [xs, xs \<equiv> as]
--- a/src/Pure/Isar/local_theory.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/local_theory.ML	Fri May 26 09:56:20 2023 +0200
@@ -34,7 +34,7 @@
   val reset_group: local_theory -> local_theory
   val standard_morphism: local_theory -> Proof.context -> morphism
   val standard_morphism_theory: local_theory -> morphism
-  val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a
+  val standard_form: local_theory -> Proof.context -> 'a Morphism.entity -> 'a
   val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
   val raw_theory: (theory -> theory) -> local_theory -> local_theory
   val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
@@ -55,7 +55,8 @@
     (string * thm list) list * local_theory
   val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory
-  val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory
+  val declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration ->
+    local_theory -> local_theory
   val theory_registration: Locale.registration -> local_theory -> local_theory
   val locale_dependency: Locale.registration -> local_theory -> local_theory
   val pretty: local_theory -> Pretty.T list
@@ -102,7 +103,8 @@
   notes: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory,
   abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory,
-  declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory,
+  declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration_entity ->
+    local_theory -> local_theory,
   theory_registration: Locale.registration -> local_theory -> local_theory,
   locale_dependency: Locale.registration -> local_theory -> local_theory,
   pretty: local_theory -> Pretty.T list};
@@ -192,9 +194,11 @@
 (* standard morphisms *)
 
 fun standard_morphism lthy ctxt =
-  Proof_Context.norm_export_morphism lthy ctxt $>
-  Morphism.binding_morphism "Local_Theory.standard_binding"
-    (Name_Space.transform_binding (Proof_Context.naming_of lthy));
+  Morphism.set_context' lthy
+   (Proof_Context.export_morphism lthy ctxt $>
+    Morphism.thm_morphism' "Local_Theory.standard" (Goal.norm_result o Proof_Context.init_global) $>
+    Morphism.binding_morphism "Local_Theory.standard_binding"
+      (Name_Space.transform_binding (Proof_Context.naming_of lthy)));
 
 fun standard_morphism_theory lthy =
   standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy));
@@ -279,7 +283,7 @@
 val define = operation2 #define false;
 val define_internal = operation2 #define true;
 val notes_kind = operation2 #notes;
-val declaration = operation2 #declaration;
+fun declaration args = operation2 #declaration args o Morphism.entity;
 val theory_registration = operation1 #theory_registration;
 fun locale_dependency registration =
   assert_bottom #> operation1 #locale_dependency registration;
@@ -297,7 +301,7 @@
   |-> (fn name =>
     map_contexts (fn _ => fn ctxt =>
       Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #>
-    declaration {syntax = false, pervasive = false} (fn phi =>
+    declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} (fn phi =>
       let val binding' = Morphism.binding phi binding in
         Context.mapping
           (Global_Theory.alias_fact binding' name)
@@ -308,7 +312,7 @@
 (* default sort *)
 
 fun set_defsort S =
-  declaration {syntax = true, pervasive = false}
+  declaration {syntax = true, pervasive = false, pos = Position.thread_data ()}
     (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S)));
 
 
@@ -320,7 +324,7 @@
     val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args;
     val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args;
   in
-    declaration {syntax = true, pervasive = false}
+    declaration {syntax = true, pervasive = false, pos = Position.thread_data ()}
       (fn _ => Proof_Context.generic_syntax add mode args') lthy
   end;
 
@@ -339,7 +343,7 @@
     val args' = map (apsnd Mixfix.reset_pos) args;
     val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args;
   in
-    declaration {syntax = true, pervasive = false}
+    declaration {syntax = true, pervasive = false, pos = Position.thread_data ()}
       (Proof_Context.generic_type_notation add mode args') lthy
   end;
 
@@ -350,7 +354,7 @@
     val args' = map (apsnd Mixfix.reset_pos) args;
     val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args;
   in
-    declaration {syntax = true, pervasive = false}
+    declaration {syntax = true, pervasive = false, pos = Position.thread_data ()}
       (Proof_Context.generic_notation add mode args') lthy
   end;
 
@@ -369,7 +373,7 @@
 (* name space aliases *)
 
 fun syntax_alias global_alias local_alias b name =
-  declaration {syntax = true, pervasive = false} (fn phi =>
+  declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (fn phi =>
     let val b' = Morphism.binding phi b
     in Context.mapping (global_alias b' name) (local_alias b' name) end);
 
--- a/src/Pure/Isar/locale.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/locale.ML	Fri May 26 09:56:20 2023 +0200
@@ -39,7 +39,7 @@
     term option * term list ->
     thm option * thm option -> thm list ->
     Element.context_i list ->
-    declaration list ->
+    Morphism.declaration_entity list ->
     (string * Attrib.fact list) list ->
     (string * morphism) list -> theory -> theory
   val locale_space: theory -> Name_Space.T
@@ -59,7 +59,8 @@
 
   (* Storing results *)
   val add_facts: string -> string -> Attrib.fact list -> Proof.context -> Proof.context
-  val add_declaration: string -> bool -> declaration -> Proof.context -> Proof.context
+  val add_declaration: string -> {syntax: bool, pos: Position.T} ->
+    Morphism.declaration_entity -> Proof.context -> Proof.context
 
   (* Activation *)
   val activate_facts: morphism option -> string * morphism -> Context.generic -> Context.generic
@@ -107,8 +108,14 @@
 type dep = {name: string, morphisms: morphism * morphism, pos: Position.T, serial: serial};
 fun eq_dep (dep1: dep, dep2: dep) = #serial dep1 = #serial dep2;
 
+fun transfer_dep thy ({name, morphisms, pos, serial}: dep) : dep =
+  {name = name, morphisms = apply2 (Morphism.set_context thy) morphisms, pos = pos, serial = serial};
+
 fun make_dep (name, morphisms) : dep =
-  {name = name, morphisms = morphisms, pos = Position.thread_data (), serial = serial ()};
+ {name = name,
+  morphisms = apply2 Morphism.reset_context morphisms,
+  pos = Position.thread_data (),
+  serial = serial ()};
 
 (*table of mixin lists, per list mixins in reverse order of declaration;
   lists indexed by registration/dependency serial,
@@ -120,7 +127,8 @@
 
 val merge_mixins: mixins * mixins -> mixins = Inttab.merge_list (eq_snd op =);
 
-fun insert_mixin serial' mixin : mixins -> mixins = Inttab.cons_list (serial', (mixin, serial ()));
+fun insert_mixin serial' (morph, b) : mixins -> mixins =
+  Inttab.cons_list (serial', ((Morphism.reset_context morph, b), serial ()));
 
 fun rename_mixin (old, new) (mixins: mixins) =
   (case Inttab.lookup mixins old of
@@ -145,7 +153,7 @@
   (* dynamic part *)
 
   (*syntax declarations*)
-  syntax_decls: (declaration * serial) list,
+  syntax_decls: (Morphism.declaration_entity * serial) list,
   (*theorem declarations*)
   notes: ((string * Attrib.fact list) * serial) list,
   (*locale dependencies (sublocale relation) in reverse order*)
@@ -212,8 +220,10 @@
   thy |> Locales.map (Name_Space.define (Context.Theory thy) true
     (binding,
       mk_locale ((parameters, spec, (apply2 o Option.map) Thm.trim_context intros,
-          map Thm.trim_context axioms, hyp_spec),
-        ((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes),
+          map Thm.trim_context axioms,
+          map Element.trim_context_ctxt hyp_spec),
+        ((map (fn decl => (Morphism.entity_reset_context decl, serial ())) syntax_decls,
+          map (fn (a, facts) => ((a, map Attrib.trim_context_fact facts), serial ())) notes),
           (map (fn (name, morph) => make_dep (name, (morph, Morphism.identity))) dependencies,
             Inttab.empty)))) #> snd);
           (* FIXME Morphism.identity *)
@@ -233,15 +243,17 @@
 fun axioms_of thy = map (Thm.transfer thy) o #axioms o the_locale thy;
 
 fun instance_of thy name morph = params_of thy name |>
-  map (Morphism.term morph o Free o #1);
+  map (Morphism.term (Morphism.set_context thy morph) o Free o #1);
 
 fun specification_of thy = #spec o the_locale thy;
 
-fun hyp_spec_of thy = #hyp_spec o the_locale thy;
+fun hyp_spec_of thy = map (Element.transfer_ctxt thy) o #hyp_spec o the_locale thy;
+
+fun dependencies_of thy = map (transfer_dep thy) o #dependencies o the_locale thy;
 
-fun dependencies_of thy = #dependencies o the_locale thy;
-
-fun mixins_of thy name serial = lookup_mixins (#mixins (the_locale thy name)) serial;
+fun mixins_of thy name serial =
+  lookup_mixins (#mixins (the_locale thy name)) serial
+  |> (map o apfst o apfst) (Morphism.set_context thy);
 
 
 (* Print instance and qualifiers *)
@@ -395,7 +407,10 @@
 
 fun add_reg thy export (name, morph) =
   let
-    val reg = {morphisms = (morph, export), pos = Position.thread_data (), serial = serial ()};
+    val reg =
+     {morphisms = (Morphism.reset_context morph, Morphism.reset_context export),
+      pos = Position.thread_data (),
+      serial = serial ()};
     val id = (name, instance_of thy name (morph $> export));
   in (map_registrations o apfst) (Idtab.insert (K false) (id, reg)) end;
 
@@ -485,34 +500,35 @@
     val _ = trace "syntax" (name, morph) context;
     val thy = Context.theory_of context;
     val {syntax_decls, ...} = the_locale thy name;
+    val form_syntax_decl =
+      Morphism.form o Morphism.transform morph o Morphism.entity_set_context thy;
   in
-    context
-    |> fold_rev (fn (decl, _) => decl morph) syntax_decls
+    fold_rev (form_syntax_decl o #1) syntax_decls context
       handle ERROR msg => activate_err msg "syntax" (name, morph) context
   end;
 
-fun activate_notes activ_elem transfer context export' (name, morph) input =
+fun activate_notes activ_elem context export' (name, morph) input =
   let
     val thy = Context.theory_of context;
     val mixin =
       (case export' of
         NONE => Morphism.identity
       | SOME export => collect_mixins context (name, morph $> export) $> export);
-    val morph' = transfer input $> morph $> mixin;
+    val morph' = Morphism.set_context thy (morph $> mixin);
     val notes' = map (Element.transform_ctxt morph') (lazy_notes thy name);
   in
     (notes', input) |-> fold (fn elem => fn res =>
-      activ_elem (Element.transform_ctxt (transfer res) elem) res)
+      activ_elem (Element.transfer_ctxt thy elem) res)
   end handle ERROR msg => activate_err msg "facts" (name, morph) context;
 
-fun activate_notes_trace activ_elem transfer context export' (name, morph) context' =
+fun activate_notes_trace activ_elem context export' (name, morph) context' =
   let
     val _ = trace "facts" (name, morph) context';
   in
-    activate_notes activ_elem transfer context export' (name, morph) context'
+    activate_notes activ_elem context export' (name, morph) context'
   end;
 
-fun activate_all name thy activ_elem transfer (marked, input) =
+fun activate_all name thy activ_elem (marked, input) =
   let
     val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
     val input' = input |>
@@ -522,7 +538,7 @@
       (case asm of SOME A => activ_elem (Assumes [(Binding.empty_atts, [(A, [])])]) | _ => I) |>
       (not (null defs) ?
         activ_elem (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs)));
-    val activate = activate_notes activ_elem transfer (Context.Theory thy) NONE;
+    val activate = activate_notes activ_elem (Context.Theory thy) NONE;
   in
     roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input')
   end;
@@ -535,8 +551,8 @@
   |> Context_Position.set_visible_generic false
   |> pair (Idents.get context)
   |> roundup (Context.theory_of context)
-      (activate_notes_trace init_element Morphism.transfer_morphism'' context export)
-      (the_default Morphism.identity export) dep
+      (activate_notes_trace init_element context export)
+      (Morphism.default export) dep
   |-> Idents.put
   |> Context_Position.restore_visible_generic context;
 
@@ -556,7 +572,7 @@
     context
     |> Context_Position.set_visible_generic false
     |> pair empty_idents
-    |> activate_all name thy init_element Morphism.transfer_morphism''
+    |> activate_all name thy init_element
     |-> (fn marked' => Idents.put (merge_idents (marked, marked')))
     |> Context_Position.restore_visible_generic context
     |> Context.proof_of
@@ -647,11 +663,11 @@
       val stored_notes = ((kind, map Attrib.trim_context_fact facts), serial ());
       val applied_notes = make_notes kind facts;
 
-      fun apply_notes morph = applied_notes |> fold (fn elem => fn context =>
-        let val elem' = Element.transform_ctxt (Morphism.transfer_morphism'' context $> morph) elem
-        in Element.init elem' context end);
-      val apply_registrations = Context.theory_map (fn context =>
-        fold_rev (apply_notes o #2) (registrations_of context loc) context);
+      fun apply_notes morph = applied_notes |> fold (fn elem => fn thy =>
+        let val elem' = Element.transform_ctxt (Morphism.set_context thy morph) elem
+        in Context.theory_map (Element.init elem') thy end);
+      fun apply_registrations thy =
+        fold_rev (apply_notes o #2) (registrations_of (Context.Theory thy) loc) thy;
     in
       ctxt
       |> Attrib.local_notes kind facts |> #2
@@ -659,10 +675,12 @@
         ((change_locale loc o apfst o apsnd) (cons stored_notes) #> apply_registrations)
     end;
 
-fun add_declaration loc syntax decl =
-  syntax ?
-    Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
-  #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration decl)];
+fun add_declaration loc {syntax, pos} decl =
+  let val decl0 = Morphism.entity_reset_context decl in
+    syntax ?
+      Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl0, serial ())))
+    #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration pos decl0)]
+  end;
 
 
 (*** Reasoning about locales ***)
@@ -732,7 +750,7 @@
       | cons_elem (elem as Lazy_Notes _) = show_facts ? cons elem
       | cons_elem elem = cons elem;
     val elems =
-      activate_all name thy cons_elem (K (Morphism.transfer_morphism thy)) (empty_idents, [])
+      activate_all name thy cons_elem (empty_idents, [])
       |> snd |> rev
       |> tap consolidate_notes
       |> map force_notes;
--- a/src/Pure/Isar/method.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/method.ML	Fri May 26 09:56:20 2023 +0200
@@ -335,8 +335,9 @@
                ML_Lex.read_source source @ ML_Lex.read ")))")
           |> the_tactic;
       in
-        fn phi => set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi))
-      end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context))));
+        Morphism.entity (fn phi =>
+          set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi)))
+      end)) >> (fn decl => Morphism.form_entity (the_tactic (Morphism.form decl context))));
 
 
 (* method facts *)
@@ -606,11 +607,18 @@
                   map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms;
                 val facts =
                   Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)]
-                  |> map (fn (_, bs) => ((Binding.empty, [Attrib.internal (K attribute)]), bs));
+                  |> map (fn (_, bs) =>
+                    ((Binding.empty, [Attrib.internal pos (K attribute)]),
+                      Attrib.trim_context_thms bs));
 
-                fun decl phi =
-                  Context.mapping I init #>
-                  Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd;
+                val decl =
+                  Morphism.entity (fn phi => fn context =>
+                    let val psi = Morphism.set_context'' context phi in
+                      context
+                      |> Context.mapping I init
+                      |> Attrib.generic_notes "" (Attrib.transform_facts psi facts)
+                      |> snd
+                    end);
 
                 val modifier_report =
                   (#1 (Token.range_of modifier_toks),
--- a/src/Pure/Isar/proof.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/proof.ML	Fri May 26 09:56:20 2023 +0200
@@ -477,7 +477,7 @@
   in
     result
     |> map (Raw_Simplifier.norm_hhf result_ctxt #> protect_prems)
-    |> Proof_Context.goal_export result_ctxt goal_ctxt
+    |> Proof_Context.export_goal result_ctxt goal_ctxt
     |> (fn rules => Seq.lift provide_goal (EVERY (map refine rules) goal) state)
   end;
 
@@ -1096,7 +1096,7 @@
       let
         val ctxt' = context_of state';
         val export0 =
-          Assumption.export false result_ctxt (Proof_Context.augment result_text ctxt') #>
+          Assumption.export result_ctxt (Proof_Context.augment result_text ctxt') #>
           fold_rev (fn (x, v) => Thm.forall_intr_name (x, Thm.cterm_of params_ctxt v)) params #>
           Raw_Simplifier.norm_hhf_protect ctxt';
         val export = map export0 #> Variable.export result_ctxt ctxt';
--- a/src/Pure/Isar/proof_context.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/proof_context.ML	Fri May 26 09:56:20 2023 +0200
@@ -102,10 +102,10 @@
   val standard_typ_check: Proof.context -> typ list -> typ list
   val standard_term_check_finish: Proof.context -> term list -> term list
   val standard_term_uncheck: Proof.context -> term list -> term list
-  val goal_export: Proof.context -> Proof.context -> thm list -> thm list
+  val export_: {goal: bool} -> Proof.context -> Proof.context -> thm list -> thm list
   val export: Proof.context -> Proof.context -> thm list -> thm list
+  val export_goal: Proof.context -> Proof.context -> thm list -> thm list
   val export_morphism: Proof.context -> Proof.context -> morphism
-  val norm_export_morphism: Proof.context -> Proof.context -> morphism
   val auto_bind_goal: term list -> Proof.context -> Proof.context
   val auto_bind_facts: term list -> Proof.context -> Proof.context
   val simult_matches: Proof.context -> term * term list -> (indexname * term) list
@@ -839,21 +839,17 @@
 
 (** export results **)
 
-fun common_export is_goal inner outer =
-  map (Assumption.export is_goal inner outer) #>
+fun export_ goal inner outer =
+  map (Assumption.export_ goal inner outer) #>
   Variable.export inner outer;
 
-val goal_export = common_export true;
-val export = common_export false;
+val export = export_{goal = false};
+val export_goal = export_{goal = true};
 
 fun export_morphism inner outer =
   Assumption.export_morphism inner outer $>
   Variable.export_morphism inner outer;
 
-fun norm_export_morphism inner outer =
-  export_morphism inner outer $>
-  Morphism.thm_morphism "Proof_Context.norm_export" (Goal.norm_result outer);
-
 
 
 (** term bindings **)
--- a/src/Pure/Isar/spec_rules.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/spec_rules.ML	Fri May 26 09:56:20 2023 +0200
@@ -117,7 +117,7 @@
   {pos = pos, name = name, rough_classification = rough_classification, terms = terms,
     rules = map f rules};
 
-structure Rules = Generic_Data
+structure Data = Generic_Data
 (
   type T = spec_rule Item_Net.T;
   val empty : T = Item_Net.init eq_spec #terms;
@@ -132,9 +132,9 @@
     val thy = Context.theory_of context;
     val transfer = Global_Theory.transfer_theories thy;
     fun imported spec =
-      imports |> exists (fn thy => Item_Net.member (Rules.get (Context.Theory thy)) spec);
+      imports |> exists (fn thy => Item_Net.member (Data.get (Context.Theory thy)) spec);
   in
-    Item_Net.content (Rules.get context)
+    Item_Net.content (Data.get context)
     |> filter_out imported
     |> (map o map_spec_rules) transfer
   end;
@@ -148,7 +148,7 @@
 (* retrieve *)
 
 fun retrieve_generic context =
-  Item_Net.retrieve (Rules.get context)
+  Item_Net.retrieve (Data.get context)
   #> (map o map_spec_rules) (Thm.transfer'' context);
 
 val retrieve = retrieve_generic o Context.Proof;
@@ -158,27 +158,28 @@
 (* add *)
 
 fun add b rough_classification terms rules lthy =
-  let val thms0 = map Thm.trim_context (map (Drule.mk_term o Thm.cterm_of lthy) terms @ rules) in
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = true}
+  let
+    val n = length terms;
+    val thms0 = map Thm.trim_context (map (Drule.mk_term o Thm.cterm_of lthy) terms @ rules);
+  in
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of b}
       (fn phi => fn context =>
         let
+          val psi = Morphism.set_trim_context'' context phi;
           val pos = Position.thread_data ();
-          val name = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding phi b);
+          val name = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding psi b);
           val (terms', rules') =
-            map (Thm.transfer (Context.theory_of context)) thms0
-            |> Morphism.fact phi
-            |> chop (length terms)
-            |>> map (Thm.term_of o Drule.dest_term)
-            ||> map Thm.trim_context;
+            chop n (Morphism.fact psi thms0)
+            |>> map (Thm.term_of o Drule.dest_term);
         in
-          context |> (Rules.map o Item_Net.update)
+          context |> (Data.map o Item_Net.update)
             {pos = pos, name = name, rough_classification = rough_classification,
               terms = terms', rules = rules'}
         end)
   end;
 
 fun add_global b rough_classification terms rules thy =
-  thy |> (Context.theory_map o Rules.map o Item_Net.update)
+  thy |> (Context.theory_map o Data.map o Item_Net.update)
    {pos = Position.thread_data (),
     name = Sign.full_name thy b,
     rough_classification = rough_classification,
--- a/src/Pure/Isar/specification.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/specification.ML	Fri May 26 09:56:20 2023 +0200
@@ -259,21 +259,19 @@
     val ((lhs, (_, raw_th)), lthy2) = lthy
       |> Local_Theory.define_internal ((b, mx), ((Binding.suffix_name "_raw" name, []), rhs));
 
-    val th = prove lthy2 raw_th;
-    val lthy3 = lthy2 |> Spec_Rules.add name Spec_Rules.equational [lhs] [th];
+    val ([(def_name, [th])], lthy3) = lthy2
+      |> Local_Theory.notes [((name, atts), [([prove lthy2 raw_th], [])])];
 
-    val ([(def_name, [th'])], lthy4) = lthy3
-      |> Local_Theory.notes [((name, atts), [([th], [])])];
+    val lthy4 = lthy3
+      |> Spec_Rules.add name Spec_Rules.equational [lhs] [th]
+      |> Code.declare_default_eqns [(th, true)];
 
-    val lthy5 = lthy4
-      |> Code.declare_default_eqns [(th', true)];
-
-    val lhs' = Morphism.term (Local_Theory.target_morphism lthy5) lhs;
+    val lhs' = Morphism.term (Local_Theory.target_morphism lthy4) lhs;
 
     val _ =
-      Proof_Display.print_consts int (Position.thread_data ()) lthy5
+      Proof_Display.print_consts int (Position.thread_data ()) lthy4
         (Frees.defined (Frees.build (Frees.add_frees lhs'))) [(x, T)];
-  in ((lhs, (def_name, th')), lthy5) end;
+  in ((lhs, (def_name, th)), lthy4) end;
 
 fun definition xs ys As B = gen_def check_spec_open (K I) xs ys As B false;
 val definition_cmd = gen_def read_spec_open Attrib.check_src;
--- a/src/Pure/Isar/token.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/Isar/token.ML	Fri May 26 09:56:20 2023 +0200
@@ -29,8 +29,8 @@
     Typ of typ |
     Term of term |
     Fact of string option * thm list |
-    Attribute of morphism -> attribute |
-    Declaration of declaration |
+    Attribute of attribute Morphism.entity |
+    Declaration of Morphism.declaration_entity |
     Files of file Exn.result list |
     Output of XML.body option
   val pos_of: T -> Position.T
@@ -82,7 +82,8 @@
   val get_name: T -> name_value option
   val declare_maxidx: T -> Proof.context -> Proof.context
   val map_facts: (string option -> thm list -> thm list) -> T -> T
-  val trim_context_src: src -> src
+  val trim_context: T -> T
+  val transfer: theory -> T -> T
   val transform: morphism -> T -> T
   val init_assignable: T -> T
   val assign: value option -> T -> T
@@ -103,6 +104,7 @@
   val print_properties: Keyword.keywords -> Properties.T -> string
   val make: (int * int) * string -> Position.T -> T * Position.T
   val make_string: string * Position.T -> T
+  val make_string0: string -> T
   val make_int: int -> T list
   val make_src: string * Position.T -> T list -> src
   type 'a parser = T list -> 'a * T list
@@ -196,8 +198,8 @@
   Typ of typ |
   Term of term |
   Fact of string option * thm list |  (*optional name for dynamic fact, i.e. fact "variable"*)
-  Attribute of morphism -> attribute |
-  Declaration of declaration |
+  Attribute of attribute Morphism.entity |
+  Declaration of Morphism.declaration_entity |
   Files of file Exn.result list |
   Output of XML.body option;
 
@@ -484,7 +486,33 @@
     | Fact (a, ths) => Fact (a, f a ths)
     | _ => v));
 
-val trim_context_src = (map o map_facts) (K (map Thm.trim_context));
+
+(* implicit context *)
+
+local
+
+fun context thm_context morphism_context attribute_context declaration_context =
+  let
+    fun token_context tok = map_value
+      (fn Source src => Source (map token_context src)
+        | Fact (a, ths) => Fact (a, map thm_context ths)
+        | Name (a, phi) => Name (a, morphism_context phi)
+        | Attribute a => Attribute (attribute_context a)
+        | Declaration a => Declaration (declaration_context a)
+        | v => v) tok;
+  in token_context end;
+
+in
+
+val trim_context =
+  context Thm.trim_context Morphism.reset_context
+    Morphism.entity_reset_context Morphism.entity_reset_context;
+
+fun transfer thy =
+  context (Thm.transfer thy) (Morphism.set_context thy)
+    (Morphism.entity_set_context thy) (Morphism.entity_set_context thy);
+
+end;
 
 
 (* transform *)
@@ -530,6 +558,10 @@
 
 (* pretty *)
 
+fun pretty_keyword3 tok =
+  let val props = Position.properties_of (pos_of tok)
+  in Pretty.mark_str (Markup.properties props Markup.keyword3, unparse tok) end;
+
 fun pretty_value ctxt tok =
   (case get_value tok of
     SOME (Literal markup) =>
@@ -540,6 +572,8 @@
   | SOME (Term t) => Syntax.pretty_term ctxt t
   | SOME (Fact (_, ths)) =>
       Pretty.enclose "(" ")" (Pretty.breaks (map (Pretty.cartouche o Thm.pretty_thm ctxt) ths))
+  | SOME (Attribute _) => pretty_keyword3 tok
+  | SOME (Declaration _) => pretty_keyword3 tok
   | _ => Pretty.marks_str (markups Keyword.empty_keywords tok, unparse tok));
 
 
@@ -744,6 +778,8 @@
     val pos' = Position.no_range_position pos;
   in Token ((x, (pos', pos')), y, z) end;
 
+fun make_string0 s = make_string (s, Position.none);
+
 val make_int = explode Keyword.empty_keywords Position.none o signed_string_of_int;
 
 fun make_src a args = make_string a :: args;
@@ -758,8 +794,9 @@
 
 (* wrapped syntax *)
 
-fun syntax_generic scan src context =
+fun syntax_generic scan src0 context =
   let
+    val src = map (transfer (Context.theory_of context)) src0;
     val (name, pos) = name_of_src src;
     val old_reports = maps reports_of_value src;
     val args1 = map init_assignable (args_of_src src);
--- a/src/Pure/assumption.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/assumption.ML	Fri May 26 09:56:20 2023 +0200
@@ -17,8 +17,10 @@
   val local_prems_of: Proof.context -> Proof.context -> thm list
   val add_assms: export -> cterm list -> Proof.context -> thm list * Proof.context
   val add_assumes: cterm list -> Proof.context -> thm list * Proof.context
-  val export: bool -> Proof.context -> Proof.context -> thm -> thm
   val export_term: Proof.context -> Proof.context -> term -> term
+  val export_: {goal: bool} -> Proof.context -> Proof.context -> thm -> thm
+  val export: Proof.context -> Proof.context -> thm -> thm
+  val export_goal: Proof.context -> Proof.context -> thm -> thm
   val export_morphism: Proof.context -> Proof.context -> morphism
 end;
 
@@ -128,24 +130,32 @@
 
 (* export *)
 
-fun export is_goal inner outer =
+fun export_term inner outer =
+  fold_rev (fn (e, As) => #2 (e false As)) (local_assumptions_of inner outer);
+
+fun export_thm is_goal inner outer =
+  fold_rev (fn (e, As) => #1 (e is_goal As)) (local_assumptions_of inner outer);
+
+fun export_{goal} inner outer =
   Raw_Simplifier.norm_hhf_protect inner #>
-  fold_rev (fn (e, As) => #1 (e is_goal As)) (local_assumptions_of inner outer) #>
+  export_thm goal inner outer #>
   Raw_Simplifier.norm_hhf_protect outer;
 
-fun export_term inner outer =
-  fold_rev (fn (e, As) => #2 (e false As)) (local_assumptions_of inner outer);
+val export = export_{goal = false};
+val export_goal = export_{goal = true};
 
 fun export_morphism inner outer =
   let
-    val thm = export false inner outer;
+    val export0 = export_thm false inner outer;
+    fun thm thy =
+      let val norm = norm_hhf_protect (Proof_Context.init_global thy)
+      in norm #> export0 #> norm end;
     val term = export_term inner outer;
     val typ = Logic.type_map term;
   in
-    Morphism.transfer_morphism' inner $>
-    Morphism.transfer_morphism' outer $>
     Morphism.morphism "Assumption.export"
-      {binding = [], typ = [typ], term = [term], fact = [map thm]}
+      {binding = [], typ = [K typ], term = [K term], fact = [map o thm o Morphism.the_theory]}
+    |> Morphism.set_context (Proof_Context.theory_of inner)
   end;
 
 end;
--- a/src/Pure/context.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/context.ML	Fri May 26 09:56:20 2023 +0200
@@ -65,6 +65,7 @@
   val certificate_theory_id: certificate -> theory_id
   val eq_certificate: certificate * certificate -> bool
   val join_certificate: certificate * certificate -> certificate
+  val join_certificate_theory: theory * theory -> theory
   (*generic context*)
   datatype generic = Theory of theory | Proof of Proof.context
   val theory_tracing: bool Unsynchronized.ref
@@ -627,16 +628,24 @@
   | eq_certificate (Certificate_Id thy_id1, Certificate_Id thy_id2) = eq_thy_id (thy_id1, thy_id2)
   | eq_certificate _ = false;
 
+fun err_join (thy_id1, thy_id2) =
+  error ("Cannot join unrelated theory certificates " ^
+    display_name thy_id1 ^ " and " ^ display_name thy_id2);
+
 fun join_certificate (cert1, cert2) =
   let val (thy_id1, thy_id2) = apply2 certificate_theory_id (cert1, cert2) in
     if eq_thy_id (thy_id1, thy_id2) then (case cert1 of Certificate _ => cert1 | _ => cert2)
     else if proper_subthy_id (thy_id2, thy_id1) then cert1
     else if proper_subthy_id (thy_id1, thy_id2) then cert2
-    else
-      error ("Cannot join unrelated theory certificates " ^
-        display_name thy_id1 ^ " and " ^ display_name thy_id2)
+    else err_join (thy_id1, thy_id2)
   end;
 
+fun join_certificate_theory (thy1, thy2) =
+  let val (thy_id1, thy_id2) = apply2 theory_id (thy1, thy2) in
+    if subthy_id (thy_id2, thy_id1) then thy1
+    else if proper_subthy_id (thy_id1, thy_id2) then thy2
+    else err_join (thy_id1, thy_id2)
+  end;
 
 
 (*** generic context ***)
--- a/src/Pure/drule.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/drule.ML	Fri May 26 09:56:20 2023 +0200
@@ -75,7 +75,7 @@
   val eta_contraction_rule: thm -> thm
   val norm_hhf_eq: thm
   val norm_hhf_eqs: thm list
-  val is_norm_hhf: term -> bool
+  val is_norm_hhf: {protect: bool} -> term -> bool
   val norm_hhf: theory -> term -> term
   val norm_hhf_cterm: Proof.context -> cterm -> cterm
   val protect: cterm -> cterm
@@ -669,15 +669,19 @@
 val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq);
 val norm_hhf_eqs = [norm_hhf_eq, sort_constraint_eq];
 
-fun is_norm_hhf (Const ("Pure.sort_constraint", _)) = false
-  | is_norm_hhf (Const ("Pure.imp", _) $ _ $ (Const ("Pure.all", _) $ _)) = false
-  | is_norm_hhf (Abs _ $ _) = false
-  | is_norm_hhf (t $ u) = is_norm_hhf t andalso is_norm_hhf u
-  | is_norm_hhf (Abs (_, _, t)) = is_norm_hhf t
-  | is_norm_hhf _ = true;
+fun is_norm_hhf {protect} =
+  let
+    fun is_norm (Const ("Pure.sort_constraint", _)) = false
+      | is_norm (Const ("Pure.imp", _) $ _ $ (Const ("Pure.all", _) $ _)) = false
+      | is_norm (Const ("Pure.prop", _) $ t) = protect orelse is_norm t
+      | is_norm (Abs _ $ _) = false
+      | is_norm (t $ u) = is_norm t andalso is_norm u
+      | is_norm (Abs (_, _, t)) = is_norm t
+      | is_norm _ = true;
+  in is_norm end;
 
 fun norm_hhf thy t =
-  if is_norm_hhf t then t
+  if is_norm_hhf {protect = false} t then t
   else Pattern.rewrite_term thy [norm_hhf_prop] [] t;
 
 fun norm_hhf_cterm ctxt raw_ct =
@@ -685,7 +689,7 @@
     val thy = Proof_Context.theory_of ctxt;
     val ct = Thm.transfer_cterm thy raw_ct;
     val t = Thm.term_of ct;
-  in if is_norm_hhf t then ct else Thm.cterm_of ctxt (norm_hhf thy t) end;
+  in if is_norm_hhf {protect = false} t then ct else Thm.cterm_of ctxt (norm_hhf thy t) end;
 
 
 (* var indexes *)
--- a/src/Pure/ex/Def.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/ex/Def.thy	Fri May 26 09:56:20 2023 +0200
@@ -25,12 +25,12 @@
 
 (* context data *)
 
-type def = {lhs: term, mk_eq: morphism -> thm};
+type def = {lhs: term, eq: thm};
 
 val eq_def : def * def -> bool = op aconv o apply2 #lhs;
 
-fun transform_def phi ({lhs, mk_eq}: def) =
-  {lhs = Morphism.term phi lhs, mk_eq = Morphism.transform phi mk_eq};
+fun transform_def phi ({lhs, eq}: def) =
+  {lhs = Morphism.term phi lhs, eq = Morphism.thm phi eq};
 
 structure Data = Generic_Data
 (
@@ -40,12 +40,11 @@
 );
 
 fun declare_def lhs eq lthy =
-  let
-    val eq0 = Thm.trim_context eq;
-    val def: def = {lhs = lhs, mk_eq = fn phi => Morphism.thm phi eq0};
-  in
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = true}
-      (fn phi => (Data.map o Item_Net.update) (transform_def phi def))
+  let val def0: def = {lhs = lhs, eq = Thm.trim_context eq} in
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
+      (fn phi => fn context =>
+        let val psi = Morphism.set_trim_context'' context phi
+        in (Data.map o Item_Net.update) (transform_def psi def0) context end)
   end;
 
 fun get_def ctxt ct =
@@ -53,15 +52,10 @@
     val thy = Proof_Context.theory_of ctxt;
     val data = Data.get (Context.Proof ctxt);
     val t = Thm.term_of ct;
-    fun match_def {lhs, mk_eq} =
+    fun match_def {lhs, eq} =
       if Pattern.matches thy (lhs, t) then
-        let
-          val inst = Thm.match (Thm.cterm_of ctxt lhs, ct);
-          val eq =
-            Morphism.form mk_eq
-            |> Thm.transfer thy
-            |> Thm.instantiate inst;
-        in SOME eq end
+        let val inst = Thm.match (Thm.cterm_of ctxt lhs, ct)
+        in SOME (Thm.instantiate inst (Thm.transfer thy eq)) end
       else NONE;
   in Item_Net.retrieve_matching data t |> get_first match_def end;
 
--- a/src/Pure/goal.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/goal.ML	Fri May 26 09:56:20 2023 +0200
@@ -221,7 +221,7 @@
     res
     |> Thm.close_derivation \<^here>
     |> Conjunction.elim_balanced (length props)
-    |> map (Assumption.export false ctxt' ctxt)
+    |> map (Assumption.export ctxt' ctxt)
     |> Variable.export ctxt' ctxt
     |> map Drule.zero_var_indexes
   end;
@@ -311,7 +311,7 @@
 fun norm_hhf_tac ctxt =
   resolve_tac ctxt [Drule.asm_rl]  (*cheap approximation -- thanks to builtin Logic.flatten_params*)
   THEN' SUBGOAL (fn (t, i) =>
-    if Drule.is_norm_hhf t then all_tac
+    if Drule.is_norm_hhf {protect = false} t then all_tac
     else rewrite_goal_tac ctxt Drule.norm_hhf_eqs i);
 
 
--- a/src/Pure/morphism.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/morphism.ML	Fri May 26 09:56:20 2023 +0200
@@ -9,7 +9,6 @@
 signature BASIC_MORPHISM =
 sig
   type morphism
-  type declaration = morphism -> Context.generic -> Context.generic
   val $> : morphism * morphism -> morphism
 end
 
@@ -17,12 +16,18 @@
 sig
   include BASIC_MORPHISM
   exception MORPHISM of string * exn
+  val the_theory: theory option -> theory
+  val set_context: theory -> morphism -> morphism
+  val set_context': Proof.context -> morphism -> morphism
+  val set_context'': Context.generic -> morphism -> morphism
+  val reset_context: morphism -> morphism
   val morphism: string ->
-   {binding: (binding -> binding) list,
-    typ: (typ -> typ) list,
-    term: (term -> term) list,
-    fact: (thm list -> thm list) list} -> morphism
+   {binding: (theory option -> binding -> binding) list,
+    typ: (theory option -> typ -> typ) list,
+    term: (theory option -> term -> term) list,
+    fact: (theory option -> thm list -> thm list) list} -> morphism
   val is_identity: morphism -> bool
+  val is_empty: morphism -> bool
   val pretty: morphism -> Pretty.T
   val binding: morphism -> binding -> binding
   val binding_prefix: morphism -> (string * bool) list
@@ -32,18 +37,36 @@
   val thm: morphism -> thm -> thm
   val cterm: morphism -> cterm -> cterm
   val identity: morphism
+  val default: morphism option -> morphism
   val compose: morphism -> morphism -> morphism
-  val transform: morphism -> (morphism -> 'a) -> morphism -> 'a
-  val form: (morphism -> 'a) -> 'a
+  type 'a entity
+  val entity: (morphism -> 'a) -> 'a entity
+  val entity_reset_context: 'a entity -> 'a entity
+  val entity_set_context: theory -> 'a entity -> 'a entity
+  val entity_set_context': Proof.context -> 'a entity -> 'a entity
+  val entity_set_context'': Context.generic -> 'a entity -> 'a entity
+  val transform: morphism -> 'a entity -> 'a entity
+  val transform_reset_context: morphism -> 'a entity -> 'a entity
+  val form: 'a entity -> 'a
+  val form_entity: (morphism -> 'a) -> 'a
+  type declaration = morphism -> Context.generic -> Context.generic
+  type declaration_entity = (Context.generic -> Context.generic) entity
   val binding_morphism: string -> (binding -> binding) -> morphism
+  val typ_morphism': string -> (theory -> typ -> typ) -> morphism
   val typ_morphism: string -> (typ -> typ) -> morphism
+  val term_morphism': string -> (theory -> term -> term) -> morphism
   val term_morphism: string -> (term -> term) -> morphism
+  val fact_morphism': string -> (theory -> thm list -> thm list) -> morphism
   val fact_morphism: string -> (thm list -> thm list) -> morphism
+  val thm_morphism': string -> (theory -> thm -> thm) -> morphism
   val thm_morphism: string -> (thm -> thm) -> morphism
   val transfer_morphism: theory -> morphism
   val transfer_morphism': Proof.context -> morphism
   val transfer_morphism'': Context.generic -> morphism
   val trim_context_morphism: morphism
+  val set_trim_context: theory -> morphism -> morphism
+  val set_trim_context': Proof.context -> morphism -> morphism
+  val set_trim_context'': Context.generic -> morphism -> morphism
   val instantiate_frees_morphism: ctyp TFrees.table * cterm Frees.table -> morphism
   val instantiate_morphism: ctyp TVars.table * cterm Vars.table -> morphism
 end;
@@ -53,30 +76,53 @@
 
 (* named functions *)
 
-type 'a funs = (string * ('a -> 'a)) list;
+type 'a funs = (string * (theory option -> 'a -> 'a)) list;
 
 exception MORPHISM of string * exn;
 
-fun app (name, f) x = f x
+fun app context (name, f) x = f context x
   handle exn =>
     if Exn.is_interrupt exn then Exn.reraise exn else raise MORPHISM (name, exn);
 
-fun apply fs = fold_rev app fs;
+
+(* optional context *)
+
+fun the_theory (SOME thy) = thy
+  | the_theory NONE = raise Fail "Morphism lacks theory context";
+
+fun join_transfer (SOME thy) = Thm.join_transfer thy
+  | join_transfer NONE = I;
+
+val join_context = join_options Context.join_certificate_theory;
 
 
 (* type morphism *)
 
 datatype morphism = Morphism of
- {names: string list,
+ {context: theory option,
+  names: string list,
   binding: binding funs,
   typ: typ funs,
   term: term funs,
   fact: thm list funs};
 
-type declaration = morphism -> Context.generic -> Context.generic;
+fun rep (Morphism args) = args;
+
+fun apply which phi =
+  let val args = rep phi
+  in fold_rev (app (#context args)) (which args) end;
+
+fun put_context context (Morphism {context = _, names, binding, typ, term, fact}) =
+  Morphism {context = context, names = names, binding = binding, typ = typ, term = term, fact = fact};
+
+val set_context = put_context o SOME;
+val set_context' = set_context o Proof_Context.theory_of;
+val set_context'' = set_context o Context.theory_of;
+val reset_context = put_context NONE;
 
 fun morphism a {binding, typ, term, fact} =
   Morphism {
+    context = NONE,
     names = if a = "" then [] else [a],
     binding = map (pair a) binding,
     typ = map (pair a) typ,
@@ -84,18 +130,20 @@
     fact = map (pair a) fact};
 
 (*syntactic test only!*)
-fun is_identity (Morphism {names, binding, typ, term, fact}) =
+fun is_identity (Morphism {context = _, names, binding, typ, term, fact}) =
   null names andalso null binding andalso null typ andalso null term andalso null fact;
 
-fun pretty (Morphism {names, ...}) = Pretty.enum ";" "{" "}" (map Pretty.str (rev names));
+fun is_empty phi = is_none (#context (rep phi)) andalso is_identity phi;
+
+fun pretty phi = Pretty.enum ";" "{" "}" (map Pretty.str (rev (#names (rep phi))));
 
 val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty);
 
-fun binding (Morphism {binding, ...}) = apply binding;
+val binding = apply #binding;
 fun binding_prefix morph = Binding.name "x" |> binding morph |> Binding.prefix_of;
-fun typ (Morphism {typ, ...}) = apply typ;
-fun term (Morphism {term, ...}) = apply term;
-fun fact (Morphism {fact, ...}) = apply fact;
+val typ = apply #typ;
+val term = apply #term;
+fun fact phi = map (join_transfer (#context (rep phi))) #> apply #fact phi;
 val thm = singleton o fact;
 val cterm = Drule.cterm_rule o thm;
 
@@ -104,36 +152,73 @@
 
 val identity = morphism "" {binding = [], typ = [], term = [], fact = []};
 
-fun compose
-    (Morphism {names = names1, binding = binding1, typ = typ1, term = term1, fact = fact1})
-    (Morphism {names = names2, binding = binding2, typ = typ2, term = term2, fact = fact2}) =
-  Morphism {
-    names = names1 @ names2,
-    binding = binding1 @ binding2,
-    typ = typ1 @ typ2,
-    term = term1 @ term2,
-    fact = fact1 @ fact2};
+val default = the_default identity;
+
+fun compose phi1 phi2 =
+  if is_empty phi1 then phi2
+  else if is_empty phi2 then phi1
+  else
+    let
+      val {context = context1, names = names1, binding = binding1,
+        typ = typ1, term = term1, fact = fact1} = rep phi1;
+      val {context = context2, names = names2, binding = binding2,
+        typ = typ2, term = term2, fact = fact2} = rep phi2;
+    in
+      Morphism {
+        context = join_context (context1, context2),
+        names = names1 @ names2,
+        binding = binding1 @ binding2,
+        typ = typ1 @ typ2,
+        term = term1 @ term2,
+        fact = fact1 @ fact2}
+    end;
 
 fun phi1 $> phi2 = compose phi2 phi1;
 
-fun transform phi f = fn psi => f (phi $> psi);
-fun form f = f identity;
+
+(* abstract entities *)
+
+datatype 'a entity = Entity of (morphism -> 'a) * morphism;
+fun entity f = Entity (f, identity);
+
+fun entity_morphism g (Entity (f, phi)) = Entity (f, g phi);
+fun entity_reset_context a = entity_morphism reset_context a;
+fun entity_set_context thy a = entity_morphism (set_context thy) a;
+fun entity_set_context' ctxt a = entity_morphism (set_context' ctxt) a;
+fun entity_set_context'' context a = entity_morphism (set_context'' context) a;
+
+fun transform phi = entity_morphism (compose phi);
+fun transform_reset_context phi = entity_morphism (reset_context o compose phi);
+
+fun form (Entity (f, phi)) = f phi;
+fun form_entity f = f identity;
+
+type declaration = morphism -> Context.generic -> Context.generic;
+type declaration_entity = (Context.generic -> Context.generic) entity;
 
 
 (* concrete morphisms *)
 
-fun binding_morphism a binding = morphism a {binding = [binding], typ = [], term = [], fact = []};
-fun typ_morphism a typ = morphism a {binding = [], typ = [typ], term = [], fact = []};
-fun term_morphism a term = morphism a {binding = [], typ = [], term = [term], fact = []};
-fun fact_morphism a fact = morphism a {binding = [], typ = [], term = [], fact = [fact]};
-fun thm_morphism a thm = morphism a {binding = [], typ = [], term = [], fact = [map thm]};
+fun binding_morphism a binding = morphism a {binding = [K binding], typ = [], term = [], fact = []};
+fun typ_morphism' a typ = morphism a {binding = [], typ = [typ o the_theory], term = [], fact = []};
+fun typ_morphism a typ = morphism a {binding = [], typ = [K typ], term = [], fact = []};
+fun term_morphism' a term = morphism a {binding = [], typ = [], term = [term o the_theory], fact = []};
+fun term_morphism a term = morphism a {binding = [], typ = [], term = [K term], fact = []};
+fun fact_morphism' a fact = morphism a {binding = [], typ = [], term = [], fact = [fact o the_theory]};
+fun fact_morphism a fact = morphism a {binding = [], typ = [], term = [], fact = [K fact]};
+fun thm_morphism' a thm = morphism a {binding = [], typ = [], term = [], fact = [map o thm o the_theory]};
+fun thm_morphism a thm = morphism a {binding = [], typ = [], term = [], fact = [K (map thm)]};
 
-val transfer_morphism = thm_morphism "transfer" o Thm.join_transfer;
+fun transfer_morphism thy = fact_morphism "transfer" I |> set_context thy;
 val transfer_morphism' = transfer_morphism o Proof_Context.theory_of;
 val transfer_morphism'' = transfer_morphism o Context.theory_of;
 
 val trim_context_morphism = thm_morphism "trim_context" Thm.trim_context;
 
+fun set_trim_context thy phi = set_context thy phi $> trim_context_morphism;
+val set_trim_context' = set_trim_context o Proof_Context.theory_of;
+val set_trim_context'' = set_trim_context o Context.theory_of;
+
 
 (* instantiate *)
 
@@ -148,9 +233,9 @@
         {binding = [],
          typ =
           if TFrees.is_empty instT then []
-          else [Term_Subst.instantiateT_frees instT],
-         term = [Term_Subst.instantiate_frees (instT, inst)],
-         fact = [map (Thm.instantiate_frees (cinstT, cinst))]}
+          else [K (Term_Subst.instantiateT_frees instT)],
+         term = [K (Term_Subst.instantiate_frees (instT, inst))],
+         fact = [K (map (Thm.instantiate_frees (cinstT, cinst)))]}
     end;
 
 fun instantiate_morphism (cinstT, cinst) =
@@ -164,9 +249,9 @@
         {binding = [],
          typ =
           if TVars.is_empty instT then []
-          else [Term_Subst.instantiateT instT],
-         term = [Term_Subst.instantiate (instT, inst)],
-         fact = [map (Thm.instantiate (cinstT, cinst))]}
+          else [K (Term_Subst.instantiateT instT)],
+         term = [K (Term_Subst.instantiate (instT, inst))],
+         fact = [K (map (Thm.instantiate (cinstT, cinst)))]}
     end;
 
 end;
--- a/src/Pure/raw_simplifier.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/raw_simplifier.ML	Fri May 26 09:56:20 2023 +0200
@@ -36,7 +36,7 @@
   type simproc
   val eq_simproc: simproc * simproc -> bool
   val cert_simproc: theory -> string ->
-    {lhss: term list, proc: morphism -> Proof.context -> cterm -> thm option} -> simproc
+    {lhss: term list, proc: (Proof.context -> cterm -> thm option) Morphism.entity} -> simproc
   val transform_simproc: morphism -> simproc -> simproc
   val simpset_of: Proof.context -> simpset
   val put_simpset: simpset -> Proof.context -> Proof.context
@@ -709,7 +709,7 @@
   Simproc of
     {name: string,
      lhss: term list,
-     proc: morphism -> Proof.context -> cterm -> thm option,
+     proc: (Proof.context -> cterm -> thm option) Morphism.entity,
      stamp: stamp};
 
 fun eq_simproc (Simproc {stamp = stamp1, ...}, Simproc {stamp = stamp2, ...}) = stamp1 = stamp2;
@@ -721,7 +721,7 @@
   Simproc
    {name = name,
     lhss = map (Morphism.term phi) lhss,
-    proc = Morphism.transform phi proc,
+    proc = Morphism.transform_reset_context phi proc,
     stamp = stamp};
 
 local
@@ -1423,11 +1423,11 @@
 
 local
 
-fun gen_norm_hhf ss ctxt0 th0 =
+fun gen_norm_hhf protect ss ctxt0 th0 =
   let
     val (ctxt, th) = Thm.join_transfer_context (ctxt0, th0);
     val th' =
-      if Drule.is_norm_hhf (Thm.prop_of th) then th
+      if Drule.is_norm_hhf protect (Thm.prop_of th) then th
       else
         Conv.fconv_rule (rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset ss ctxt)) th;
   in th' |> Thm.adjust_maxidx_thm ~1 |> Variable.gen_all ctxt end;
@@ -1445,8 +1445,8 @@
 
 in
 
-val norm_hhf = gen_norm_hhf hhf_ss;
-val norm_hhf_protect = gen_norm_hhf hhf_protect_ss;
+val norm_hhf = gen_norm_hhf {protect = false} hhf_ss;
+val norm_hhf_protect = gen_norm_hhf {protect = true} hhf_protect_ss;
 
 end;
 
--- a/src/Pure/simplifier.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/simplifier.ML	Fri May 26 09:56:20 2023 +0200
@@ -128,7 +128,8 @@
     val ctxt' = fold Proof_Context.augment lhss ctxt;
     val lhss' = Variable.export_terms ctxt' ctxt lhss;
   in
-    cert_simproc (Proof_Context.theory_of ctxt) name {lhss = lhss', proc = proc}
+    cert_simproc (Proof_Context.theory_of ctxt) name
+      {lhss = lhss', proc = Morphism.entity proc}
   end;
 
 local
@@ -138,15 +139,16 @@
     val simproc =
       make_simproc lthy (Local_Theory.full_name lthy b) {lhss = prep lthy lhss, proc = proc};
   in
-    lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context =>
-      let
-        val b' = Morphism.binding phi b;
-        val simproc' = transform_simproc phi simproc;
-      in
-        context
-        |> Simprocs.map (#2 o Name_Space.define context true (b', simproc'))
-        |> map_ss (fn ctxt => ctxt addsimprocs [simproc'])
-      end)
+    lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of b}
+      (fn phi => fn context =>
+        let
+          val b' = Morphism.binding phi b;
+          val simproc' = transform_simproc phi simproc;
+        in
+          context
+          |> Simprocs.map (#2 o Name_Space.define context true (b', simproc'))
+          |> map_ss (fn ctxt => ctxt addsimprocs [simproc'])
+        end)
   end;
 
 in
@@ -161,28 +163,27 @@
 
 local
 
-fun make_cong ctxt = Thm.close_derivation \<^here> o Thm.reflexive
-  o Thm.cterm_of ctxt o Logic.varify_global o list_comb;
-
-fun add_cong (const_binding, (const, target_params)) gthy =
-  if null target_params
-  then gthy
+fun add_foundation_cong (binding, (const, target_params)) gthy =
+  if null target_params then gthy
   else
     let
-      val cong = make_cong (Context.proof_of gthy) (const, target_params)
-      val cong_binding = Binding.qualify_name true const_binding "cong"
+      val thy = Context.theory_of gthy;
+      val cong =
+        list_comb (const, target_params)
+        |> Logic.varify_global
+        |> Thm.global_cterm_of thy
+        |> Thm.reflexive
+        |> Thm.close_derivation \<^here>;
+      val cong_binding = Binding.qualify_name true binding "cong";
     in
       gthy
-      |> Attrib.generic_notes Thm.theoremK
-           [((cong_binding, []), [([cong], [])])]
-      |> snd
+      |> Attrib.generic_notes Thm.theoremK [((cong_binding, []), [([cong], [])])]
+      |> #2
     end;
 
-in
+val _ = Theory.setup (Generic_Target.add_foundation_interpretation add_foundation_cong);
 
-val _ = Theory.setup (Generic_Target.add_foundation_interpretation add_cong);
-
-end;
+in end;
 
 
 
@@ -311,8 +312,8 @@
 val add_del =
   (Args.del -- Args.colon >> K (op delsimprocs) ||
     Scan.option (Args.add -- Args.colon) >> K (op addsimprocs))
-  >> (fn f => fn simproc => fn phi => Thm.declaration_attribute
-      (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc])))));
+  >> (fn f => fn simproc => Morphism.entity (fn phi => Thm.declaration_attribute
+      (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc]))))));
 
 in
 
--- a/src/Pure/variable.ML	Fri May 26 09:49:45 2023 +0200
+++ b/src/Pure/variable.ML	Fri May 26 09:56:20 2023 +0200
@@ -583,9 +583,8 @@
     val term = singleton (export_terms inner outer);
     val typ = Logic.type_map term;
   in
-    Morphism.transfer_morphism' inner $>
-    Morphism.transfer_morphism' outer $>
-    Morphism.morphism "Variable.export" {binding = [], typ = [typ], term = [term], fact = [fact]}
+    Morphism.morphism "Variable.export"
+      {binding = [], typ = [K typ], term = [K term], fact = [K fact]}
   end;
 
 
--- a/src/ZF/OrdQuant.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/ZF/OrdQuant.thy	Fri May 26 09:56:20 2023 +0200
@@ -350,13 +350,11 @@
 text \<open>Setting up the one-point-rule simproc\<close>
 
 simproc_setup defined_rex ("\<exists>x[M]. P(x) \<and> Q(x)") = \<open>
-  fn _ => Quantifier1.rearrange_Bex
-    (fn ctxt => unfold_tac ctxt @{thms rex_def})
+  K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms rex_def}))
 \<close>
 
 simproc_setup defined_rall ("\<forall>x[M]. P(x) \<longrightarrow> Q(x)") = \<open>
-  fn _ => Quantifier1.rearrange_Ball
-    (fn ctxt => unfold_tac ctxt @{thms rall_def})
+  K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms rall_def}))
 \<close>
 
 end
--- a/src/ZF/pair.thy	Fri May 26 09:49:45 2023 +0200
+++ b/src/ZF/pair.thy	Fri May 26 09:56:20 2023 +0200
@@ -19,13 +19,11 @@
 ML \<open>val ZF_ss = simpset_of \<^context>\<close>
 
 simproc_setup defined_Bex ("\<exists>x\<in>A. P(x) \<and> Q(x)") = \<open>
-  fn _ => Quantifier1.rearrange_Bex
-    (fn ctxt => unfold_tac ctxt @{thms Bex_def})
+  K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}))
 \<close>
 
 simproc_setup defined_Ball ("\<forall>x\<in>A. P(x) \<longrightarrow> Q(x)") = \<open>
-  fn _ => Quantifier1.rearrange_Ball
-    (fn ctxt => unfold_tac ctxt @{thms Ball_def})
+  K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def}))
 \<close>