Theory Abstract_Metric_Spaces
section ‹Abstract Metric Spaces›
theory Abstract_Metric_Spaces
  imports Elementary_Metric_Spaces Abstract_Limits Abstract_Topological_Spaces
begin
locale Metric_space =
  fixes M :: "'a set" and d :: "'a ⇒ 'a ⇒ real"
  assumes nonneg [simp]: "⋀x y. 0 ≤ d x y"
  assumes commute: "⋀x y. d x y = d y x"
  assumes zero [simp]: "⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ d x y = 0 ⟷ x=y"
  assumes triangle: "⋀x y z. ⟦x ∈ M; y ∈ M; z ∈ M⟧ ⟹ d x z ≤ d x y + d y z"
text ‹Link with the type class version›
interpretation Met_TC: Metric_space UNIV dist
  by (simp add: dist_commute dist_triangle Metric_space.intro)
context Metric_space
begin
lemma subspace: "M' ⊆ M ⟹ Metric_space M' d"
  by (simp add: commute in_mono Metric_space.intro triangle)
lemma abs_mdist [simp] : "¦d x y¦ = d x y"
  by simp
lemma mdist_pos_less: "⟦x ≠ y; x ∈ M; y ∈ M⟧ ⟹ 0 < d x y"
  by (metis less_eq_real_def nonneg zero)
lemma mdist_zero [simp]: "x ∈ M ⟹ d x x = 0"
  by simp
lemma mdist_pos_eq [simp]: "⟦x ∈ M; y ∈ M⟧ ⟹ 0 < d x y ⟷ x ≠ y"
  using mdist_pos_less zero by fastforce
lemma triangle': "⟦x ∈ M; y ∈ M; z ∈ M⟧ ⟹ d x z ≤ d x y + d z y"
  by (simp add: commute triangle)
lemma triangle'': "⟦x ∈ M; y ∈ M; z ∈ M⟧ ⟹ d x z ≤ d y x + d y z"
  by (simp add: commute triangle)
lemma mdist_reverse_triangle: "⟦x ∈ M; y ∈ M; z ∈ M⟧ ⟹ ¦d x y - d y z¦ ≤ d x z"
  by (smt (verit) commute triangle)
text‹ Open and closed balls                                                                ›
definition mball where "mball x r ≡ {y. x ∈ M ∧ y ∈ M ∧ d x y < r}"
definition mcball where "mcball x r ≡ {y. x ∈ M ∧ y ∈ M ∧ d x y ≤ r}"
lemma in_mball [simp]: "y ∈ mball x r ⟷ x ∈ M ∧ y ∈ M ∧ d x y < r"
  by (simp add: mball_def)
lemma centre_in_mball_iff [iff]: "x ∈ mball x r ⟷ x ∈ M ∧ 0 < r"
  using in_mball mdist_zero by force
lemma mball_subset_mspace: "mball x r ⊆ M"
  by auto
lemma mball_eq_empty: "mball x r = {} ⟷ (x ∉ M) ∨ r ≤ 0"
  by (smt (verit, best) Collect_empty_eq centre_in_mball_iff mball_def nonneg)
lemma mball_subset: "⟦d x y + a ≤ b; y ∈ M⟧ ⟹ mball x a ⊆ mball y b"
  by (smt (verit) commute in_mball subsetI triangle)
lemma disjoint_mball: "r + r' ≤ d x x' ⟹ disjnt (mball x r) (mball x' r')"
  by (smt (verit) commute disjnt_iff in_mball triangle)
lemma mball_subset_concentric: "r ≤ s ⟹ mball x r ⊆ mball x s"
  by auto
lemma in_mcball [simp]: "y ∈ mcball x r ⟷ x ∈ M ∧ y ∈ M ∧ d x y ≤ r"
  by (simp add: mcball_def)
lemma centre_in_mcball_iff [iff]: "x ∈ mcball x r ⟷ x ∈ M ∧ 0 ≤ r"
  using mdist_zero by force
lemma mcball_eq_empty: "mcball x r = {} ⟷ (x ∉ M) ∨ r < 0"
  by (smt (verit, best) Collect_empty_eq centre_in_mcball_iff empty_iff mcball_def nonneg)
lemma mcball_subset_mspace: "mcball x r ⊆ M"
  by auto
lemma mball_subset_mcball: "mball x r ⊆ mcball x r"
  by auto
lemma mcball_subset: "⟦d x y + a ≤ b; y ∈ M⟧ ⟹ mcball x a ⊆ mcball y b"
  by (smt (verit) in_mcball mdist_reverse_triangle subsetI)
lemma mcball_subset_concentric: "r ≤ s ⟹ mcball x r ⊆ mcball x s"
  by force
lemma mcball_subset_mball: "⟦d x y + a < b; y ∈ M⟧ ⟹ mcball x a ⊆ mball y b"
  by (smt (verit) commute in_mball in_mcball subsetI triangle)
lemma mcball_subset_mball_concentric: "a < b ⟹ mcball x a ⊆ mball x b"
  by force
end
subsection ‹Metric topology                                                           ›
context Metric_space
begin
definition mopen where 
  "mopen U ≡ U ⊆ M ∧ (∀x. x ∈ U ⟶ (∃r>0. mball x r ⊆ U))"
definition mtopology :: "'a topology" where 
  "mtopology ≡ topology mopen"
lemma is_topology_metric_topology [iff]: "istopology mopen"
proof -
  have "⋀S T. ⟦mopen S; mopen T⟧ ⟹ mopen (S ∩ T)"
    by (smt (verit, del_insts) Int_iff in_mball mopen_def subset_eq)
  moreover have "⋀𝒦. (∀K∈𝒦. mopen K) ⟶ mopen (⋃𝒦)"
    using mopen_def by fastforce
  ultimately show ?thesis
    by (simp add: istopology_def)
qed
lemma openin_mtopology: "openin mtopology U ⟷ U ⊆ M ∧ (∀x. x ∈ U ⟶ (∃r>0. mball x r ⊆ U))"
  by (simp add: mopen_def mtopology_def)
lemma topspace_mtopology [simp]: "topspace mtopology = M"
  by (meson order.refl mball_subset_mspace openin_mtopology openin_subset openin_topspace subset_antisym zero_less_one)
lemma subtopology_mspace [simp]: "subtopology mtopology M = mtopology"
  by (metis subtopology_topspace topspace_mtopology)
lemma open_in_mspace [iff]: "openin mtopology M"
  by (metis openin_topspace topspace_mtopology)
lemma closedin_mspace [iff]: "closedin mtopology M"
  by (metis closedin_topspace topspace_mtopology)
lemma openin_mball [iff]: "openin mtopology (mball x r)"
proof -
  have "⋀y. ⟦x ∈ M; d x y < r⟧ ⟹ ∃s>0. mball y s ⊆ mball x r"
    by (metis add_diff_cancel_left' add_diff_eq commute less_add_same_cancel1 mball_subset order_refl)
  then show ?thesis
    by (auto simp: openin_mtopology)
qed
lemma mtopology_base:
   "mtopology = topology(arbitrary union_of (λU. ∃x ∈ M. ∃r>0. U = mball x r))"
proof -
  have "⋀S. ∃x r. x ∈ M ∧ 0 < r ∧ S = mball x r ⟹ openin mtopology S"
    using openin_mball by blast
  moreover have "⋀U x. ⟦openin mtopology U; x ∈ U⟧ ⟹ ∃B. (∃x r. x ∈ M ∧ 0 < r ∧ B = mball x r) ∧ x ∈ B ∧ B ⊆ U"
    by (metis centre_in_mball_iff in_mono openin_mtopology)
  ultimately show ?thesis
    by (smt (verit) topology_base_unique)
qed
lemma closedin_metric:
   "closedin mtopology C ⟷ C ⊆ M ∧ (∀x. x ∈ M - C ⟶ (∃r>0. disjnt C (mball x r)))"  (is "?lhs = ?rhs")
proof
  show "?lhs ⟹ ?rhs"
    unfolding closedin_def openin_mtopology
    by (metis Diff_disjoint disjnt_def disjnt_subset2 topspace_mtopology)
  show "?rhs ⟹ ?lhs"
    unfolding closedin_def openin_mtopology disjnt_def
    by (metis Diff_subset Diff_triv Int_Diff Int_commute inf.absorb_iff2 mball_subset_mspace topspace_mtopology)
qed
lemma closedin_mcball [iff]: "closedin mtopology (mcball x r)"
proof -
  have "∃ra>0. disjnt (mcball x r) (mball y ra)" if "x ∉ M" for y
    by (metis disjnt_empty1 gt_ex mcball_eq_empty that)
  moreover have "disjnt (mcball x r) (mball y (d x y - r))" if "y ∈ M" "d x y > r" for y
    using that disjnt_iff in_mball in_mcball mdist_reverse_triangle by force
  ultimately show ?thesis
    using closedin_metric mcball_subset_mspace by fastforce
qed
lemma mball_iff_mcball: "(∃r>0. mball x r ⊆ U) = (∃r>0. mcball x r ⊆ U)"
  by (meson dense mball_subset_mcball mcball_subset_mball_concentric order_trans)
lemma openin_mtopology_mcball:
  "openin mtopology U ⟷ U ⊆ M ∧ (∀x. x ∈ U ⟶ (∃r. 0 < r ∧ mcball x r ⊆ U))"
  by (simp add: mball_iff_mcball openin_mtopology)
lemma metric_derived_set_of:
  "mtopology derived_set_of S = {x ∈ M. ∀r>0. ∃y∈S. y≠x ∧ y ∈ mball x r}" (is "?lhs=?rhs")
proof
  show "?lhs ⊆ ?rhs"
    unfolding openin_mtopology derived_set_of_def
    by clarsimp (metis in_mball openin_mball openin_mtopology zero)
  show "?rhs ⊆ ?lhs"
    unfolding openin_mtopology derived_set_of_def 
    by clarify (metis subsetD topspace_mtopology)
qed
lemma metric_closure_of:
   "mtopology closure_of S = {x ∈ M. ∀r>0. ∃y ∈ S. y ∈ mball x r}"
proof -
  have "⋀x r. ⟦0 < r; x ∈ mtopology closure_of S⟧ ⟹ ∃y∈S. y ∈ mball x r"
    by (metis centre_in_mball_iff in_closure_of openin_mball topspace_mtopology)
  moreover have "⋀x T. ⟦x ∈ M; ∀r>0. ∃y∈S. y ∈ mball x r⟧ ⟹ x ∈ mtopology closure_of S"
    by (smt (verit) in_closure_of in_mball openin_mtopology subsetD topspace_mtopology)
  ultimately show ?thesis
    by (auto simp: in_closure_of)
qed
lemma metric_closure_of_alt:
  "mtopology closure_of S = {x ∈ M. ∀r>0. ∃y ∈ S. y ∈ mcball x r}"
proof -
  have "⋀x r. ⟦∀r>0. x ∈ M ∧ (∃y∈S. y ∈ mcball x r); 0 < r⟧ ⟹ ∃y∈S. y ∈ M ∧ d x y < r"
    by (meson dense in_mcball le_less_trans)
  then show ?thesis
    by (fastforce simp: metric_closure_of in_closure_of)
qed
lemma metric_interior_of:
   "mtopology interior_of S = {x ∈ M. ∃ε>0. mball x ε ⊆ S}" (is "?lhs=?rhs")
proof
  show "?lhs ⊆ ?rhs"
    using interior_of_maximal_eq openin_mtopology by fastforce
  show "?rhs ⊆ ?lhs"
    using interior_of_def openin_mball by fastforce
qed
lemma metric_interior_of_alt:
   "mtopology interior_of S = {x ∈ M. ∃ε>0. mcball x ε ⊆ S}"
  by (fastforce simp: mball_iff_mcball metric_interior_of)
lemma in_interior_of_mball:
   "x ∈ mtopology interior_of S ⟷ x ∈ M ∧ (∃ε>0. mball x ε ⊆ S)"
  using metric_interior_of by force
lemma in_interior_of_mcball:
   "x ∈ mtopology interior_of S ⟷ x ∈ M ∧ (∃ε>0. mcball x ε ⊆ S)"
  using metric_interior_of_alt by force
lemma Hausdorff_space_mtopology: "Hausdorff_space mtopology"
  unfolding Hausdorff_space_def
proof clarify
  fix x y
  assume x: "x ∈ topspace mtopology" and y: "y ∈ topspace mtopology" and "x ≠ y"
  then have gt0: "d x y / 2 > 0"
    by auto
  have "disjnt (mball x (d x y / 2)) (mball y (d x y / 2))"
    by (simp add: disjoint_mball)
  then show "∃U V. openin mtopology U ∧ openin mtopology V ∧ x ∈ U ∧ y ∈ V ∧ disjnt U V"
    by (metis centre_in_mball_iff gt0 openin_mball topspace_mtopology x y)
qed
subsection‹Bounded sets›
definition mbounded where "mbounded S ⟷ (∃x B. S ⊆ mcball x B)"
lemma mbounded_pos: "mbounded S ⟷ (∃x B. 0 < B ∧ S ⊆ mcball x B)"
proof -
  have "∃x' r'. 0 < r' ∧ S ⊆ mcball x' r'" if "S ⊆ mcball x r" for x r
    by (metis gt_ex less_eq_real_def linorder_not_le mcball_subset_concentric order_trans that)
  then show ?thesis
    by (auto simp: mbounded_def)
qed
lemma mbounded_alt:
  "mbounded S ⟷ S ⊆ M ∧ (∃B. ∀x ∈ S. ∀y ∈ S. d x y ≤ B)"
proof -
  have "⋀x B. S ⊆ mcball x B ⟹ ∀x∈S. ∀y∈S. d x y ≤ 2 * B"
    by (smt (verit, best) commute in_mcball subsetD triangle)
  then show ?thesis
    unfolding mbounded_def by (metis in_mcball in_mono subsetI)
qed
lemma mbounded_alt_pos:
  "mbounded S ⟷ S ⊆ M ∧ (∃B>0. ∀x ∈ S. ∀y ∈ S. d x y ≤ B)"
  by (smt (verit, del_insts) gt_ex mbounded_alt)
lemma mbounded_subset: "⟦mbounded T; S ⊆ T⟧ ⟹ mbounded S"
  by (meson mbounded_def order_trans)
lemma mbounded_subset_mspace: "mbounded S ⟹ S ⊆ M"
  by (simp add: mbounded_alt)
lemma mbounded:
   "mbounded S ⟷ S = {} ∨ (∀x ∈ S. x ∈ M) ∧ (∃y B. y ∈ M ∧ (∀x ∈ S. d y x ≤ B))"
  by (meson all_not_in_conv in_mcball mbounded_def subset_iff)
lemma mbounded_empty [iff]: "mbounded {}"
  by (simp add: mbounded)
lemma mbounded_mcball: "mbounded (mcball x r)"
  using mbounded_def by auto
lemma mbounded_mball [iff]: "mbounded (mball x r)"
  by (meson mball_subset_mcball mbounded_def)
lemma mbounded_insert: "mbounded (insert a S) ⟷ a ∈ M ∧ mbounded S"
proof -
  have "⋀y B. ⟦y ∈ M; ∀x∈S. d y x ≤ B⟧
           ⟹ ∃y. y ∈ M ∧ (∃B ≥ d y a. ∀x∈S. d y x ≤ B)"
    by (metis order.trans nle_le)
  then show ?thesis
    by (auto simp: mbounded)
qed
lemma mbounded_Int: "mbounded S ⟹ mbounded (S ∩ T)"
  by (meson inf_le1 mbounded_subset)
lemma mbounded_Un: "mbounded (S ∪ T) ⟷ mbounded S ∧ mbounded T" (is "?lhs=?rhs")
proof
  assume R: ?rhs
  show ?lhs
  proof (cases "S={} ∨ T={}")
    case True then show ?thesis
      using R by auto
  next
    case False
    obtain x y B C where "S ⊆ mcball x B" "T ⊆ mcball y C" "B > 0" "C > 0" "x ∈ M" "y ∈ M"
      using R mbounded_pos
      by (metis False mcball_eq_empty subset_empty)
    then have "S ∪ T ⊆ mcball x (B + C + d x y)"
      by (smt (verit) commute dual_order.trans le_supI mcball_subset mdist_pos_eq)
    then show ?thesis
      using mbounded_def by blast
  qed
next
  show "?lhs ⟹ ?rhs"
    using mbounded_def by auto
qed
lemma mbounded_Union:
  "⟦finite ℱ; ⋀X. X ∈ ℱ ⟹ mbounded X⟧ ⟹ mbounded (⋃ℱ)"
  by (induction ℱ rule: finite_induct) (auto simp: mbounded_Un)
lemma mbounded_closure_of:
   "mbounded S ⟹ mbounded (mtopology closure_of S)"
  by (meson closedin_mcball closure_of_minimal mbounded_def)
lemma mbounded_closure_of_eq:
   "S ⊆ M ⟹ (mbounded (mtopology closure_of S) ⟷ mbounded S)"
  by (metis closure_of_subset mbounded_closure_of mbounded_subset topspace_mtopology)
lemma maxdist_thm:
  assumes "mbounded S"
      and "x ∈ S"
      and "y ∈ S"
    shows  "d x y = (SUP z∈S. ¦d x z - d z y¦)"
proof -
  have "¦d x z - d z y¦ ≤ d x y" if "z ∈ S" for z
    by (metis all_not_in_conv assms mbounded mdist_reverse_triangle that) 
  moreover have "d x y ≤ r"
    if "⋀z. z ∈ S ⟹ ¦d x z - d z y¦ ≤ r" for r :: real
    using that assms mbounded_subset_mspace mdist_zero by fastforce
  ultimately show ?thesis
    by (intro cSup_eq [symmetric]) auto
qed
lemma metric_eq_thm: "⟦S ⊆ M; x ∈ S; y ∈ S⟧ ⟹ (x = y) = (∀z∈S. d x z = d y z)"
  by (metis commute  subset_iff zero)
lemma compactin_imp_mbounded:
  assumes "compactin mtopology S"
  shows "mbounded S"
proof -
  have "S ⊆ M"
    and com: "⋀𝒰. ⟦∀U∈𝒰. openin mtopology U; S ⊆ ⋃𝒰⟧ ⟹ ∃ℱ. finite ℱ ∧ ℱ ⊆ 𝒰 ∧ S ⊆ ⋃ℱ"
    using assms by (auto simp: compactin_def mbounded_def)
  show ?thesis
  proof (cases "S = {}")
    case False
    with ‹S ⊆ M› obtain a where "a ∈ S" "a ∈ M"
      by blast
    with ‹S ⊆ M› gt_ex have "S ⊆ ⋃(range (mball a))"
      by force
    then obtain ℱ where "finite ℱ" "ℱ ⊆ range (mball a)" "S ⊆ ⋃ℱ"
      by (metis (no_types, opaque_lifting) com imageE openin_mball)
  then show ?thesis
      using mbounded_Union mbounded_subset by fastforce
  qed auto
qed
end 
lemma mcball_eq_cball [simp]: "Met_TC.mcball = cball"
  by force
lemma mball_eq_ball [simp]: "Met_TC.mball = ball"
  by force
lemma mopen_eq_open [simp]: "Met_TC.mopen = open"
  by (force simp: open_contains_ball Met_TC.mopen_def)
lemma limitin_iff_tendsto [iff]: "limitin Met_TC.mtopology σ x F = tendsto σ x F"
  by (simp add: Met_TC.mtopology_def)
lemma mtopology_is_euclidean [simp]: "Met_TC.mtopology = euclidean"
  by (simp add: Met_TC.mtopology_def)
lemma mbounded_iff_bounded [iff]: "Met_TC.mbounded A ⟷ bounded A"
  by (metis Met_TC.mbounded UNIV_I all_not_in_conv bounded_def)
subsection‹Subspace of a metric space›
locale Submetric = Metric_space + 
  fixes A
  assumes subset: "A ⊆ M"
sublocale Submetric ⊆ sub: Metric_space A d
  by (simp add: subset subspace)
context Submetric
begin 
lemma mball_submetric_eq: "sub.mball a r = (if a ∈ A then A ∩ mball a r else {})"
and mcball_submetric_eq: "sub.mcball a r = (if a ∈ A then A ∩ mcball a r else {})"
  using subset by force+
lemma mtopology_submetric: "sub.mtopology = subtopology mtopology A"
  unfolding topology_eq
proof (intro allI iffI)
  fix S
  assume "openin sub.mtopology S"
  then have "∃T. openin (subtopology mtopology A) T ∧ x ∈ T ∧ T ⊆ S" if "x ∈ S" for x
    by (metis mball_submetric_eq openin_mball openin_subtopology_Int2 sub.centre_in_mball_iff sub.openin_mtopology subsetD that)
  then show "openin (subtopology mtopology A) S"
    by (meson openin_subopen)
next
  fix S
  assume "openin (subtopology mtopology A) S"
  then obtain T where "openin mtopology T" "S = T ∩ A"
    by (meson openin_subtopology)
  then have "mopen T"
    by (simp add: mopen_def openin_mtopology)
  then have "sub.mopen (T ∩ A)"
    unfolding sub.mopen_def mopen_def
    by (metis inf.coboundedI2 mball_submetric_eq Int_iff ‹S = T ∩ A› inf.bounded_iff subsetI)
  then show "openin sub.mtopology S"
    using ‹S = T ∩ A› sub.mopen_def sub.openin_mtopology by force
qed
lemma mbounded_submetric: "sub.mbounded T ⟷ mbounded T ∧ T ⊆ A"
  by (meson mbounded_alt sub.mbounded_alt subset subset_trans)
end
lemma (in Metric_space) submetric_empty [iff]: "Submetric M d {}"
proof qed auto
subsection ‹Abstract type of metric spaces›
typedef 'a metric = "{(M::'a set,d). Metric_space M d}"
  morphisms "dest_metric" "metric"
proof -
  have "Metric_space {} (λx y. 0)"
    by (auto simp: Metric_space_def)
  then show ?thesis
    by blast
qed
definition mspace where "mspace m ≡ fst (dest_metric m)"
definition mdist where "mdist m ≡ snd (dest_metric m)"
lemma Metric_space_mspace_mdist [iff]: "Metric_space (mspace m) (mdist m)"
  by (metis Product_Type.Collect_case_prodD dest_metric mdist_def mspace_def)
lemma mdist_nonneg [simp]: "⋀x y. 0 ≤ mdist m x y"
  by (metis Metric_space_def Metric_space_mspace_mdist)
lemma mdist_commute: "⋀x y. mdist m x y = mdist m y x"
  by (metis Metric_space_def Metric_space_mspace_mdist)
lemma mdist_zero [simp]: "⋀x y. ⟦x ∈ mspace m; y ∈ mspace m⟧ ⟹ mdist m x y = 0 ⟷ x=y"
  by (meson Metric_space.zero Metric_space_mspace_mdist)
lemma mdist_triangle: "⋀x y z. ⟦x ∈ mspace m; y ∈ mspace m; z ∈ mspace m⟧ ⟹ mdist m x z ≤ mdist m x y + mdist m y z"
  by (meson Metric_space.triangle Metric_space_mspace_mdist)
lemma (in Metric_space) mspace_metric[simp]: 
  "mspace (metric (M,d)) = M"
  by (simp add: metric_inverse mspace_def subspace)
lemma (in Metric_space) mdist_metric[simp]: 
  "mdist (metric (M,d)) = d"
  by (simp add: mdist_def metric_inverse subspace)
lemma metric_collapse [simp]: "metric (mspace m, mdist m) = m"
  by (simp add: dest_metric_inverse mdist_def mspace_def)
definition mtopology_of :: "'a metric ⇒ 'a topology"
  where "mtopology_of ≡ λm. Metric_space.mtopology (mspace m) (mdist m)"
lemma topspace_mtopology_of [simp]: "topspace (mtopology_of m) = mspace m"
  by (simp add: Metric_space.topspace_mtopology Metric_space_mspace_mdist mtopology_of_def)
lemma (in Metric_space) mtopology_of [simp]:
  "mtopology_of (metric (M,d)) = mtopology"
  by (simp add: mtopology_of_def)
definition "mball_of ≡ λm. Metric_space.mball (mspace m) (mdist m)"
lemma in_mball_of [simp]: "y ∈ mball_of m x r ⟷ x ∈ mspace m ∧ y ∈ mspace m ∧ mdist m x y < r"
  by (simp add: Metric_space.in_mball mball_of_def)
lemma (in Metric_space) mball_of [simp]:
  "mball_of (metric (M,d)) = mball"
  by (simp add: mball_of_def)
definition "mcball_of ≡ λm. Metric_space.mcball (mspace m) (mdist m)"
lemma in_mcball_of [simp]: "y ∈ mcball_of m x r ⟷ x ∈ mspace m ∧ y ∈ mspace m ∧ mdist m x y ≤ r"
  by (simp add: Metric_space.in_mcball mcball_of_def)
lemma (in Metric_space) mcball_of [simp]:
  "mcball_of (metric (M,d)) = mcball"
  by (simp add: mcball_of_def)
definition "euclidean_metric ≡ metric (UNIV,dist)"
lemma mspace_euclidean_metric [simp]: "mspace euclidean_metric = UNIV"
  by (simp add: euclidean_metric_def)
lemma mdist_euclidean_metric [simp]: "mdist euclidean_metric = dist"
  by (simp add: euclidean_metric_def)
lemma mtopology_of_euclidean [simp]: "mtopology_of euclidean_metric = euclidean"
  by (simp add: Met_TC.mtopology_def mtopology_of_def)
text ‹Allows reference to the current metric space within the locale as a value›
definition (in Metric_space) "Self ≡ metric (M,d)"
lemma (in Metric_space) mspace_Self [simp]: "mspace Self = M"
  by (simp add: Self_def)
lemma (in Metric_space) mdist_Self [simp]: "mdist Self = d"
  by (simp add: Self_def)
text‹ Subspace of a metric space›
definition submetric where
  "submetric ≡ λm S. metric (S ∩ mspace m, mdist m)"
lemma mspace_submetric [simp]: "mspace (submetric m S) = S ∩ mspace m" 
  unfolding submetric_def
  by (meson Metric_space.subspace inf_le2 Metric_space_mspace_mdist Metric_space.mspace_metric)
lemma mdist_submetric [simp]: "mdist (submetric m S) = mdist m"
  unfolding submetric_def
  by (meson Metric_space.subspace inf_le2 Metric_space.mdist_metric Metric_space_mspace_mdist)
lemma submetric_UNIV [simp]: "submetric m UNIV = m"
  by (simp add: submetric_def dest_metric_inverse mdist_def mspace_def)
lemma submetric_submetric [simp]:
   "submetric (submetric m S) T = submetric m (S ∩ T)"
  by (metis submetric_def Int_assoc inf_commute mdist_submetric mspace_submetric)
lemma submetric_mspace [simp]:
   "submetric m (mspace m) = m"
  by (simp add: submetric_def dest_metric_inverse mdist_def mspace_def)
lemma submetric_restrict:
   "submetric m S = submetric m (mspace m ∩ S)"
  by (metis submetric_mspace submetric_submetric)
lemma mtopology_of_submetric: "mtopology_of (submetric m A) = subtopology (mtopology_of m) A"
proof -
  interpret Submetric "mspace m" "mdist m" "A ∩ mspace m"
    using Metric_space_mspace_mdist Submetric.intro Submetric_axioms.intro inf_le2 by blast
  have "sub.mtopology = subtopology (mtopology_of m) A"
    by (metis inf_commute mtopology_of_def mtopology_submetric subtopology_mspace subtopology_subtopology)
  then show ?thesis
    by (simp add: submetric_def)
qed
subsection‹The discrete metric›
locale discrete_metric =
  fixes M :: "'a set"
definition (in discrete_metric) dd :: "'a ⇒ 'a ⇒ real"
  where "dd ≡ λx y::'a. if x=y then 0 else 1"
lemma metric_M_dd: "Metric_space M discrete_metric.dd"
  by (simp add: discrete_metric.dd_def Metric_space.intro)
sublocale discrete_metric ⊆ disc: Metric_space M dd
  by (simp add: metric_M_dd)
lemma (in discrete_metric) mopen_singleton:
  assumes "x ∈ M" shows "disc.mopen {x}"
proof -
  have "disc.mball x (1/2) ⊆ {x}"
    by (smt (verit) dd_def disc.in_mball less_divide_eq_1_pos singleton_iff subsetI)
  with assms show ?thesis
    using disc.mopen_def half_gt_zero_iff zero_less_one by blast
qed
lemma (in discrete_metric) mtopology_discrete_metric:
   "disc.mtopology = discrete_topology M"
proof -
  have "⋀x. x ∈ M ⟹ openin disc.mtopology {x}"
    by (simp add: disc.mtopology_def mopen_singleton)
  then show ?thesis
    by (metis disc.topspace_mtopology discrete_topology_unique)
qed
lemma (in discrete_metric) discrete_ultrametric:
   "dd x z ≤ max (dd x y) (dd y z)"
  by (simp add: dd_def)
lemma (in discrete_metric) dd_le1: "dd x y ≤ 1"
  by (simp add: dd_def)
lemma (in discrete_metric) mbounded_discrete_metric: "disc.mbounded S ⟷ S ⊆ M"
  by (meson dd_le1 disc.mbounded_alt)
subsection‹Metrizable spaces›
definition metrizable_space where
  "metrizable_space X ≡ ∃M d. Metric_space M d ∧ X = Metric_space.mtopology M d"
lemma (in Metric_space) metrizable_space_mtopology: "metrizable_space mtopology"
  using local.Metric_space_axioms metrizable_space_def by blast
lemma (in Metric_space) first_countable_mtopology: "first_countable mtopology"
proof (clarsimp simp add: first_countable_def)
  fix x
  assume "x ∈ M"
  define ℬ where "ℬ ≡ mball x ` {r ∈ ℚ. 0 < r}"
  show "∃ℬ. countable ℬ ∧ (∀V∈ℬ. openin mtopology V) ∧ (∀U. openin mtopology U ∧ x ∈ U ⟶ (∃V∈ℬ. x ∈ V ∧ V ⊆ U))"
  proof (intro exI conjI ballI)
    show "countable ℬ"
      by (simp add: ℬ_def countable_rat)
    show "∀U. openin mtopology U ∧ x ∈ U ⟶ (∃V∈ℬ. x ∈ V ∧ V ⊆ U)"
    proof clarify
      fix U
      assume "openin mtopology U" and "x ∈ U"
      then obtain r where "r>0" and r: "mball x r ⊆ U"
        by (meson openin_mtopology)
      then obtain q where "q ∈ Rats" "0 < q" "q < r"
        using Rats_dense_in_real by blast
      then show "∃V∈ℬ. x ∈ V ∧ V ⊆ U"
        unfolding ℬ_def using ‹x ∈ M› r by fastforce
    qed
  qed (auto simp: ℬ_def)
qed
lemma metrizable_imp_first_countable:
   "metrizable_space X ⟹ first_countable X"
  by (force simp: metrizable_space_def Metric_space.first_countable_mtopology)
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 [simp]:
   "metrizable_space(discrete_topology U)"
  by (metis discrete_metric.mtopology_discrete_metric metric_M_dd metrizable_space_def)
lemma empty_metrizable_space: "metrizable_space trivial_topology"
  by simp
lemma metrizable_space_subtopology:
  assumes "metrizable_space X"
  shows "metrizable_space(subtopology X S)"
proof -
  obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d"
    using assms metrizable_space_def by blast
  then interpret Submetric M d "M ∩ S"
    by (simp add: Submetric.intro Submetric_axioms_def)
  show ?thesis
    unfolding metrizable_space_def
    by (metis X mtopology_submetric sub.Metric_space_axioms subtopology_restrict topspace_mtopology)
qed
lemma homeomorphic_metrizable_space_aux:
  assumes "X homeomorphic_space Y" "metrizable_space X"
  shows "metrizable_space Y"
proof -
  obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d"
    using assms by (auto simp: metrizable_space_def)
  then interpret m: Metric_space M d 
    by simp
  obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g"
    and fg: "(∀x ∈ M. g(f x) = x) ∧ (∀y ∈ topspace Y. f(g y) = y)"
    using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce
  define d' where "d' x y ≡ d (g x) (g y)" for x y
  interpret m': Metric_space "topspace Y" "d'"
    unfolding d'_def
  proof
    show "(d (g x) (g y) = 0) = (x = y)" if "x ∈ topspace Y" "y ∈ topspace Y" for x y
      by (metis fg X hmg homeomorphic_imp_surjective_map imageI m.topspace_mtopology m.zero that)
    show "d (g x) (g z) ≤ d (g x) (g y) + d (g y) (g z)"
      if "x ∈ topspace Y" and "y ∈ topspace Y" and "z ∈ topspace Y" for x y z
      by (metis X that hmg homeomorphic_eq_everything_map imageI m.topspace_mtopology m.triangle)
  qed (auto simp: m.nonneg m.commute)
  have "Y = Metric_space.mtopology (topspace Y) d'"
    unfolding topology_eq
  proof (intro allI)
    fix S
    have "openin m'.mtopology S" if S: "S ⊆ topspace Y" and "openin X (g ` S)"
      unfolding m'.openin_mtopology
    proof (intro conjI that strip)
      fix y
      assume "y ∈ S"
      then obtain r where "r>0" and r: "m.mball (g y) r ⊆ g ` S" 
        using X ‹openin X (g ` S)› m.openin_mtopology using ‹y ∈ S› by auto
      then have "g ` m'.mball y r ⊆ m.mball (g y) r"
        using X d'_def hmg homeomorphic_imp_surjective_map by fastforce
      with S fg have "m'.mball y r ⊆ S"
        by (smt (verit, del_insts) image_iff m'.in_mball r subset_iff)
      then show "∃r>0. m'.mball y r ⊆ S"
        using ‹0 < r› by blast 
    qed
    moreover have "openin X (g ` S)" if ope': "openin m'.mtopology S"
    proof -
      have "∃r>0. m.mball (g y) r ⊆ g ` S" if "y ∈ S" for y
      proof -
        have y: "y ∈ topspace Y"
          using m'.openin_mtopology ope' that by blast
        obtain r where "r > 0" and r: "m'.mball y r ⊆ S"
          using ope' by (meson ‹y ∈ S› m'.openin_mtopology)
        moreover have "⋀x. ⟦x ∈ M; d (g y) x < r⟧ ⟹ ∃u. u ∈ topspace Y ∧ d' y u < r ∧ x = g u"
          using fg X d'_def hmf homeomorphic_imp_surjective_map by fastforce
        ultimately have "m.mball (g y) r ⊆ g ` m'.mball y r"
          using y by (force simp: m'.openin_mtopology)
        then show ?thesis
          using ‹0 < r› r by blast
      qed
      then show ?thesis
        using X hmg homeomorphic_imp_surjective_map m.openin_mtopology ope' openin_subset by fastforce
    qed
    ultimately have "(S ⊆ topspace Y ∧ openin X (g ` S)) = openin m'.mtopology S"
      using m'.topspace_mtopology openin_subset by blast
    then show "openin Y S = openin m'.mtopology S"
      by (simp add: m'.mopen_def homeomorphic_map_openness_eq [OF hmg])
  qed
  then show ?thesis
    using m'.metrizable_space_mtopology by force
qed
lemma homeomorphic_metrizable_space:
  assumes "X homeomorphic_space Y"
  shows "metrizable_space X ⟷ metrizable_space Y"
  using assms homeomorphic_metrizable_space_aux homeomorphic_space_sym by metis
lemma metrizable_space_retraction_map_image:
   "retraction_map X Y r ∧ metrizable_space X
        ⟹ metrizable_space Y"
  using hereditary_imp_retractive_property metrizable_space_subtopology homeomorphic_metrizable_space
  by blast
lemma metrizable_imp_Hausdorff_space:
   "metrizable_space X ⟹ Hausdorff_space X"
  by (metis Metric_space.Hausdorff_space_mtopology metrizable_space_def)
lemma metrizable_imp_t1_space:
   "metrizable_space X ⟹ t1_space X"
  by (simp add: Hausdorff_imp_t1_space metrizable_imp_Hausdorff_space)
lemma closed_imp_gdelta_in:
  assumes X: "metrizable_space X" and S: "closedin X S"
  shows "gdelta_in X S"
proof -
  obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d"
    using X metrizable_space_def by blast
  then interpret M: Metric_space M d
    by blast
  have "S ⊆ M"
    using M.closedin_metric ‹X = M.mtopology› S by blast
  show ?thesis
  proof (cases "S = {}")
    case True
    then show ?thesis
      by simp
  next
    case False
    have "∃y∈S. d x y < inverse (1 + real n)" if "x ∈ S" for x n
      using ‹S ⊆ M› M.mdist_zero [of x] that by force
    moreover
    have "x ∈ S" if "x ∈ M" and §: "⋀n. ∃y∈S. d x y < inverse(Suc n)" for x
    proof -
      have *: "∃y∈S. d x y < ε" if "ε > 0" for ε
        by (metis § that not0_implies_Suc order_less_le order_less_le_trans real_arch_inverse)
      have "closedin M.mtopology S"
        using S by (simp add: Xeq)
      with * ‹x ∈ M› show ?thesis
        by (force simp: M.closedin_metric disjnt_iff)
    qed
    ultimately have Seq: "S = ⋂(range (λn. {x∈M. ∃y∈S. d x y < inverse(Suc n)}))"
      using ‹S ⊆ M› by force
    have "openin M.mtopology {xa ∈ M. ∃y∈S. d xa y < inverse (1 + real n)}" for n
    proof (clarsimp simp: M.openin_mtopology)
      fix x y
      assume "x ∈ M" "y ∈ S" and dxy: "d x y < inverse (1 + real n)"
      then have "⋀z. ⟦z ∈ M; d x z < inverse (1 + real n) - d x y⟧ ⟹ ∃y∈S. d z y < inverse (1 + real n)"
        by (smt (verit) M.commute M.triangle ‹S ⊆ M› in_mono)
      with dxy show "∃r>0. M.mball x r ⊆ {z ∈ M. ∃y∈S. d z y < inverse (1 + real n)}"
        by (rule_tac x="inverse(Suc n) - d x y" in exI) auto
    qed
    then have "gdelta_in X (⋂(range (λn. {x∈M. ∃y∈S. d x y < inverse(Suc n)})))"
      by (force simp: Xeq intro: gdelta_in_Inter open_imp_gdelta_in)
    with Seq show ?thesis
      by presburger
  qed
qed
lemma open_imp_fsigma_in:
   "⟦metrizable_space X; openin X S⟧ ⟹ fsigma_in X S"
  by (meson closed_imp_gdelta_in fsigma_in_gdelta_in openin_closedin openin_subset)
lemma metrizable_space_euclidean:
  "metrizable_space (euclidean :: 'a::metric_space topology)"
  using Met_TC.metrizable_space_mtopology by auto
lemma (in Metric_space) regular_space_mtopology:
   "regular_space mtopology"
unfolding regular_space_def
proof clarify
  fix C a
  assume C: "closedin mtopology C" and a: "a ∈ topspace mtopology" and "a ∉ C"
  have "openin mtopology (topspace mtopology - C)"
    by (simp add: C openin_diff)
  then obtain r where "r>0" and r: "mball a r ⊆ topspace mtopology - C"
    unfolding openin_mtopology using ‹a ∉ C› a by auto
  show "∃U V. openin mtopology U ∧ openin mtopology V ∧ a ∈ U ∧ C ⊆ V ∧ disjnt U V"
  proof (intro exI conjI)
    show "a ∈ mball a (r/2)"
      using ‹0 < r› a by force
    show "C ⊆ topspace mtopology - mcball a (r/2)"
      using C ‹0 < r› r by (fastforce simp: closedin_metric)
  qed (auto simp: openin_mball closedin_mcball openin_diff disjnt_iff)
qed
lemma metrizable_imp_regular_space:
   "metrizable_space X ⟹ regular_space X"
  by (metis Metric_space.regular_space_mtopology metrizable_space_def)
lemma regular_space_euclidean:
 "regular_space (euclidean :: 'a::metric_space topology)"
  by (simp add: metrizable_imp_regular_space metrizable_space_euclidean)
subsection‹Limits at a point in a topological space›
lemma (in Metric_space) eventually_atin_metric:
   "eventually P (atin mtopology a) ⟷
        (a ∈ M ⟶ (∃δ>0. ∀x. x ∈ M ∧ 0 < d x a ∧ d x a < δ ⟶ P x))"  (is "?lhs=?rhs")
proof (cases "a ∈ M")
  case True
  show ?thesis
  proof
    assume L: ?lhs 
    with True obtain U where "openin mtopology U" "a ∈ U" and U: "∀x∈U - {a}. P x"
      by (auto simp: eventually_atin)
    then obtain r where "r>0" and "mball a r ⊆ U"
      by (meson openin_mtopology)
    with U show ?rhs
      by (smt (verit, ccfv_SIG) commute in_mball insert_Diff_single insert_iff subset_iff)
  next
    assume ?rhs 
    then obtain δ where "δ>0" and δ: "∀x. x ∈ M ∧ 0 < d x a ∧ d x a < δ ⟶ P x"
      using True by blast
    then have "∀x ∈ mball a δ - {a}. P x"
      by (simp add: commute)
    then show ?lhs
      unfolding eventually_atin openin_mtopology
      by (metis True ‹0 < δ› centre_in_mball_iff openin_mball openin_mtopology) 
  qed
qed auto
subsection ‹Normal spaces and metric spaces›
lemma (in Metric_space) normal_space_mtopology:
   "normal_space mtopology"
  unfolding normal_space_def
proof clarify
  fix S T
  assume "closedin mtopology S"
  then have "⋀x. x ∈ M - S ⟹ (∃r>0. mball x r ⊆ M - S)"
    by (simp add: closedin_def openin_mtopology)
  then obtain δ where d0: "⋀x. x ∈ M - S ⟹ δ x > 0 ∧ mball x (δ x) ⊆ M - S"
    by metis
  assume "closedin mtopology T"
  then have "⋀x. x ∈ M - T ⟹ (∃r>0. mball x r ⊆ M - T)"
    by (simp add: closedin_def openin_mtopology)
  then obtain ε where e: "⋀x. x ∈ M - T ⟹ ε x > 0 ∧ mball x (ε x) ⊆ M - T"
    by metis
  assume "disjnt S T"
  have "S ⊆ M" "T ⊆ M"
    using ‹closedin mtopology S› ‹closedin mtopology T› closedin_metric by blast+
  have δ: "⋀x. x ∈ T ⟹ δ x > 0 ∧ mball x (δ x) ⊆ M - S"
    by (meson DiffI ‹T ⊆ M› ‹disjnt S T› d0 disjnt_iff subsetD)
  have ε: "⋀x. x ∈ S ⟹ ε x > 0 ∧ mball x (ε x) ⊆ M - T"
    by (meson Diff_iff ‹S ⊆ M› ‹disjnt S T› disjnt_iff e subsetD)
  show "∃U V. openin mtopology U ∧ openin mtopology V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
  proof (intro exI conjI)
    show "openin mtopology (⋃x∈S. mball x (ε x / 2))" "openin mtopology (⋃x∈T. mball x (δ x / 2))"
      by force+
    show "S ⊆ (⋃x∈S. mball x (ε x / 2))"
      using ε ‹S ⊆ M› by force
    show "T ⊆ (⋃x∈T. mball x (δ x / 2))"
      using δ ‹T ⊆ M› by force
    show "disjnt (⋃x∈S. mball x (ε x / 2)) (⋃x∈T. mball x (δ x / 2))"
      using ε δ 
      apply (clarsimp simp: disjnt_iff subset_iff)
      by (smt (verit, ccfv_SIG) field_sum_of_halves triangle')
  qed 
qed
lemma metrizable_imp_normal_space:
   "metrizable_space X ⟹ normal_space X"
  by (metis Metric_space.normal_space_mtopology metrizable_space_def)
subsection‹Topological limitin in metric spaces›
lemma (in Metric_space) limitin_mspace:
   "limitin mtopology f l F ⟹ l ∈ M"
  using limitin_topspace by fastforce
lemma (in Metric_space) limitin_metric_unique:
   "⟦limitin mtopology f l1 F; limitin mtopology f l2 F; F ≠ bot⟧ ⟹ l1 = l2"
  by (meson Hausdorff_space_mtopology limitin_Hausdorff_unique)
lemma (in Metric_space) limitin_metric:
   "limitin mtopology f l F ⟷ l ∈ M ∧ (∀ε>0. eventually (λx. f x ∈ M ∧ d (f x) l < ε) F)"  
   (is "?lhs=?rhs")
proof
  assume L: ?lhs
  show ?rhs
    unfolding limitin_def
  proof (intro conjI strip)
    show "l ∈ M"
      using L limitin_mspace by blast
    fix ε::real
    assume "ε>0"
    then have "∀⇩F x in F. f x ∈ mball l ε"
      using L openin_mball by (fastforce simp: limitin_def)
    then show "∀⇩F x in F. f x ∈ M ∧ d (f x) l < ε"
      using commute eventually_mono by fastforce
  qed
next
  assume R: ?rhs 
  then show ?lhs
    by (force simp: limitin_def commute openin_mtopology subset_eq elim: eventually_mono)
qed
lemma (in Metric_space) limit_metric_sequentially:
   "limitin mtopology f l sequentially ⟷
     l ∈ M ∧ (∀ε>0. ∃N. ∀n≥N. f n ∈ M ∧ d (f n) l < ε)"
  by (auto simp: limitin_metric eventually_sequentially)
lemma (in Submetric) limitin_submetric_iff:
   "limitin sub.mtopology f l F ⟷
     l ∈ A ∧ eventually (λx. f x ∈ A) F ∧ limitin mtopology f l F" (is "?lhs=?rhs")
  by (simp add: limitin_subtopology mtopology_submetric)
lemma (in Metric_space) metric_closedin_iff_sequentially_closed:
   "closedin mtopology S ⟷
    S ⊆ M ∧ (∀σ l. range σ ⊆ S ∧ limitin mtopology σ l sequentially ⟶ l ∈ S)" (is "?lhs=?rhs")
proof
  assume ?lhs then show ?rhs
    by (force simp: closedin_metric limitin_closedin range_subsetD)
next
  assume R: ?rhs
  show ?lhs
    unfolding closedin_metric
  proof (intro conjI strip)
    show "S ⊆ M"
      using R by blast
    fix x
    assume "x ∈ M - S"
    have False if "∀r>0. ∃y. y ∈ M ∧ y ∈ S ∧ d x y < r"
    proof -
      have "∀n. ∃y. y ∈ M ∧ y ∈ S ∧ d x y < inverse(Suc n)"
        using that by auto
      then obtain σ where σ: "⋀n. σ n ∈ M ∧ σ n ∈ S ∧ d x (σ n) < inverse(Suc n)"
        by metis
      then have "range σ ⊆ M"
        by blast
      have "∃N. ∀n≥N. d x (σ n) < ε" if "ε>0" for ε
      proof -
        have "real (Suc (nat ⌈inverse ε⌉)) ≥ inverse ε"
          by linarith
        then have "∀n ≥ nat ⌈inverse ε⌉. d x (σ n) < ε"
          by (metis σ inverse_inverse_eq inverse_le_imp_le nat_ceiling_le_eq nle_le not_less_eq_eq order.strict_trans2 that)
        then show ?thesis ..
      qed
      with σ have "limitin mtopology σ x sequentially"
        using ‹x ∈ M - S› commute limit_metric_sequentially by auto
      then show ?thesis
        by (metis R DiffD2 σ image_subset_iff ‹x ∈ M - S›)
    qed
    then show "∃r>0. disjnt S (mball x r)"
      by (meson disjnt_iff in_mball)
  qed
qed
lemma (in Metric_space) limit_atin_metric:
   "limitin X f y (atin mtopology x) ⟷
      y ∈ topspace X ∧
      (x ∈ M
       ⟶ (∀V. openin X V ∧ y ∈ V
               ⟶ (∃δ>0.  ∀x'. x' ∈ M ∧ 0 < d x' x ∧ d x' x < δ ⟶ f x' ∈ V)))"
  by (force simp: limitin_def eventually_atin_metric)
lemma (in Metric_space) limitin_metric_dist_null:
   "limitin mtopology f l F ⟷ l ∈ M ∧ eventually (λx. f x ∈ M) F ∧ ((λx. d (f x) l) ⤏ 0) F"
  by (simp add: limitin_metric tendsto_iff eventually_conj_iff all_conj_distrib imp_conjR gt_ex)
subsection‹Cauchy sequences and complete metric spaces›
context Metric_space
begin
definition MCauchy :: "(nat ⇒ 'a) ⇒ bool"
  where "MCauchy σ ≡ range σ ⊆ M ∧ (∀ε>0. ∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε)"
definition mcomplete
  where "mcomplete ≡ (∀σ. MCauchy σ ⟶ (∃x. limitin mtopology σ x sequentially))"
lemma mcomplete_empty [iff]: "Metric_space.mcomplete {} d"
  by (simp add: Metric_space.MCauchy_def Metric_space.mcomplete_def subspace)
lemma MCauchy_imp_MCauchy_suffix: "MCauchy σ ⟹ MCauchy (σ ∘ (+)n)"
  unfolding MCauchy_def image_subset_iff comp_apply
  by (metis UNIV_I add.commute trans_le_add1) 
lemma mcomplete:
   "mcomplete ⟷
    (∀σ. (∀⇩F n in sequentially. σ n ∈ M) ∧
     (∀ε>0. ∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε) ⟶
     (∃x. limitin mtopology σ x sequentially))" (is "?lhs=?rhs")
proof
  assume L: ?lhs 
  show ?rhs
  proof clarify
    fix σ
    assume "∀⇩F n in sequentially. σ n ∈ M"
      and σ: "∀ε>0. ∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
    then obtain N where "⋀n. n≥N ⟹ σ n ∈ M"
      by (auto simp: eventually_sequentially)
    with σ have "MCauchy (σ ∘ (+)N)"
      unfolding MCauchy_def image_subset_iff comp_apply by (meson le_add1 trans_le_add2)
    then obtain x where "limitin mtopology (σ ∘ (+)N) x sequentially"
      using L MCauchy_imp_MCauchy_suffix mcomplete_def by blast
    then have "limitin mtopology σ x sequentially"
      unfolding o_def by (auto simp: add.commute limitin_sequentially_offset_rev)
    then show "∃x. limitin mtopology σ x sequentially" ..
  qed
qed (simp add: mcomplete_def MCauchy_def image_subset_iff)
lemma mcomplete_empty_mspace: "M = {} ⟹ mcomplete"
  using MCauchy_def mcomplete_def by blast
lemma MCauchy_const [simp]: "MCauchy (λn. a) ⟷ a ∈ M"
  using MCauchy_def mdist_zero by auto
lemma convergent_imp_MCauchy:
  assumes "range σ ⊆ M" and lim: "limitin mtopology σ l sequentially"
  shows "MCauchy σ"
  unfolding MCauchy_def image_subset_iff
proof (intro conjI strip)
  fix ε::real
  assume "ε > 0"
  then have "∀⇩F n in sequentially. σ n ∈ M ∧ d (σ n) l < ε/2"
    using half_gt_zero lim limitin_metric by blast
  then obtain N where "⋀n. n≥N ⟹ σ n ∈ M ∧ d (σ n) l < ε/2"
    by (force simp: eventually_sequentially)
  then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
    by (smt (verit) limitin_mspace mdist_reverse_triangle field_sum_of_halves lim)
qed (use assms in blast)
lemma mcomplete_alt:
   "mcomplete ⟷ (∀σ. MCauchy σ ⟷ range σ ⊆ M ∧ (∃x. limitin mtopology σ x sequentially))"
  using MCauchy_def convergent_imp_MCauchy mcomplete_def by blast
lemma MCauchy_subsequence:
  assumes "strict_mono r" "MCauchy σ"
  shows "MCauchy (σ ∘ r)"
proof -
  have "d (σ (r n)) (σ (r n')) < ε"
    if "N ≤ n" "N ≤ n'" "strict_mono r" "∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
    for ε N n n'
    using that by (meson le_trans strict_mono_imp_increasing)
  moreover have "range (λx. σ (r x)) ⊆ M"
    using MCauchy_def assms by blast
  ultimately show ?thesis
    using assms by (simp add: MCauchy_def) metis
qed
lemma MCauchy_offset:
  assumes cau: "MCauchy (σ ∘ (+)k)" and σ: "⋀n. n < k ⟹ σ n ∈ M" 
  shows "MCauchy σ"
  unfolding MCauchy_def image_subset_iff
proof (intro conjI strip)
  fix n
  show "σ n ∈ M"
    using assms
    unfolding MCauchy_def image_subset_iff
    by (metis UNIV_I comp_apply le_iff_add linorder_not_le)
next
  fix ε :: real
  assume "ε > 0"
  obtain N where "∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d ((σ ∘ (+)k) n) ((σ ∘ (+)k) n') < ε"
    using cau ‹ε > 0› by (fastforce simp: MCauchy_def)
  then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
    unfolding o_def
    by (intro exI [where x="k+N"]) (smt (verit, del_insts) add.assoc le_add1 less_eqE)
qed
lemma MCauchy_convergent_subsequence:
  assumes cau: "MCauchy σ" and "strict_mono r" 
     and lim: "limitin mtopology (σ ∘ r) a sequentially"
  shows "limitin mtopology σ a sequentially"
  unfolding limitin_metric
proof (intro conjI strip)
  show "a ∈ M"
    by (meson assms limitin_mspace)
  fix ε :: real
  assume "ε > 0"
  then obtain N1 where N1: "⋀n n'. ⟦n≥N1; n'≥N1⟧ ⟹ d (σ n) (σ n') < ε/2"
    using cau unfolding MCauchy_def by (meson half_gt_zero)
  obtain N2 where N2: "⋀n. n ≥ N2 ⟹ (σ ∘ r) n ∈ M ∧ d ((σ ∘ r) n) a < ε/2"
    by (metis (no_types, lifting) lim ‹ε > 0› half_gt_zero limit_metric_sequentially)
  have "σ n ∈ M ∧ d (σ n) a < ε" if "n ≥ max N1 N2" for n
  proof (intro conjI)
    show "σ n ∈ M"
      using MCauchy_def cau by blast
    have "N1 ≤ r n"
      by (meson ‹strict_mono r› le_trans max.cobounded1 strict_mono_imp_increasing that)
    then show "d (σ n) a < ε"
      using N1[of n "r n"] N2[of n] ‹σ n ∈ M› ‹a ∈ M› triangle that by fastforce
  qed
  then show "∀⇩F n in sequentially. σ n ∈ M ∧ d (σ n) a < ε"
    using eventually_sequentially by blast
qed
lemma MCauchy_interleaving_gen:
  "MCauchy (λn. if even n then x(n div 2) else y(n div 2)) ⟷
    (MCauchy x ∧ MCauchy y ∧ (λn. d (x n) (y n)) ⇢ 0)" (is "?lhs=?rhs")
proof
  assume L: ?lhs
  have evens: "strict_mono (λn::nat. 2 * n)" and odds: "strict_mono (λn::nat. Suc (2 * n))"
    by (auto simp: strict_mono_def)
  show ?rhs
  proof (intro conjI)
    show "MCauchy x" "MCauchy y"
      using MCauchy_subsequence [OF evens L] MCauchy_subsequence [OF odds L] by (auto simp: o_def)
    show "(λn. d (x n) (y n)) ⇢ 0"
      unfolding LIMSEQ_iff
    proof (intro strip)
      fix ε :: real
      assume "ε > 0"
      then obtain N where N: 
        "⋀n n'. ⟦n≥N; n'≥N⟧ ⟹ d (if even n then x (n div 2) else y (n div 2))
                                   (if even n' then x (n' div 2) else y (n' div 2))  < ε"
        using L MCauchy_def by fastforce
      have "d (x n) (y n) < ε" if "n≥N" for n
        using N [of "2*n" "Suc(2*n)"] that by auto
      then show "∃N. ∀n≥N. norm (d (x n) (y n) - 0) < ε"
        by auto
    qed
  qed
next
  assume R: ?rhs
  show ?lhs
    unfolding MCauchy_def
  proof (intro conjI strip)
    show "range (λn. if even n then x (n div 2) else y (n div 2)) ⊆ M"
      using R by (auto simp: MCauchy_def)
    fix ε :: real
    assume "ε > 0"
    obtain Nx where Nx: "⋀n n'. ⟦n≥Nx; n'≥Nx⟧ ⟹ d (x n) (x n')  < ε/2"
      by (meson half_gt_zero MCauchy_def R ‹ε > 0›)
    obtain Ny where Ny: "⋀n n'. ⟦n≥Ny; n'≥Ny⟧ ⟹ d (y n) (y n')  < ε/2"
      by (meson half_gt_zero MCauchy_def R ‹ε > 0›)
    obtain Nxy where Nxy: "⋀n. n≥Nxy ⟹ d (x n) (y n) < ε/2"
      using R ‹ε > 0› half_gt_zero unfolding LIMSEQ_iff
      by (metis abs_mdist diff_zero real_norm_def)
    define N where "N ≡ 2 * Max{Nx,Ny,Nxy}"
    show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < ε"
    proof (intro exI strip)
      fix n n'
      assume "N ≤ n" and "N ≤ n'"
      then have "n div 2 ≥ Nx" "n div 2 ≥ Ny" "n div 2 ≥ Nxy" "n' div 2 ≥ Nx" "n' div 2 ≥ Ny" 
        by (auto simp: N_def)
      then have dxyn: "d (x (n div 2)) (y (n div 2)) < ε/2" 
            and dxnn': "d (x (n div 2)) (x (n' div 2)) < ε/2"
            and dynn': "d (y (n div 2)) (y (n' div 2)) < ε/2"
        using Nx Ny Nxy by blast+
      have inM: "x (n div 2) ∈ M" "x (n' div 2) ∈ M""y (n div 2) ∈ M" "y (n' div 2) ∈ M"
        using MCauchy_def R by blast+
      show "d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < ε"
      proof (cases "even n")
        case nt: True
        show ?thesis
        proof (cases "even n'")
          case True
          with ‹ε > 0› nt dxnn' show ?thesis by auto
        next
          case False
          with nt dxyn dynn' inM triangle show ?thesis
            by fastforce
        qed
      next
        case nf: False
        show ?thesis 
        proof (cases "even n'")
          case True
          then show ?thesis
            by (smt (verit) ‹ε > 0› dxyn dxnn' triangle commute inM field_sum_of_halves)
        next
          case False
          with ‹ε > 0› nf dynn' show ?thesis by auto
        qed
      qed
    qed
  qed
qed
lemma MCauchy_interleaving:
   "MCauchy (λn. if even n then σ(n div 2) else a) ⟷
    range σ ⊆ M ∧ limitin mtopology σ a sequentially"  (is "?lhs=?rhs")
proof -
  have "?lhs ⟷ (MCauchy σ ∧ a ∈ M ∧ (λn. d (σ n) a) ⇢ 0)"
    by (simp add: MCauchy_interleaving_gen [where y = "λn. a"])
  also have "... = ?rhs"
    by (metis MCauchy_def always_eventually convergent_imp_MCauchy limitin_metric_dist_null range_subsetD)
  finally show ?thesis .
qed
lemma mcomplete_nest:
   "mcomplete ⟷
      (∀C::nat ⇒'a set. (∀n. closedin mtopology (C n)) ∧
          (∀n. C n ≠ {}) ∧ decseq C ∧ (∀ε>0. ∃n a. C n ⊆ mcball a ε)
          ⟶ ⋂ (range C) ≠ {})" (is "?lhs=?rhs")
proof
  assume L: ?lhs 
  show ?rhs
    unfolding imp_conjL
  proof (intro strip)
    fix C :: "nat ⇒ 'a set"
    assume clo: "∀n. closedin mtopology (C n)"
      and ne: "∀n. C n ≠ ({}::'a set)"
      and dec: "decseq C"
      and cover [rule_format]: "∀ε>0. ∃n a. C n ⊆ mcball a ε"
    obtain σ where σ: "⋀n. σ n ∈ C n"
      by (meson ne empty_iff set_eq_iff)
    have "MCauchy σ"
      unfolding MCauchy_def
    proof (intro conjI strip)
      show "range σ ⊆ M"
        using σ clo metric_closedin_iff_sequentially_closed by auto 
      fix ε :: real
      assume "ε > 0"
      then obtain N a where N: "C N ⊆ mcball a (ε/3)"
        using cover by fastforce
      have "d (σ m) (σ n) < ε" if "N ≤ m" "N ≤ n" for m n
      proof -
        have "d a (σ m) ≤ ε/3" "d a (σ n) ≤ ε/3"
          using dec N σ that by (fastforce simp: decseq_def)+
        then have "d (σ m) (σ n) ≤ ε/3 + ε/3"
          using triangle σ commute dec decseq_def subsetD that N
          by (smt (verit, ccfv_threshold) in_mcball)
        also have "... < ε"
          using ‹ε > 0› by auto
        finally show ?thesis .
      qed
      then show "∃N. ∀m n. N ≤ m ⟶ N ≤ n ⟶ d (σ m) (σ n) < ε"
        by blast
    qed
    then obtain x where x: "limitin mtopology σ x sequentially"
      using L mcomplete_def by blast
    have "x ∈ C n" for n
    proof (rule limitin_closedin [OF x])
      show "closedin mtopology (C n)"
        by (simp add: clo)
      show "∀⇩F x in sequentially. σ x ∈ C n"
        by (metis σ dec decseq_def eventually_sequentiallyI subsetD)
    qed auto
    then show "⋂ (range C) ≠ {}"
      by blast
qed
next
  assume R: ?rhs  
  show ?lhs
    unfolding mcomplete_def
  proof (intro strip)
    fix σ
    assume "MCauchy σ"
    then have "range σ ⊆ M"
      using MCauchy_def by blast
    define C where "C ≡ λn. mtopology closure_of (σ ` {n..})"
    have "∀n. closedin mtopology (C n)" 
      by (auto simp: C_def)
    moreover
    have ne: "⋀n. C n ≠ {}"
      using ‹MCauchy σ› by (auto simp: C_def MCauchy_def disjnt_iff closure_of_eq_empty_gen)
    moreover
    have dec: "decseq C"
      unfolding monotone_on_def
    proof (intro strip)
      fix m n::nat
      assume "m ≤ n"
      then have "{n..} ⊆ {m..}"
        by auto
      then show "C n ⊆ C m"
        unfolding C_def by (meson closure_of_mono image_mono)
    qed
    moreover
    have C: "∃N u. C N ⊆ mcball u ε" if "ε>0" for ε
    proof -
      obtain N where "⋀m n. N ≤ m ∧ N ≤ n ⟹ d (σ m) (σ n) < ε"
        by (meson MCauchy_def ‹0 < ε› ‹MCauchy σ›)
      then have "σ ` {N..} ⊆ mcball (σ N) ε"
        using MCauchy_def ‹MCauchy σ› by (force simp: less_eq_real_def)
      then have "C N ⊆ mcball (σ N) ε"
        by (simp add: C_def closure_of_minimal)
      then show ?thesis
        by blast
    qed
    ultimately obtain l where x: "l ∈ ⋂ (range C)"
      by (metis R ex_in_conv)
    then have *: "⋀ε N. 0 < ε ⟹ ∃n'. N ≤ n' ∧ l ∈ M ∧ σ n' ∈ M ∧ d l (σ n') < ε"
      by (force simp: C_def metric_closure_of)
    then have "l ∈ M"
      using gt_ex by blast
    show "∃l. limitin mtopology σ l sequentially"
      unfolding limitin_metric
    proof (intro conjI strip exI)
      show "l ∈ M"
        using ‹∀n. closedin mtopology (C n)› closedin_subset x by fastforce
      fix ε::real
      assume "ε > 0"
      obtain N where N: "⋀m n. N ≤ m ∧ N ≤ n ⟹ d (σ m) (σ n) < ε/2"
        by (meson MCauchy_def ‹0 < ε› ‹MCauchy σ› half_gt_zero)
      with * [of "ε/2" N]
      have "∀n≥N. σ n ∈ M ∧ d (σ n) l < ε"
        by (smt (verit) ‹range σ ⊆ M› commute field_sum_of_halves range_subsetD triangle)
      then show "∀⇩F n in sequentially. σ n ∈ M ∧ d (σ n) l < ε"
        using eventually_sequentially by blast
    qed
  qed
qed
lemma mcomplete_nest_sing:
   "mcomplete ⟷
    (∀C. (∀n. closedin mtopology (C n)) ∧
          (∀n. C n ≠ {}) ∧ decseq C ∧ (∀e>0. ∃n a. C n ⊆ mcball a e)
         ⟶ (∃l. l ∈ M ∧ ⋂ (range C) = {l}))"
proof -
  have *: False
    if clo: "∀n. closedin mtopology (C n)"
      and cover: "∀ε>0. ∃n a. C n ⊆ mcball a ε"
      and no_sing: "⋀y. y ∈ M ⟹ ⋂ (range C) ≠ {y}"
      and l: "∀n. l ∈ C n"
    for C :: "nat ⇒ 'a set" and l
  proof -
    have inM: "⋀x. x ∈ ⋂ (range C) ⟹ x ∈ M"
      using closedin_metric clo by fastforce
    then have "l ∈ M"
      by (simp add: l)
    have False if l': "l' ∈ ⋂ (range C)" and "l' ≠ l" for l'
    proof -
      have "l' ∈ M"
        using inM l' by blast
      obtain n a where na: "C n ⊆ mcball a (d l l' / 3)"
        using inM ‹l ∈ M› l' ‹l' ≠ l› cover by force
      then have "d a l ≤ (d l l' / 3)" "d a l' ≤ (d l l' / 3)" "a ∈ M"
        using l l' na in_mcball by auto
      then have "d l l' ≤ (d l l' / 3) + (d l l' / 3)"
        using ‹l ∈ M› ‹l' ∈ M› mdist_reverse_triangle by fastforce
      then show False
        using nonneg [of l l'] ‹l' ≠ l› ‹l ∈ M› ‹l' ∈ M› zero by force
    qed
    then show False
      by (metis l ‹l ∈ M› no_sing INT_I empty_iff insertI1 is_singletonE is_singletonI')
  qed
  show ?thesis
    unfolding mcomplete_nest imp_conjL 
    apply (intro all_cong1 imp_cong refl)
    using * 
    by (smt (verit) Inter_iff ex_in_conv range_constant range_eqI)
qed
lemma mcomplete_fip:
   "mcomplete ⟷
    (∀𝒞. (∀C ∈ 𝒞. closedin mtopology C) ∧
         (∀e>0. ∃C a. C ∈ 𝒞 ∧ C ⊆ mcball a e) ∧ (∀ℱ. finite ℱ ∧ ℱ ⊆ 𝒞 ⟶ ⋂ ℱ ≠ {})
        ⟶ ⋂ 𝒞 ≠ {})" 
   (is "?lhs = ?rhs")
proof
  assume L: ?lhs 
  show ?rhs
    unfolding mcomplete_nest_sing imp_conjL
  proof (intro strip)
    fix 𝒞 :: "'a set set"
    assume clo: "∀C∈𝒞. closedin mtopology C"
      and cover: "∀e>0. ∃C a. C ∈ 𝒞 ∧ C ⊆ mcball a e"
      and fip: "∀ℱ. finite ℱ ⟶ ℱ ⊆ 𝒞 ⟶ ⋂ ℱ ≠ {}"
    then have "∀n. ∃C. C ∈ 𝒞 ∧ (∃a. C ⊆ mcball a (inverse (Suc n)))"
      by simp
    then obtain C where C: "⋀n. C n ∈ 𝒞" 
          and coverC: "⋀n. ∃a. C n ⊆ mcball a (inverse (Suc n))"
      by metis
    define D where "D ≡ λn. ⋂ (C ` {..n})"
    have cloD: "closedin mtopology (D n)" for n
      unfolding D_def using clo C by blast
    have neD: "D n ≠ {}" for n
      using fip C by (simp add: D_def image_subset_iff)
    have decD: "decseq D"
      by (force simp: D_def decseq_def)
    have coverD: "∃n a. D n ⊆ mcball a ε" if " ε >0" for ε
    proof -
      obtain n where "inverse (Suc n) < ε"
        using ‹0 < ε› reals_Archimedean by blast
      then obtain a where "C n ⊆ mcball a ε"
        by (meson coverC less_eq_real_def mcball_subset_concentric order_trans)
      then show ?thesis
        unfolding D_def by blast
    qed
    have *: "a ∈ ⋂𝒞" if a: "⋂ (range D) = {a}" and "a ∈ M" for a
    proof -
      have aC: "a ∈ C n" for n
        using that by (auto simp: D_def)
      have eqa: "⋀u. (∀n. u ∈ C n) ⟹ a = u"
        using that by (auto simp: D_def)
      have "a ∈ T" if "T ∈ 𝒞" for T
      proof -
        have cloT: "closedin mtopology (T ∩ D n)" for n
          using clo cloD that by blast
        have "⋂ (insert T (C ` {..n})) ≠ {}" for n
          using that C by (intro fip [rule_format]) auto
        then have neT: "T ∩ D n ≠ {}" for n
          by (simp add: D_def)
        have decT: "decseq (λn. T ∩ D n)"
          by (force simp: D_def decseq_def)
        have coverT: "∃n a. T ∩ D n ⊆ mcball a ε" if " ε >0" for ε
          by (meson coverD le_infI2 that)
        show ?thesis
          using L [unfolded mcomplete_nest_sing, rule_format, of "λn. T ∩ D n"] a
          by (force simp: cloT neT decT coverT)
      qed
      then show ?thesis by auto
    qed
    show "⋂ 𝒞 ≠ {}"
      by (metis L cloD neD decD coverD * empty_iff mcomplete_nest_sing)
  qed
next
  assume R [rule_format]: ?rhs
  show ?lhs
    unfolding mcomplete_nest imp_conjL
  proof (intro strip)
    fix C :: "nat ⇒ 'a set"
    assume clo: "∀n. closedin mtopology (C n)"
      and ne: "∀n. C n ≠ {}"
      and dec: "decseq C"
      and cover: "∀ε>0. ∃n a. C n ⊆ mcball a ε"
    have "⋂(C ` N) ≠ {}" if "finite N" for N
    proof -
      obtain k where "N ⊆ {..k}"
        using ‹finite N› finite_nat_iff_bounded_le by auto
      with dec have "C k ⊆ ⋂(C ` N)" by (auto simp: decseq_def)
      then show ?thesis
        using ne by force
    qed
    with clo cover R [of "range C"] show "⋂ (range C) ≠ {}"
      by (metis (no_types, opaque_lifting) finite_subset_image image_iff UNIV_I)
  qed
qed
lemma mcomplete_fip_sing:
   "mcomplete ⟷
    (∀𝒞. (∀C∈𝒞. closedin mtopology C) ∧
     (∀e>0. ∃c a. c ∈ 𝒞 ∧ c ⊆ mcball a e) ∧
     (∀ℱ. finite ℱ ∧ ℱ ⊆ 𝒞 ⟶ ⋂ ℱ ≠ {}) ⟶
     (∃l. l ∈ M ∧ ⋂ 𝒞 = {l}))"
   (is "?lhs = ?rhs")
proof
  have *: "l ∈ M" "⋂ 𝒞 = {l}"
    if clo: "Ball 𝒞 (closedin mtopology)"
      and cover: "∀e>0. ∃C a. C ∈ 𝒞 ∧ C ⊆ mcball a e"
      and fin: "∀ℱ. finite ℱ ⟶ ℱ ⊆ 𝒞 ⟶ ⋂ ℱ ≠ {}"
      and l: "l ∈ ⋂ 𝒞"
    for 𝒞 :: "'a set set" and l
  proof -
    show "l ∈ M"
      by (meson Inf_lower2 clo cover gt_ex metric_closedin_iff_sequentially_closed subsetD that(4))
    show  "⋂ 𝒞 = {l}"
    proof (cases "𝒞 = {}")
      case True
      then show ?thesis
        using cover mbounded_pos by auto
    next
      case False
      have CM: "⋀a. a ∈ ⋂𝒞 ⟹ a ∈ M"
        using False clo closedin_subset by fastforce
      have "l' ∉ ⋂ 𝒞" if "l' ≠ l" for l'
      proof 
        assume l': "l' ∈ ⋂ 𝒞"
        with CM have "l' ∈ M" by blast
        with that ‹l ∈ M› have gt0: "0 < d l l'"
          by simp
        then obtain C a where "C ∈ 𝒞" and C: "C ⊆ mcball a (d l l' / 3)"
          using cover [rule_format, of "d l l' / 3"] by auto
        then have "d a l ≤ (d l l' / 3)" "d a l' ≤ (d l l' / 3)" "a ∈ M"
          using l l' in_mcball by auto
        then have "d l l' ≤ (d l l' / 3) + (d l l' / 3)"
          using ‹l ∈ M› ‹l' ∈ M› mdist_reverse_triangle by fastforce
        with gt0 show False by auto
      qed
      then show ?thesis
        using l by fastforce
    qed
  qed
  assume L: ?lhs
  with * show ?rhs
    unfolding mcomplete_fip imp_conjL ex_in_conv [symmetric]
    by (elim all_forward imp_forward2 asm_rl) (blast intro:  elim: )
next
  assume ?rhs then show ?lhs
    unfolding mcomplete_fip by (force elim!: all_forward)
qed
end
definition mcomplete_of :: "'a metric ⇒ bool"
  where "mcomplete_of ≡ λm. Metric_space.mcomplete (mspace m) (mdist m)"
lemma (in Metric_space) mcomplete_of [simp]: "mcomplete_of (metric (M,d)) = mcomplete"
  by (simp add: mcomplete_of_def)
lemma mcomplete_trivial: "Metric_space.mcomplete {} (λx y. 0)"
  using Metric_space.intro Metric_space.mcomplete_empty_mspace by force
lemma mcomplete_trivial_singleton: "Metric_space.mcomplete {λx. a} (λx y. 0)"
proof -
  interpret Metric_space "{λx. a}" "λx y. 0"
    by unfold_locales auto
  show ?thesis
    unfolding mcomplete_def MCauchy_def image_subset_iff by (metis UNIV_I limit_metric_sequentially)
qed
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 TYPE('a::metric_space) ⟷ complete (UNIV::'a set)"
  by (auto simp: Met_TC.mcomplete_def complete_def)
context Submetric
begin 
lemma MCauchy_submetric:
   "sub.MCauchy σ ⟷ range σ ⊆ A ∧ MCauchy σ"
  using MCauchy_def sub.MCauchy_def subset by force
lemma closedin_mcomplete_imp_mcomplete:
  assumes clo: "closedin mtopology A" and "mcomplete"
  shows "sub.mcomplete"
  unfolding sub.mcomplete_def
proof (intro strip)
  fix σ
  assume "sub.MCauchy σ"
  then have σ: "MCauchy σ" "range σ ⊆ A"
    using MCauchy_submetric by blast+
  then obtain x where x: "limitin mtopology σ x sequentially"
    using ‹mcomplete› unfolding mcomplete_def by blast
  then have "x ∈ A"
    using σ clo metric_closedin_iff_sequentially_closed by force
  with σ x show "∃x. limitin sub.mtopology σ x sequentially"
    using limitin_submetric_iff range_subsetD by fastforce
qed
lemma sequentially_closedin_mcomplete_imp_mcomplete:
  assumes "mcomplete" and "⋀σ l. range σ ⊆ A ∧ limitin mtopology σ l sequentially ⟹ l ∈ A"
  shows "sub.mcomplete"
  using assms closedin_mcomplete_imp_mcomplete metric_closedin_iff_sequentially_closed subset by blast
end
context Metric_space
begin
lemma mcomplete_Un:
  assumes A: "Submetric M d A" "Metric_space.mcomplete A d" 
      and B: "Submetric M d B" "Metric_space.mcomplete B d"
  shows   "Submetric M d (A ∪ B)" "Metric_space.mcomplete (A ∪ B) d" 
proof -
  show "Submetric M d (A ∪ B)"
    by (meson assms le_sup_iff Submetric_axioms_def Submetric_def) 
  then interpret MAB: Metric_space "A ∪ B" d
    by (meson Submetric.subset subspace)
  interpret MA: Metric_space A d
    by (meson A Submetric.subset subspace)
  interpret MB: Metric_space B d
    by (meson B Submetric.subset subspace)
  show "Metric_space.mcomplete (A ∪ B) d"
    unfolding MAB.mcomplete_def
  proof (intro strip)
    fix σ
    assume "MAB.MCauchy σ"
    then have "range σ ⊆ A ∪ B"
      using MAB.MCauchy_def by blast
    then have "UNIV ⊆ σ -` A ∪ σ -` B"
      by blast
    then consider "infinite (σ -` A)" | "infinite (σ -` B)"
      using finite_subset by auto
    then show "∃x. limitin MAB.mtopology σ x sequentially"
    proof cases
      case 1
      then obtain r where "strict_mono r" and r: "⋀n::nat. r n ∈ σ -` A"
        using infinite_enumerate by blast 
      then have "MA.MCauchy (σ ∘ r)"
        using MA.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence ‹MAB.MCauchy σ› by auto
      with A obtain x where "limitin MA.mtopology (σ ∘ r) x sequentially"
        using MA.mcomplete_def by blast
      then have "limitin MAB.mtopology (σ ∘ r) x sequentially"
        by (metis MA.limit_metric_sequentially MAB.limit_metric_sequentially UnCI)
      then show ?thesis
        using MAB.MCauchy_convergent_subsequence ‹MAB.MCauchy σ› ‹strict_mono r› by blast
    next
      case 2
      then obtain r where "strict_mono r" and r: "⋀n::nat. r n ∈ σ -` B"
        using infinite_enumerate by blast 
      then have "MB.MCauchy (σ ∘ r)"
        using MB.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence ‹MAB.MCauchy σ› by auto
      with B obtain x where "limitin MB.mtopology (σ ∘ r) x sequentially"
        using MB.mcomplete_def by blast
      then have "limitin MAB.mtopology (σ ∘ r) x sequentially"
        by (metis MB.limit_metric_sequentially MAB.limit_metric_sequentially UnCI)
      then show ?thesis
        using MAB.MCauchy_convergent_subsequence ‹MAB.MCauchy σ› ‹strict_mono r› by blast
    qed
  qed
qed
lemma mcomplete_Union:
  assumes "finite 𝒮"
    and "⋀A. A ∈ 𝒮 ⟹ Submetric M d A" "⋀A. A ∈ 𝒮 ⟹ Metric_space.mcomplete A d"
  shows   "Submetric M d (⋃𝒮)" "Metric_space.mcomplete (⋃𝒮) d" 
  using assms
  by (induction rule: finite_induct) (auto simp: mcomplete_Un)
lemma mcomplete_Inter:
  assumes "finite 𝒮" "𝒮 ≠ {}"
    and sub: "⋀A. A ∈ 𝒮 ⟹ Submetric M d A" 
    and comp: "⋀A. A ∈ 𝒮 ⟹ Metric_space.mcomplete A d"
  shows "Submetric M d (⋂𝒮)" "Metric_space.mcomplete (⋂𝒮) d"
proof -
  show "Submetric M d (⋂𝒮)"
    using assms unfolding Submetric_def Submetric_axioms_def
    by (metis Inter_lower equals0I inf.orderE le_inf_iff) 
  then interpret MS: Submetric M d "⋂𝒮" 
    by (meson Submetric.subset subspace)
  show "Metric_space.mcomplete (⋂𝒮) d"
    unfolding MS.sub.mcomplete_def
  proof (intro strip)
    fix σ
    assume "MS.sub.MCauchy σ"
    then have "range σ ⊆ ⋂𝒮"
      using MS.MCauchy_submetric by blast
    obtain A where "A ∈ 𝒮" and A: "Metric_space.mcomplete A d"
      using assms by blast
    then have "range σ ⊆ A"
      using ‹range σ ⊆ ⋂𝒮› by blast
    interpret SA: Submetric M d A
      by (meson ‹A ∈ 𝒮› sub Submetric.subset subspace)
    have "MCauchy σ"
      using MS.MCauchy_submetric ‹MS.sub.MCauchy σ› by blast
    then obtain x where x: "limitin SA.sub.mtopology σ x sequentially"
      by (metis A SA.sub.MCauchy_def SA.sub.mcomplete_alt MCauchy_def ‹range σ ⊆ A›)
    show "∃x. limitin MS.sub.mtopology σ x sequentially"
      unfolding MS.limitin_submetric_iff
    proof (intro exI conjI)
      show "x ∈ ⋂ 𝒮"
      proof clarsimp
        fix U
        assume "U ∈ 𝒮"
        interpret SU: Submetric M d U 
          by (meson ‹U ∈ 𝒮› sub Submetric.subset subspace)
        have "range σ ⊆ U"
          using ‹U ∈ 𝒮› ‹range σ ⊆ ⋂ 𝒮› by blast
        moreover have "Metric_space.mcomplete U d"
          by (simp add: ‹U ∈ 𝒮› comp)
        ultimately obtain x' where x': "limitin SU.sub.mtopology σ x' sequentially"
          using MCauchy_def SU.sub.MCauchy_def SU.sub.mcomplete_alt ‹MCauchy σ› by meson 
        have "x' = x"
        proof (intro limitin_metric_unique)
          show "limitin mtopology σ x' sequentially"
            by (meson SU.Submetric_axioms Submetric.limitin_submetric_iff x')
          show "limitin mtopology σ x sequentially"
            by (meson SA.Submetric_axioms Submetric.limitin_submetric_iff x)
        qed auto
        then show "x ∈ U"
          using SU.sub.limitin_mspace x' by blast
      qed
      show "∀⇩F n in sequentially. σ n ∈ ⋂𝒮"
        by (meson ‹range σ ⊆ ⋂ 𝒮› always_eventually range_subsetD)
      show "limitin mtopology σ x sequentially"
        by (meson SA.Submetric_axioms Submetric.limitin_submetric_iff x)
    qed
  qed
qed
lemma mcomplete_Int:
  assumes A: "Submetric M d A" "Metric_space.mcomplete A d" 
      and B: "Submetric M d B" "Metric_space.mcomplete B d"
    shows   "Submetric M d (A ∩ B)" "Metric_space.mcomplete (A ∩ B) d"
  using mcomplete_Inter [of "{A,B}"] assms by force+
subsection‹Totally bounded subsets of metric spaces›
definition mtotally_bounded 
  where "mtotally_bounded S ≡ ∀ε>0. ∃K. finite K ∧ K ⊆ S ∧ S ⊆ (⋃x∈K. mball x ε)"
lemma mtotally_bounded_empty [iff]: "mtotally_bounded {}"
by (simp add: mtotally_bounded_def)
lemma finite_imp_mtotally_bounded:
   "⟦finite S; S ⊆ M⟧ ⟹ mtotally_bounded S"
  by (auto simp: mtotally_bounded_def)
lemma mtotally_bounded_imp_subset: "mtotally_bounded S ⟹ S ⊆ M"
  by (force simp: mtotally_bounded_def intro!: zero_less_one)
lemma mtotally_bounded_sing [simp]:
   "mtotally_bounded {x} ⟷ x ∈ M"
  by (meson empty_subsetI finite.simps finite_imp_mtotally_bounded insert_subset mtotally_bounded_imp_subset)
lemma mtotally_bounded_Un:
  assumes  "mtotally_bounded S" "mtotally_bounded T"
  shows "mtotally_bounded (S ∪ T)"
proof -
  have "∃K. finite K ∧ K ⊆ S ∪ T ∧ S ∪ T ⊆ (⋃x∈K. mball x e)"
    if  "e>0" and K: "finite K ∧ K ⊆ S ∧ S ⊆ (⋃x∈K. mball x e)"
      and L: "finite L ∧ L ⊆ T ∧ T ⊆ (⋃x∈L. mball x e)" for K L e
    using that by (rule_tac x="K ∪ L" in exI) auto
  with assms show ?thesis
    unfolding mtotally_bounded_def by presburger
qed
 
lemma mtotally_bounded_Union:
  assumes "finite f" "⋀S. S ∈ f ⟹ mtotally_bounded S"
  shows "mtotally_bounded (⋃f)"
  using assms by (induction f) (auto simp: mtotally_bounded_Un)
lemma mtotally_bounded_imp_mbounded:
  assumes "mtotally_bounded S"
  shows "mbounded S"
proof -
  obtain K where "finite K ∧ K ⊆ S ∧ S ⊆ (⋃x∈K. mball x 1)" 
    using assms by (force simp: mtotally_bounded_def)
  then show ?thesis
    by (smt (verit) finite_imageI image_iff mbounded_Union mbounded_mball mbounded_subset)
qed
lemma mtotally_bounded_sequentially:
  "mtotally_bounded S ⟷
    S ⊆ M ∧ (∀σ::nat ⇒ 'a. range σ ⊆ S ⟶ (∃r. strict_mono r ∧ MCauchy (σ ∘ r)))"
  (is "_ ⟷ _ ∧ ?rhs")
proof (cases "S ⊆ M")
  case True
  show ?thesis
  proof -
    { fix σ :: "nat ⇒ 'a"                                                            
      assume L: "mtotally_bounded S" and σ: "range σ ⊆ S"
      have "∃j > i. d (σ i) (σ j) < 3*ε/2 ∧ infinite (σ -` mball (σ j) (ε/2))"
        if inf: "infinite (σ -` mball (σ i) ε)" and "ε > 0" for i ε
      proof -
        obtain K where "finite K" "K ⊆ S" and K: "S ⊆ (⋃x∈K. mball x (ε/4))"
          by (metis L mtotally_bounded_def ‹ε > 0› zero_less_divide_iff zero_less_numeral)
        then have K_imp_ex: "⋀y. y ∈ S ⟹ ∃x∈K. d x y < ε/4"
          by fastforce
        have False if "∀x∈K. d x (σ i) < ε + ε/4 ⟶ finite (σ -` mball x (ε/4))"
        proof -
          have "∃w. w ∈ K ∧ d w (σ i) < 5 * ε/4 ∧ d w (σ j) < ε/4"
            if "d (σ i) (σ j) < ε" for j
          proof -
            obtain w where w: "d w (σ j) < ε/4" "w ∈ K"
              using K_imp_ex σ by blast
            then have "d w (σ i) < ε + ε/4"
              by (smt (verit, ccfv_SIG) True ‹K ⊆ S› σ rangeI subset_eq that triangle')
            with w show ?thesis
              using in_mball by auto
          qed
          then have "(σ -` mball (σ i) ε) ⊆ (⋃x∈K. if d x (σ i) < ε + ε/4 then σ -` mball x (ε/4) else {})"
            using True ‹K ⊆ S› by force
          then show False
            using finite_subset inf ‹finite K› that by fastforce
        qed
        then obtain x where "x ∈ K" and dxi: "d x (σ i) < ε + ε/4" and infx: "infinite (σ -` mball x (ε/4))"
          by blast
        then obtain j where "j ∈ (σ -` mball x (ε/4)) - {..i}"
          using bounded_nat_set_is_finite by (meson Diff_infinite_finite finite_atMost)
        then have "j > i" and dxj: "d x (σ j) < ε/4" 
          by auto
        have "(σ -` mball x (ε/4)) ⊆ (σ -` mball y (ε/2))" if "d x y < ε/4" "y ∈ M" for y
          using that by (simp add: mball_subset vimage_mono)
        then have infj: "infinite (σ -` mball (σ j) (ε/2))"
          by (meson True ‹d x (σ j) < ε/4› σ in_mono infx rangeI finite_subset)
        have "σ i ∈ M" "σ j ∈ M" "x ∈ M"  
          using True ‹K ⊆ S› ‹x ∈ K› σ by force+
        then have "d (σ i) (σ j) ≤ d x (σ i) + d x (σ j)"
          using triangle'' by blast
        also have "… < 3*ε/2"
          using dxi dxj by auto
        finally have "d (σ i) (σ j) < 3*ε/2" .
        with ‹i < j› infj show ?thesis by blast
      qed
      then obtain nxt where nxt: "⋀i ε. ⟦ε > 0; infinite (σ -` mball (σ i) ε)⟧ ⟹ 
                 nxt i ε > i ∧ d (σ i) (σ (nxt i ε)) < 3*ε/2 ∧ infinite (σ -` mball (σ (nxt i ε)) (ε/2))"
        by metis
      have "mbounded S"
        using L by (simp add: mtotally_bounded_imp_mbounded)
      then obtain B where B: "∀y ∈ S. d (σ 0) y ≤ B" and "B > 0"
        by (meson σ mbounded_alt_pos range_subsetD)
      define eps where "eps ≡ λn. (B+1) / 2^n"
      have [simp]: "eps (Suc n) = eps n / 2" "eps n > 0" for n
        using ‹B > 0› by (auto simp: eps_def)
      have "UNIV ⊆ σ -` mball (σ 0) (B+1)"
        using B True σ unfolding image_iff subset_iff
        by (smt (verit, best) UNIV_I in_mball vimageI)
      then have inf0: "infinite (σ -` mball (σ 0) (eps 0))"
        using finite_subset by (auto simp: eps_def)
      define r where "r ≡ rec_nat 0 (λn rec. nxt rec (eps n))"
      have [simp]: "r 0 = 0" "r (Suc n) = nxt (r n) (eps n)" for n
        by (auto simp: r_def)
      have σrM[simp]: "σ (r n) ∈ M" for n
        using True σ by blast
      have inf: "infinite (σ -` mball (σ (r n)) (eps n))" for n
      proof (induction n)
        case 0 then show ?case  
          by (simp add: inf0)
      next
        case (Suc n) then show ?case
          using nxt [of "eps n" "r n"] by simp
      qed
      then have "r (Suc n) > r n" for n
        by (simp add: nxt)
      then have "strict_mono r"
        by (simp add: strict_mono_Suc_iff)
      have d_less: "d (σ (r n)) (σ (r (Suc n))) < 3 * eps n / 2" for n
        using nxt [OF _ inf] by simp
      have eps_plus: "eps (k + n) = eps n * (1/2)^k" for k n
        by (simp add: eps_def power_add field_simps)
      have *: "d (σ (r n)) (σ (r (k + n))) < 3 * eps n" for n k
      proof -
        have "d (σ (r n)) (σ (r (k+n))) ≤ 3/2 * eps n * (∑i<k. (1/2)^i)"
        proof (induction k)
          case 0 then show ?case 
            by simp
        next
          case (Suc k)
          have "d (σ (r n)) (σ (r (Suc k + n))) ≤ d (σ (r n)) (σ (r (k + n))) + d (σ (r (k + n))) (σ (r (Suc (k + n))))"
            by (metis σrM add.commute add_Suc_right triangle)
          with d_less[of "k+n"] Suc show ?case
            by (simp add: algebra_simps eps_plus)
        qed
        also have "… < 3/2 * eps n * 2"
          using geometric_sum [of "1/2::real" k] by simp
        finally show ?thesis by simp
      qed
      have "∃N. ∀n≥N. ∀n'≥N. d (σ (r n)) (σ (r n')) < ε" if "ε > 0" for ε
      proof -
        define N where "N ≡ nat ⌈(log 2 (6*(B+1) / ε))⌉"
        have §: "b ≤ 2 ^ nat ⌈log 2 b⌉" for b
          by (smt (verit) less_log_of_power real_nat_ceiling_ge)
        have N: "6 * eps N ≤ ε"
          using § [of "(6*(B+1) / ε)"] that by (auto simp: N_def eps_def field_simps)
        have "d (σ (r N)) (σ (r n)) < 3 * eps N" if "n ≥ N" for n
          by (metis * add.commute nat_le_iff_add that)
        then have "∀n≥N. ∀n'≥N. d (σ (r n)) (σ (r n')) < 3 * eps N + 3 * eps N"
          by (smt (verit, best) σrM triangle'')
        with N show ?thesis
          by fastforce
      qed
      then have "MCauchy (σ ∘ r)"
        unfolding MCauchy_def using True σ by auto
      then have "∃r. strict_mono r ∧ MCauchy (σ ∘ r)"
        using ‹strict_mono r› by blast      
    }
    moreover
    { assume R: ?rhs
      have "mtotally_bounded S"
        unfolding mtotally_bounded_def
      proof (intro strip)
        fix ε :: real
        assume "ε > 0"
        have False if §: "⋀K. ⟦finite K; K ⊆ S⟧ ⟹ ∃s∈S. s ∉ (⋃x∈K. mball x ε)"
        proof -
          obtain f where f: "⋀K. ⟦finite K; K ⊆ S⟧ ⟹ f K ∈ S ∧ f K ∉ (⋃x∈K. mball x ε)"
            using § by metis
          define σ where "σ ≡ wfrec less_than (λseq n. f (seq ` {..<n}))"
          have σ_eq: "σ n = f (σ ` {..<n})" for n
            by (simp add: cut_apply def_wfrec [OF σ_def])
          have [simp]: "σ n ∈ S" for n
            using wf_less_than
          proof (induction n rule: wf_induct_rule)
            case (less n) with f show ?case
              by (auto simp: σ_eq [of n])
          qed
          then have "range σ ⊆ S" by blast
          have σ: "p < n ⟹ ε ≤ d (σ p) (σ n)" for n p
            using f[of "σ ` {..<n}"] True by (fastforce simp: σ_eq [of n] Ball_def)
          then obtain r where "strict_mono r" "MCauchy (σ ∘ r)"
            by (meson R ‹range σ ⊆ S›)
          with ‹0 < ε› obtain N 
            where N: "⋀n n'. ⟦n≥N; n'≥N⟧ ⟹ d (σ (r n)) (σ (r n')) < ε"
            by (force simp: MCauchy_def)
          show ?thesis
            using N [of N "Suc (r N)"] ‹strict_mono r›
            by (smt (verit) Suc_le_eq σ le_SucI order_refl strict_mono_imp_increasing)
        qed
        then show "∃K. finite K ∧ K ⊆ S ∧ S ⊆ (⋃x∈K. mball x ε)"
          by blast
      qed
    }
    ultimately show ?thesis 
      using True by blast
  qed
qed (use mtotally_bounded_imp_subset in auto)
lemma mtotally_bounded_subset:
   "⟦mtotally_bounded S; T ⊆ S⟧ ⟹ mtotally_bounded T"
  by (meson mtotally_bounded_sequentially order_trans) 
lemma mtotally_bounded_submetric:
  assumes "mtotally_bounded S" "S ⊆ T" "T ⊆ M"
  shows "Metric_space.mtotally_bounded T d S"
proof -
  interpret Submetric M d T
    using ‹T ⊆ M› by unfold_locales
  show ?thesis
    using assms
    unfolding sub.mtotally_bounded_def mtotally_bounded_def
    by (force simp: subset_iff elim!: all_forward ex_forward)
qed
lemma mtotally_bounded_absolute:
   "mtotally_bounded S ⟷ S ⊆ M ∧ Metric_space.mtotally_bounded S d S "
proof -
  have "mtotally_bounded S" if "S ⊆ M" "Metric_space.mtotally_bounded S d S"
  proof -
    interpret Submetric M d S
      using ‹S ⊆ M› by unfold_locales
    show ?thesis
      using that
      by (meson MCauchy_submetric mtotally_bounded_sequentially sub.mtotally_bounded_sequentially)
  qed
  moreover have "mtotally_bounded S ⟹ Metric_space.mtotally_bounded S d S"
    by (simp add: mtotally_bounded_imp_subset mtotally_bounded_submetric)
  ultimately show ?thesis
    using mtotally_bounded_imp_subset by blast
qed
lemma mtotally_bounded_closure_of:
  assumes "mtotally_bounded S"
  shows "mtotally_bounded (mtopology closure_of S)"
proof -
  have "S ⊆ M"
    by (simp add: assms mtotally_bounded_imp_subset)
  have "mtotally_bounded(mtopology closure_of S)"
    unfolding mtotally_bounded_def
  proof (intro strip)
    fix ε::real
    assume "ε > 0"
    then obtain K where "finite K" "K ⊆ S" and K: "S ⊆ (⋃x∈K. mball x (ε/2))"
      by (metis assms mtotally_bounded_def half_gt_zero)
    have "mtopology closure_of S ⊆ (⋃x∈K. mball x ε)"
      unfolding metric_closure_of
    proof clarsimp
      fix x
      assume "x ∈ M" and x: "∀r>0. ∃y∈S. y ∈ M ∧ d x y < r"
      then obtain y where "y ∈ S" and y: "d x y < ε/2"
        using ‹0 < ε› half_gt_zero by blast
      then obtain x' where "x' ∈ K" "y ∈ mball x' (ε/2)"
        using K by auto
      then have "d x' x < ε/2 + ε/2"
        using triangle y ‹x ∈ M› commute by fastforce
      then show "∃x'∈K. x' ∈ M ∧ d x' x < ε"
        using ‹K ⊆ S› ‹S ⊆ M› ‹x' ∈ K› by force
    qed
    then show "∃K. finite K ∧ K ⊆ mtopology closure_of S ∧ mtopology closure_of S ⊆ (⋃x∈K. mball x ε)"
      using closure_of_subset_Int  ‹K ⊆ S› ‹finite K› K by fastforce
  qed
  then show ?thesis
    by (simp add: assms inf.absorb2 mtotally_bounded_imp_subset)
qed
lemma mtotally_bounded_closure_of_eq:
   "S ⊆ M ⟹ mtotally_bounded (mtopology closure_of S) ⟷ mtotally_bounded S"
  by (metis closure_of_subset mtotally_bounded_closure_of mtotally_bounded_subset topspace_mtopology)
lemma mtotally_bounded_cauchy_sequence:
  assumes "MCauchy σ"
  shows "mtotally_bounded (range σ)"
  unfolding MCauchy_def mtotally_bounded_def
proof (intro strip)
  fix ε::real
  assume "ε > 0"
  then obtain N where "⋀n. N ≤ n ⟹ d (σ N) (σ n) < ε"
    using assms by (force simp: MCauchy_def)
  then have "⋀m. ∃n≤N. σ n ∈ M ∧ σ m ∈ M ∧ d (σ n) (σ m) < ε"
    by (metis MCauchy_def assms mdist_zero nle_le range_subsetD)
  then
  show "∃K. finite K ∧ K ⊆ range σ ∧ range σ ⊆ (⋃x∈K. mball x ε)"
    by (rule_tac x="σ ` {0..N}" in exI) force
qed
lemma MCauchy_imp_mbounded:
   "MCauchy σ ⟹ mbounded (range σ)"
  by (simp add: mtotally_bounded_cauchy_sequence mtotally_bounded_imp_mbounded)
subsection‹Compactness in metric spaces›
lemma Bolzano_Weierstrass_property:
  assumes "S ⊆ U" "S ⊆ M"
  shows
   "(∀σ::nat⇒'a. range σ ⊆ S
         ⟶ (∃l r. l ∈ U ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially)) ⟷
    (∀T. T ⊆ S ∧ infinite T ⟶ U ∩ mtopology derived_set_of T ≠ {})"  (is "?lhs=?rhs")
proof
  assume L: ?lhs 
  show ?rhs
  proof clarify
    fix T
    assume "T ⊆ S" and "infinite T"
      and T: "U ∩ mtopology derived_set_of T = {}"
    then obtain σ :: "nat⇒'a" where "inj σ" "range σ ⊆ T"
      by (meson infinite_countable_subset)
    with L obtain l r where "l ∈ U" "strict_mono r" 
           and lr: "limitin mtopology (σ ∘ r) l sequentially"
      by (meson ‹T ⊆ S› subset_trans)
    then obtain ε where "ε > 0" and ε: "⋀y. y ∈ T ⟹ y = l ∨ ¬ d l y < ε"
      using T ‹T ⊆ S› ‹S ⊆ M› 
      by (force simp: metric_derived_set_of limitin_metric disjoint_iff)
    with lr have "∀⇩F n in sequentially. σ (r n) ∈ M ∧ d (σ (r n)) l < ε"
      by (auto simp: limitin_metric)
    then obtain N where N: "d (σ (r N)) l < ε" "d (σ (r (Suc N))) l < ε"
      using less_le_not_le by (auto simp: eventually_sequentially)
    moreover have "σ (r N) ≠ l ∨ σ (r (Suc N)) ≠ l"
      by (meson ‹inj σ› ‹strict_mono r› injD n_not_Suc_n strict_mono_eq)
    ultimately
    show False
       using ε ‹range σ ⊆ T› commute by fastforce
  qed
next
  assume R: ?rhs 
  show ?lhs
  proof (intro strip)
    fix σ :: "nat ⇒ 'a"
    assume "range σ ⊆ S"
    show "∃l r. l ∈ U ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially"
    proof (cases "finite (range σ)")
      case True
      then obtain m where "infinite (σ -` {σ m})"
        by (metis image_iff inf_img_fin_dom nat_not_finite)
      then obtain r where [iff]: "strict_mono r" and r: "⋀n::nat. r n ∈ σ -` {σ m}"
        using infinite_enumerate by blast 
      have [iff]: "σ m ∈ U" "σ m ∈ M"
        using ‹range σ ⊆ S› assms by blast+
      show ?thesis
      proof (intro conjI exI)
        show "limitin mtopology (σ ∘ r) (σ m) sequentially"
          using r by (simp add: limitin_metric)
      qed auto
    next
      case False
      then obtain l where "l ∈ U" and l: "l ∈ mtopology derived_set_of (range σ)"
        by (meson R ‹range σ ⊆ S› disjoint_iff)
      then obtain g where g: "⋀ε. ε>0 ⟹ σ (g ε) ≠ l ∧ d l (σ (g ε)) < ε"
        by (simp add: metric_derived_set_of) metis
      have "range σ ⊆ M"
        using ‹range σ ⊆ S› assms by auto
      have "l ∈ M"
        using l metric_derived_set_of by auto
      define E where  
        "E ≡ λrec n. insert (inverse (Suc n)) ((λi. d l (σ i)) ` (⋃k<n. {0..rec k})) - {0}"
      define r where "r ≡ wfrec less_than (λrec n. g (Min (E rec n)))"
      have "(⋃k<n. {0..cut r less_than n k}) = (⋃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 (σ (r n)) > 0" for n
        using wf_less_than
      proof (induction n rule: wf_induct_rule)
        case (less n) 
        then have *: "Min (E r n) > 0"
          using ‹l ∈ M› ‹range σ ⊆ M› by (auto simp: E_def image_subset_iff)
        show ?case
          using g [OF *] r_eq [of n]
          by (metis ‹l ∈ M› ‹range σ ⊆ M› mdist_pos_less range_subsetD)
      qed
      then have non_l: "σ (r n) ≠ l" for n
        using ‹range σ ⊆ M› mdist_pos_eq by blast
      have Min_pos: "Min (E r n) > 0" for n
        using dl_pos ‹l ∈ M› ‹range σ ⊆ M› by (auto simp: E_def image_subset_iff)
      have d_small: "d (σ(r n)) l < inverse(Suc n)" for n
      proof -
        have "d (σ(r n)) l < Min (E r n)"
          by (simp add: ‹0 < Min (E r n)› commute g r_eq) 
        also have "... ≤ inverse(Suc n)"
          by (simp add: E_def)
        finally show ?thesis .
      qed
      have d_lt_d: "d l (σ (r n)) < d l (σ i)" if §: "p < n" "i ≤ r p" "σ i ≠ l" for i p n
      proof -
        have 1: "d l (σ i) ∈ E r n"
          using § ‹l ∈ M› ‹range σ ⊆ M› 
          by (force simp: E_def image_subset_iff image_iff)
        have "d l (σ (g (Min (E r n)))) < Min (E r n)"
          by (rule conjunct2 [OF g [OF Min_pos]])
        also have "Min (E r n) ≤ d l (σ i)"
          using 1 unfolding E_def by (force intro!: Min.coboundedI)
        finally show ?thesis
          by (simp add: r_eq) 
      qed
      have r: "r p < r n" if "p < n" for p n
        using d_lt_d [OF that] non_l by (meson linorder_not_le order_less_irrefl) 
      show ?thesis
      proof (intro exI conjI)
        show "strict_mono r"
          by (simp add: r strict_monoI)
        show "limitin mtopology (σ ∘ r) l sequentially"
          unfolding limitin_metric
        proof (intro conjI strip ‹l ∈ M›)
          fix ε :: real
          assume "ε > 0"
          then have "∀⇩F n in sequentially. inverse(Suc n) < ε"
            using Archimedean_eventually_inverse by auto
          then show "∀⇩F n in sequentially. (σ ∘ r) n ∈ M ∧ d ((σ ∘ r) n) l < ε"
            by (smt (verit) ‹range σ ⊆ M› commute comp_apply d_small eventually_mono range_subsetD)
        qed
      qed (use ‹l ∈ U› in auto)
    qed
  qed
qed
subsubsection ‹More on Bolzano Weierstrass›
lemma Bolzano_Weierstrass_A: 
  assumes "compactin mtopology S" "T ⊆ S" "infinite T"
  shows "S ∩ mtopology derived_set_of T ≠ {}"
  by (simp add: assms compactin_imp_Bolzano_Weierstrass)
lemma Bolzano_Weierstrass_B:
  fixes σ :: "nat ⇒ 'a"
  assumes "S ⊆ M" "range σ ⊆ S"
    and "⋀T. ⟦T ⊆ S ∧ infinite T⟧ ⟹ S ∩ mtopology derived_set_of T ≠ {}"
  shows "∃l r. l ∈ S ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially"
  using Bolzano_Weierstrass_property assms by blast
lemma Bolzano_Weierstrass_C:
  assumes "S ⊆ M"
  assumes "⋀σ:: nat ⇒ 'a. range σ ⊆ S ⟹
                (∃l r. l ∈ S ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially)"
  shows "mtotally_bounded S"
  unfolding mtotally_bounded_sequentially
  by (metis convergent_imp_MCauchy assms image_comp image_mono subset_UNIV subset_trans)
lemma Bolzano_Weierstrass_D:
  assumes "S ⊆ M" "S ⊆ ⋃𝒞" and opeU: "⋀U. U ∈ 𝒞 ⟹ openin mtopology U"
  assumes §: "(∀σ::nat⇒'a. range σ ⊆ S
         ⟶ (∃l r. l ∈ S ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially))"
  shows "∃ε>0. ∀x ∈ S. ∃U ∈ 𝒞. mball x ε ⊆ U"
proof (rule ccontr)
  assume "¬ (∃ε>0. ∀x ∈ S. ∃U ∈ 𝒞. mball x ε ⊆ U)"
  then have "∀n. ∃x∈S. ∀U∈𝒞. ¬ mball x (inverse (Suc n)) ⊆ U"
    by simp
  then obtain σ where "⋀n. σ n ∈ S" 
       and σ: "⋀n U. U ∈ 𝒞 ⟹ ¬ mball (σ n) (inverse (Suc n)) ⊆ U"
    by metis
  then obtain l r where "l ∈ S" "strict_mono r" 
         and lr: "limitin mtopology (σ ∘ r) l sequentially"
    by (meson § image_subsetI)
  with ‹S ⊆ ⋃𝒞› obtain B where "l ∈ B" "B ∈ 𝒞"
    by auto
  then obtain ε where "ε > 0" and ε: "⋀z. ⟦z ∈ M; d z l < ε⟧ ⟹ z ∈ B"
    by (metis opeU [OF ‹B ∈ 𝒞›] commute in_mball openin_mtopology subset_iff)
  then have "∀⇩F n in sequentially. σ (r n) ∈ M ∧ d (σ (r n)) l < ε/2"
    using lr half_gt_zero unfolding limitin_metric o_def by blast
  moreover have "∀⇩F n in sequentially. inverse (real (Suc n)) < ε/2"
    using Archimedean_eventually_inverse ‹0 < ε› half_gt_zero by blast
  ultimately obtain n where n: "d (σ (r n)) l < ε/2" "inverse (real (Suc n)) < ε/2"
    by (smt (verit, del_insts) eventually_sequentially le_add1 le_add2)
  have "x ∈ B" if "d (σ (r n)) x < inverse (Suc(r n))" "x ∈ M" for x
  proof -
    have rle: "inverse (real (Suc (r n))) ≤ inverse (real (Suc n))"
      using ‹strict_mono r› strict_mono_imp_increasing by auto
    have "d x l ≤ d (σ (r n)) x + d (σ (r n)) l"
      using that by (metis triangle ‹⋀n. σ n ∈ S› ‹l ∈ S› ‹S ⊆ M› commute subsetD)
    also have "... < ε"
      using that n rle by linarith
    finally show ?thesis
      by (simp add: ε that)
  qed
  then show False
    using σ [of B "r n"] by (simp add: ‹B ∈ 𝒞› subset_iff)
qed
lemma Bolzano_Weierstrass_E:
  assumes "mtotally_bounded S" "S ⊆ M"
  and S: "⋀𝒞. ⟦⋀U. U ∈ 𝒞 ⟹ openin mtopology U; S ⊆ ⋃𝒞⟧ ⟹ ∃ε>0. ∀x ∈ S. ∃U ∈ 𝒞. mball x ε ⊆ U"
  shows "compactin mtopology S"
proof (clarsimp simp: compactin_def assms)
  fix 𝒰 :: "'a set set"
  assume 𝒰: "∀x∈𝒰. openin mtopology x" and "S ⊆ ⋃𝒰"
  then obtain ε where "ε>0" and ε: "⋀x. x ∈ S ⟹ ∃U ∈ 𝒰. mball x ε ⊆ U"
    by (metis S)
  then obtain f where f: "⋀x. x ∈ S ⟹ f x ∈ 𝒰 ∧ mball x ε ⊆ f x"
    by metis
  then obtain K where "finite K" "K ⊆ S" and K: "S ⊆ (⋃x∈K. mball x ε)"
    by (metis ‹0 < ε› ‹mtotally_bounded S› mtotally_bounded_def)
  show "∃ℱ. finite ℱ ∧ ℱ ⊆ 𝒰 ∧ S ⊆ ⋃ℱ"
  proof (intro conjI exI)
    show "finite (f ` K)"
      by (simp add: ‹finite K›)
    show "f ` K ⊆ 𝒰"
      using ‹K ⊆ S› f by blast
    show "S ⊆ ⋃(f ` K)"
      using K ‹K ⊆ S› by (force dest: f)
  qed
qed
lemma compactin_eq_Bolzano_Weierstrass:
  "compactin mtopology S ⟷
   S ⊆ M ∧ (∀T. T ⊆ S ∧ infinite T ⟶ S ∩ mtopology derived_set_of T ≠ {})"
  using Bolzano_Weierstrass_C Bolzano_Weierstrass_D Bolzano_Weierstrass_E
  by (smt (verit, del_insts) Bolzano_Weierstrass_property compactin_imp_Bolzano_Weierstrass compactin_subspace subset_refl topspace_mtopology)
lemma compactin_sequentially:
  shows "compactin mtopology S ⟷
        S ⊆ M ∧
        ((∀σ::nat⇒'a. range σ ⊆ S
         ⟶ (∃l r. l ∈ S ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially)))"
  by (metis Bolzano_Weierstrass_property compactin_eq_Bolzano_Weierstrass subset_refl)
lemma compactin_imp_mtotally_bounded: 
  "compactin mtopology S ⟹ mtotally_bounded S"
  by (simp add: Bolzano_Weierstrass_C compactin_sequentially)
lemma lebesgue_number:
    "⟦compactin mtopology S; S ⊆ ⋃𝒞; ⋀U. U ∈ 𝒞 ⟹ openin mtopology U⟧
    ⟹ ∃ε>0. ∀x ∈ S. ∃U ∈ 𝒞. mball x ε ⊆ U"
  by (simp add: Bolzano_Weierstrass_D compactin_sequentially)
lemma compact_space_sequentially:
   "compact_space mtopology ⟷
    (∀σ::nat⇒'a. range σ ⊆ M
         ⟶ (∃l r. l ∈ M ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially))"
  by (simp add: compact_space_def compactin_sequentially)
lemma compact_space_eq_Bolzano_Weierstrass:
   "compact_space mtopology ⟷
    (∀S. S ⊆ M ∧ infinite S ⟶ mtopology derived_set_of S ≠ {})"
  using Int_absorb1 [OF derived_set_of_subset_topspace [of mtopology]]
  by (force simp: compact_space_def compactin_eq_Bolzano_Weierstrass)
lemma compact_space_nest:
   "compact_space mtopology ⟷
    (∀C. (∀n::nat. closedin mtopology (C n)) ∧ (∀n. C n ≠ {}) ∧ decseq C ⟶ ⋂(range C) ≠ {})"
   (is "?lhs=?rhs")
proof
  assume L: ?lhs
  show ?rhs
  proof clarify
    fix C :: "nat ⇒ 'a set"
    assume "∀n. closedin mtopology (C n)"
      and "∀n. C n ≠ {}"
      and "decseq C"
      and "⋂ (range C) = {}"
    then obtain K where K: "finite K" "⋂(C ` K) = {}"
      by (metis L compact_space_imp_nest)
    then obtain k where "K ⊆ {..k}"
      using finite_nat_iff_bounded_le by auto
    then have "C k ⊆ ⋂(C ` K)"
      using ‹decseq C› by (auto simp:decseq_def)
    then show False
      by (simp add: K ‹∀n. C n ≠ {}›)
  qed
next
  assume R [rule_format]: ?rhs
  show ?lhs
    unfolding compact_space_sequentially
  proof (intro strip)
    fix σ :: "nat ⇒ 'a"
    assume σ: "range σ ⊆ M"
    have "mtopology closure_of σ ` {n..} ≠ {}" for n
      using ‹range σ ⊆ M› by (auto simp: closure_of_eq_empty image_subset_iff)
    moreover have "decseq (λn. mtopology closure_of σ ` {n..})"
      using closure_of_mono image_mono by (smt (verit) atLeast_subset_iff decseq_def) 
    ultimately obtain l where l: "⋀n. l ∈ mtopology closure_of σ ` {n..}"
      using R [of "λn. mtopology closure_of (σ ` {n..})"] by auto
    then have "l ∈ M" and "⋀n. ∀r>0. ∃k≥n. σ k ∈ M ∧ d l (σ k) < r"
      using metric_closure_of by fastforce+
    then obtain f where f: "⋀n r. r>0 ⟹ f n r ≥ n ∧ σ (f n r) ∈ M ∧ d l (σ (f n r)) < r"
      by metis
    define r where "r = rec_nat (f 0 1) (λn rec. (f (Suc rec) (inverse (Suc (Suc n)))))"
    have r: "d l (σ(r n)) < inverse(Suc n)" for n
      by (induction n) (auto simp: rec_nat_0_imp [OF r_def] rec_nat_Suc_imp [OF r_def] f)
    have "r n < r(Suc n)" for n
      by (simp add: Suc_le_lessD f r_def)
    then have "strict_mono r"
      by (simp add: strict_mono_Suc_iff)
    moreover have "limitin mtopology (σ ∘ r) l sequentially"
      proof (clarsimp simp: limitin_metric ‹l ∈ M›)
        fix ε :: real
        assume "ε > 0"
        then have "(∀⇩F n in sequentially. inverse (real (Suc n)) < ε)"
          using Archimedean_eventually_inverse by blast
        then show "∀⇩F n in sequentially. σ (r n) ∈ M ∧ d (σ (r n)) l < ε"
          by eventually_elim (metis commute ‹range σ ⊆ M›  order_less_trans r range_subsetD)
      qed
    ultimately show "∃l r. l ∈ M ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially"
      using ‹l ∈ M› by blast
  qed
qed
lemma (in discrete_metric) mcomplete_discrete_metric: "disc.mcomplete"
proof (clarsimp simp: disc.mcomplete_def)
  fix σ :: "nat ⇒ 'a"
  assume "disc.MCauchy σ"
  then obtain N where "⋀n. N ≤ n ⟹ σ N = σ n"
    unfolding disc.MCauchy_def by (metis dd_def dual_order.refl order_less_irrefl zero_less_one)
  moreover have "range σ ⊆ M"
    using ‹disc.MCauchy σ› disc.MCauchy_def by blast
  ultimately have "limitin disc.mtopology σ (σ N) sequentially"
    by (metis disc.limit_metric_sequentially disc.zero range_subsetD)
  then show "∃x. limitin disc.mtopology σ x sequentially" ..
qed
lemma compact_space_imp_mcomplete: "compact_space mtopology ⟹ mcomplete"
  by (simp add: compact_space_nest mcomplete_nest)
lemma (in Submetric) compactin_imp_mcomplete:
   "compactin mtopology A ⟹ sub.mcomplete"
  by (simp add: compactin_subspace mtopology_submetric sub.compact_space_imp_mcomplete)
lemma (in Submetric) mcomplete_imp_closedin:
  assumes "sub.mcomplete"
  shows "closedin mtopology A"
proof -
  have "l ∈ A"
    if "range σ ⊆ A" and l: "limitin mtopology σ l sequentially"
    for σ :: "nat ⇒ 'a" and l
  proof -
    have "sub.MCauchy σ"
      using convergent_imp_MCauchy subset that by (force simp: MCauchy_submetric)
    then have "limitin sub.mtopology σ l sequentially"
      using assms unfolding sub.mcomplete_def
      using l limitin_metric_unique limitin_submetric_iff trivial_limit_sequentially by blast
    then show ?thesis
      using limitin_submetric_iff by blast
  qed
  then show ?thesis
    using metric_closedin_iff_sequentially_closed subset by auto
qed
lemma (in Submetric) closedin_eq_mcomplete:
   "mcomplete ⟹ (closedin mtopology A ⟷ sub.mcomplete)"
  using closedin_mcomplete_imp_mcomplete mcomplete_imp_closedin by blast
lemma compact_space_eq_mcomplete_mtotally_bounded:
   "compact_space mtopology ⟷ mcomplete ∧ mtotally_bounded M"
  by (meson Bolzano_Weierstrass_C compact_space_imp_mcomplete compact_space_sequentially limitin_mspace 
            mcomplete_alt mtotally_bounded_sequentially subset_refl)
lemma compact_closure_of_imp_mtotally_bounded:
   "⟦compactin mtopology (mtopology closure_of S); S ⊆ M⟧
      ⟹ mtotally_bounded S"
  using compactin_imp_mtotally_bounded mtotally_bounded_closure_of_eq by blast
lemma mtotally_bounded_eq_compact_closure_of:
  assumes "mcomplete"
  shows "mtotally_bounded S ⟷ S ⊆ M ∧ compactin mtopology (mtopology closure_of S)"
  (is "?lhs=?rhs")
proof
  assume L: ?lhs 
  show ?rhs
    unfolding compactin_subspace
  proof (intro conjI)
    show "S ⊆ M"
      using L by (simp add: mtotally_bounded_imp_subset)
    show "mtopology closure_of S ⊆ topspace mtopology"
      by (simp add: ‹S ⊆ M› closure_of_minimal)
    then have MSM: "mtopology closure_of S ⊆ M"
      by auto
    interpret S: Submetric M d "mtopology closure_of S"
    proof qed (use MSM in auto)
    have "S.sub.mtotally_bounded (mtopology closure_of S)"
      using L mtotally_bounded_absolute mtotally_bounded_closure_of by blast
    then
    show "compact_space (subtopology mtopology (mtopology closure_of S))"
      using S.closedin_mcomplete_imp_mcomplete S.mtopology_submetric S.sub.compact_space_eq_mcomplete_mtotally_bounded assms by force
  qed
qed (auto simp: compact_closure_of_imp_mtotally_bounded)
lemma compact_closure_of_eq_Bolzano_Weierstrass:
   "compactin mtopology (mtopology closure_of S) ⟷
    (∀T. infinite T ∧ T ⊆ S ∧ T ⊆ M ⟶ mtopology derived_set_of T ≠ {})"  (is "?lhs=?rhs")
proof
  assume L: ?lhs 
  show ?rhs
  proof (intro strip)
    fix T
    assume T: "infinite T ∧ T ⊆ S ∧ T ⊆ M"
    show "mtopology derived_set_of T ≠ {}"
    proof (intro compact_closure_of_imp_Bolzano_Weierstrass)
      show "compactin mtopology (mtopology closure_of S)"
        by (simp add: L)
    qed (use T in auto)
  qed
next
  have "compactin mtopology (mtopology closure_of S)"
    if §: "⋀T. ⟦infinite T; T ⊆ S⟧ ⟹ mtopology derived_set_of T ≠ {}" and "S ⊆ M" for S
    unfolding compactin_sequentially
  proof (intro conjI strip)
    show MSM: "mtopology closure_of S ⊆ M"
      using closure_of_subset_topspace by fastforce
    fix σ :: "nat ⇒ 'a"
    assume σ: "range σ ⊆ mtopology closure_of S"
    then have "∃y ∈ S. d (σ n) y < inverse(Suc n)" for n
      by (simp add: metric_closure_of image_subset_iff) (metis inverse_Suc of_nat_Suc)
    then obtain τ where τ: "⋀n. τ n ∈ S ∧ d (σ n) (τ n) < inverse(Suc n)"
      by metis
    then have "range τ ⊆ S"
      by blast
    moreover
    have *: "∀T. T ⊆ S ∧ infinite T ⟶ mtopology closure_of S ∩ mtopology derived_set_of T ≠ {}"
      using "§"(1) derived_set_of_mono derived_set_of_subset_closure_of by fastforce
    moreover have "S ⊆ mtopology closure_of S"
      by (simp add: ‹S ⊆ M› closure_of_subset)
    ultimately obtain l r where lr:
      "l ∈ mtopology closure_of S" "strict_mono r" "limitin mtopology (τ ∘ r) l sequentially"
      using Bolzano_Weierstrass_property ‹S ⊆ M› by metis
    then have "l ∈ M"
      using limitin_mspace by blast
    have dr_less: "d ((σ ∘ r) n) ((τ ∘ r) n) < inverse(Suc n)" for n
    proof -
      have "d ((σ ∘ r) n) ((τ ∘ r) n) < inverse(Suc (r n))"
        using τ by auto
      also have "... ≤ inverse(Suc n)"
        using lr strict_mono_imp_increasing by auto
      finally show ?thesis .
    qed
    have "limitin mtopology (σ ∘ r) l sequentially"
      unfolding limitin_metric
    proof (intro conjI strip)
      show "l ∈ M"
        using limitin_mspace lr by blast
      fix ε :: real
      assume "ε > 0"
      then have "∀⇩F n in sequentially. (τ ∘ r) n ∈ M ∧ d ((τ ∘ r) n) l < ε/2"
        using lr half_gt_zero limitin_metric by blast 
      moreover have "∀⇩F n in sequentially. inverse (real (Suc n)) < ε/2"
        using Archimedean_eventually_inverse ‹0 < ε› half_gt_zero by blast
      then have "∀⇩F n in sequentially. d ((σ ∘ r) n) ((τ ∘ r) n) < ε/2"
        by eventually_elim (smt (verit, del_insts) dr_less)
      ultimately have "∀⇩F n in sequentially. d ((σ ∘ r) n) l < ε/2 + ε/2"
        by eventually_elim (smt (verit) triangle ‹l ∈ M› MSM σ comp_apply order_trans range_subsetD)      
      then show "∀⇩F n in sequentially. (σ ∘ r) n ∈ M ∧ d ((σ ∘ r) n) l < ε"
        apply eventually_elim
        using ‹mtopology closure_of S ⊆ M› σ by auto
    qed
    with lr show "∃l r. l ∈ mtopology closure_of S ∧ strict_mono r ∧ limitin mtopology (σ ∘ r) l sequentially"
      by blast
  qed
  then show "?rhs ⟹ ?lhs"
    by (metis Int_subset_iff closure_of_restrict inf_le1 topspace_mtopology)
qed
end
lemma (in discrete_metric) mtotally_bounded_discrete_metric:
   "disc.mtotally_bounded S ⟷ finite S ∧ S ⊆ M" (is "?lhs=?rhs")
proof
  assume L: ?lhs 
  show ?rhs
  proof
    show "finite S"
      by (metis (no_types) L closure_of_subset_Int compactin_discrete_topology disc.mtotally_bounded_eq_compact_closure_of
          disc.topspace_mtopology discrete_metric.mcomplete_discrete_metric inf.absorb_iff2 mtopology_discrete_metric finite_subset)
    show "S ⊆ M"
      by (simp add: L disc.mtotally_bounded_imp_subset)
  qed
qed (simp add: disc.finite_imp_mtotally_bounded)
context Metric_space
begin
lemma derived_set_of_infinite_openin_metric:
   "mtopology derived_set_of S =
    {x ∈ M. ∀U. x ∈ U ∧ openin mtopology U ⟶ infinite(S ∩ U)}"
  by (simp add: derived_set_of_infinite_openin Hausdorff_space_mtopology)
lemma derived_set_of_infinite_1: 
  assumes "infinite (S ∩ mball x ε)" 
  shows "infinite (S ∩ mcball x ε)"
  by (meson Int_mono assms finite_subset mball_subset_mcball subset_refl)
lemma derived_set_of_infinite_2:
  assumes "openin mtopology U" "⋀ε. 0 < ε ⟹ infinite (S ∩ mcball x ε)" and "x ∈ U"
  shows "infinite (S ∩ U)"
  by (metis assms openin_mtopology_mcball finite_Int inf.absorb_iff2 inf_assoc)
lemma derived_set_of_infinite_mball:
  "mtopology derived_set_of S = {x ∈ M. ∀e>0. infinite(S ∩ mball x e)}"
  unfolding derived_set_of_infinite_openin_metric
  by (metis (no_types, opaque_lifting) centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2)
lemma derived_set_of_infinite_mcball:
  "mtopology derived_set_of S = {x ∈ M. ∀e>0. infinite(S ∩ mcball x e)}"
  unfolding derived_set_of_infinite_openin_metric
  by (metis (no_types, opaque_lifting) centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2)
end
subsection‹Continuous functions on metric spaces›
context Metric_space
begin
lemma continuous_map_to_metric:
   "continuous_map X mtopology f ⟷
    (∀x ∈ topspace X. ∀ε>0. ∃U. openin X U ∧ x ∈ U ∧ (∀y∈U. f y ∈ mball (f x) ε))"
   (is "?lhs=?rhs")
proof
  show "?lhs ⟹ ?rhs"
    unfolding continuous_map_eq_topcontinuous_at topcontinuous_at_def
    by (metis PiE centre_in_mball_iff openin_mball topspace_mtopology)
next
  assume R: ?rhs
  then have "∀x∈topspace X. f x ∈ M"
    by (meson gt_ex in_mball)
  moreover 
  have "⋀x V. ⟦x ∈ topspace X; openin mtopology V; f x ∈ V⟧ ⟹ ∃U. openin X U ∧ x ∈ U ∧ (∀y∈U. f y ∈ V)"
    unfolding openin_mtopology by (metis Int_iff R inf.orderE)
  ultimately
  show ?lhs
    by (simp add: continuous_map_eq_topcontinuous_at topcontinuous_at_def)
qed 
lemma continuous_map_from_metric:
   "continuous_map mtopology X f ⟷
    f ∈ M → topspace X ∧
    (∀a ∈ M. ∀U. openin X U ∧ f a ∈ U ⟶ (∃r>0. ∀x. x ∈ M ∧ d a x < r ⟶ f x ∈ U))"
proof (cases "f ` M ⊆ topspace X")
  case True
  then show ?thesis
    by (fastforce simp: continuous_map openin_mtopology subset_eq)
next
  case False
  then show ?thesis
    by (simp add: continuous_map_def image_subset_iff_funcset)
qed
text ‹An abstract formulation, since the limits do not have to be sequential›
lemma continuous_map_uniform_limit:
  assumes contf: "∀⇩F ξ in F. continuous_map X mtopology (f ξ)"
    and dfg: "⋀ε. 0 < ε ⟹ ∀⇩F ξ in F. ∀x ∈ topspace X. g x ∈ M ∧ d (f ξ x) (g x) < ε"
    and nontriv: "¬ trivial_limit F"
  shows "continuous_map X mtopology g"
  unfolding continuous_map_to_metric
proof (intro strip)
  fix x and ε::real
  assume "x ∈ topspace X" and "ε > 0"
  then obtain ξ where k: "continuous_map X mtopology (f ξ)" 
    and gM: "∀x ∈ topspace X. g x ∈ M" 
    and third: "∀x ∈ topspace X. d (f ξ x) (g x) < ε/3"
    using eventually_conj [OF contf] contf dfg [of "ε/3"] eventually_happens' [OF nontriv]
    by (smt (verit, ccfv_SIG) zero_less_divide_iff)
  then obtain U where U: "openin X U" "x ∈ U" and Uthird: "∀y∈U. d (f ξ y) (f ξ x) < ε/3"
    unfolding continuous_map_to_metric
    by (metis ‹0 < ε› ‹x ∈ topspace X› commute divide_pos_pos in_mball zero_less_numeral)
  have f_inM: "f ξ y ∈ M" if "y∈U" for y
    using U k openin_subset that by (fastforce simp: continuous_map_def)
  have "d (g y) (g x) < ε" if "y∈U" for y
  proof -
    have "g y ∈ M"
      using U gM openin_subset that by blast
    have "d (g y) (g x) ≤ d (g y) (f ξ x) + d (f ξ x) (g x)"
      by (simp add: U ‹g y ∈ M› ‹x ∈ topspace X› f_inM gM triangle)
    also have "… ≤ d (g y) (f ξ y) + d (f ξ y) (f ξ x) + d (f ξ x) (g x)"
      by (simp add: U ‹g y ∈ M› commute f_inM that triangle')
    also have "… < ε/3 + ε/3 + ε/3"
      by (smt (verit) U(1) Uthird ‹x ∈ topspace X› commute openin_subset subsetD that third)
    finally show ?thesis by simp
  qed
  with U gM show "∃U. openin X U ∧ x ∈ U ∧ (∀y∈U. g y ∈ mball (g x) ε)"
    by (metis commute in_mball in_mono openin_subset)
qed
lemma continuous_map_uniform_limit_alt:
  assumes contf: "∀⇩F ξ in F. continuous_map X mtopology (f ξ)"
    and gim: "g ∈ topspace X → M"
    and dfg: "⋀ε. 0 < ε ⟹ ∀⇩F ξ in F. ∀x ∈ topspace X. d (f ξ x) (g x) < ε"
    and nontriv: "¬ trivial_limit F"
  shows "continuous_map X mtopology g"
proof (rule continuous_map_uniform_limit [OF contf])
  fix ε :: real
  assume "ε > 0"
  with gim dfg show "∀⇩F ξ in F. ∀x∈topspace X. g x ∈ M ∧ d (f ξ x) (g x) < ε"
    by (simp add: Pi_iff)
qed (use nontriv in auto)
lemma continuous_map_uniformly_Cauchy_limit:
  assumes "mcomplete"
  assumes contf: "∀⇩F n in sequentially. continuous_map X mtopology (f n)"
    and Cauchy': "⋀ε. ε > 0 ⟹ ∃N. ∀m n x. N ≤ m ⟶ N ≤ n ⟶ x ∈ topspace X ⟶ d (f m x) (f n x) < ε"
  obtains g where
    "continuous_map X mtopology g"
    "⋀ε. 0 < ε ⟹ ∀⇩F n in sequentially. ∀x∈topspace X. d (f n x) (g x) < ε"
proof -
  have "⋀x. x ∈ topspace X ⟹ ∃l. limitin mtopology (λn. f n x) l sequentially"
    using ‹mcomplete› [unfolded mcomplete, rule_format] assms
    unfolding continuous_map_def Pi_iff topspace_mtopology
    by (smt (verit, del_insts) eventually_mono)
  then obtain g where g: "⋀x. x ∈ topspace X ⟹ limitin mtopology (λn. f n x) (g x) sequentially"
    by metis
  show thesis
  proof
    show "∀⇩F n in sequentially. ∀x∈topspace X. d (f n x) (g x) < ε"
      if "ε > 0" for ε :: real
    proof -
      obtain N where N: "⋀m n x. ⟦N ≤ m; N ≤ n; x ∈ topspace X⟧ ⟹ d (f m x) (f n x) < ε/2"
        by (meson Cauchy' ‹0 < ε› half_gt_zero)
      obtain P where P: "⋀n x. ⟦n ≥ P; x ∈ topspace X⟧ ⟹ f n x ∈ M"
        using contf by (auto simp: eventually_sequentially continuous_map_def)
      show ?thesis
      proof (intro eventually_sequentiallyI strip)
        fix n x
        assume "max N P ≤ n" and x: "x ∈ topspace X"
        obtain L where "g x ∈ M" and L: "∀n≥L. f n x ∈ M ∧ d (f n x) (g x) < ε/2"
          using g [OF x] ‹ε > 0› unfolding limitin_metric
          by (metis (no_types, lifting) eventually_sequentially half_gt_zero)
        define n' where "n' ≡ Max{L,N,P}"
        have L': "∀m ≥ n'. f m x ∈ M ∧ d (f m x) (g x) < ε/2"
          using L by (simp add: n'_def)
        moreover
        have "d (f n x) (f n' x) < ε/2"
          using N [of n n' x] ‹max N P ≤ n› n'_def x by fastforce
        ultimately have "d (f n x) (g x) < ε/2 + ε/2"
          by (smt (verit, ccfv_SIG) P ‹g x ∈ M› ‹max N P ≤ n› le_refl max.bounded_iff mdist_zero triangle' x)
        then show "d (f n x) (g x) < ε" by simp
      qed
    qed
    then show "continuous_map X mtopology g"
      by (smt (verit, del_insts) eventually_mono g limitin_mspace trivial_limit_sequentially continuous_map_uniform_limit [OF contf])
  qed
qed
lemma metric_continuous_map:
  assumes "Metric_space M' d'"
  shows
   "continuous_map mtopology (Metric_space.mtopology M' d') f ⟷
    f ` M ⊆ M' ∧ (∀a ∈ M. ∀ε>0. ∃δ>0.  (∀x. x ∈ M ∧ d a x < δ ⟶ d' (f a) (f x) < ε))"
   (is "?lhs = ?rhs")
proof -
  interpret M': Metric_space M' d'
    by (simp add: assms)
  show ?thesis
  proof
    assume L: ?lhs
    show ?rhs
    proof (intro conjI strip)
      show "f ` M ⊆ M'"
        using L by (auto simp: continuous_map_def)
      fix a and ε :: real
      assume "a ∈ M" and "ε > 0"
      then have "openin mtopology {x ∈ M. f x ∈ M'.mball (f a) ε}" "f a ∈ M'"
        using L unfolding continuous_map_def by fastforce+
      then obtain δ where "δ > 0" "mball a δ ⊆ {x ∈ M. f x ∈ M' ∧ d' (f a) (f x) < ε}"
        using ‹0 < ε› ‹a ∈ M› openin_mtopology by auto
      then show "∃δ>0. ∀x. x ∈ M ∧ d a x < δ ⟶ d' (f a) (f x) < ε"
        using ‹a ∈ M› in_mball by blast
    qed
  next
    assume R: ?rhs    
    show ?lhs
      unfolding continuous_map_def
    proof (intro conjI strip)
      fix U
      assume "openin M'.mtopology U"
      then show "openin mtopology {x ∈ topspace mtopology. f x ∈ U}"
        using R
        by (force simp: continuous_map_def openin_mtopology M'.openin_mtopology subset_iff)
    qed (use R in auto)
  qed
qed
end 
subsection ‹Completely metrizable spaces›
text ‹These spaces are topologically complete›
definition completely_metrizable_space where
 "completely_metrizable_space X ≡ 
     ∃M d. Metric_space M d ∧ Metric_space.mcomplete M d ∧ X = Metric_space.mtopology M d"
lemma empty_completely_metrizable_space: 
  "completely_metrizable_space trivial_topology"
  unfolding completely_metrizable_space_def subtopology_eq_discrete_topology_empty [symmetric]
  by (metis Metric_space.mcomplete_empty_mspace discrete_metric.mtopology_discrete_metric metric_M_dd)
lemma completely_metrizable_imp_metrizable_space:
   "completely_metrizable_space X ⟹ metrizable_space X"
  using completely_metrizable_space_def metrizable_space_def by auto
lemma (in Metric_space) completely_metrizable_space_mtopology:
   "mcomplete ⟹ completely_metrizable_space mtopology"
  using Metric_space_axioms completely_metrizable_space_def by blast
lemma completely_metrizable_space_discrete_topology:
   "completely_metrizable_space (discrete_topology U)"
  unfolding completely_metrizable_space_def
  by (metis discrete_metric.mcomplete_discrete_metric discrete_metric.mtopology_discrete_metric metric_M_dd)
lemma completely_metrizable_space_euclidean:
    "completely_metrizable_space (euclidean:: 'a::complete_space topology)"
  using Met_TC.completely_metrizable_space_mtopology complete_UNIV by auto
lemma completely_metrizable_space_closedin:
  assumes X: "completely_metrizable_space X" and S: "closedin X S"
  shows "completely_metrizable_space(subtopology X S)"
proof -
  obtain M d where "Metric_space M d" and comp: "Metric_space.mcomplete M d" 
                and Xeq: "X = Metric_space.mtopology M d"
    using assms completely_metrizable_space_def by blast
  then interpret Metric_space M d
    by blast
  show ?thesis
    unfolding completely_metrizable_space_def
  proof (intro conjI exI)
    show "Metric_space S d"
      using S Xeq closedin_subset subspace by force
    have sub: "Submetric_axioms M S"
      by (metis S Xeq closedin_metric Submetric_axioms_def)
    then show "Metric_space.mcomplete S d"
      using S Submetric.closedin_mcomplete_imp_mcomplete Submetric_def Xeq comp by blast
    show "subtopology X S = Metric_space.mtopology S d"
      by (metis Metric_space_axioms Xeq sub Submetric.intro Submetric.mtopology_submetric)
  qed
qed
lemma completely_metrizable_space_cbox: "completely_metrizable_space (top_of_set (cbox a b))"
    using closed_closedin completely_metrizable_space_closedin completely_metrizable_space_euclidean by blast
lemma homeomorphic_completely_metrizable_space_aux:
  assumes homXY: "X homeomorphic_space Y" and X: "completely_metrizable_space X"
  shows "completely_metrizable_space Y"
proof -
  obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g"
    and fg: "⋀x. x ∈ topspace X ⟹ g(f x) = x" "⋀y. y ∈ topspace Y ⟹ f(g y) = y"
    and fim: "f ∈ topspace X → topspace Y" and gim: "g ∈ topspace Y → topspace X"
    using homXY
    using homeomorphic_space_unfold by blast
  obtain M d where Md: "Metric_space M d" "Metric_space.mcomplete M d" and Xeq: "X = Metric_space.mtopology M d"
    using X by (auto simp: completely_metrizable_space_def)
  then interpret MX: Metric_space M d by metis
  define D where "D ≡ λx y. d (g x) (g y)"
  have "Metric_space (topspace Y) D"
  proof
    show "(D x y = 0) ⟷ (x = y)" if "x ∈ topspace Y" "y ∈ topspace Y" for x y
      unfolding D_def
      by (metis that MX.topspace_mtopology MX.zero Xeq fg gim Pi_iff)
    show "D x z ≤ D x y +D y z"
      if "x ∈ topspace Y" "y ∈ topspace Y" "z ∈ topspace Y" for x y z
      using that MX.triangle Xeq gim by (auto simp: D_def)
  qed (auto simp: D_def MX.commute)
  then interpret MY: Metric_space "topspace Y" "λx y. D x y" by metis
  show ?thesis
    unfolding completely_metrizable_space_def
  proof (intro exI conjI)
    show "Metric_space (topspace Y) D"
      using MY.Metric_space_axioms by blast
    have gball: "g ` MY.mball y r = MX.mball (g y) r" if "y ∈ topspace Y" for y r
      using that MX.topspace_mtopology Xeq gim hmg homeomorphic_imp_surjective_map
      unfolding MX.mball_def MY.mball_def  by (fastforce simp: D_def)
    have "∃r>0. MY.mball y r ⊆ S" if "openin Y S" and "y ∈ S" for S y
    proof -
      have "openin X (g`S)"
        using hmg homeomorphic_map_openness_eq that by auto
      then obtain r where "r>0" "MX.mball (g y) r ⊆ g`S"
        using MX.openin_mtopology Xeq ‹y ∈ S› by auto
      then show ?thesis
        by (smt (verit, ccfv_SIG) MY.in_mball gball fg image_iff in_mono openin_subset subsetI that(1))
    qed
    moreover have "openin Y S"
      if "S ⊆ topspace Y" and "⋀y. y ∈ S ⟹ ∃r>0. MY.mball y r ⊆ S" for S
    proof -
      have "⋀x. x ∈ g`S ⟹ ∃r>0. MX.mball x r ⊆ g`S"
        by (smt (verit) gball imageE image_mono subset_iff that)
      then have "openin X (g`S)"
        using MX.openin_mtopology Xeq gim that(1) by auto
      then show ?thesis
        using hmg homeomorphic_map_openness_eq that(1) by blast
    qed
    ultimately show Yeq: "Y = MY.mtopology"
      unfolding topology_eq MY.openin_mtopology by (metis openin_subset)
    show "MY.mcomplete"
      unfolding MY.mcomplete_def
    proof (intro strip)
      fix σ
      assume σ: "MY.MCauchy σ"
      have "MX.MCauchy (g ∘ σ)"
        unfolding MX.MCauchy_def 
      proof (intro conjI strip)
        show "range (g ∘ σ) ⊆ M"
          using MY.MCauchy_def Xeq σ gim by auto
        fix ε :: real
        assume "ε > 0"
        then obtain N where "∀n n'. N ≤ n ⟶ N ≤ n' ⟶ D (σ n) (σ n') < ε"
          using MY.MCauchy_def σ by presburger
        then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d ((g ∘ σ) n) ((g ∘ σ) n') < ε"
          by (auto simp: o_def D_def)
      qed
      then obtain x where x: "limitin MX.mtopology (g ∘ σ) x sequentially" "x ∈ topspace X"
        using MX.limitin_mspace MX.topspace_mtopology Md Xeq unfolding MX.mcomplete_def
        by blast
      with x have "limitin MY.mtopology (f ∘ (g ∘ σ)) (f x) sequentially"
        by (metis Xeq Yeq continuous_map_limit hmf homeomorphic_imp_continuous_map)
      moreover have "f ∘ (g ∘ σ) = σ"
        using ‹MY.MCauchy σ›  by (force simp: fg MY.MCauchy_def subset_iff)
      ultimately have "limitin MY.mtopology σ (f x) sequentially" by simp
      then show "∃y. limitin MY.mtopology σ y sequentially"
        by blast 
    qed
  qed
qed
lemma homeomorphic_completely_metrizable_space:
   "X homeomorphic_space Y
        ⟹ completely_metrizable_space X ⟷ completely_metrizable_space Y"
  by (meson homeomorphic_completely_metrizable_space_aux homeomorphic_space_sym)
lemma completely_metrizable_space_retraction_map_image:
  assumes r: "retraction_map X Y r" and X: "completely_metrizable_space X"
  shows "completely_metrizable_space Y"
proof -
  obtain s where s: "retraction_maps X Y r s"
    using r retraction_map_def by blast
  then have "subtopology X (s ` topspace Y) homeomorphic_space Y"
    using retraction_maps_section_image2 by blast
  then show ?thesis
    by (metis X retract_of_space_imp_closedin retraction_maps_section_image1 
        homeomorphic_completely_metrizable_space completely_metrizable_space_closedin 
        completely_metrizable_imp_metrizable_space metrizable_imp_Hausdorff_space s)
qed
subsection ‹Product metric›
text‹For the nicest fit with the main Euclidean theories, we choose the Euclidean product, 
though other definitions of the product work.›
definition "prod_dist ≡ λd1 d2 (x,y) (x',y'). sqrt(d1 x x' ^ 2 + d2 y y' ^ 2)"
locale Metric_space12 = M1: Metric_space M1 d1 + M2: Metric_space M2 d2 for M1 d1 M2 d2
lemma (in Metric_space12) prod_metric: "Metric_space (M1 × M2) (prod_dist d1 d2)"
proof
  fix x y z
  assume xyz: "x ∈ M1 × M2" "y ∈ M1 × M2" "z ∈ M1 × M2"
  have "sqrt ((d1 x1 z1)⇧2 + (d2 x2 z2)⇧2) ≤ sqrt ((d1 x1 y1)⇧2 + (d2 x2 y2)⇧2) + sqrt ((d1 y1 z1)⇧2 + (d2 y2 z2)⇧2)"
      (is "sqrt ?L ≤ ?R")
    if "x = (x1, x2)" "y = (y1, y2)" "z = (z1, z2)"
    for x1 x2 y1 y2 z1 z2
  proof -
    have tri: "d1 x1 z1 ≤ d1 x1 y1 + d1 y1 z1" "d2 x2 z2 ≤ d2 x2 y2 + d2 y2 z2"
      using that xyz M1.triangle [of x1 y1 z1] M2.triangle [of x2 y2 z2] by auto
    show ?thesis
    proof (rule real_le_lsqrt)
      have "?L ≤ (d1 x1 y1 + d1 y1 z1)⇧2 + (d2 x2 y2 + d2 y2 z2)⇧2"
        using tri by (smt (verit) M1.nonneg M2.nonneg power_mono)
      also have "... ≤ ?R⇧2"
        by (metis real_sqrt_sum_squares_triangle_ineq sqrt_le_D)
      finally show "?L ≤ ?R⇧2" .
    qed auto
  qed
  then show "prod_dist d1 d2 x z ≤ prod_dist d1 d2 x y + prod_dist d1 d2 y z"
    by (simp add: prod_dist_def case_prod_unfold)
qed (auto simp: M1.commute M2.commute case_prod_unfold prod_dist_def)
sublocale Metric_space12 ⊆ Prod_metric: Metric_space "M1×M2" "prod_dist d1 d2" 
  by (simp add: prod_metric)
text ‹For easy reference to theorems outside of the locale›
lemma Metric_space12_mspace_mdist:
  "Metric_space12 (mspace m1) (mdist m1) (mspace m2) (mdist m2)"
  by (simp add: Metric_space12_def)
definition prod_metric where
 "prod_metric ≡ λm1 m2. metric (mspace m1 × mspace m2, prod_dist (mdist m1) (mdist m2))"
lemma submetric_prod_metric:
   "submetric (prod_metric m1 m2) (S × T) = prod_metric (submetric m1 S) (submetric m2 T)"
  apply (simp add: prod_metric_def)
  by (simp add: submetric_def Metric_space.mspace_metric Metric_space.mdist_metric Metric_space12.prod_metric Metric_space12_def Times_Int_Times)
lemma mspace_prod_metric [simp]:"
  mspace (prod_metric m1 m2) = mspace m1 × mspace m2"
  by (simp add: prod_metric_def Metric_space.mspace_metric Metric_space12.prod_metric Metric_space12_mspace_mdist)
lemma mdist_prod_metric [simp]: 
  "mdist (prod_metric m1 m2) = prod_dist (mdist m1) (mdist m2)"
  by (metis Metric_space.mdist_metric Metric_space12.prod_metric Metric_space12_mspace_mdist prod_metric_def)
lemma prod_dist_dist [simp]: "prod_dist dist dist = dist"
  by (simp add: prod_dist_def dist_prod_def fun_eq_iff)
lemma prod_metric_euclidean [simp]:
  "prod_metric euclidean_metric euclidean_metric = euclidean_metric"
  by (simp add: prod_metric_def euclidean_metric_def)
context Metric_space12
begin
lemma component_le_prod_metric:
   "d1 x1 x2 ≤ prod_dist d1 d2 (x1,y1) (x2,y2)" "d2 y1 y2 ≤ prod_dist d1 d2 (x1,y1) (x2,y2)"
  by (auto simp: prod_dist_def)
lemma prod_metric_le_components:
  "⟦x1 ∈ M1; y1 ∈ M1; x2 ∈ M2; y2 ∈ M2⟧
    ⟹ prod_dist d1 d2 (x1,x2) (y1,y2) ≤ d1 x1 y1 + d2 x2 y2"
  by (auto simp: prod_dist_def sqrt_sum_squares_le_sum)
lemma mball_prod_metric_subset:
   "Prod_metric.mball (x,y) r ⊆ M1.mball x r × M2.mball y r"
  by clarsimp (smt (verit, best) component_le_prod_metric)
lemma mcball_prod_metric_subset:
   "Prod_metric.mcball (x,y) r ⊆ M1.mcball x r × M2.mcball y r"
  by clarsimp (smt (verit, best) component_le_prod_metric)
lemma mball_subset_prod_metric:
   "M1.mball x1 r1 × M2.mball x2 r2 ⊆ Prod_metric.mball (x1,x2) (r1 + r2)"
  using prod_metric_le_components by force
lemma mcball_subset_prod_metric:
   "M1.mcball x1 r1 × M2.mcball x2 r2 ⊆ Prod_metric.mcball (x1,x2) (r1 + r2)"
  using prod_metric_le_components by force
lemma mtopology_prod_metric:
  "Prod_metric.mtopology = prod_topology M1.mtopology M2.mtopology"
  unfolding prod_topology_def
proof (rule topology_base_unique [symmetric])
  fix U
  assume "U ∈ {S × T |S T. openin M1.mtopology S ∧ openin M2.mtopology T}"
  then obtain S T where Ueq: "U = S × T"
    and S: "openin M1.mtopology S" and T: "openin M2.mtopology T"
    by auto
  have "S ⊆ M1"
    using M1.openin_mtopology S by auto
  have "T ⊆ M2"
    using M2.openin_mtopology T by auto
  show "openin Prod_metric.mtopology U"
    unfolding Prod_metric.openin_mtopology
  proof (intro conjI strip)
    show "U ⊆ M1 × M2"
      using Ueq by (simp add: Sigma_mono ‹S ⊆ M1› ‹T ⊆ M2›)
    fix z
    assume "z ∈ U"
    then obtain x1 x2 where "x1 ∈ S" "x2 ∈ T" and zeq: "z = (x1,x2)"
      using Ueq by blast
    obtain r1 where "r1>0" and r1: "M1.mball x1 r1 ⊆ S"
      by (meson M1.openin_mtopology ‹openin M1.mtopology S› ‹x1 ∈ S›)
    obtain r2 where "r2>0" and r2: "M2.mball x2 r2 ⊆ T"
      by (meson M2.openin_mtopology ‹openin M2.mtopology T› ‹x2 ∈ T›)
    have "Prod_metric.mball (x1,x2) (min r1 r2) ⊆ U"
    proof (rule order_trans [OF mball_prod_metric_subset])
      show "M1.mball x1 (min r1 r2) × M2.mball x2 (min r1 r2) ⊆ U"
        using Ueq r1 r2 by force
    qed
    then show "∃r>0. Prod_metric.mball z r ⊆ U"
      by (smt (verit, del_insts) zeq ‹0 < r1› ‹0 < r2›)
  qed
next
  fix U z
  assume "openin Prod_metric.mtopology U" and "z ∈ U"
  then have "U ⊆ M1 × M2"
    by (simp add: Prod_metric.openin_mtopology)
  then obtain x y where "x ∈ M1" "y ∈ M2" and zeq: "z = (x,y)"
    using ‹z ∈ U› by blast
  obtain r where "r>0" and r: "Prod_metric.mball (x,y) r ⊆ U"
    by (metis Prod_metric.openin_mtopology ‹openin Prod_metric.mtopology U› ‹z ∈ U› zeq)
  define B1 where "B1 ≡ M1.mball x (r/2)"
  define B2 where "B2 ≡ M2.mball y (r/2)"
  have "openin M1.mtopology B1" "openin M2.mtopology B2"
    by (simp_all add: B1_def B2_def)
  moreover have "(x,y) ∈ B1 × B2"
    using ‹r > 0› by (simp add: ‹x ∈ M1› ‹y ∈ M2› B1_def B2_def)
  moreover have "B1 × B2 ⊆ U"
    using r prod_metric_le_components by (force simp: B1_def B2_def)
  ultimately show "∃B. B ∈ {S × T |S T. openin M1.mtopology S ∧ openin M2.mtopology T} ∧ z ∈ B ∧ B ⊆ U"
    by (auto simp: zeq)
qed
lemma MCauchy_prod_metric:
   "Prod_metric.MCauchy σ ⟷ M1.MCauchy (fst ∘ σ) ∧ M2.MCauchy (snd ∘ σ)"
   (is "?lhs ⟷ ?rhs")
proof safe
  assume L: ?lhs
  then have "range σ ⊆ M1 × M2"
    using Prod_metric.MCauchy_def by blast
  then have 1: "range (fst ∘ σ) ⊆ M1" and 2: "range (snd ∘ σ) ⊆ M2"
    by auto
  have N1: "∃N. ∀n≥N. ∀n'≥N. d1 (fst (σ n)) (fst (σ n')) < ε" 
    and N2: "∃N. ∀n≥N. ∀n'≥N. d2 (snd (σ n)) (snd (σ n')) < ε" if "ε>0" for ε :: real
    using that L unfolding Prod_metric.MCauchy_def
    by (smt (verit, del_insts) add.commute add_less_imp_less_left add_right_mono 
        component_le_prod_metric prod.collapse)+
  show "M1.MCauchy (fst ∘ σ)"
    using 1 N1 M1.MCauchy_def by auto
  have "∃N. ∀n≥N. ∀n'≥N. d2 (snd (σ n)) (snd (σ n')) < ε" if "ε>0" for ε :: real
    using that L unfolding Prod_metric.MCauchy_def
    by (smt (verit, del_insts) add.commute add_less_imp_less_left add_right_mono 
        component_le_prod_metric prod.collapse)
  show "M2.MCauchy (snd ∘ σ)"
    using 2 N2 M2.MCauchy_def by auto
next
  assume M1: "M1.MCauchy (fst ∘ σ)" and M2: "M2.MCauchy (snd ∘ σ)"
  then have subM12: "range (fst ∘ σ) ⊆ M1" "range (snd ∘ σ) ⊆ M2"
    using M1.MCauchy_def M2.MCauchy_def by blast+
  show ?lhs
    unfolding Prod_metric.MCauchy_def
  proof (intro conjI strip)
    show "range σ ⊆ M1 × M2"
      using subM12 by (smt (verit, best) SigmaI image_subset_iff o_apply prod.collapse) 
    fix ε :: real
    assume "ε > 0"
    obtain N1 where N1: "⋀n n'. N1 ≤ n ⟹ N1 ≤ n' ⟹ d1 ((fst ∘ σ) n) ((fst ∘ σ) n') < ε/2"
      by (meson M1.MCauchy_def ‹0 < ε› M1 zero_less_divide_iff zero_less_numeral)
    obtain N2 where N2: "⋀n n'. N2 ≤ n ⟹ N2 ≤ n' ⟹ d2 ((snd ∘ σ) n) ((snd ∘ σ) n') < ε/2"
      by (meson M2.MCauchy_def ‹0 < ε› M2 zero_less_divide_iff zero_less_numeral)
    have "prod_dist d1 d2 (σ n) (σ n') < ε"
      if "N1 ≤ n" and "N2 ≤ n" and "N1 ≤ n'" and "N2 ≤ n'" for n n'
    proof -
      obtain a b a' b' where σ: "σ n = (a,b)" "σ n' = (a',b')"
        by fastforce+
      have "prod_dist d1 d2 (a,b) (a',b') ≤ d1 a a' + d2 b b'"
        by (metis ‹range σ ⊆ M1 × M2› σ mem_Sigma_iff prod_metric_le_components range_subsetD)
      also have "… < ε/2 + ε/2"
        using N1 N2 σ that by fastforce
      finally show ?thesis
        by (simp add: σ)
    qed
    then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ prod_dist d1 d2 (σ n) (σ n') < ε"
      by (metis order.trans linorder_le_cases)
  qed
qed
lemma mcomplete_prod_metric:
  "Prod_metric.mcomplete ⟷ M1 = {} ∨ M2 = {} ∨ M1.mcomplete ∧ M2.mcomplete"
  (is "?lhs ⟷ ?rhs")
proof (cases "M1 = {} ∨ M2 = {}")
  case False
  then obtain x y where "x ∈ M1" "y ∈ M2"
    by blast
  have "M1.mcomplete ∧ M2.mcomplete ⟹ Prod_metric.mcomplete"
    by (simp add: Prod_metric.mcomplete_def M1.mcomplete_def M2.mcomplete_def 
        mtopology_prod_metric MCauchy_prod_metric limitin_pairwise)
  moreover
  { assume L: "Prod_metric.mcomplete"
    have "M1.mcomplete"
      unfolding M1.mcomplete_def
    proof (intro strip)
      fix σ
      assume "M1.MCauchy σ"
      then have "Prod_metric.MCauchy (λn. (σ n, y))"
        using ‹y ∈ M2› by (simp add: M1.MCauchy_def M2.MCauchy_def MCauchy_prod_metric)
      then obtain z where "limitin Prod_metric.mtopology (λn. (σ n, y)) z sequentially"
        using L Prod_metric.mcomplete_def by blast
      then show "∃x. limitin M1.mtopology σ x sequentially"
        by (auto simp: Prod_metric.mcomplete_def M1.mcomplete_def 
             mtopology_prod_metric limitin_pairwise o_def)
    qed
  }
  moreover
  { assume L: "Prod_metric.mcomplete"
    have "M2.mcomplete"
      unfolding M2.mcomplete_def
    proof (intro strip)
      fix σ
      assume "M2.MCauchy σ"
      then have "Prod_metric.MCauchy (λn. (x, σ n))"
        using ‹x ∈ M1› by (simp add: M2.MCauchy_def M1.MCauchy_def MCauchy_prod_metric)
      then obtain z where "limitin Prod_metric.mtopology (λn. (x, σ n)) z sequentially"
        using L Prod_metric.mcomplete_def by blast
      then show "∃x. limitin M2.mtopology σ x sequentially"
        by (auto simp: Prod_metric.mcomplete_def M2.mcomplete_def 
             mtopology_prod_metric limitin_pairwise o_def)
    qed
  }
  ultimately show ?thesis
    using False by blast 
qed auto
lemma mbounded_prod_metric:
   "Prod_metric.mbounded U ⟷ M1.mbounded  (fst ` U) ∧ M2.mbounded (snd ` U)"
proof -
  have "(∃B. U ⊆ Prod_metric.mcball (x,y) B) 
    ⟷ ((∃B. (fst ` U) ⊆ M1.mcball x B) ∧ (∃B. (snd ` U) ⊆ M2.mcball y B))" (is "?lhs ⟷ ?rhs")
    for x y
  proof safe
    fix B
    assume "U ⊆ Prod_metric.mcball (x, y) B"
    then have "(fst ` U) ⊆ M1.mcball x B" "(snd ` U) ⊆ M2.mcball y B"
      using mcball_prod_metric_subset by fastforce+
    then show "∃B. (fst ` U) ⊆ M1.mcball x B" "∃B. (snd ` U) ⊆ M2.mcball y B"
      by auto
  next
    fix B1 B2
    assume "(fst ` U) ⊆ M1.mcball x B1" "(snd ` U) ⊆ M2.mcball y B2"
    then have "fst ` U × snd ` U ⊆  M1.mcball x B1 × M2.mcball y B2"
      by blast
    also have "… ⊆ Prod_metric.mcball (x, y) (B1+B2)"
      by (intro mcball_subset_prod_metric)
    finally show "∃B. U ⊆ Prod_metric.mcball (x, y) B"
      by (metis subsetD subsetI subset_fst_snd)
  qed
  then show ?thesis
    by (simp add: M1.mbounded_def M2.mbounded_def Prod_metric.mbounded_def)
qed 
lemma mbounded_Times:
   "Prod_metric.mbounded (S × T) ⟷ S = {} ∨ T = {} ∨ M1.mbounded S ∧ M2.mbounded T"
  by (auto simp: mbounded_prod_metric)
lemma mtotally_bounded_Times:
   "Prod_metric.mtotally_bounded (S × T) ⟷
    S = {} ∨ T = {} ∨ M1.mtotally_bounded S ∧ M2.mtotally_bounded T"
    (is "?lhs ⟷ _")
proof (cases "S = {} ∨ T = {}")
  case False
  then obtain x y where "x ∈ S" "y ∈ T"
    by auto
  have "M1.mtotally_bounded S" if L: ?lhs
    unfolding M1.mtotally_bounded_sequentially
  proof (intro conjI strip)
    show "S ⊆ M1"
      using Prod_metric.mtotally_bounded_imp_subset ‹y ∈ T› that by blast
    fix σ :: "nat ⇒ 'a"
    assume "range σ ⊆ S"
    with L obtain r where "strict_mono r" "Prod_metric.MCauchy ((λn. (σ n,y)) ∘ r)"
      unfolding Prod_metric.mtotally_bounded_sequentially
      by (smt (verit) SigmaI ‹y ∈ T› image_subset_iff)
    then have "M1.MCauchy (fst ∘ (λn. (σ n,y)) ∘ r)"
      by (simp add: MCauchy_prod_metric o_def)
    with ‹strict_mono r› show "∃r. strict_mono r ∧ M1.MCauchy (σ ∘ r)"
      by (auto simp: o_def)
  qed
  moreover
  have "M2.mtotally_bounded T" if L: ?lhs
    unfolding M2.mtotally_bounded_sequentially
  proof (intro conjI strip)
    show "T ⊆ M2"
      using Prod_metric.mtotally_bounded_imp_subset ‹x ∈ S› that by blast
    fix σ :: "nat ⇒ 'b"
    assume "range σ ⊆ T"
    with L obtain r where "strict_mono r" "Prod_metric.MCauchy ((λn. (x,σ n)) ∘ r)"
      unfolding Prod_metric.mtotally_bounded_sequentially
      by (smt (verit) SigmaI ‹x ∈ S› image_subset_iff)
    then have "M2.MCauchy (snd ∘ (λn. (x,σ n)) ∘ r)"
      by (simp add: MCauchy_prod_metric o_def)
    with ‹strict_mono r› show "∃r. strict_mono r ∧ M2.MCauchy (σ ∘ r)"
      by (auto simp: o_def)
  qed
  moreover have ?lhs if 1: "M1.mtotally_bounded S" and 2: "M2.mtotally_bounded T"
    unfolding Prod_metric.mtotally_bounded_sequentially
  proof (intro conjI strip)
    show "S × T ⊆ M1 × M2"
      using that 
      by (auto simp: M1.mtotally_bounded_sequentially M2.mtotally_bounded_sequentially)
    fix σ :: "nat ⇒ 'a × 'b"
    assume σ: "range σ ⊆ S × T"
    with 1 obtain r1 where r1: "strict_mono r1" "M1.MCauchy (fst ∘ σ ∘ r1)"
      by (metis M1.mtotally_bounded_sequentially comp_apply image_subset_iff mem_Sigma_iff prod.collapse)
    from σ 2 obtain r2 where r2: "strict_mono r2" "M2.MCauchy (snd ∘ σ ∘ r1 ∘ r2)"
      apply (clarsimp simp: M2.mtotally_bounded_sequentially image_subset_iff)
      by (smt (verit, best) comp_apply mem_Sigma_iff prod.collapse)
    then have "M1.MCauchy (fst ∘ σ ∘ r1 ∘ r2)"
      by (simp add: M1.MCauchy_subsequence r1)
    with r2 have "Prod_metric.MCauchy (σ ∘ (r1 ∘ r2))"
      by (simp add: MCauchy_prod_metric o_def)
    then show "∃r. strict_mono r ∧ Prod_metric.MCauchy (σ ∘ r)"
      using r1 r2 strict_mono_o by blast
  qed
  ultimately show ?thesis
    using False by blast
qed auto
lemma mtotally_bounded_prod_metric:
   "Prod_metric.mtotally_bounded U ⟷
    M1.mtotally_bounded (fst ` U) ∧ M2.mtotally_bounded (snd ` U)" (is "?lhs ⟷ ?rhs")
proof
  assume L: ?lhs
  then have "U ⊆ M1 × M2" 
    and *: "⋀σ. range σ ⊆ U ⟹ ∃r::nat⇒nat. strict_mono r ∧ Prod_metric.MCauchy (σ∘r)"
    by (simp_all add: Prod_metric.mtotally_bounded_sequentially)
  show ?rhs
    unfolding M1.mtotally_bounded_sequentially M2.mtotally_bounded_sequentially
  proof (intro conjI strip)
    show "fst ` U ⊆ M1" "snd ` U ⊆ M2"
      using ‹U ⊆ M1 × M2› by auto
  next
    fix σ :: "nat ⇒ 'a"
    assume "range σ ⊆ fst ` U"
    then obtain ζ where ζ: "⋀n. σ n = fst (ζ n) ∧ ζ n ∈ U"
      unfolding image_subset_iff image_iff by (meson UNIV_I)
    then obtain r where "strict_mono r ∧ Prod_metric.MCauchy (ζ∘r)"
      by (metis "*" image_subset_iff)
    with ζ show "∃r. strict_mono r ∧ M1.MCauchy (σ ∘ r)"
      by (auto simp: MCauchy_prod_metric o_def)
  next
    fix σ:: "nat ⇒ 'b"
    assume "range σ ⊆ snd ` U"
    then obtain ζ where ζ: "⋀n. σ n = snd (ζ n) ∧ ζ n ∈ U"
      unfolding image_subset_iff image_iff by (meson UNIV_I)
    then obtain r where "strict_mono r ∧ Prod_metric.MCauchy (ζ∘r)"
      by (metis "*" image_subset_iff)
    with ζ show "∃r. strict_mono r ∧ M2.MCauchy (σ ∘ r)"
      by (auto simp: MCauchy_prod_metric o_def)
  qed
next
  assume ?rhs
  then have "Prod_metric.mtotally_bounded ((fst ` U) × (snd ` U))"
    by (simp add: mtotally_bounded_Times)
  then show ?lhs
    by (metis Prod_metric.mtotally_bounded_subset subset_fst_snd)
qed
end
lemma metrizable_space_prod_topology:
   "metrizable_space (prod_topology X Y) ⟷
    (prod_topology X Y) = trivial_topology ∨ metrizable_space X ∧ metrizable_space Y"
   (is "?lhs ⟷ ?rhs")
proof (cases "(prod_topology X Y) = trivial_topology")
  case False
  then obtain x y where "x ∈ topspace X" "y ∈ topspace Y"
    by fastforce
  show ?thesis
  proof
    show "?rhs ⟹ ?lhs"
      unfolding metrizable_space_def
      using Metric_space12.mtopology_prod_metric
      by (metis False Metric_space12.prod_metric Metric_space12_def) 
  next
    assume L: ?lhs 
    have "metrizable_space (subtopology (prod_topology X Y) (topspace X × {y}))"
      "metrizable_space (subtopology (prod_topology X Y) ({x} × topspace Y))"
      using L metrizable_space_subtopology by auto
    moreover
    have "(subtopology (prod_topology X Y) (topspace X × {y})) homeomorphic_space X"
      by (metis ‹y ∈ topspace Y› homeomorphic_space_prod_topology_sing1 homeomorphic_space_sym prod_topology_subtopology(2))
    moreover
    have "(subtopology (prod_topology X Y) ({x} × topspace Y)) homeomorphic_space Y"
      by (metis ‹x ∈ topspace X› homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym prod_topology_subtopology(1))
    ultimately show ?rhs
      by (simp add: homeomorphic_metrizable_space)
  qed
qed auto
lemma completely_metrizable_space_prod_topology:
   "completely_metrizable_space (prod_topology X Y) ⟷
    (prod_topology X Y) = trivial_topology ∨
    completely_metrizable_space X ∧ completely_metrizable_space Y"
   (is "?lhs ⟷ ?rhs")
proof (cases "(prod_topology X Y) = trivial_topology")
  case False
  then obtain x y where "x ∈ topspace X" "y ∈ topspace Y"
    by fastforce
  show ?thesis
  proof
    show "?rhs ⟹ ?lhs"
      unfolding completely_metrizable_space_def
      by (metis False Metric_space12.mtopology_prod_metric Metric_space12.mcomplete_prod_metric
          Metric_space12.prod_metric Metric_space12_def)
  next
    assume L: ?lhs 
    then have "Hausdorff_space (prod_topology X Y)"
      by (simp add: completely_metrizable_imp_metrizable_space metrizable_imp_Hausdorff_space)
    then have H: "Hausdorff_space X ∧ Hausdorff_space Y"
      using False Hausdorff_space_prod_topology by blast
    then have "closedin (prod_topology X Y) (topspace X × {y}) ∧ closedin (prod_topology X Y) ({x} × topspace Y)"
      using ‹x ∈ topspace X› ‹y ∈ topspace Y›
      by (auto simp: closedin_Hausdorff_sing_eq closedin_prod_Times_iff)
    with L have "completely_metrizable_space(subtopology (prod_topology X Y) (topspace X × {y}))
               ∧ completely_metrizable_space(subtopology (prod_topology X Y) ({x} × topspace Y))"
      by (simp add: completely_metrizable_space_closedin)
    moreover
    have "(subtopology (prod_topology X Y) (topspace X × {y})) homeomorphic_space X"
      by (metis ‹y ∈ topspace Y› homeomorphic_space_prod_topology_sing1 homeomorphic_space_sym prod_topology_subtopology(2))
    moreover
    have "(subtopology (prod_topology X Y) ({x} × topspace Y)) homeomorphic_space Y"
      by (metis ‹x ∈ topspace X› homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym prod_topology_subtopology(1))
    ultimately show ?rhs
      by (simp add: homeomorphic_completely_metrizable_space)
  qed
next
  case True then show ?thesis
    using empty_completely_metrizable_space by auto
qed 
    
subsection‹More sequential characterizations in a metric space›
context Metric_space
begin
  
definition decreasing_dist :: "(nat ⇒ 'a) ⇒ 'a ⇒ bool"
    where "decreasing_dist σ x ≡ (∀m n. m < n ⟶ d (σ n) x < d (σ m) x)"
lemma decreasing_dist_imp_inj: "decreasing_dist σ a ⟹ inj σ"
  by (metis decreasing_dist_def dual_order.irrefl linorder_inj_onI')
lemma eventually_atin_within_metric:
   "eventually P (atin_within mtopology a S) ⟷
    (a ∈ M ⟶ (∃δ>0. ∀x. x ∈ M ∧ x ∈ S ∧ 0 < d x a ∧ d x a < δ ⟶ P x))" (is "?lhs=?rhs")
proof
  assume ?lhs then show ?rhs
unfolding eventually_atin_within openin_mtopology subset_iff
  by (metis commute in_mball mdist_zero order_less_irrefl topspace_mtopology)
next
  assume R: ?rhs 
  show ?lhs
  proof (cases "a ∈ M")
    case True
    then obtain δ where "δ > 0" and δ: "⋀x. ⟦x ∈ M; x ∈ S; 0 < d x a; d x a < δ⟧ ⟹ P x"
      using R by blast
    then have "openin mtopology (mball a δ) ∧ (∀x ∈ mball a δ. x ∈ S ∧ x ≠ a ⟶ P x)"
      by (simp add: commute openin_mball)
    then show ?thesis
      by (metis True ‹0 < δ› centre_in_mball_iff eventually_atin_within) 
  next
    case False
    with R show ?thesis
      by (simp add: eventually_atin_within)
  qed
qed
lemma eventually_atin_within_A:
  assumes 
    "(⋀σ. ⟦range σ ⊆ (S ∩ M) - {a}; decreasing_dist σ a;
          inj σ; limitin mtopology σ a sequentially⟧
      ⟹ eventually (λn. P (σ n)) sequentially)"
  shows "eventually P (atin_within mtopology a S)"
proof -
  have False if SP: "⋀δ. δ>0 ⟹ ∃x ∈ M-{a}. d x a < δ ∧ x ∈ S ∧ ¬ P x" and "a ∈ M"
  proof -
    define Φ where "Φ ≡ λn x. x ∈ M-{a} ∧ d x a < inverse (Suc n) ∧ x ∈ S ∧ ¬ P x"
    obtain σ where σ: "⋀n. Φ n (σ n)" and dless: "⋀n. d (σ(Suc n)) a < d (σ n) a"
    proof -
      obtain x0 where x0: "Φ 0 x0"
        using SP [OF zero_less_one] by (force simp: Φ_def)
      have "∃y. Φ (Suc n) y ∧ d y a < d x a" if "Φ n x" for n x
        using SP [of "min (inverse (Suc (Suc n))) (d x a)"] ‹a ∈ M› that
        by (auto simp: Φ_def)
      then obtain f where f: "⋀n x. Φ n x ⟹ Φ (Suc n) (f n x) ∧ d (f n x) a < d x a" 
        by metis
      show thesis
        proof
          show "Φ n (rec_nat x0 f n)" for n
            by (induction n) (auto simp: x0 dest: f)
          with f show "d (rec_nat x0 f (Suc n)) a < d (rec_nat x0 f n) a" for n
            by auto 
        qed
    qed
    have 1: "range σ ⊆ (S ∩ M) - {a}"
      using σ by (auto simp: Φ_def)
    have "d (σ(Suc (m+n))) a < d (σ n) a" for m n
      by (induction m) (auto intro: order_less_trans dless)
    then have 2: "decreasing_dist σ a"
      unfolding decreasing_dist_def by (metis add.commute less_imp_Suc_add)
    have "∀⇩F xa in sequentially. d (σ xa) a < ε" if "ε > 0" for ε
    proof -
      obtain N where "inverse (Suc N) < ε"
        using ‹ε > 0› reals_Archimedean by blast
      with σ 2 show ?thesis
        unfolding decreasing_dist_def by (smt (verit, best) Φ_def eventually_at_top_dense)
    qed
    then have 4: "limitin mtopology σ a sequentially"
      using σ ‹a ∈ M› by (simp add: Φ_def limitin_metric)
    show False
      using 2 assms [OF 1 _ decreasing_dist_imp_inj 4] σ by (force simp: Φ_def)
  qed
  then show ?thesis
    by (fastforce simp: eventually_atin_within_metric)
qed
lemma eventually_atin_within_B:
  assumes ev: "eventually P (atin_within mtopology a S)" 
    and ran: "range σ ⊆ (S ∩ M) - {a}"
    and lim: "limitin mtopology σ a sequentially"
  shows "eventually (λn. P (σ n)) sequentially"
proof -
  have "a ∈ M"
    using lim limitin_mspace by auto
  with ev obtain δ where "0 < δ" 
    and δ: "⋀σ. ⟦σ ∈ M; σ ∈ S; 0 < d σ a; d σ a < δ⟧ ⟹ P σ"
    by (auto simp: eventually_atin_within_metric)
  then have *: "⋀n. σ n ∈ M ∧ d (σ n) a < δ ⟹ P (σ n)"
    using ‹a ∈ M› ran by auto
  have "∀⇩F n in sequentially. σ n ∈ M ∧ d (σ n) a < δ"
    using lim ‹0 < δ› by (auto simp: limitin_metric)
  then show ?thesis
    by (simp add: "*" eventually_mono)
qed
lemma eventually_atin_within_sequentially:
     "eventually P (atin_within mtopology a S) ⟷
        (∀σ. range σ ⊆ (S ∩ M) - {a} ∧
            limitin mtopology σ a sequentially
            ⟶ eventually (λn. P(σ n)) sequentially)"
  by (metis eventually_atin_within_A eventually_atin_within_B)
lemma eventually_atin_within_sequentially_inj:
     "eventually P (atin_within mtopology a S) ⟷
        (∀σ. range σ ⊆ (S ∩ M) - {a} ∧ inj σ ∧
            limitin mtopology σ a sequentially
            ⟶ eventually (λn. P(σ n)) sequentially)"
  by (metis eventually_atin_within_A eventually_atin_within_B)
lemma eventually_atin_within_sequentially_decreasing:
     "eventually P (atin_within mtopology a S) ⟷
        (∀σ. range σ ⊆ (S ∩ M) - {a} ∧ decreasing_dist σ a ∧
            limitin mtopology σ a sequentially
            ⟶ eventually (λn. P(σ n)) sequentially)"
  by (metis eventually_atin_within_A eventually_atin_within_B)
lemma eventually_atin_sequentially:
   "eventually P (atin mtopology a) ⟷
    (∀σ. range σ ⊆ M - {a} ∧ limitin mtopology σ a sequentially
         ⟶ eventually (λn. P(σ n)) sequentially)"
  using eventually_atin_within_sequentially [where S=UNIV] by simp
lemma eventually_atin_sequentially_inj:
   "eventually P (atin mtopology a) ⟷
    (∀σ. range σ ⊆ M - {a} ∧ inj σ ∧
        limitin mtopology σ a sequentially
        ⟶ eventually (λn. P(σ n)) sequentially)"
  using eventually_atin_within_sequentially_inj [where S=UNIV] by simp
lemma eventually_atin_sequentially_decreasing:
   "eventually P (atin mtopology a) ⟷
    (∀σ. range σ ⊆ M - {a} ∧ decreasing_dist σ a ∧
         limitin mtopology σ a sequentially
        ⟶ eventually (λn. P(σ n)) sequentially)"
  using eventually_atin_within_sequentially_decreasing [where S=UNIV] by simp
end
context Metric_space12
begin
lemma limit_atin_sequentially_within:
  "limitin M2.mtopology f l (atin_within M1.mtopology a S) ⟷
     l ∈ M2 ∧
     (∀σ. range σ ⊆ S ∩ M1 - {a} ∧
          limitin M1.mtopology σ a sequentially
          ⟶ limitin M2.mtopology (f ∘ σ) l sequentially)"
    by (auto simp: M1.eventually_atin_within_sequentially limitin_def)
lemma limit_atin_sequentially_within_inj:
  "limitin M2.mtopology f l (atin_within M1.mtopology a S) ⟷
     l ∈ M2 ∧
     (∀σ. range σ ⊆ S ∩ M1 - {a} ∧ inj σ ∧
          limitin M1.mtopology σ a sequentially
          ⟶ limitin M2.mtopology (f ∘ σ) l sequentially)"
    by (auto simp: M1.eventually_atin_within_sequentially_inj limitin_def)
lemma limit_atin_sequentially_within_decreasing:
  "limitin M2.mtopology f l (atin_within M1.mtopology a S) ⟷
     l ∈ M2 ∧
     (∀σ. range σ ⊆ S ∩ M1 - {a} ∧ M1.decreasing_dist σ a ∧ 
          limitin M1.mtopology σ a sequentially
          ⟶ limitin M2.mtopology (f ∘ σ) l sequentially)"
    by (auto simp: M1.eventually_atin_within_sequentially_decreasing limitin_def)
lemma limit_atin_sequentially:
   "limitin M2.mtopology f l (atin M1.mtopology a) ⟷
        l ∈ M2 ∧
        (∀σ. range σ ⊆ M1 - {a} ∧
            limitin M1.mtopology σ a sequentially
            ⟶ limitin M2.mtopology (f ∘ σ) l sequentially)"
  using limit_atin_sequentially_within [where S=UNIV] by simp
lemma limit_atin_sequentially_inj:
   "limitin M2.mtopology f l (atin M1.mtopology a) ⟷
        l ∈ M2 ∧
        (∀σ. range σ ⊆ M1 - {a} ∧ inj σ ∧
            limitin M1.mtopology σ a sequentially
            ⟶ limitin M2.mtopology (f ∘ σ) l sequentially)"
  using limit_atin_sequentially_within_inj [where S=UNIV] by simp
lemma limit_atin_sequentially_decreasing:
  "limitin M2.mtopology f l (atin M1.mtopology a) ⟷
     l ∈ M2 ∧
     (∀σ. range σ ⊆ M1 - {a} ∧ M1.decreasing_dist σ a ∧ 
          limitin M1.mtopology σ a sequentially
          ⟶ limitin M2.mtopology (f ∘ σ) l sequentially)"
  using limit_atin_sequentially_within_decreasing [where S=UNIV] by simp
end
text ‹An experiment: same result as within the locale, but using metric space variables›
lemma limit_atin_sequentially_within:
  "limitin (mtopology_of m2) f l (atin_within (mtopology_of m1) a S) ⟷
     l ∈ mspace m2 ∧
     (∀σ. range σ ⊆ S ∩ mspace m1 - {a} ∧
          limitin (mtopology_of m1) σ a sequentially
          ⟶ limitin (mtopology_of m2) (f ∘ σ) l sequentially)"
  using Metric_space12.limit_atin_sequentially_within [OF Metric_space12_mspace_mdist]
  by (metis mtopology_of_def)
context Metric_space
begin
lemma atin_within_imp_M:
   "atin_within mtopology x S ≠ bot ⟹ x ∈ M"
  by (metis derived_set_of_trivial_limit in_derived_set_of topspace_mtopology)
lemma atin_within_sequentially_sequence:
  assumes "atin_within mtopology x S ≠ bot"
  obtains σ where "range σ ⊆ S ∩ M - {x}" 
          "decreasing_dist σ x" "inj σ" "limitin mtopology σ x sequentially"
  by (metis eventually_atin_within_A eventually_False assms)
lemma derived_set_of_sequentially:
  "mtopology derived_set_of S =
   {x ∈ M. ∃σ. range σ ⊆ S ∩ M - {x} ∧ limitin mtopology σ x sequentially}"
proof -
  have False
    if "range σ ⊆ S ∩ M - {x}"
      and "limitin mtopology σ x sequentially"
      and "atin_within mtopology x S = bot"
    for x σ
  proof -
    have "∀⇩F n in sequentially. P (σ n)" for P
      using that by (metis eventually_atin_within_B eventually_bot)
    then show False
      by (meson eventually_False_sequentially eventually_mono)
  qed
  then show ?thesis
    using derived_set_of_trivial_limit 
    by (fastforce elim!: atin_within_sequentially_sequence intro: atin_within_imp_M)
qed
lemma derived_set_of_sequentially_alt:
  "mtopology derived_set_of S =
   {x. ∃σ. range σ ⊆ S - {x} ∧ limitin mtopology σ x sequentially}"
proof -
  have *: "∃σ. range σ ⊆ S ∩ M - {x} ∧ limitin mtopology σ x sequentially"
    if σ: "range σ ⊆ S - {x}" and lim: "limitin mtopology σ x sequentially" for x σ
  proof -
    obtain N where "∀n≥N. σ n ∈ M ∧ d (σ n) x < 1"
      using lim limit_metric_sequentially by fastforce
    with σ obtain a where a:"a ∈ S ∩ M - {x}" by auto
    show ?thesis
    proof (intro conjI exI)
      show "range (λn. if σ n ∈ M then σ n else a) ⊆ S ∩ M - {x}"
        using a σ by fastforce
      show "limitin mtopology (λn. if σ n ∈ M then σ n else a) x sequentially"
        using lim limit_metric_sequentially by fastforce
    qed
  qed
  show ?thesis
    by (auto simp: limitin_mspace derived_set_of_sequentially intro!: *)
qed
lemma derived_set_of_sequentially_inj:
   "mtopology derived_set_of S =
    {x ∈ M. ∃σ. range σ ⊆ S ∩ M - {x} ∧ inj σ ∧ limitin mtopology σ x sequentially}"
proof -
  have False
    if "x ∈ M" and "range σ ⊆ S ∩ M - {x}"
      and "limitin mtopology σ x sequentially"
      and "atin_within mtopology x S = bot"
    for x σ
  proof -
    have "∀⇩F n in sequentially. P (σ n)" for P
      using that derived_set_of_sequentially_alt derived_set_of_trivial_limit by fastforce
    then show False
      by (meson eventually_False_sequentially eventually_mono)
  qed
  then show ?thesis
    using derived_set_of_trivial_limit 
    by (fastforce elim!: atin_within_sequentially_sequence intro: atin_within_imp_M)
qed
lemma derived_set_of_sequentially_inj_alt:
   "mtopology derived_set_of S =
    {x. ∃σ. range σ ⊆ S - {x} ∧ inj σ ∧ limitin mtopology σ x sequentially}"
proof -
  have "∃σ. range σ ⊆ S - {x} ∧ inj σ ∧ limitin mtopology σ x sequentially"
    if "atin_within mtopology x S ≠ bot" for x
    by (metis Diff_empty Int_subset_iff atin_within_sequentially_sequence subset_Diff_insert that)
  moreover have False
    if "range (λx. σ (x::nat)) ⊆ S - {x}"
      and "limitin mtopology σ x sequentially"
      and "atin_within mtopology x S = bot"
    for x σ
  proof -
    have "∀⇩F n in sequentially. P (σ n)" for P
      using that derived_set_of_sequentially_alt derived_set_of_trivial_limit by fastforce
    then show False
      by (meson eventually_False_sequentially eventually_mono)
  qed
  ultimately show ?thesis
    using derived_set_of_trivial_limit by (fastforce intro: atin_within_imp_M)
qed
lemma derived_set_of_sequentially_decreasing:
   "mtopology derived_set_of S =
    {x ∈ M. ∃σ. range σ ⊆ S - {x} ∧ decreasing_dist σ x ∧ limitin mtopology σ x sequentially}"
proof -
  have "∃σ. range σ ⊆ S - {x} ∧ decreasing_dist σ x ∧ limitin mtopology σ x sequentially"
    if "atin_within mtopology x S ≠ bot" for x
    by (metis Diff_empty atin_within_sequentially_sequence le_infE subset_Diff_insert that)
  moreover have False
    if "x ∈ M" and "range σ ⊆ S - {x}"
      and "limitin mtopology σ x sequentially"
      and "atin_within mtopology x S = bot"
    for x σ
  proof -
    have "∀⇩F n in sequentially. P (σ n)" for P
      using that derived_set_of_sequentially_alt derived_set_of_trivial_limit by fastforce
    then show False
      by (meson eventually_False_sequentially eventually_mono)
  qed
  ultimately show ?thesis
    using derived_set_of_trivial_limit by (fastforce intro: atin_within_imp_M)
qed
lemma derived_set_of_sequentially_decreasing_alt:
   "mtopology derived_set_of S =
    {x. ∃σ. range σ ⊆ S - {x} ∧ decreasing_dist σ x ∧ limitin mtopology σ x sequentially}"
  using derived_set_of_sequentially_alt derived_set_of_sequentially_decreasing by auto
lemma closure_of_sequentially:
   "mtopology closure_of S = 
    {x ∈ M. ∃σ. range σ ⊆ S ∩ M ∧ limitin mtopology σ x sequentially}"
  by (auto simp: closure_of derived_set_of_sequentially)
end 
    
subsection ‹Three strong notions of continuity for metric spaces›
subsubsection ‹Lipschitz continuity›
definition Lipschitz_continuous_map 
  where "Lipschitz_continuous_map ≡ 
      λm1 m2 f. f ∈ mspace m1 → mspace m2 ∧
        (∃B. ∀x ∈ mspace m1. ∀y ∈ mspace m1. mdist m2 (f x) (f y) ≤ B * mdist m1 x y)"
lemma Lipschitz_continuous_map_image: 
  "Lipschitz_continuous_map m1 m2 f ⟹ f ∈ mspace m1 → mspace m2"
  by (simp add: Lipschitz_continuous_map_def)
lemma Lipschitz_continuous_map_pos:
   "Lipschitz_continuous_map m1 m2 f ⟷
    f ∈ mspace m1 → mspace m2 ∧
        (∃B>0. ∀x ∈ mspace m1. ∀y ∈ mspace m1. mdist m2 (f x) (f y) ≤ B * mdist m1 x y)"
proof -
  have "B * mdist m1 x y ≤ (¦B¦ + 1) * mdist m1 x y" "¦B¦ + 1 > 0" for x y B
    by (auto simp: mult_right_mono)
  then show ?thesis
    unfolding Lipschitz_continuous_map_def by (meson dual_order.trans)
qed
lemma Lipschitz_continuous_map_eq:
  assumes "Lipschitz_continuous_map m1 m2 f" "⋀x. x ∈ mspace m1 ⟹ f x = g x"
  shows "Lipschitz_continuous_map m1 m2 g"
  using Lipschitz_continuous_map_def assms by (simp add: Lipschitz_continuous_map_pos Pi_iff)
lemma Lipschitz_continuous_map_from_submetric:
  assumes "Lipschitz_continuous_map m1 m2 f"
  shows "Lipschitz_continuous_map (submetric m1 S) m2 f"
  unfolding Lipschitz_continuous_map_def 
proof
  show "f ∈ mspace (submetric m1 S) → mspace m2"
    using Lipschitz_continuous_map_pos assms by fastforce
qed (use assms in ‹fastforce simp: Lipschitz_continuous_map_def›)
lemma Lipschitz_continuous_map_from_submetric_mono:
   "⟦Lipschitz_continuous_map (submetric m1 T) m2 f; S ⊆ T⟧
           ⟹ Lipschitz_continuous_map (submetric m1 S) m2 f"
  by (metis Lipschitz_continuous_map_from_submetric inf.absorb_iff2 submetric_submetric)
lemma Lipschitz_continuous_map_into_submetric:
   "Lipschitz_continuous_map m1 (submetric m2 S) f ⟷
        f ∈ mspace m1 → S ∧ Lipschitz_continuous_map m1 m2 f"
  by (auto simp: Lipschitz_continuous_map_def)
lemma Lipschitz_continuous_map_const:
  "Lipschitz_continuous_map m1 m2 (λx. c) ⟷
        mspace m1 = {} ∨ c ∈ mspace m2"
  unfolding Lipschitz_continuous_map_def Pi_iff
  by (metis all_not_in_conv mdist_nonneg mdist_zero mult_1)
lemma Lipschitz_continuous_map_id:
   "Lipschitz_continuous_map m1 m1 (λx. x)"
  unfolding Lipschitz_continuous_map_def by (metis funcset_id mult_1 order_refl)
lemma Lipschitz_continuous_map_compose:
  assumes f: "Lipschitz_continuous_map m1 m2 f" and g: "Lipschitz_continuous_map m2 m3 g"
  shows "Lipschitz_continuous_map m1 m3 (g ∘ f)"
  unfolding Lipschitz_continuous_map_def 
proof
  show "g ∘ f ∈ mspace m1 → mspace m3"
    by (smt (verit, best) Lipschitz_continuous_map_image Pi_iff comp_apply f g)
  obtain B where B: "∀x∈mspace m1. ∀y∈mspace m1. mdist m2 (f x) (f y) ≤ B * mdist m1 x y"
    using assms unfolding Lipschitz_continuous_map_def by presburger
  obtain C where "C>0" and C: "∀x∈mspace m2. ∀y∈mspace m2. mdist m3 (g x) (g y) ≤ C * mdist m2 x y"
    using assms unfolding Lipschitz_continuous_map_pos by metis
  show "∃B. ∀x∈mspace m1. ∀y∈mspace m1. mdist m3 ((g ∘ f) x) ((g ∘ f) y) ≤ B * mdist m1 x y"
  proof (intro strip exI)
    fix x y
    assume §: "x ∈ mspace m1" "y ∈ mspace m1"
    then have "mdist m3 ((g ∘ f) x) ((g ∘ f) y) ≤ C * mdist m2 (f x) (f y)"
      using C Lipschitz_continuous_map_image f by fastforce
    also have "… ≤ C * B * mdist m1 x y"
      by (simp add: "§" B ‹0 < C›)
    finally show "mdist m3 ((g ∘ f) x) ((g ∘ f) y) ≤ C * B * mdist m1 x y" .
  qed
qed
subsubsection ‹Uniform continuity›
definition uniformly_continuous_map 
  where "uniformly_continuous_map ≡ 
      λm1 m2 f. f ∈ mspace m1 → mspace m2 ∧
        (∀ε>0. ∃δ>0. ∀x ∈ mspace m1. ∀y ∈ mspace m1. 
                           mdist m1 y x < δ ⟶ mdist m2 (f y) (f x) < ε)"
lemma uniformly_continuous_map_funspace: 
  "uniformly_continuous_map m1 m2 f ⟹ f ∈ mspace m1 → mspace m2"
  by (simp add: uniformly_continuous_map_def)
lemma ucmap_A:
  assumes "uniformly_continuous_map m1 m2 f"
  and "(λn. mdist m1 (ρ n) (σ n)) ⇢ 0"
    and "range ρ ⊆ mspace m1" "range σ ⊆ mspace m1"
  shows "(λn. mdist m2 (f (ρ n)) (f (σ n))) ⇢ 0"
  using assms 
  unfolding uniformly_continuous_map_def image_subset_iff tendsto_iff
  apply clarsimp
  by (metis (mono_tags, lifting) eventually_sequentially)
lemma ucmap_B:
  assumes §: "⋀ρ σ. ⟦range ρ ⊆ mspace m1; range σ ⊆ mspace m1;
              (λn. mdist m1 (ρ n) (σ n)) ⇢ 0⟧
              ⟹ (λn. mdist m2 (f (ρ n)) (f (σ n))) ⇢ 0"
    and "0 < ε"
    and ρ: "range ρ ⊆ mspace m1"
    and σ: "range σ ⊆ mspace m1"
    and "(λn. mdist m1 (ρ n) (σ n)) ⇢ 0"
  shows "∃n. mdist m2 (f (ρ (n::nat))) (f (σ n)) < ε"
  using § [OF ρ σ] assms by (meson LIMSEQ_le_const linorder_not_less)
lemma ucmap_C: 
  assumes §: "⋀ρ σ ε. ⟦ε > 0; range ρ ⊆ mspace m1; range σ ⊆ mspace m1;
                       ((λn. mdist m1 (ρ n) (σ n)) ⇢ 0)⟧
                      ⟹ ∃n. mdist m2 (f (ρ n)) (f (σ n)) < ε"
    and fim: "f ∈ mspace m1 → mspace m2"
  shows "uniformly_continuous_map m1 m2 f"
proof -
  {assume "¬ (∀ε>0. ∃δ>0. ∀x∈mspace m1. ∀y∈mspace m1. mdist m1 y x < δ ⟶ mdist m2 (f y) (f x) < ε)" 
    then obtain ε where "ε > 0" 
      and "⋀n. ∃x∈mspace m1. ∃y∈mspace m1. mdist m1 y x < inverse(Suc n) ∧ mdist m2 (f y) (f x) ≥ ε"
      by (meson inverse_Suc linorder_not_le)
    then obtain ρ σ where space: "range ρ ⊆ mspace m1" "range σ ⊆ mspace m1"
         and dist: "⋀n. mdist m1 (σ n) (ρ n) < inverse(Suc n) ∧ mdist m2 (f(σ n)) (f(ρ n)) ≥ ε"
      by (metis image_subset_iff)
    have False 
      using § [OF ‹ε > 0› space] dist Lim_null_comparison
      by (smt (verit) LIMSEQ_norm_0 inverse_eq_divide mdist_commute mdist_nonneg real_norm_def)
  }
  moreover
  have "t ∈ mspace m2" if "t ∈ f ` mspace m1" for t
    using fim that by blast
  ultimately show ?thesis
    by (fastforce simp: uniformly_continuous_map_def)
qed
lemma uniformly_continuous_map_sequentially:
  "uniformly_continuous_map m1 m2 f ⟷
    f ∈ mspace m1 → mspace m2 ∧
    (∀ρ σ. range ρ ⊆ mspace m1 ∧ range σ ⊆ mspace m1 ∧ (λn. mdist m1 (ρ n) (σ n)) ⇢ 0
          ⟶ (λn. mdist m2 (f (ρ n)) (f (σ n))) ⇢ 0)"
   (is "?lhs ⟷ ?rhs")
proof
  show "?lhs ⟹ ?rhs"
    by (simp add: ucmap_A uniformly_continuous_map_funspace)
  show "?rhs ⟹ ?lhs"
    by (intro ucmap_B ucmap_C) auto
qed
lemma uniformly_continuous_map_sequentially_alt:
    "uniformly_continuous_map m1 m2 f ⟷
      f ∈ mspace m1 → mspace m2 ∧
      (∀ε>0. ∀ρ σ. range ρ ⊆ mspace m1 ∧ range σ ⊆ mspace m1 ∧
            ((λn. mdist m1 (ρ n) (σ n)) ⇢ 0)
            ⟶ (∃n. mdist m2 (f (ρ n)) (f (σ n)) < ε))"
   (is "?lhs ⟷ ?rhs")
proof
  show "?lhs ⟹ ?rhs"
    using uniformly_continuous_map_funspace by (intro conjI strip ucmap_B | fastforce simp: ucmap_A)+
  show "?rhs ⟹ ?lhs"
    by (intro ucmap_C) auto
qed
lemma uniformly_continuous_map_eq:
   "⟦⋀x. x ∈ mspace m1 ⟹ f x = g x; uniformly_continuous_map m1 m2 f⟧
      ⟹ uniformly_continuous_map m1 m2 g"
  by (simp add: uniformly_continuous_map_def Pi_iff)
lemma uniformly_continuous_map_from_submetric:
  assumes "uniformly_continuous_map m1 m2 f"
  shows "uniformly_continuous_map (submetric m1 S) m2 f"
  unfolding uniformly_continuous_map_def 
proof
  show "f ∈ mspace (submetric m1 S) → mspace m2"
    using assms by (auto simp: uniformly_continuous_map_def)
qed (use assms in ‹force simp: uniformly_continuous_map_def›)
lemma uniformly_continuous_map_from_submetric_mono:
   "⟦uniformly_continuous_map (submetric m1 T) m2 f; S ⊆ T⟧
           ⟹ uniformly_continuous_map (submetric m1 S) m2 f"
  by (metis uniformly_continuous_map_from_submetric inf.absorb_iff2 submetric_submetric)
lemma uniformly_continuous_map_into_submetric:
   "uniformly_continuous_map m1 (submetric m2 S) f ⟷
        f ∈ mspace m1 → S ∧ uniformly_continuous_map m1 m2 f"
  by (auto simp: uniformly_continuous_map_def)
lemma uniformly_continuous_map_const:
  "uniformly_continuous_map m1 m2 (λx. c) ⟷
        mspace m1 = {} ∨ c ∈ mspace m2"
  unfolding uniformly_continuous_map_def Pi_iff
  by (metis empty_iff equals0I mdist_zero)
lemma uniformly_continuous_map_id [simp]:
   "uniformly_continuous_map m1 m1 (λx. x)"
  by (metis funcset_id uniformly_continuous_map_def)
lemma uniformly_continuous_map_compose:
  assumes f: "uniformly_continuous_map m1 m2 f" and g: "uniformly_continuous_map m2 m3 g"
  shows "uniformly_continuous_map m1 m3 (g ∘ f)"
  using f g unfolding uniformly_continuous_map_def comp_apply Pi_iff
  by metis
lemma uniformly_continuous_map_real_const [simp]:
   "uniformly_continuous_map m euclidean_metric (λx. c)"
  by (simp add: euclidean_metric_def uniformly_continuous_map_const)
text ‹Equivalence between "abstract" and "type class" notions›
lemma uniformly_continuous_map_euclidean [simp]:
  "uniformly_continuous_map (submetric euclidean_metric S) euclidean_metric f
     = uniformly_continuous_on S f"
  by (auto simp: uniformly_continuous_map_def uniformly_continuous_on_def)
subsubsection ‹Cauchy continuity›
definition Cauchy_continuous_map where
 "Cauchy_continuous_map ≡
  λm1 m2 f. ∀σ. Metric_space.MCauchy (mspace m1) (mdist m1) σ 
            ⟶ Metric_space.MCauchy (mspace m2) (mdist m2) (f ∘ σ)"
lemma Cauchy_continuous_map_euclidean [simp]:
  "Cauchy_continuous_map (submetric euclidean_metric S) euclidean_metric f
     = Cauchy_continuous_on S f"
  by (auto simp: Cauchy_continuous_map_def Cauchy_continuous_on_def Cauchy_def Met_TC.subspace Metric_space.MCauchy_def)
lemma Cauchy_continuous_map_funspace:
  assumes "Cauchy_continuous_map m1 m2 f"
  shows "f ∈ mspace m1 → mspace m2"
proof clarsimp
  fix x
  assume "x ∈ mspace m1"
  then have "Metric_space.MCauchy (mspace m1) (mdist m1) (λn. x)"
    by (simp add: Metric_space.MCauchy_const Metric_space_mspace_mdist)
  then have "Metric_space.MCauchy (mspace m2) (mdist m2) (f ∘ (λn. x))"
    by (meson Cauchy_continuous_map_def assms)
  then show "f x ∈ mspace m2"
    by (simp add: Metric_space.MCauchy_def [OF Metric_space_mspace_mdist])
qed
lemma Cauchy_continuous_map_eq:
   "⟦⋀x. x ∈ mspace m1 ⟹ f x = g x; Cauchy_continuous_map m1 m2 f⟧
      ⟹ Cauchy_continuous_map m1 m2 g"
  by (simp add: image_subset_iff Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist])
lemma Cauchy_continuous_map_from_submetric:
  assumes "Cauchy_continuous_map m1 m2 f"
  shows "Cauchy_continuous_map (submetric m1 S) m2 f"
  using assms
  by (simp add: image_subset_iff Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist])
lemma Cauchy_continuous_map_from_submetric_mono:
   "⟦Cauchy_continuous_map (submetric m1 T) m2 f; S ⊆ T⟧
           ⟹ Cauchy_continuous_map (submetric m1 S) m2 f"
  by (metis Cauchy_continuous_map_from_submetric inf.absorb_iff2 submetric_submetric)
lemma Cauchy_continuous_map_into_submetric:
   "Cauchy_continuous_map m1 (submetric m2 S) f ⟷
   f ∈ mspace m1 → S ∧ Cauchy_continuous_map m1 m2 f"  (is "?lhs ⟷ ?rhs")
proof -
  have "?lhs ⟹ Cauchy_continuous_map m1 m2 f"
    by (simp add: Cauchy_continuous_map_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist])
  moreover have "?rhs ⟹ ?lhs"
    by (auto simp: Cauchy_continuous_map_def  Metric_space.MCauchy_def [OF Metric_space_mspace_mdist])
  ultimately show ?thesis
    by (metis Cauchy_continuous_map_funspace Int_iff funcsetI funcset_mem mspace_submetric)
qed
lemma Cauchy_continuous_map_const [simp]:
  "Cauchy_continuous_map m1 m2 (λx. c) ⟷ mspace m1 = {} ∨ c ∈ mspace m2"
proof -
   have "mspace m1 = {} ⟹ Cauchy_continuous_map m1 m2 (λx. c)"
    by (simp add: Cauchy_continuous_map_def Metric_space.MCauchy_def Metric_space_mspace_mdist)
  moreover have "c ∈ mspace m2 ⟹ Cauchy_continuous_map m1 m2 (λx. c)"
    by (simp add: Cauchy_continuous_map_def o_def Metric_space.MCauchy_const [OF Metric_space_mspace_mdist])
  ultimately show ?thesis
    using Cauchy_continuous_map_funspace by blast
qed
lemma Cauchy_continuous_map_id [simp]:
   "Cauchy_continuous_map m1 m1 (λx. x)"
  by (simp add: Cauchy_continuous_map_def o_def)
lemma Cauchy_continuous_map_compose:
  assumes f: "Cauchy_continuous_map m1 m2 f" and g: "Cauchy_continuous_map m2 m3 g"
  shows "Cauchy_continuous_map m1 m3 (g ∘ f)"
  by (metis (no_types, lifting) Cauchy_continuous_map_def f fun.map_comp g)
lemma Lipschitz_imp_uniformly_continuous_map:
  assumes "Lipschitz_continuous_map m1 m2 f"
  shows "uniformly_continuous_map m1 m2 f"
  proof -
  have "f ∈ mspace m1 → mspace m2"
    by (simp add: Lipschitz_continuous_map_image assms)
  moreover have "∃δ>0. ∀x∈mspace m1. ∀y∈mspace m1. mdist m1 y x < δ ⟶ mdist m2 (f y) (f x) < ε"
    if "ε > 0" for ε
  proof -
    obtain B where "∀x∈mspace m1. ∀y∈mspace m1. mdist m2 (f x) (f y) ≤ B * mdist m1 x y"
             and "B>0"
      using that assms by (force simp: Lipschitz_continuous_map_pos)
    then have "∀x∈mspace m1. ∀y∈mspace m1. mdist m1 y x < ε/B ⟶ mdist m2 (f y) (f x) < ε"
      by (smt (verit, ccfv_SIG) less_divide_eq mdist_nonneg mult.commute that zero_less_divide_iff)
    with ‹B>0› show ?thesis
      by (metis divide_pos_pos that)
  qed
  ultimately show ?thesis
    by (auto simp: uniformly_continuous_map_def)
qed
lemma uniformly_imp_Cauchy_continuous_map:
   "uniformly_continuous_map m1 m2 f ⟹ Cauchy_continuous_map m1 m2 f"
  unfolding uniformly_continuous_map_def Cauchy_continuous_map_def
  apply (simp add: image_subset_iff o_def Metric_space.MCauchy_def [OF Metric_space_mspace_mdist])
  by (metis funcset_mem)
lemma locally_Cauchy_continuous_map:
  assumes "ε > 0"
    and §: "⋀x. x ∈ mspace m1 ⟹ Cauchy_continuous_map (submetric m1 (mball_of m1 x ε)) m2 f"
  shows "Cauchy_continuous_map m1 m2 f"
  unfolding Cauchy_continuous_map_def
proof (intro strip)
  interpret M1: Metric_space "mspace m1" "mdist m1"
    by (simp add: Metric_space_mspace_mdist)
  interpret M2: Metric_space "mspace m2" "mdist m2"
    by (simp add: Metric_space_mspace_mdist)
  fix σ
  assume σ: "M1.MCauchy σ"
  with ‹ε > 0› obtain N where N: "⋀n n'. ⟦n≥N; n'≥N⟧ ⟹ mdist m1 (σ n) (σ n') < ε"
    using M1.MCauchy_def by fastforce
  then have "M1.mball (σ N) ε ⊆ mspace m1"
    by (auto simp: image_subset_iff M1.mball_def)
  then interpret MS1: Metric_space "mball_of m1 (σ N) ε ∩ mspace m1" "mdist m1"
    by (simp add: M1.subspace)
  show "M2.MCauchy (f ∘ σ)"
  proof (rule M2.MCauchy_offset)
    have "M1.MCauchy (σ ∘ (+) N)"
      by (simp add: M1.MCauchy_imp_MCauchy_suffix σ)
    moreover have "range (σ ∘ (+) N) ⊆ M1.mball (σ N) ε"
      using N [OF order_refl] M1.MCauchy_def σ by fastforce
    ultimately have "MS1.MCauchy (σ ∘ (+) N)"
      unfolding M1.MCauchy_def MS1.MCauchy_def by (simp add: mball_of_def)
    moreover have "σ N ∈ mspace m1"
      using M1.MCauchy_def σ by auto
    ultimately show "M2.MCauchy (f ∘ σ ∘ (+) N)"
      unfolding comp_assoc
      by (metis "§" Cauchy_continuous_map_def mdist_submetric mspace_submetric)
  next
    fix n
    have "σ n ∈ mspace m1"
      by (meson Metric_space.MCauchy_def Metric_space_mspace_mdist σ range_subsetD)
    then have "σ n ∈ mball_of m1 (σ n) ε"
      by (simp add: Metric_space.centre_in_mball_iff Metric_space_mspace_mdist assms(1) mball_of_def)
    then show "(f ∘ σ) n ∈ mspace m2"
      using Cauchy_continuous_map_funspace [OF § [of "σ n"]] ‹σ n ∈ mspace m1› by auto
  qed
qed
context Metric_space12
begin
lemma Cauchy_continuous_imp_continuous_map:
  assumes "Cauchy_continuous_map (metric (M1,d1)) (metric (M2,d2)) f"
  shows "continuous_map M1.mtopology M2.mtopology f"
proof (clarsimp simp: continuous_map_atin)
  fix x
  assume "x ∈ M1"
  show "limitin M2.mtopology f (f x) (atin M1.mtopology x)"
    unfolding limit_atin_sequentially
  proof (intro conjI strip)
    show "f x ∈ M2"
      using Cauchy_continuous_map_funspace ‹x ∈ M1› assms by fastforce
    fix σ
    assume "range σ ⊆ M1 - {x} ∧ limitin M1.mtopology σ x sequentially"
    then have "M1.MCauchy (λn. if even n then σ (n div 2) else x)"
      by (force simp: M1.MCauchy_interleaving)
    then have "M2.MCauchy (f ∘ (λn. if even n then σ (n div 2) else x))"
      using assms by (simp add: Cauchy_continuous_map_def)
    then show "limitin M2.mtopology (f ∘ σ) (f x) sequentially"
      using M2.MCauchy_interleaving [of "f ∘ σ" "f x"]
      by (simp add: o_def if_distrib cong: if_cong)
  qed
qed
lemma continuous_imp_Cauchy_continuous_map:
  assumes "M1.mcomplete"
    and f: "continuous_map M1.mtopology M2.mtopology f"
  shows "Cauchy_continuous_map (metric (M1,d1)) (metric (M2,d2)) f"
  unfolding Cauchy_continuous_map_def
proof clarsimp
  fix σ
  assume σ: "M1.MCauchy σ"
  then obtain y where y: "limitin M1.mtopology σ y sequentially"
    using M1.mcomplete_def assms by blast
  have "range (f ∘ σ) ⊆ M2"
    using σ f by (simp add: M2.subspace M1.MCauchy_def M1.metric_continuous_map image_subset_iff)
  then show "M2.MCauchy (f ∘ σ)"
    using continuous_map_limit [OF f y] M2.convergent_imp_MCauchy
    by blast
qed
end
text ‹The same outside the locale›
lemma Cauchy_continuous_imp_continuous_map:
  assumes "Cauchy_continuous_map m1 m2 f"
  shows "continuous_map (mtopology_of m1) (mtopology_of m2) f"
  using assms Metric_space12.Cauchy_continuous_imp_continuous_map [OF Metric_space12_mspace_mdist]
  by (auto simp: mtopology_of_def)
lemma continuous_imp_Cauchy_continuous_map:
  assumes "Metric_space.mcomplete (mspace m1) (mdist m1)"
    and "continuous_map (mtopology_of m1) (mtopology_of m2) f"
  shows "Cauchy_continuous_map m1 m2 f"
  using assms Metric_space12.continuous_imp_Cauchy_continuous_map [OF Metric_space12_mspace_mdist]
  by (auto simp: mtopology_of_def)
lemma uniformly_continuous_imp_continuous_map:
   "uniformly_continuous_map m1 m2 f
        ⟹ continuous_map (mtopology_of m1) (mtopology_of m2) f"
  by (simp add: Cauchy_continuous_imp_continuous_map uniformly_imp_Cauchy_continuous_map)
lemma Lipschitz_continuous_imp_continuous_map:
   "Lipschitz_continuous_map m1 m2 f
     ⟹ continuous_map (mtopology_of m1) (mtopology_of m2) f"
  by (simp add: Lipschitz_imp_uniformly_continuous_map uniformly_continuous_imp_continuous_map)
lemma Lipschitz_imp_Cauchy_continuous_map:
   "Lipschitz_continuous_map m1 m2 f
        ⟹ Cauchy_continuous_map m1 m2 f"
  by (simp add: Lipschitz_imp_uniformly_continuous_map uniformly_imp_Cauchy_continuous_map)
lemma Cauchy_imp_uniformly_continuous_map:
  assumes f: "Cauchy_continuous_map m1 m2 f"
    and tbo: "Metric_space.mtotally_bounded (mspace m1) (mdist m1) (mspace m1)"
  shows "uniformly_continuous_map m1 m2 f"
  unfolding uniformly_continuous_map_sequentially_alt imp_conjL
proof (intro conjI strip)
  show "f ∈ mspace m1 → mspace m2"
    by (simp add: Cauchy_continuous_map_funspace f)
  interpret M1: Metric_space "mspace m1" "mdist m1"
    by (simp add: Metric_space_mspace_mdist)
  interpret M2: Metric_space "mspace m2" "mdist m2"
    by (simp add: Metric_space_mspace_mdist)
  fix ε :: real and ρ σ 
  assume "ε > 0"
    and ρ: "range ρ ⊆ mspace m1"
    and σ: "range σ ⊆ mspace m1"
    and 0: "(λn. mdist m1 (ρ n) (σ n)) ⇢ 0"
  then obtain r1 where "strict_mono r1" and r1: "M1.MCauchy (ρ ∘ r1)"
    using M1.mtotally_bounded_sequentially tbo by meson
  then obtain r2 where "strict_mono r2" and r2: "M1.MCauchy (σ ∘ r1 ∘ r2)"
    by (metis M1.mtotally_bounded_sequentially tbo σ image_comp image_subset_iff range_subsetD)
  define r where "r ≡ r1 ∘ r2"
  have r: "strict_mono r"
    by (simp add: r_def ‹strict_mono r1› ‹strict_mono r2› strict_mono_o)
  define η where "η ≡ λn. if even n then (ρ ∘ r) (n div 2) else (σ ∘ r) (n div 2)"
  have "M1.MCauchy η"
    unfolding η_def M1.MCauchy_interleaving_gen
  proof (intro conjI)
    show "M1.MCauchy (ρ ∘ r)"
      by (simp add: M1.MCauchy_subsequence ‹strict_mono r2› fun.map_comp r1 r_def)
    show "M1.MCauchy (σ ∘ r)"
      by (simp add: fun.map_comp r2 r_def)
    show "(λn. mdist m1 ((ρ ∘ r) n) ((σ ∘ r) n)) ⇢ 0"
      using LIMSEQ_subseq_LIMSEQ [OF 0 r] by (simp add: o_def)
  qed
  then have "Metric_space.MCauchy (mspace m2) (mdist m2) (f ∘ η)"
    by (meson Cauchy_continuous_map_def f)
  then have "(λn. mdist m2 (f (ρ (r n))) (f (σ (r n)))) ⇢ 0"
    using M2.MCauchy_interleaving_gen [of "f ∘ ρ ∘ r" "f ∘ σ ∘ r"]
    by (simp add: η_def o_def if_distrib cong: if_cong)
  then show "∃n. mdist m2 (f (ρ n)) (f (σ n)) < ε"
    by (meson LIMSEQ_le_const ‹0 < ε› linorder_not_le)
qed
lemma continuous_imp_uniformly_continuous_map:
   "compact_space (mtopology_of m1) ∧
        continuous_map (mtopology_of m1) (mtopology_of m2) f
        ⟹ uniformly_continuous_map m1 m2 f"
  by (simp add: Cauchy_imp_uniformly_continuous_map continuous_imp_Cauchy_continuous_map
                Metric_space.compact_space_eq_mcomplete_mtotally_bounded
                Metric_space_mspace_mdist mtopology_of_def)
lemma continuous_eq_Cauchy_continuous_map:
   "Metric_space.mcomplete (mspace m1) (mdist m1) ⟹ 
    continuous_map (mtopology_of m1) (mtopology_of m2) f ⟷ Cauchy_continuous_map m1 m2 f"
  using Cauchy_continuous_imp_continuous_map continuous_imp_Cauchy_continuous_map by blast
lemma continuous_eq_uniformly_continuous_map:
   "compact_space (mtopology_of m1) 
    ⟹ continuous_map (mtopology_of m1) (mtopology_of m2) f ⟷
        uniformly_continuous_map m1 m2 f"
  using continuous_imp_uniformly_continuous_map uniformly_continuous_imp_continuous_map by blast
lemma Cauchy_eq_uniformly_continuous_map:
   "Metric_space.mtotally_bounded (mspace m1) (mdist m1) (mspace m1)
    ⟹ Cauchy_continuous_map m1 m2 f ⟷ uniformly_continuous_map m1 m2 f"
  using Cauchy_imp_uniformly_continuous_map uniformly_imp_Cauchy_continuous_map by blast
lemma Lipschitz_continuous_map_projections:
  "Lipschitz_continuous_map (prod_metric m1 m2) m1 fst"
  "Lipschitz_continuous_map (prod_metric m1 m2) m2 snd"
  by (simp add: Lipschitz_continuous_map_def prod_dist_def fst_Pi snd_Pi; 
      metis mult_numeral_1 real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2)+
lemma Lipschitz_continuous_map_pairwise:
   "Lipschitz_continuous_map m (prod_metric m1 m2) f ⟷
    Lipschitz_continuous_map m m1 (fst ∘ f) ∧ Lipschitz_continuous_map m m2 (snd ∘ f)"
   (is "?lhs ⟷ ?rhs")
proof 
  show "?lhs ⟹ ?rhs"
    by (simp add: Lipschitz_continuous_map_compose Lipschitz_continuous_map_projections)
  have "Lipschitz_continuous_map m (prod_metric m1 m2) (λx. (f1 x, f2 x))"
    if f1: "Lipschitz_continuous_map m m1 f1" and f2: "Lipschitz_continuous_map m m2 f2" for f1 f2
  proof -
    obtain B1 where "B1 > 0" 
      and B1: "⋀x y. ⟦x ∈ mspace m; y ∈ mspace m⟧ ⟹ mdist m1 (f1 x) (f1 y) ≤ B1 * mdist m x y"
      by (meson Lipschitz_continuous_map_pos f1)
    obtain B2 where "B2 > 0" 
      and B2: "⋀x y. ⟦x ∈ mspace m; y ∈ mspace m⟧ ⟹ mdist m2 (f2 x) (f2 y) ≤ B2 * mdist m x y"
      by (meson Lipschitz_continuous_map_pos f2)
    show ?thesis
      unfolding Lipschitz_continuous_map_pos
    proof (intro exI conjI strip)
      have f1im: "f1 ∈ mspace m → mspace m1"
        by (simp add: Lipschitz_continuous_map_image f1)
      moreover have f2im: "f2 ∈ mspace m → mspace m2"
        by (simp add: Lipschitz_continuous_map_image f2)
      ultimately show "(λx. (f1 x, f2 x)) ∈ mspace m → mspace (prod_metric m1 m2)"
        by auto
      show "B1+B2 > 0"
        using ‹0 < B1› ‹0 < B2› by linarith
      fix x y
      assume xy: "x ∈ mspace m" "y ∈ mspace m"
      with f1im f2im have "mdist (prod_metric m1 m2) (f1 x, f2 x) (f1 y, f2 y) ≤ mdist m1 (f1 x) (f1 y) + mdist m2 (f2 x) (f2 y)"
        unfolding mdist_prod_metric
        by (intro Metric_space12.prod_metric_le_components [OF Metric_space12_mspace_mdist]) auto
      also have "... ≤ (B1+B2) * mdist m x y"
        using B1 [OF xy] B2 [OF xy] by (simp add: vector_space_over_itself.scale_left_distrib) 
      finally show "mdist (prod_metric m1 m2) (f1 x, f2 x) (f1 y, f2 y) ≤ (B1+B2) * mdist m x y" .
    qed
  qed
  then show "?rhs ⟹ ?lhs"
    by force
qed
lemma uniformly_continuous_map_pairwise:
   "uniformly_continuous_map m (prod_metric m1 m2) f ⟷ 
    uniformly_continuous_map m m1 (fst ∘ f) ∧ uniformly_continuous_map m m2 (snd ∘ f)"
   (is "?lhs ⟷ ?rhs")
proof 
  show "?lhs ⟹ ?rhs"
    by (simp add: Lipschitz_continuous_map_projections Lipschitz_imp_uniformly_continuous_map uniformly_continuous_map_compose)
  have "uniformly_continuous_map m (prod_metric m1 m2) (λx. (f1 x, f2 x))"
    if f1: "uniformly_continuous_map m m1 f1" and f2: "uniformly_continuous_map m m2 f2" for f1 f2
  proof -
    show ?thesis
      unfolding uniformly_continuous_map_def
    proof (intro conjI strip)
      have f1im: "f1 ∈ mspace m → mspace m1"
        by (simp add: uniformly_continuous_map_funspace f1)
      moreover have f2im: "f2 ∈ mspace m → mspace m2"
        by (simp add: uniformly_continuous_map_funspace f2)
      ultimately show "(λx. (f1 x, f2 x)) ∈ mspace m → mspace (prod_metric m1 m2)"
        by auto
      fix ε:: real
      assume "ε > 0"
      obtain δ1 where "δ1>0" 
        and δ1: "⋀x y. ⟦x ∈ mspace m; y ∈ mspace m; mdist m y x < δ1⟧ ⟹ mdist m1 (f1 y) (f1 x) < ε/2"
        by (metis ‹0 < ε› f1 half_gt_zero uniformly_continuous_map_def)
      obtain δ2 where "δ2>0" 
        and δ2: "⋀x y. ⟦x ∈ mspace m; y ∈ mspace m; mdist m y x < δ2⟧ ⟹ mdist m2 (f2 y) (f2 x) < ε/2"
        by (metis ‹0 < ε› f2 half_gt_zero uniformly_continuous_map_def)
      show "∃δ>0. ∀x∈mspace m. ∀y∈mspace m. mdist m y x < δ ⟶ mdist (prod_metric m1 m2) (f1 y, f2 y) (f1 x, f2 x) < ε"
      proof (intro exI conjI strip)
        show "min δ1 δ2>0"
          using ‹0 < δ1› ‹0 < δ2› by auto
        fix x y
        assume xy: "x ∈ mspace m" "y ∈ mspace m" and d: "mdist m y x < min δ1 δ2"
        have *: "mdist m1 (f1 y) (f1 x) < ε/2" "mdist m2 (f2 y) (f2 x) < ε/2"
          using δ1 δ2 d xy by auto
        have "mdist (prod_metric m1 m2) (f1 y, f2 y) (f1 x, f2 x) ≤ mdist m1 (f1 y) (f1 x) + mdist m2 (f2 y) (f2 x)"
          unfolding mdist_prod_metric using f1im f2im xy
          by (intro Metric_space12.prod_metric_le_components [OF Metric_space12_mspace_mdist]) auto
        also have "... < ε/2 + ε/2"
          using * by simp
        finally show "mdist (prod_metric m1 m2) (f1 y, f2 y) (f1 x, f2 x) < ε"
          by simp
      qed
    qed
  qed
  then show "?rhs ⟹ ?lhs"
    by force
qed
lemma Cauchy_continuous_map_pairwise:
   "Cauchy_continuous_map m (prod_metric m1 m2) f ⟷ Cauchy_continuous_map m m1 (fst ∘ f) ∧ Cauchy_continuous_map m m2 (snd ∘ f)"
  by (auto simp: Cauchy_continuous_map_def Metric_space12.MCauchy_prod_metric[OF Metric_space12_mspace_mdist] comp_assoc)
lemma Lipschitz_continuous_map_paired:
   "Lipschitz_continuous_map m (prod_metric m1 m2) (λx. (f x, g x)) ⟷
        Lipschitz_continuous_map m m1 f ∧ Lipschitz_continuous_map m m2 g"
  by (simp add: Lipschitz_continuous_map_pairwise o_def)
lemma uniformly_continuous_map_paired:
   "uniformly_continuous_map m (prod_metric m1 m2) (λx. (f x, g x)) ⟷
        uniformly_continuous_map m m1 f ∧ uniformly_continuous_map m m2 g"
  by (simp add: uniformly_continuous_map_pairwise o_def)
lemma Cauchy_continuous_map_paired:
   "Cauchy_continuous_map m (prod_metric m1 m2) (λx. (f x, g x)) ⟷
        Cauchy_continuous_map m m1 f ∧ Cauchy_continuous_map m m2 g"
  by (simp add: Cauchy_continuous_map_pairwise o_def)
lemma mbounded_Lipschitz_continuous_image:
  assumes f: "Lipschitz_continuous_map m1 m2 f" and S: "Metric_space.mbounded (mspace m1) (mdist m1) S"
  shows "Metric_space.mbounded (mspace m2) (mdist m2) (f`S)"
proof -
  obtain B where fim: "f ∈ mspace m1 → mspace m2"
    and "B>0" and B: "⋀x y. ⟦x ∈ mspace m1; y ∈ mspace m1⟧ ⟹ mdist m2 (f x) (f y) ≤ B * mdist m1 x y"
    by (metis Lipschitz_continuous_map_pos f)
  show ?thesis
    unfolding Metric_space.mbounded_alt_pos [OF Metric_space_mspace_mdist]
  proof
    obtain C where "S ⊆ mspace m1" and "C>0" and C: "⋀x y. ⟦x ∈ S; y ∈ S⟧ ⟹ mdist m1 x y ≤ C"
      using S by (auto simp: Metric_space.mbounded_alt_pos [OF Metric_space_mspace_mdist])
    show "f ` S ⊆ mspace m2"
      using fim ‹S ⊆ mspace m1› by blast
    have "⋀x y. ⟦x ∈ S; y ∈ S⟧ ⟹ mdist m2 (f x) (f y) ≤ B * C"
      by (smt (verit) B C ‹0 < B› ‹S ⊆ mspace m1› mdist_nonneg mult_mono subsetD)
    moreover have "B*C > 0"
      by (simp add: ‹0 < B› ‹0 < C›)
    ultimately show "∃B>0. ∀x∈f ` S. ∀y∈f ` S. mdist m2 x y ≤ B"
      by auto
  qed
qed
lemma mtotally_bounded_Cauchy_continuous_image:
  assumes f: "Cauchy_continuous_map m1 m2 f" and S: "Metric_space.mtotally_bounded (mspace m1) (mdist m1) S"
  shows "Metric_space.mtotally_bounded (mspace m2) (mdist m2) (f ` S)"
  unfolding Metric_space.mtotally_bounded_sequentially[OF Metric_space_mspace_mdist]
proof (intro conjI strip)
  have "S ⊆ mspace m1"
    using S by (simp add: Metric_space.mtotally_bounded_sequentially[OF Metric_space_mspace_mdist])
  then show "f ` S ⊆ mspace m2"
    using Cauchy_continuous_map_funspace f by blast
  fix σ :: "nat ⇒ 'b"
  assume "range σ ⊆ f ` S"
  then have "∀n. ∃x. σ n = f x ∧ x ∈ S"
    by (meson imageE range_subsetD)
  then obtain ρ where ρ: "⋀n. σ n = f (ρ n)" "range ρ ⊆ S"
    by (metis image_subset_iff)
  then have "σ = f ∘ ρ"
    by fastforce
  obtain r where "strict_mono r" "Metric_space.MCauchy (mspace m1) (mdist m1) (ρ ∘ r)"
    by (meson ρ S Metric_space.mtotally_bounded_sequentially[OF Metric_space_mspace_mdist])
  then have "Metric_space.MCauchy (mspace m2) (mdist m2) (f ∘ ρ ∘ r)"
    using f unfolding Cauchy_continuous_map_def by (metis fun.map_comp)
  then show "∃r. strict_mono r ∧ Metric_space.MCauchy (mspace m2) (mdist m2) (σ ∘ r)"
    using ‹σ = f ∘ ρ› ‹strict_mono r› by blast
qed
lemma Lipschitz_coefficient_pos:
  assumes "x ∈ mspace m" "y ∈ mspace m" "f x ≠ f y" 
    and "f ∈ mspace m → mspace m2" 
    and "⋀x y. ⟦x ∈ mspace m; y ∈ mspace m⟧
            ⟹ mdist m2 (f x) (f y) ≤ k * mdist m x y"
  shows  "0 < k"
  using assms by (smt (verit, best) Pi_iff mdist_nonneg mdist_zero mult_nonpos_nonneg)
lemma Lipschitz_continuous_map_metric:
   "Lipschitz_continuous_map (prod_metric m m) euclidean_metric (λ(x,y). mdist m x y)"
proof -
  have "⋀x y x' y'. ⟦x ∈ mspace m; y ∈ mspace m; x' ∈ mspace m; y' ∈ mspace m⟧
       ⟹ ¦mdist m x y - mdist m x' y'¦ ≤ 2 * sqrt ((mdist m x x')⇧2 + (mdist m y y')⇧2)"
    by (smt (verit, del_insts) mdist_commute mdist_triangle real_sqrt_sum_squares_ge2)
  then show ?thesis
    by (fastforce simp: Lipschitz_continuous_map_def prod_dist_def dist_real_def)
qed
lemma Lipschitz_continuous_map_mdist:
  assumes f: "Lipschitz_continuous_map m m' f" 
    and g: "Lipschitz_continuous_map m m' g"
  shows "Lipschitz_continuous_map m euclidean_metric (λx. mdist m' (f x) (g x))"
    (is "Lipschitz_continuous_map m _ ?h")
proof -
  have eq: "?h = ((λ(x,y). mdist m' x y) ∘ (λx. (f x,g x)))"
    by force
  show ?thesis
    unfolding eq
  proof (rule Lipschitz_continuous_map_compose)
    show "Lipschitz_continuous_map m (prod_metric m' m') (λx. (f x, g x))"
      by (simp add: Lipschitz_continuous_map_paired f g)
    show "Lipschitz_continuous_map (prod_metric m' m') euclidean_metric (λ(x,y). mdist m' x y)"
      by (simp add: Lipschitz_continuous_map_metric)
  qed
qed
lemma uniformly_continuous_map_mdist:
  assumes f: "uniformly_continuous_map m m' f" 
    and g: "uniformly_continuous_map m m' g"
  shows "uniformly_continuous_map m euclidean_metric (λx. mdist m' (f x) (g x))"
    (is "uniformly_continuous_map m _ ?h")
proof -
  have eq: "?h = ((λ(x,y). mdist m' x y) ∘ (λx. (f x,g x)))"
    by force
  show ?thesis
    unfolding eq
  proof (rule uniformly_continuous_map_compose)
    show "uniformly_continuous_map m (prod_metric m' m') (λx. (f x, g x))"
      by (simp add: uniformly_continuous_map_paired f g)
    show "uniformly_continuous_map (prod_metric m' m') euclidean_metric (λ(x,y). mdist m' x y)"
      by (simp add: Lipschitz_continuous_map_metric Lipschitz_imp_uniformly_continuous_map)
  qed
qed
lemma Cauchy_continuous_map_mdist:
  assumes f: "Cauchy_continuous_map m m' f" 
    and g: "Cauchy_continuous_map m m' g"
  shows "Cauchy_continuous_map m euclidean_metric (λx. mdist m' (f x) (g x))"
    (is "Cauchy_continuous_map m _ ?h")
proof -
  have eq: "?h = ((λ(x,y). mdist m' x y) ∘ (λx. (f x,g x)))"
    by force
  show ?thesis
    unfolding eq
  proof (rule Cauchy_continuous_map_compose)
    show "Cauchy_continuous_map m (prod_metric m' m') (λx. (f x, g x))"
      by (simp add: Cauchy_continuous_map_paired f g)
    show "Cauchy_continuous_map (prod_metric m' m') euclidean_metric (λ(x,y). mdist m' x y)"
      by (simp add: Lipschitz_continuous_map_metric Lipschitz_imp_Cauchy_continuous_map)
  qed
qed
lemma mtopology_of_prod_metric [simp]:
    "mtopology_of (prod_metric m1 m2) = prod_topology (mtopology_of m1) (mtopology_of m2)"
  by (simp add: mtopology_of_def Metric_space12.mtopology_prod_metric[OF Metric_space12_mspace_mdist])
lemma continuous_map_metric:
   "continuous_map (prod_topology (mtopology_of m) (mtopology_of m)) euclidean
      (λ(x,y). mdist m x y)"
  using Lipschitz_continuous_imp_continuous_map [OF Lipschitz_continuous_map_metric]
  by auto
lemma continuous_map_mdist_alt:
  assumes "continuous_map X (prod_topology (mtopology_of m) (mtopology_of m)) f"
  shows "continuous_map X euclidean (λx. case_prod (mdist m) (f x))"
    (is "continuous_map _ _ ?h")
proof -
  have eq: "?h = case_prod (mdist m) ∘ f"
    by force
  show ?thesis
    unfolding eq
    using assms continuous_map_compose continuous_map_metric by blast
qed
lemma continuous_map_mdist [continuous_intros]:
  assumes f: "continuous_map X (mtopology_of m) f" 
      and g: "continuous_map X (mtopology_of m) g"
  shows "continuous_map X euclidean (λx. mdist m (f x) (g x))"
    (is "continuous_map X _ ?h")
proof -
  have eq: "?h = ((λ(x,y). mdist m x y) ∘ (λx. (f x,g x)))"
    by force
  show ?thesis
    unfolding eq
  proof (rule continuous_map_compose)
    show "continuous_map X (prod_topology (mtopology_of m) (mtopology_of m)) (λx. (f x, g x))"
      by (simp add: continuous_map_paired f g)
  qed (simp add: continuous_map_metric)
qed
lemma continuous_on_mdist:
   "a ∈ mspace m ⟹ continuous_map (mtopology_of m) euclidean (mdist m a)"
  by (simp add: continuous_map_mdist)
subsection ‹Isometries›
lemma (in Metric_space12) isometry_imp_embedding_map:
  assumes fim: "f ∈ M1 → M2" and d: "⋀x y. ⟦x ∈ M1; y ∈ M1⟧ ⟹ d2 (f x) (f y) = d1 x y"
  shows "embedding_map M1.mtopology M2.mtopology f"
proof -
  have "inj_on f M1"
    by (metis M1.zero d inj_onI)
  then obtain g where g: "⋀x. x ∈ M1 ⟹ g (f x) = x"
    by (metis inv_into_f_f)
  have "homeomorphic_maps M1.mtopology (subtopology M2.mtopology (f ` topspace M1.mtopology)) f g"
    unfolding homeomorphic_maps_def
  proof (intro conjI; clarsimp)
    show "continuous_map M1.mtopology (subtopology M2.mtopology (f ` M1)) f"
    proof (rule continuous_map_into_subtopology)
      show "continuous_map M1.mtopology M2.mtopology f"
        by (metis M1.metric_continuous_map M2.Metric_space_axioms d fim image_subset_iff_funcset)
    qed simp
    have "Lipschitz_continuous_map (submetric (metric(M2,d2)) (f ` M1)) (metric(M1,d1)) g"
      unfolding Lipschitz_continuous_map_def
    proof (intro conjI exI strip; simp)
      show "d1 (g x) (g y) ≤ 1 * d2 x y" if "x ∈ f ` M1 ∧ x ∈ M2" and "y ∈ f ` M1 ∧ y ∈ M2" for x y
        using that d g by force
    qed (use g in auto)
    then have "continuous_map (mtopology_of (submetric (metric(M2,d2)) (f ` M1))) M1.mtopology g"
      using Lipschitz_continuous_imp_continuous_map by force
    moreover have "mtopology_of (submetric (metric(M2,d2)) (f ` M1)) = subtopology M2.mtopology (f ` M1)"
      by (simp add: mtopology_of_submetric)
    ultimately show "continuous_map (subtopology M2.mtopology (f ` M1)) M1.mtopology g"
       by simp
  qed (use g in auto)
  then show ?thesis
    by (auto simp: embedding_map_def homeomorphic_map_maps)
qed
lemma (in Metric_space12) isometry_imp_homeomorphic_map:
  assumes fim: "f ` M1 = M2" and d: "⋀x y. ⟦x ∈ M1; y ∈ M1⟧ ⟹ d2 (f x) (f y) = d1 x y"
  shows "homeomorphic_map M1.mtopology M2.mtopology f"
  by (metis image_eqI M1.topspace_mtopology M2.subtopology_mspace d embedding_map_def fim isometry_imp_embedding_map Pi_iff)
subsection‹"Capped" equivalent bounded metrics and general product metrics›
definition (in Metric_space) capped_dist where
  "capped_dist ≡ λδ x y. if δ > 0 then min δ (d x y) else d x y"
lemma (in Metric_space) capped_dist: "Metric_space M (capped_dist δ)"
proof
  fix x y
  assume "x ∈ M" "y ∈ M"
  then show "(capped_dist δ x y = 0) = (x = y)"
    by (smt (verit, best) capped_dist_def zero)
  fix z 
  assume "z ∈ M"
  show "capped_dist δ x z ≤ capped_dist δ x y + capped_dist δ y z"
    unfolding capped_dist_def using ‹x ∈ M› ‹y ∈ M› ‹z ∈ M› 
    by (smt (verit, del_insts) Metric_space.mdist_pos_eq Metric_space_axioms mdist_reverse_triangle)
qed (use capped_dist_def commute in auto)
definition capped_metric where
  "capped_metric δ m ≡ metric(mspace m, Metric_space.capped_dist (mdist m) δ)"
lemma capped_metric:
  "capped_metric δ m = (if δ ≤ 0 then m else metric(mspace m, λx y. min δ (mdist m x y)))"
proof -
  interpret Metric_space "mspace m" "mdist m"
    by (simp add: Metric_space_mspace_mdist)
  show ?thesis
    by (auto simp: capped_metric_def capped_dist_def)
qed
lemma capped_metric_mspace [simp]:
  "mspace (capped_metric δ m) = mspace m"
  by (simp add: Metric_space.capped_dist Metric_space.mspace_metric capped_metric_def)
lemma capped_metric_mdist:
  "mdist (capped_metric δ m) = (λx y. if δ ≤ 0 then mdist m x y else min δ (mdist m x y))"
  by (metis Metric_space.capped_dist Metric_space.capped_dist_def Metric_space.mdist_metric 
      Metric_space_mspace_mdist capped_metric capped_metric_def leI)
lemma mdist_capped_le: "mdist (capped_metric δ m) x y ≤ mdist m x y"
  by (simp add: capped_metric_mdist)
lemma mdist_capped: "δ > 0 ⟹ mdist (capped_metric δ m) x y ≤ δ"
  by (simp add: capped_metric_mdist)
lemma mball_of_capped_metric [simp]: 
  assumes "x ∈ mspace m" "r > δ" "δ > 0" 
  shows "mball_of (capped_metric δ m) x r = mspace m"
proof -
  interpret Metric_space "mspace m" "mdist m"
    by auto
  have "Metric_space.mball (mspace m) (mdist (capped_metric δ m)) x r ⊆ mspace m"
    by (metis Metric_space.mball_subset_mspace Metric_space_mspace_mdist capped_metric_mspace)
  moreover have "mspace m ⊆ Metric_space.mball (mspace m) (mdist (capped_metric δ m)) x r"
    by (smt (verit) Metric_space.in_mball Metric_space_mspace_mdist assms capped_metric_mspace mdist_capped subset_eq)
  ultimately show ?thesis
    by (simp add: mball_of_def)
qed
lemma Metric_space_capped_dist[simp]:
  "Metric_space (mspace m) (Metric_space.capped_dist (mdist m) δ)"
  using Metric_space.capped_dist Metric_space_mspace_mdist by blast
lemma mtopology_capped_metric:
  "mtopology_of(capped_metric δ m) = mtopology_of m"
proof (cases "δ > 0")
  case True
  interpret Metric_space "mspace m" "mdist m"
    by (simp add: Metric_space_mspace_mdist)
  interpret Cap: Metric_space "mspace m" "mdist (capped_metric δ m)"
    by (metis Metric_space_mspace_mdist capped_metric_mspace)
  show ?thesis
    unfolding topology_eq
  proof
    fix S
    show "openin (mtopology_of (capped_metric δ m)) S = openin (mtopology_of m) S"
    proof (cases "S ⊆ mspace m")
      case True
      have "mball x r ⊆ Cap.mball x r" for x r
        by (smt (verit, ccfv_SIG) Cap.in_mball in_mball mdist_capped_le subsetI)
      moreover have "∃r>0. Cap.mball x r ⊆ S" if "r>0" "x ∈ S" and r: "mball x r ⊆ S" for r x
      proof (intro exI conjI)
        show "min (δ/2) r > 0"
          using ‹r>0› ‹δ>0› by force
        show "Cap.mball x (min (δ/2) r) ⊆ S"
          using that
          by clarsimp (smt (verit) capped_metric_mdist field_sum_of_halves in_mball subsetD)
      qed
      ultimately have "(∃r>0. Cap.mball x r ⊆ S) = (∃r>0. mball x r ⊆ S)" if "x ∈ S" for x
        by (meson subset_trans that)
      then show ?thesis
        by (simp add: mtopology_of_def openin_mtopology Cap.openin_mtopology)
    qed (simp add: openin_closedin_eq)
  qed
qed (simp add: capped_metric)
text ‹Might have been easier to prove this within the locale to start with (using Self)›
lemma (in Metric_space) mtopology_capped_metric:
  "Metric_space.mtopology M (capped_dist δ) = mtopology"
  using mtopology_capped_metric [of δ "metric(M,d)"]
  by (simp add: Metric_space.mtopology_of capped_dist capped_metric_def)
lemma (in Metric_space) MCauchy_capped_metric:
  "Metric_space.MCauchy M (capped_dist δ) σ ⟷ MCauchy σ"
proof (cases "δ > 0")
  case True
  interpret Cap: Metric_space "M" "capped_dist δ"
    by (simp add: capped_dist)
  show ?thesis
  proof
    assume σ: "Cap.MCauchy σ"
    show "MCauchy σ"
      unfolding MCauchy_def
    proof (intro conjI strip)
      show "range σ ⊆ M"
        using Cap.MCauchy_def σ by presburger
      fix ε :: real
      assume "ε > 0"
      with True σ
      obtain N where "∀n n'. N ≤ n ⟶ N ≤ n' ⟶ capped_dist δ (σ n) (σ n') < min δ ε"
        unfolding Cap.MCauchy_def by (metis min_less_iff_conj)
      with True show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
        by (force simp: capped_dist_def)
    qed
  next
    assume "MCauchy σ"
    then show "Cap.MCauchy σ"
      unfolding MCauchy_def Cap.MCauchy_def by (force simp: capped_dist_def)
  qed
qed (simp add: capped_dist_def)
lemma (in Metric_space) mcomplete_capped_metric:
   "Metric_space.mcomplete M (capped_dist δ) ⟷ mcomplete"
  by (simp add: MCauchy_capped_metric Metric_space.mcomplete_def capped_dist mtopology_capped_metric mcomplete_def)
lemma bounded_equivalent_metric:
  assumes "δ > 0"
  obtains m' where "mspace m' = mspace m" "mtopology_of m' = mtopology_of m" "⋀x y. mdist m' x y < δ"
proof
  let ?m = "capped_metric (δ/2) m"
  fix x y
  show "mdist ?m x y < δ"
    by (smt (verit, best) assms field_sum_of_halves mdist_capped)    
qed (auto simp: mtopology_capped_metric)
text ‹A technical lemma needed below›
lemma Sup_metric_cartesian_product:
  fixes I m
  defines "S ≡ PiE I (mspace ∘ m)"
  defines "D ≡ λx y. if x ∈ S ∧ y ∈ S then SUP i∈I. mdist (m i) (x i) (y i) else 0"
  defines "m' ≡ metric(S,D)"
  assumes "I ≠ {}"
    and c: "⋀i x y. ⟦i ∈ I; x ∈ mspace(m i); y ∈ mspace(m i)⟧ ⟹ mdist (m i) x y ≤ c"
  shows "Metric_space S D" 
    and "∀x ∈ S. ∀y ∈ S. ∀b. D x y ≤ b ⟷ (∀i ∈ I. mdist (m i) (x i) (y i) ≤ b)"  (is "?the2")
proof -
  have bdd: "bdd_above ((λi. mdist (m i) (x i) (y i)) ` I)"
    if "x ∈ S" "y ∈ S" for x y 
    using c that by (force simp: S_def bdd_above_def)
  have D_iff: "D x y ≤ b ⟷ (∀i ∈ I. mdist (m i) (x i) (y i) ≤ b)"
    if "x ∈ S" "y ∈ S" for x y b
    using that ‹I ≠ {}› by (simp add: D_def PiE_iff cSup_le_iff bdd)
  show "Metric_space S D"
  proof
    fix x y
    show D0: "0 ≤ D x y"
      using bdd ‹I ≠ {}›
      by (metis D_def D_iff Orderings.order_eq_iff dual_order.trans ex_in_conv mdist_nonneg) 
    show "D x y = D y x"
      by (simp add: D_def mdist_commute)
    assume "x ∈ S" and "y ∈ S"
    then
    have "D x y = 0 ⟷ (∀i∈I. mdist (m i) (x i) (y i) = 0)"
      using D0 D_iff [of x y 0] nle_le by fastforce
    also have "... ⟷ x = y"
      using ‹x ∈ S› ‹y ∈ S› by (fastforce simp: S_def PiE_iff extensional_def)
    finally show "(D x y = 0) ⟷ (x = y)" .
    fix z
    assume "z ∈ S"
    have "mdist (m i) (x i) (z i) ≤ D x y + D y z" if "i ∈ I" for i
    proof -
      have "mdist (m i) (x i) (z i) ≤ mdist (m i) (x i) (y i) + mdist (m i) (y i) (z i)"
        by (metis PiE_E S_def ‹x ∈ S› ‹y ∈ S› ‹z ∈ S› comp_apply mdist_triangle that)
      also have "... ≤ D x y + D y z"
        using ‹x ∈ S› ‹y ∈ S› ‹z ∈ S› by (meson D_iff add_mono order_refl that)
      finally show ?thesis .
    qed
    then show "D x z ≤ D x y + D y z"
      by (simp add: D_iff ‹x ∈ S› ‹z ∈ S›)
  qed
  then interpret Metric_space S D .
  show ?the2
  proof (intro strip)
    show "(D x y ≤ b) = (∀i∈I. mdist (m i) (x i) (y i) ≤ b)"
      if "x ∈ S" and "y ∈ S" for x y b
      using that by (simp add: D_iff m'_def)
  qed
qed
lemma metrizable_topology_A:
  assumes "metrizable_space (product_topology X I)"
  shows "(product_topology X I) = trivial_topology ∨ (∀i ∈ I. metrizable_space (X i))"
    by (meson assms metrizable_space_retraction_map_image retraction_map_product_projection)
lemma metrizable_topology_C:
  assumes "completely_metrizable_space (product_topology X I)"
  shows "(product_topology X I) = trivial_topology ∨ (∀i ∈ I. completely_metrizable_space (X i))"
    by (meson assms completely_metrizable_space_retraction_map_image retraction_map_product_projection)
lemma metrizable_topology_B:
  fixes a X I
  defines "L ≡ {i ∈ I. ∄a. topspace (X i) ⊆ {a}}"
  assumes "topspace (product_topology X I) ≠ {}"
    and met: "metrizable_space (product_topology X I)"
    and "⋀i. i ∈ I ⟹ metrizable_space (X i)"
  shows  "countable L"
proof -
  have "⋀i. ∃p q. i ∈ L ⟶ p ∈ topspace(X i) ∧ q ∈ topspace(X i) ∧ p ≠ q"
    unfolding L_def by blast
  then obtain φ ψ where φ: "⋀i. i ∈ L ⟹ φ i ∈ topspace(X i) ∧ ψ i ∈ topspace(X i) ∧ φ i ≠ ψ i"
    by metis
  obtain z where z: "z ∈ (Π⇩E i∈I. topspace (X i))"
    using assms(2) by fastforce
  define p where "p ≡ λi. if i ∈ L then φ i else z i"
  define q where "q ≡ λi j. if j = i then ψ i else p j"
  have p: "p ∈ topspace(product_topology X I)"
    using z φ by (auto simp: p_def L_def)
  then have q: "⋀i. i ∈ L ⟹ q i ∈ topspace (product_topology X I)" 
    by (auto simp: L_def q_def φ)
  have fin: "finite {i ∈ L. q i ∉ U}" if U: "openin (product_topology X I) U" "p ∈ U" for U
  proof -
    obtain V where V: "finite {i ∈ I. V i ≠ topspace (X i)}" "(∀i∈I. openin (X i) (V i))" "p ∈ Pi⇩E I V" "Pi⇩E I V ⊆ U"
      using U by (force simp: openin_product_topology_alt)
    moreover 
    have "V x ≠ topspace (X x)" if "x ∈ L" and "q x ∉ U" for x
      using that V q
      by (smt (verit, del_insts) PiE_iff q_def subset_eq topspace_product_topology)
    then have "{i ∈ L. q i ∉ U} ⊆ {i ∈ I. V i ≠ topspace (X i)}"
      by (force simp: L_def)
    ultimately show ?thesis
      by (meson finite_subset)
  qed
  obtain M d where "Metric_space M d" and XI: "product_topology X I = Metric_space.mtopology M d"
    using met metrizable_space_def by blast
  then interpret Metric_space M d
    by blast
  define C where "C ≡ ⋃n::nat. {i ∈ L. q i ∉ mball p (inverse (Suc n))}"
  have "finite {i ∈ L. q i ∉ mball p (inverse (real (Suc n)))}" for n
    using XI p  by (intro fin; force)
  then have "countable C"
    unfolding C_def
    by (meson countableI_type countable_UN countable_finite)
  moreover have "L ⊆ C"
  proof (clarsimp simp: C_def)
    fix i
    assume "i ∈ L" and "q i ∈ M" and "p ∈ M"
    then show "∃n. ¬ d p (q i) < inverse (1 + real n)"
      using reals_Archimedean [of "d p (q i)"]
      by (metis φ mdist_pos_eq not_less_iff_gr_or_eq of_nat_Suc p_def q_def)
  qed
  ultimately show ?thesis
    using countable_subset by blast
qed
lemma metrizable_topology_DD:
  assumes "topspace (product_topology X I) ≠ {}"
    and co: "countable {i ∈ I. ∄a. topspace (X i) ⊆ {a}}"
    and m: "⋀i. i ∈ I ⟹ X i = mtopology_of (m i)"
  obtains M d where "Metric_space M d" "product_topology X I = Metric_space.mtopology M d"
                    "(⋀i. i ∈ I ⟹ mcomplete_of (m i)) ⟹ Metric_space.mcomplete M d"
proof (cases "I = {}")
  case True
  then show ?thesis
    by (metis discrete_metric.mcomplete_discrete_metric discrete_metric.mtopology_discrete_metric metric_M_dd product_topology_empty_discrete that)
next
  case False
  obtain nk and C:: "nat set" where nk: "{i ∈ I. ∄a. topspace (X i) ⊆ {a}} = nk ` C" and "inj_on nk C"
    using co by (force simp: countable_as_injective_image_subset)
  then obtain kn where kn: "⋀w. w ∈ C ⟹ kn (nk w) = w"
    by (metis inv_into_f_f)
  define cm where "cm ≡ λi. capped_metric (inverse(Suc(kn i))) (m i)"
  have mspace_cm: "mspace (cm i) = mspace (m i)" for i
    by (simp add: cm_def)
  have c1: "⋀i x y. mdist (cm i) x y ≤ 1"
    by (simp add: cm_def capped_metric_mdist min_le_iff_disj divide_simps)
  then have bdd: "bdd_above ((λi. mdist (cm i) (x i) (y i)) ` I)" for x y
    by (meson bdd_above.I2)
  define M where "M ≡ Pi⇩E I (mspace ∘ cm)"
  define d where "d ≡ λx y. if x ∈ M ∧ y ∈ M then SUP i∈I. mdist (cm i) (x i) (y i) else 0"
  have d_le1: "d x y ≤ 1" for x y
    using ‹I ≠ {}› c1 by (simp add: d_def bdd cSup_le_iff)
  with ‹I ≠ {}› Sup_metric_cartesian_product [of I cm]
  have "Metric_space M d" 
    and *: "∀x∈M. ∀y∈M. ∀b. (d x y ≤ b) ⟷ (∀i∈I. mdist (cm i) (x i) (y i) ≤ b)"
    by (auto simp: False bdd M_def d_def cSUP_le_iff intro: c1) 
  then interpret Metric_space M d 
    by metis
  have le_d: "mdist (cm i) (x i) (y i) ≤ d x y" if "i ∈ I" "x ∈ M" "y ∈ M" for i x y
    using "*" that by blast
  have product_m: "PiE I (λi. mspace (m i)) = topspace(product_topology X I)"
    using m by force
  define m' where "m' = metric (M,d)"
  define J where "J ≡ λU. {i ∈ I. U i ≠ topspace (X i)}"
  have 1: "∃U. finite (J U) ∧ (∀i∈I. openin (X i) (U i)) ∧ x ∈ Pi⇩E I U ∧ Pi⇩E I U ⊆ mball z r"
    if "x ∈ M" "z ∈ M" and r: "0 < r" "d z x < r" for x z r
  proof -
    have x: "⋀i. i ∈ I ⟹ x i ∈ topspace(X i)"
      using M_def m mspace_cm that(1) by auto
    have z: "⋀i. i ∈ I ⟹ z i ∈ topspace(X i)"
      using M_def m mspace_cm that(2) by auto
    obtain R where "0 < R" "d z x < R" "R < r"
      using r dense by (smt (verit, ccfv_threshold))
    define U where "U ≡ λi. if R ≤ inverse(Suc(kn i)) then mball_of (m i) (z i) R else topspace(X i)"
    show ?thesis
    proof (intro exI conjI)
      obtain n where n: "real n * R > 1"
        using ‹0 < R› ex_less_of_nat_mult by blast
      have "finite (nk ` (C ∩ {..n}))"
        by force
      moreover 
      have "∃m. m ∈ C ∧ m ≤ n ∧ i = nk m"
        if R: "R ≤ inverse (1 + real (kn i))" and "i ∈ I" 
          and neq: "mball_of (m i) (z i) R ≠ topspace (X i)" for i 
      proof -
        interpret MI: Metric_space "mspace (m i)" "mdist (m i)"
          by auto
        have "MI.mball (z i) R ≠ topspace (X i)"
          by (metis mball_of_def neq)
        then have "∄a. topspace (X i) ⊆ {a}"
          using ‹0 < R› m subset_antisym ‹i ∈ I› z by fastforce
        then have "i ∈ nk ` C"
          using nk ‹i ∈ I› by auto
        then show ?thesis
          by (smt (verit, ccfv_SIG) R ‹0 < R› image_iff kn lift_Suc_mono_less_iff mult_mono n not_le_imp_less of_nat_0_le_iff of_nat_Suc right_inverse)
      qed
      then have "J U ⊆ nk ` (C ∩ {..n})"
        by (auto simp: image_iff Bex_def J_def U_def split: if_split_asm)
      ultimately show "finite (J U)"
        using finite_subset by blast
      show "∀i∈I. openin (X i) (U i)"
        by (simp add: Metric_space.openin_mball U_def mball_of_def mtopology_of_def m)
      have xin: "x ∈ Pi⇩E I (topspace ∘ X)"
        using M_def ‹x ∈ M› x by auto
      moreover 
      have "⋀i. ⟦i ∈ I; R ≤ inverse (1 + real (kn i))⟧ ⟹ mdist (m i) (z i) (x i) < R"
        by (smt (verit, ccfv_SIG) ‹d z x < R› capped_metric_mdist cm_def le_d of_nat_Suc that)
      ultimately show "x ∈ Pi⇩E I U"
        using m z by (auto simp: U_def PiE_iff)
      show "Pi⇩E I U ⊆ mball z r"
      proof
        fix y
        assume y: "y ∈ Pi⇩E I U"
        then have "y ∈ M"
          by (force simp: PiE_iff M_def U_def m mspace_cm split: if_split_asm)
        moreover 
        have "∀i∈I. mdist (cm i) (z i) (y i) ≤ R"
          by (smt (verit) PiE_mem U_def cm_def in_mball_of inverse_Suc mdist_capped mdist_capped_le y)
        then have "d z y ≤ R"
          by (simp add: ‹y ∈ M› ‹z ∈ M› *)
        ultimately show "y ∈ mball z r"
          using ‹R < r› ‹z ∈ M› by force
      qed
    qed
  qed
  have 2: "∃r>0. mball x r ⊆ S"
    if "finite (J U)" and x: "x ∈ Pi⇩E I U" and S: "Pi⇩E I U ⊆ S"
      and U: "⋀i. i∈I ⟹ openin (X i) (U i)" 
      and "x ∈ S" for U S x
  proof -
    { fix i
      assume "i ∈ J U"
      then have "i ∈ I"
        by (auto simp: J_def)
      then have "openin (mtopology_of (m i)) (U i)"
        using U m by force
      then have "openin (mtopology_of (cm i)) (U i)"
        by (simp add: Abstract_Metric_Spaces.mtopology_capped_metric cm_def)
      then have "∃r>0. mball_of (cm i) (x i) r ⊆ U i"
        using x
        by (simp add: Metric_space.openin_mtopology PiE_mem ‹i ∈ I› mball_of_def mtopology_of_def) 
    }
    then obtain rf where rf: "⋀j. j ∈ J U ⟹ rf j >0 ∧ mball_of (cm j) (x j) (rf j) ⊆ U j"
      by metis
    define r where "r ≡ Min (insert 1 (rf ` J U))"
    show ?thesis
    proof (intro exI conjI)
      show "r > 0"
        by (simp add: ‹finite (J U)› r_def rf)
      have r [simp]: "⋀j. j ∈ J U ⟹ r ≤ rf j" "r ≤ 1"
        by (auto simp: r_def that(1))
      have *: "mball_of (cm i) (x i) r ⊆ U i" if "i ∈ I" for i
      proof (cases "i ∈ J U")
        case True
        with r show ?thesis
          by (smt (verit) Metric_space.in_mball Metric_space_mspace_mdist mball_of_def rf subset_eq)
      next
        case False
        then show ?thesis
          by (simp add: J_def cm_def m subset_eq that)
      qed
      show "mball x r ⊆ S"
        by (smt (verit) x * in_mball_of M_def Metric_space.in_mball Metric_space_axioms PiE_iff le_d o_apply subset_eq S)
    qed
  qed
  have 3: "x ∈ M"
    if §: "⋀x. x∈S ⟹ ∃U. finite (J U) ∧ (∀i∈I. openin (X i) (U i)) ∧ x ∈ Pi⇩E I U ∧ Pi⇩E I U ⊆ S"
      and "x ∈ S" for S x
    using § [OF ‹x ∈ S›] m openin_subset by (fastforce simp: M_def PiE_iff cm_def)
  show thesis
  proof
    show "Metric_space M d"
      using Metric_space_axioms by blast
    show eq: "product_topology X I = Metric_space.mtopology M d"
      unfolding topology_eq openin_mtopology openin_product_topology_alt  
      using J_def 1 2 3 subset_iff zero by (smt (verit, ccfv_threshold))
    show "mcomplete" if "⋀i. i ∈ I ⟹ mcomplete_of (m i)" 
      unfolding mcomplete_def
    proof (intro strip)
      fix σ
      assume σ: "MCauchy σ"
      have "∃y. i ∈ I ⟶ limitin (X i) (λn. σ n i) y sequentially" for i
      proof (cases "i ∈ I")
        case True
        interpret MI: Metric_space "mspace (m i)" "mdist (m i)"
          by auto
        have "⋀σ. MI.MCauchy σ ⟶ (∃x. limitin MI.mtopology σ x sequentially)"
          by (meson MI.mcomplete_def True mcomplete_of_def that)
        moreover have "MI.MCauchy (λn. σ n i)"
          unfolding MI.MCauchy_def
        proof (intro conjI strip)
          show "range (λn. σ n i) ⊆ mspace (m i)"
            by (smt (verit, ccfv_threshold) MCauchy_def PiE_iff True σ eq image_subset_iff m topspace_mtopology topspace_mtopology_of topspace_product_topology)
          fix ε::real
          define r where "r ≡ min ε (inverse(Suc (kn i)))"
          assume "ε > 0"
          then have "r > 0"
            by (simp add: r_def)
          then obtain N where N: "⋀n n'. N ≤ n ∧ N ≤ n' ⟹ d (σ n) (σ n') < r"
            using σ unfolding MCauchy_def by meson
          show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ mdist (m i) (σ n i) (σ n' i) < ε"
          proof (intro strip exI)
            fix n n'
            assume "N ≤ n" and "N ≤ n'"
            then have "mdist (cm i) (σ n i) (σ n' i) < r"
              using *
              by (smt (verit) Metric_space.MCauchy_def Metric_space_axioms N True σ rangeI subsetD)
            then
            show "mdist (m i) (σ n i) (σ n' i) < ε"
              unfolding cm_def r_def
              by (smt (verit, ccfv_SIG) capped_metric_mdist)
          qed
        qed
        ultimately show ?thesis
          by (simp add: m mtopology_of_def)
      qed auto
      then obtain y where "⋀i. i ∈ I ⟹ limitin (X i) (λn. σ n i) (y i) sequentially"
        by metis
      with σ show "∃x. limitin mtopology σ x sequentially"
        apply (rule_tac x="λi∈I. y i" in exI)
        apply (simp add: MCauchy_def limitin_componentwise flip: eq)
        by (metis eq eventually_at_top_linorder range_subsetD topspace_mtopology topspace_product_topology)
    qed
  qed
qed
lemma metrizable_topology_D:
  assumes "topspace (product_topology X I) ≠ {}"
    and co: "countable {i ∈ I. ∄a. topspace (X i) ⊆ {a}}"
    and met: "⋀i. i ∈ I ⟹ metrizable_space (X i)"
  shows "metrizable_space (product_topology X I)"
proof -
  have "⋀i. i ∈ I ⟹ ∃m. X i = mtopology_of m"
    by (metis Metric_space.mtopology_of met metrizable_space_def)
  then obtain m where m: "⋀i. i ∈ I ⟹ X i = mtopology_of (m i)"
    by metis 
  then show ?thesis
    using metrizable_topology_DD [of X I m] assms by (force simp: metrizable_space_def)
qed
lemma metrizable_topology_E:
  assumes "topspace (product_topology X I) ≠ {}"
    and "countable {i ∈ I. ∄a. topspace (X i) ⊆ {a}}"
    and met: "⋀i. i ∈ I ⟹ completely_metrizable_space (X i)"
  shows "completely_metrizable_space (product_topology X I)"
proof -
  have "⋀i. i ∈ I ⟹ ∃m. mcomplete_of m ∧ X i = mtopology_of m"
    using met Metric_space.mtopology_of Metric_space.mcomplete_of unfolding completely_metrizable_space_def
    by metis 
  then obtain m where "⋀i. i ∈ I ⟹ mcomplete_of (m i) ∧ X i = mtopology_of (m i)"
    by metis 
  then show ?thesis
    using  metrizable_topology_DD [of X I m] assms unfolding metrizable_space_def
    by (metis (full_types) completely_metrizable_space_def)
qed
proposition metrizable_space_product_topology:
  "metrizable_space (product_topology X I) ⟷
        (product_topology X I) = trivial_topology ∨
        countable {i ∈ I. ¬ (∃a. topspace(X i) ⊆ {a})} ∧
        (∀i ∈ I. metrizable_space (X i))"
  by (metis (mono_tags, lifting) empty_metrizable_space metrizable_topology_A metrizable_topology_B metrizable_topology_D subtopology_eq_discrete_topology_empty)
proposition completely_metrizable_space_product_topology:
  "completely_metrizable_space (product_topology X I) ⟷
        (product_topology X I) = trivial_topology ∨
        countable {i ∈ I. ¬ (∃a. topspace(X i) ⊆ {a})} ∧
        (∀i ∈ I. completely_metrizable_space (X i))"
  by (smt (verit, del_insts) Collect_cong completely_metrizable_imp_metrizable_space empty_completely_metrizable_space metrizable_topology_B metrizable_topology_C metrizable_topology_E subtopology_eq_discrete_topology_empty)
lemma completely_metrizable_Euclidean_space:
  "completely_metrizable_space(Euclidean_space n)"
  unfolding Euclidean_space_def
proof (rule completely_metrizable_space_closedin)
  show "completely_metrizable_space (powertop_real (UNIV::nat set))"
    by (simp add: completely_metrizable_space_product_topology completely_metrizable_space_euclidean)
  show "closedin (powertop_real UNIV) {x. ∀i≥n. x i = 0}"
    using closedin_Euclidean_space topspace_Euclidean_space by auto
qed
lemma metrizable_Euclidean_space:
   "metrizable_space(Euclidean_space n)"
  by (simp add: completely_metrizable_Euclidean_space completely_metrizable_imp_metrizable_space)
lemma locally_connected_Euclidean_space:
   "locally_connected_space(Euclidean_space n)"
  by (simp add: locally_path_connected_Euclidean_space locally_path_connected_imp_locally_connected_space)
end