Theory Urysohn
section ‹The Urysohn lemma, its consequences and other advanced material about metric spaces›
theory Urysohn
imports Abstract_Topological_Spaces Abstract_Metric_Spaces Infinite_Sum Arcwise_Connected
begin
subsection ‹Urysohn lemma and Tietze's theorem›
proposition Urysohn_lemma:
fixes a b :: real
assumes "normal_space X" "closedin X S" "closedin X T" "disjnt S T" "a ≤ b"
obtains f where "continuous_map X (top_of_set {a..b}) f" "f ` S ⊆ {a}" "f ` T ⊆ {b}"
proof -
obtain U where "openin X U" "S ⊆ U" "X closure_of U ⊆ topspace X - T"
using assms unfolding normal_space_alt disjnt_def
by (metis Diff_mono Un_Diff_Int closedin_def subset_eq sup_bot_right)
have "∃G :: real ⇒ 'a set. G 0 = U ∧ G 1 = topspace X - T ∧
(∀x ∈ dyadics ∩ {0..1}. ∀y ∈ dyadics ∩ {0..1}. x < y ⟶ openin X (G x) ∧ openin X (G y) ∧ X closure_of (G x) ⊆ G y)"
proof (rule recursion_on_dyadic_fractions)
show "openin X U ∧ openin X (topspace X - T) ∧ X closure_of U ⊆ topspace X - T"
using ‹X closure_of U ⊆ topspace X - T› ‹openin X U› ‹closedin X T› by blast
show "∃z. (openin X x ∧ openin X z ∧ X closure_of x ⊆ z) ∧ openin X z ∧ openin X y ∧ X closure_of z ⊆ y"
if "openin X x ∧ openin X y ∧ X closure_of x ⊆ y" for x y
by (meson that closedin_closure_of normal_space_alt ‹normal_space X›)
show "openin X x ∧ openin X z ∧ X closure_of x ⊆ z"
if "openin X x ∧ openin X y ∧ X closure_of x ⊆ y" and "openin X y ∧ openin X z ∧ X closure_of y ⊆ z" for x y z
by (meson that closure_of_subset openin_subset subset_trans)
qed
then obtain G :: "real ⇒ 'a set"
where G0: "G 0 = U" and G1: "G 1 = topspace X - T"
and G: "⋀x y. ⟦x ∈ dyadics; y ∈ dyadics; 0 ≤ x; x < y; y ≤ 1⟧
⟹ openin X (G x) ∧ openin X (G y) ∧ X closure_of (G x) ⊆ G y"
by (smt (verit, del_insts) Int_iff atLeastAtMost_iff)
define f where "f ≡ λx. Inf(insert 1 {r. r ∈ dyadics ∩ {0..1} ∧ x ∈ G r})"
have f_ge: "f x ≥ 0" if "x ∈ topspace X" for x
unfolding f_def by (force intro: cInf_greatest)
moreover have f_le1: "f x ≤ 1" if "x ∈ topspace X" for x
proof -
have "bdd_below {r ∈ dyadics ∩ {0..1}. x ∈ G r}"
by (auto simp: bdd_below_def)
then show ?thesis
by (auto simp: f_def cInf_lower)
qed
ultimately have fim: "f ` topspace X ⊆ {0..1}"
by (auto simp: f_def)
have 0: "0 ∈ dyadics ∩ {0..1::real}" and 1: "1 ∈ dyadics ∩ {0..1::real}"
by (force simp: dyadics_def)+
then have opeG: "openin X (G r)" if "r ∈ dyadics ∩ {0..1}" for r
using G G0 ‹openin X U› less_eq_real_def that by auto
have "x ∈ G 0" if "x ∈ S" for x
using G0 ‹S ⊆ U› that by blast
with 0 have fimS: "f ` S ⊆ {0}"
unfolding f_def by (force intro!: cInf_eq_minimum)
have False if "r ∈ dyadics" "0 ≤ r" "r < 1" "x ∈ G r" "x ∈ T" for r x
using G [of r 1] 1
by (smt (verit, best) DiffD2 G1 Int_iff closure_of_subset inf.orderE openin_subset that)
then have "r≥1" if "r ∈ dyadics" "0 ≤ r" "r ≤ 1" "x ∈ G r" "x ∈ T" for r x
using linorder_not_le that by blast
then have fimT: "f ` T ⊆ {1}"
unfolding f_def by (force intro!: cInf_eq_minimum)
have fle1: "f z ≤ 1" for z
by (force simp: f_def intro: cInf_lower)
have fle: "f z ≤ x" if "x ∈ dyadics ∩ {0..1}" "z ∈ G x" for z x
using that by (force simp: f_def intro: cInf_lower)
have *: "b ≤ f z" if "b ≤ 1" "⋀x. ⟦x ∈ dyadics ∩ {0..1}; z ∈ G x⟧ ⟹ b ≤ x" for z b
using that by (force simp: f_def intro: cInf_greatest)
have **: "r ≤ f x" if r: "r ∈ dyadics ∩ {0..1}" "x ∉ G r" for r x
proof (rule *)
show "r ≤ s" if "s ∈ dyadics ∩ {0..1}" and "x ∈ G s" for s :: real
using that r G [of s r] by (force simp: dest: closure_of_subset openin_subset)
qed (use that in force)
have "∃U. openin X U ∧ x ∈ U ∧ (∀y ∈ U. ¦f y - f x¦ < ε)"
if "x ∈ topspace X" and "0 < ε" for x ε
proof -
have A: "∃r. r ∈ dyadics ∩ {0..1} ∧ r < y ∧ ¦r - y¦ < d" if "0 < y" "y ≤ 1" "0 < d" for y d::real
proof -
obtain n q r
where "of_nat q / 2^n < y" "y < of_nat r / 2^n" "¦q / 2^n - r / 2^n ¦ < d"
by (smt (verit, del_insts) padic_rational_approximation_straddle_pos ‹0 < d› ‹0 < y›)
then show ?thesis
unfolding dyadics_def
using divide_eq_0_iff that(2) by fastforce
qed
have B: "∃r. r ∈ dyadics ∩ {0..1} ∧ y < r ∧ ¦r - y¦ < d" if "0 ≤ y" "y < 1" "0 < d" for y d::real
proof -
obtain n q r
where "of_nat q / 2^n ≤ y" "y < of_nat r / 2^n" "¦q / 2^n - r / 2^n ¦ < d"
using padic_rational_approximation_straddle_pos_le
by (smt (verit, del_insts) ‹0 < d› ‹0 ≤ y›)
then show ?thesis
apply (clarsimp simp: dyadics_def)
using divide_eq_0_iff ‹y < 1›
by (smt (verit) divide_nonneg_nonneg divide_self of_nat_0_le_iff of_nat_1 power_0 zero_le_power)
qed
show ?thesis
proof (cases "f x = 0")
case True
with B[of 0] obtain r where r: "r ∈ dyadics ∩ {0..1}" "0 < r" "¦r¦ < ε/2"
by (smt (verit) ‹0 < ε› half_gt_zero)
show ?thesis
proof (intro exI conjI)
show "openin X (G r)"
using opeG r(1) by blast
show "x ∈ G r"
using True ** r by force
show "∀y ∈ G r. ¦f y - f x¦ < ε"
using f_ge ‹openin X (G r)› fle openin_subset r by (fastforce simp: True)
qed
next
case False
show ?thesis
proof (cases "f x = 1")
case True
with A[of 1] obtain r where r: "r ∈ dyadics ∩ {0..1}" "r < 1" "¦r-1¦ < ε/2"
by (smt (verit) ‹0 < ε› half_gt_zero)
define G' where "G' ≡ topspace X - X closure_of G r"
show ?thesis
proof (intro exI conjI)
show "openin X G'"
unfolding G'_def by fastforce
obtain r' where "r' ∈ dyadics ∧ 0 ≤ r' ∧ r' ≤ 1 ∧ r < r' ∧ ¦r' - r¦ < 1 - r"
using B r by force
moreover
have "1 - r ∈ dyadics" "0 ≤ r"
using 1 r dyadics_diff by force+
ultimately have "x ∉ X closure_of G r"
using G True r fle by force
then show "x ∈ G'"
by (simp add: G'_def that)
show "∀y ∈ G'. ¦f y - f x¦ < ε"
using ** f_le1 in_closure_of r by (fastforce simp: True G'_def)
qed
next
case False
have "0 < f x" "f x < 1"
using fle1 f_ge that(1) ‹f x ≠ 0› ‹f x ≠ 1› by (metis order_le_less) +
obtain r where r: "r ∈ dyadics ∩ {0..1}" "r < f x" "¦r - f x¦ < ε / 2"
using A ‹0 < ε› ‹0 < f x› ‹f x < 1› by (smt (verit, best) half_gt_zero)
obtain r' where r': "r' ∈ dyadics ∩ {0..1}" "f x < r'" "¦r' - f x¦ < ε / 2"
using B ‹0 < ε› ‹0 < f x› ‹f x < 1› by (smt (verit, best) half_gt_zero)
have "r < 1"
using ‹f x < 1› r(2) by force
show ?thesis
proof (intro conjI exI)
show "openin X (G r' - X closure_of G r)"
using closedin_closure_of opeG r' by blast
have "x ∈ X closure_of G r ⟹ False"
using B [of r "f x - r"] r ‹r < 1› G [of r] fle by force
then show "x ∈ G r' - X closure_of G r"
using ** r' by fastforce
show "∀y∈G r' - X closure_of G r. ¦f y - f x¦ < ε"
using r r' ** G closure_of_subset field_sum_of_halves fle openin_subset subset_eq
by (smt (verit) DiffE opeG)
qed
qed
qed
qed
then have contf: "continuous_map X (top_of_set {0..1}) f"
by (force simp: Met_TC.continuous_map_to_metric dist_real_def continuous_map_in_subtopology fim simp flip: mtopology_is_euclidean)
define g where "g ≡ λx. a + (b - a) * f x"
show thesis
proof
have "continuous_map X euclideanreal g"
using contf ‹a ≤ b› unfolding g_def by (auto simp: continuous_intros continuous_map_in_subtopology)
moreover have "g ` (topspace X) ⊆ {a..b}"
using mult_left_le [of "f _" "b-a"] contf ‹a ≤ b›
by (simp add: g_def add.commute continuous_map_in_subtopology image_subset_iff le_diff_eq)
ultimately show "continuous_map X (top_of_set {a..b}) g"
by (meson continuous_map_in_subtopology)
show "g ` S ⊆ {a}" "g ` T ⊆ {b}"
using fimS fimT by (auto simp: g_def)
qed
qed
lemma Urysohn_lemma_alt:
fixes a b :: real
assumes "normal_space X" "closedin X S" "closedin X T" "disjnt S T"
obtains f where "continuous_map X euclideanreal f" "f ` S ⊆ {a}" "f ` T ⊆ {b}"
by (metis Urysohn_lemma assms continuous_map_in_subtopology disjnt_sym linear)
lemma normal_space_iff_Urysohn_gen_alt:
assumes "a ≠ b"
shows "normal_space X ⟷
(∀S T. closedin X S ∧ closedin X T ∧ disjnt S T
⟶ (∃f. continuous_map X euclideanreal f ∧ f ` S ⊆ {a} ∧ f ` T ⊆ {b}))"
(is "?lhs=?rhs")
proof
show "?lhs ⟹ ?rhs"
by (metis Urysohn_lemma_alt)
next
assume R: ?rhs
show ?lhs
unfolding normal_space_def
proof clarify
fix S T
assume "closedin X S" and "closedin X T" and "disjnt S T"
with R obtain f where contf: "continuous_map X euclideanreal f" and "f ` S ⊆ {a}" "f ` T ⊆ {b}"
by meson
show "∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
proof (intro conjI exI)
show "openin X {x ∈ topspace X. f x ∈ ball a (¦a - b¦ / 2)}"
by (force intro!: openin_continuous_map_preimage [OF contf])
show "openin X {x ∈ topspace X. f x ∈ ball b (¦a - b¦ / 2)}"
by (force intro!: openin_continuous_map_preimage [OF contf])
show "S ⊆ {x ∈ topspace X. f x ∈ ball a (¦a - b¦ / 2)}"
using ‹closedin X S› closedin_subset ‹f ` S ⊆ {a}› assms by force
show "T ⊆ {x ∈ topspace X. f x ∈ ball b (¦a - b¦ / 2)}"
using ‹closedin X T› closedin_subset ‹f ` T ⊆ {b}› assms by force
have "⋀x. ⟦x ∈ topspace X; dist a (f x) < ¦a-b¦/2; dist b (f x) < ¦a-b¦/2⟧ ⟹ False"
by (smt (verit, best) dist_real_def dist_triangle_half_l)
then show "disjnt {x ∈ topspace X. f x ∈ ball a (¦a-b¦ / 2)} {x ∈ topspace X. f x ∈ ball b (¦a-b¦ / 2)}"
using disjnt_iff by fastforce
qed
qed
qed
lemma normal_space_iff_Urysohn_gen:
fixes a b::real
shows
"a < b ⟹
normal_space X ⟷
(∀S T. closedin X S ∧ closedin X T ∧ disjnt S T
⟶ (∃f. continuous_map X (top_of_set {a..b}) f ∧
f ` S ⊆ {a} ∧ f ` T ⊆ {b}))"
by (metis linear not_le Urysohn_lemma normal_space_iff_Urysohn_gen_alt continuous_map_in_subtopology)
lemma normal_space_iff_Urysohn_alt:
"normal_space X ⟷
(∀S T. closedin X S ∧ closedin X T ∧ disjnt S T
⟶ (∃f. continuous_map X euclideanreal f ∧
f ` S ⊆ {0} ∧ f ` T ⊆ {1}))"
by (rule normal_space_iff_Urysohn_gen_alt) auto
lemma normal_space_iff_Urysohn:
"normal_space X ⟷
(∀S T. closedin X S ∧ closedin X T ∧ disjnt S T
⟶ (∃f::'a⇒real. continuous_map X (top_of_set {0..1}) f ∧
f ` S ⊆ {0} ∧ f ` T ⊆ {1}))"
by (rule normal_space_iff_Urysohn_gen) auto
lemma normal_space_perfect_map_image:
"⟦normal_space X; perfect_map X Y f⟧ ⟹ normal_space Y"
unfolding perfect_map_def proper_map_def
using normal_space_continuous_closed_map_image by fastforce
lemma Hausdorff_normal_space_closed_continuous_map_image:
"⟦normal_space X; closed_map X Y f; continuous_map X Y f;
f ` topspace X = topspace Y; t1_space Y⟧
⟹ Hausdorff_space Y"
by (metis normal_space_continuous_closed_map_image normal_t1_imp_Hausdorff_space)
lemma normal_Hausdorff_space_closed_continuous_map_image:
"⟦normal_space X; Hausdorff_space X; closed_map X Y f;
continuous_map X Y f; f ` topspace X = topspace Y⟧
⟹ normal_space Y ∧ Hausdorff_space Y"
by (meson normal_space_continuous_closed_map_image normal_t1_eq_Hausdorff_space t1_space_closed_map_image)
lemma Lindelof_cover:
assumes "regular_space X" and "Lindelof_space X" and "S ≠ {}"
and clo: "closedin X S" "closedin X T" "disjnt S T"
obtains h :: "nat ⇒ 'a set" where
"⋀n. openin X (h n)" "⋀n. disjnt T (X closure_of (h n))" and "S ⊆ ⋃(range h)"
proof -
have "∃U. openin X U ∧ x ∈ U ∧ disjnt T (X closure_of U)"
if "x ∈ S" for x
using ‹regular_space X› unfolding regular_space
by (metis (full_types) Diff_iff ‹disjnt S T› clo closure_of_eq disjnt_iff in_closure_of that)
then obtain h where oh: "⋀x. x ∈ S ⟹ openin X (h x)"
and xh: "⋀x. x ∈ S ⟹ x ∈ h x"
and dh: "⋀x. x ∈ S ⟹ disjnt T (X closure_of h x)"
by metis
have "Lindelof_space(subtopology X S)"
by (simp add: Lindelof_space_closedin_subtopology ‹Lindelof_space X› ‹closedin X S›)
then obtain 𝒰 where 𝒰: "countable 𝒰 ∧ 𝒰 ⊆ h ` S ∧ S ⊆ ⋃𝒰"
unfolding Lindelof_space_subtopology_subset [OF closedin_subset [OF ‹closedin X S›]]
by (smt (verit, del_insts) oh xh UN_I image_iff subsetI)
with ‹S ≠ {}› have "𝒰 ≠ {}"
by blast
show ?thesis
proof
show "openin X (from_nat_into 𝒰 n)" for n
by (metis 𝒰 from_nat_into image_iff ‹𝒰 ≠ {}› oh subsetD)
show "disjnt T (X closure_of (from_nat_into 𝒰) n)" for n
using dh from_nat_into [OF ‹𝒰 ≠ {}›]
by (metis 𝒰 f_inv_into_f inv_into_into subset_eq)
show "S ⊆ ⋃(range (from_nat_into 𝒰))"
by (simp add: 𝒰 ‹𝒰 ≠ {}›)
qed
qed
lemma regular_Lindelof_imp_normal_space:
assumes "regular_space X" and "Lindelof_space X"
shows "normal_space X"
unfolding normal_space_def
proof clarify
fix S T
assume clo: "closedin X S" "closedin X T" and "disjnt S T"
show "∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
proof (cases "S={} ∨ T={}")
case True
with clo show ?thesis
by (meson closedin_def disjnt_empty1 disjnt_empty2 openin_empty openin_topspace subset_empty)
next
case False
obtain h :: "nat ⇒ 'a set" where
opeh: "⋀n. openin X (h n)" and dish: "⋀n. disjnt T (X closure_of (h n))"
and Sh: "S ⊆ ⋃(range h)"
by (metis Lindelof_cover False ‹disjnt S T› assms clo)
obtain k :: "nat ⇒ 'a set" where
opek: "⋀n. openin X (k n)" and disk: "⋀n. disjnt S (X closure_of (k n))"
and Tk: "T ⊆ ⋃(range k)"
by (metis Lindelof_cover False ‹disjnt S T› assms clo disjnt_sym)
define U where "U ≡ ⋃i. h i - (⋃j<i. X closure_of k j)"
define V where "V ≡ ⋃i. k i - (⋃j≤i. X closure_of h j)"
show ?thesis
proof (intro exI conjI)
show "openin X U" "openin X V"
unfolding U_def V_def
by (force intro!: opek opeh closedin_Union closedin_closure_of)+
show "S ⊆ U" "T ⊆ V"
using Sh Tk dish disk by (fastforce simp: U_def V_def disjnt_iff)+
have "⋀x i j. ⟦x ∈ k i; x ∈ h j; ∀j≤i. x ∉ X closure_of h j⟧
⟹ ∃i<j. x ∈ X closure_of k i"
by (metis in_closure_of linorder_not_less opek openin_subset subsetD)
then show "disjnt U V"
by (force simp: U_def V_def disjnt_iff)
qed
qed
qed
theorem Tietze_extension_closed_real_interval:
assumes "normal_space X" and "closedin X S"
and contf: "continuous_map (subtopology X S) euclideanreal f"
and fim: "f ` S ⊆ {a..b}" and "a ≤ b"
obtains g
where "continuous_map X euclideanreal g"
"⋀x. x ∈ S ⟹ g x = f x" "g ` topspace X ⊆ {a..b}"
proof -
define c where "c ≡ max ¦a¦ ¦b¦ + 1"
have "0 < c" and c: "⋀x. x ∈ S ⟹ ¦f x¦ ≤ c"
using fim by (auto simp: c_def image_subset_iff)
define good where
"good ≡ λg n. continuous_map X euclideanreal g ∧ (∀x ∈ S. ¦f x - g x¦ ≤ c * (2/3)^n)"
have step: "∃g. good g (Suc n) ∧
(∀x ∈ topspace X. ¦g x - h x¦ ≤ c * (2/3)^n / 3)"
if h: "good h n" for n h
proof -
have pos: "0 < c * (2/3) ^ n"
by (simp add: ‹0 < c›)
have S_eq: "S = topspace(subtopology X S)" and "S ⊆ topspace X"
using ‹closedin X S› closedin_subset by auto
define d where "d ≡ c/3 * (2/3) ^ n"
define SA where "SA ≡ {x ∈ S. f x - h x ∈ {..-d}}"
define SB where "SB ≡ {x ∈ S. f x - h x ∈ {d..}}"
have contfh: "continuous_map (subtopology X S) euclideanreal (λx. f x - h x)"
using that
by (simp add: contf good_def continuous_map_diff continuous_map_from_subtopology)
then have "closedin (subtopology X S) SA"
unfolding SA_def
by (smt (verit, del_insts) closed_closedin continuous_map_closedin Collect_cong S_eq closed_real_atMost)
then have "closedin X SA"
using ‹closedin X S› closedin_trans_full by blast
moreover have "closedin (subtopology X S) SB"
unfolding SB_def
using closedin_continuous_map_preimage_gen [OF contfh]
by (metis (full_types) S_eq closed_atLeast closed_closedin closedin_topspace)
then have "closedin X SB"
using ‹closedin X S› closedin_trans_full by blast
moreover have "disjnt SA SB"
using pos by (auto simp: d_def disjnt_def SA_def SB_def)
moreover have "-d ≤ d"
using pos by (auto simp: d_def)
ultimately
obtain g where contg: "continuous_map X (top_of_set {- d..d}) g"
and ga: "g ` SA ⊆ {- d}" and gb: "g ` SB ⊆ {d}"
using Urysohn_lemma ‹normal_space X› by metis
then have g_le_d: "⋀x. x ∈ topspace X ⟹ ¦g x¦ ≤ d"
by (fastforce simp: abs_le_iff continuous_map_def minus_le_iff)
have g_eq_d: "⋀x. ⟦x ∈ S; f x - h x ≤ -d⟧ ⟹ g x = -d"
using ga by (auto simp: SA_def)
have g_eq_negd: "⋀x. ⟦x ∈ S; f x - h x ≥ d⟧ ⟹ g x = d"
using gb by (auto simp: SB_def)
show ?thesis
unfolding good_def
proof (intro conjI strip exI)
show "continuous_map X euclideanreal (λx. h x + g x)"
using contg continuous_map_add continuous_map_in_subtopology that
unfolding good_def by blast
show "¦f x - (h x + g x)¦ ≤ c * (2 / 3) ^ Suc n" if "x ∈ S" for x
proof -
have x: "x ∈ topspace X"
using ‹S ⊆ topspace X› that by auto
have "¦f x - h x¦ ≤ c * (2/3) ^ n"
using good_def h that by blast
with g_eq_d [OF that] g_eq_negd [OF that] g_le_d [OF x]
have "¦f x - (h x + g x)¦ ≤ d + d"
unfolding d_def by linarith
then show ?thesis
by (simp add: d_def)
qed
show "¦h x + g x - h x¦ ≤ c * (2 / 3) ^ n / 3" if "x ∈ topspace X" for x
using that d_def g_le_d by auto
qed
qed
then obtain nxtg where nxtg: "⋀h n. good h n ⟹
good (nxtg h n) (Suc n) ∧ (∀x ∈ topspace X. ¦nxtg h n x - h x¦ ≤ c * (2/3)^n / 3)"
by metis
define g where "g ≡ rec_nat (λx. 0) (λn r. nxtg r n)"
have [simp]: "g 0 x = 0" for x
by (auto simp: g_def)
have g_Suc: "g(Suc n) = nxtg (g n) n" for n
by (auto simp: g_def)
have good: "good (g n) n" for n
proof (induction n)
case 0
with c show ?case
by (auto simp: good_def)
qed (simp add: g_Suc nxtg)
have *: "⋀n x. x ∈ topspace X ⟹ ¦g(Suc n) x - g n x¦ ≤ c * (2/3) ^ n / 3"
using nxtg g_Suc good by force
obtain h where conth: "continuous_map X euclideanreal h"
and h: "⋀ε. 0 < ε ⟹ ∀⇩F n in sequentially. ∀x∈topspace X. dist (g n x) (h x) < ε"
proof (rule Met_TC.continuous_map_uniformly_Cauchy_limit)
show "∀⇩F n in sequentially. continuous_map X (Met_TC.mtopology) (g n)"
using good good_def by fastforce
show "∃N. ∀m n x. N ≤ m ⟶ N ≤ n ⟶ x ∈ topspace X ⟶ dist (g m x) (g n x) < ε"
if "ε > 0" for ε
proof -
have "∀⇩F n in sequentially. ¦(2/3) ^ n¦ < ε/c"
proof (rule Archimedean_eventually_pow_inverse)
show "0 < ε / c"
by (simp add: ‹0 < c› that)
qed auto
then obtain N where N: "⋀n. n ≥ N ⟹ ¦(2/3) ^ n¦ < ε/c"
by (meson eventually_sequentially order_le_less_trans)
have "¦g m x - g n x¦ < ε"
if "N ≤ m" "N ≤ n" and x: "x ∈ topspace X" "m ≤ n" for m n x
proof (cases "m < n")
case True
have 23: "(∑k = m..<n. (2/3)^k) = 3 * ((2/3) ^ m - (2/3::real) ^ n)"
using ‹m ≤ n›
by (induction n) (auto simp: le_Suc_eq)
have "¦g m x - g n x¦ ≤ ¦∑k = m..<n. g (Suc k) x - g k x¦"
by (subst sum_Suc_diff' [OF ‹m ≤ n›]) linarith
also have "… ≤ (∑k = m..<n. ¦g (Suc k) x - g k x¦)"
by (rule sum_abs)
also have "… ≤ (∑k = m..<n. c * (2/3)^k / 3)"
by (meson "*" sum_mono x(1))
also have "… = (c/3) * (∑k = m..<n. (2/3)^k)"
by (simp add: sum_distrib_left)
also have "… = (c/3) * 3 * ((2/3) ^ m - (2/3) ^ n)"
by (simp add: sum_distrib_left 23)
also have "... < (c/3) * 3 * ((2/3) ^ m)"
using ‹0 < c› by auto
also have "… < ε"
using N [OF ‹N ≤ m›] ‹0 < c› by (simp add: field_simps)
finally show ?thesis .
qed (use ‹0 < ε› ‹m ≤ n› in auto)
then show ?thesis
by (metis dist_commute_lessI dist_real_def nle_le)
qed
qed auto
define φ where "φ ≡ λx. max a (min (h x) b)"
show thesis
proof
show "continuous_map X euclidean φ"
unfolding φ_def using conth by (intro continuous_intros) auto
show "φ x = f x" if "x ∈ S" for x
proof -
have x: "x ∈ topspace X"
using ‹closedin X S› closedin_subset that by blast
have "h x = f x"
proof (rule Met_TC.limitin_metric_unique)
show "limitin Met_TC.mtopology (λn. g n x) (h x) sequentially"
using h x by (force simp: tendsto_iff eventually_sequentially)
show "limitin Met_TC.mtopology (λn. g n x) (f x) sequentially"
proof (clarsimp simp: tendsto_iff)
fix ε::real
assume "ε > 0"
then have "∀⇩F n in sequentially. ¦(2/3) ^ n¦ < ε/c"
by (intro Archimedean_eventually_pow_inverse) (auto simp: ‹c > 0›)
then show "∀⇩F n in sequentially. dist (g n x) (f x) < ε"
apply eventually_elim
by (smt (verit) good x good_def ‹c > 0› dist_real_def mult.commute pos_less_divide_eq that)
qed
qed auto
then show ?thesis
using that fim by (auto simp: φ_def)
qed
then show "φ ` topspace X ⊆ {a..b}"
using fim ‹a ≤ b› by (auto simp: φ_def)
qed
qed
theorem Tietze_extension_realinterval:
assumes XS: "normal_space X" "closedin X S" and T: "is_interval T" "T ≠ {}"
and contf: "continuous_map (subtopology X S) euclideanreal f"
and "f ` S ⊆ T"
obtains g where "continuous_map X euclideanreal g" "g ` topspace X ⊆ T" "⋀x. x ∈ S ⟹ g x = f x"
proof -
define Φ where
"Φ ≡ λT::real set. ∀f. continuous_map (subtopology X S) euclidean f ⟶ f`S ⊆ T
⟶ (∃g. continuous_map X euclidean g ∧ g ` topspace X ⊆ T ∧ (∀x ∈ S. g x = f x))"
have "Φ T"
if *: "⋀T. ⟦bounded T; is_interval T; T ≠ {}⟧ ⟹ Φ T"
and "is_interval T" "T ≠ {}" for T
unfolding Φ_def
proof (intro strip)
fix f
assume contf: "continuous_map (subtopology X S) euclideanreal f"
and "f ` S ⊆ T"
have ΦT: "Φ ((λx. x / (1 + ¦x¦)) ` T)"
proof (rule *)
show "bounded ((λx. x / (1 + ¦x¦)) ` T)"
using shrink_range [of T] by (force intro: boundedI [where B=1])
show "is_interval ((λx. x / (1 + ¦x¦)) ` T)"
using connected_shrink that(2) is_interval_connected_1 by blast
show "(λx. x / (1 + ¦x¦)) ` T ≠ {}"
using ‹T ≠ {}› by auto
qed
moreover have "continuous_map (subtopology X S) euclidean ((λx. x / (1 + ¦x¦)) ∘ f)"
by (metis contf continuous_map_compose continuous_map_into_fulltopology continuous_map_real_shrink)
moreover have "((λx. x / (1 + ¦x¦)) ∘ f) ` S ⊆ (λx. x / (1 + ¦x¦)) ` T"
using ‹f ` S ⊆ T› by auto
ultimately obtain g
where contg: "continuous_map X euclidean g"
and gim: "g ` topspace X ⊆ (λx. x / (1 + ¦x¦)) ` T"
and geq: "⋀x. x ∈ S ⟹ g x = ((λx. x / (1 + ¦x¦)) ∘ f) x"
using ΦT unfolding Φ_def by force
show "∃g. continuous_map X euclideanreal g ∧ g ` topspace X ⊆ T ∧ (∀x∈S. g x = f x)"
proof (intro conjI exI)
have "continuous_map X (top_of_set {-1<..<1}) g"
using contg continuous_map_in_subtopology gim shrink_range by blast
then show "continuous_map X euclideanreal ((λx. x / (1 - ¦x¦)) ∘ g)"
by (rule continuous_map_compose) (auto simp: continuous_on_real_grow)
show "((λx. x / (1 - ¦x¦)) ∘ g) ` topspace X ⊆ T"
using gim real_grow_shrink by fastforce
show "∀x∈S. ((λx. x / (1 - ¦x¦)) ∘ g) x = f x"
using geq real_grow_shrink by force
qed
qed
moreover have "Φ T"
if "bounded T" "is_interval T" "T ≠ {}" for T
unfolding Φ_def
proof (intro strip)
fix f
assume contf: "continuous_map (subtopology X S) euclideanreal f"
and "f ` S ⊆ T"
obtain a b where ab: "closure T = {a..b}"
by (meson ‹bounded T› ‹is_interval T› compact_closure connected_compact_interval_1
connected_imp_connected_closure is_interval_connected)
with ‹T ≠ {}› have "a ≤ b" by auto
have "f ` S ⊆ {a..b}"
using ‹f ` S ⊆ T› ab closure_subset by auto
then obtain g where contg: "continuous_map X euclideanreal g"
and gf: "⋀x. x ∈ S ⟹ g x = f x" and gim: "g ` topspace X ⊆ {a..b}"
using Tietze_extension_closed_real_interval [OF XS contf _ ‹a ≤ b›] by metis
define W where "W ≡ {x ∈ topspace X. g x ∈ closure T - T}"
have "{a..b} - {a, b} ⊆ T"
using that
by (metis ab atLeastAtMost_diff_ends convex_interior_closure interior_atLeastAtMost_real
interior_subset is_interval_convex)
with finite_imp_compact have "compact (closure T - T)"
by (metis Diff_eq_empty_iff Diff_insert2 ab finite.emptyI finite_Diff_insert)
then have "closedin X W"
unfolding W_def using closedin_continuous_map_preimage [OF contg] compact_imp_closed by force
moreover have "disjnt W S"
unfolding W_def disjnt_iff using ‹f ` S ⊆ T› gf by blast
ultimately obtain h :: "'a ⇒ real"
where conth: "continuous_map X (top_of_set {0..1}) h"
and him: "h ` W ⊆ {0}" "h ` S ⊆ {1}"
by (metis XS normal_space_iff_Urysohn)
then have him01: "h ` topspace X ⊆ {0..1}"
by (meson continuous_map_in_subtopology)
obtain z where "z ∈ T"
using ‹T ≠ {}› by blast
define g' where "g' ≡ λx. z + h x * (g x - z)"
show "∃g. continuous_map X euclidean g ∧ g ` topspace X ⊆ T ∧ (∀x∈S. g x = f x)"
proof (intro exI conjI)
show "continuous_map X euclideanreal g'"
unfolding g'_def using contg conth continuous_map_in_subtopology
by (intro continuous_intros) auto
show "g' ` topspace X ⊆ T"
unfolding g'_def
proof clarify
fix x
assume "x ∈ topspace X"
show "z + h x * (g x - z) ∈ T"
proof (cases "g x ∈ T")
case True
define w where "w ≡ z + h x * (g x - z)"
have "¦h x¦ * ¦g x - z¦ ≤ ¦g x - z¦" "¦1 - h x¦ * ¦g x - z¦ ≤ ¦g x - z¦"
using him01 ‹x ∈ topspace X› by (force simp: intro: mult_left_le_one_le)+
then consider "z ≤ w ∧ w ≤ g x" | "g x ≤ w ∧ w ≤ z"
unfolding w_def by (smt (verit) left_diff_distrib mult_cancel_right2 mult_minus_right zero_less_mult_iff)
then show ?thesis
using ‹is_interval T› unfolding w_def is_interval_1 by (metis True ‹z ∈ T›)
next
case False
then have "g x ∈ closure T"
using ‹x ∈ topspace X› ab gim by blast
then have "h x = 0"
using him False ‹x ∈ topspace X› by (auto simp: W_def image_subset_iff)
then show ?thesis
by (simp add: ‹z ∈ T›)
qed
qed
show "∀x∈S. g' x = f x"
using gf him by (auto simp: W_def g'_def)
qed
qed
ultimately show thesis
using assms that unfolding Φ_def by best
qed
lemma normal_space_iff_Tietze:
"normal_space X ⟷
(∀f S. closedin X S ∧
continuous_map (subtopology X S) euclidean f
⟶ (∃g:: 'a ⇒ real. continuous_map X euclidean g ∧ (∀x ∈ S. g x = f x)))"
(is "?lhs ⟷ ?rhs")
proof
assume ?lhs
then show ?rhs
by (metis Tietze_extension_realinterval empty_not_UNIV is_interval_univ subset_UNIV)
next
assume R: ?rhs
show ?lhs
unfolding normal_space_iff_Urysohn_alt
proof clarify
fix S T
assume "closedin X S"
and "closedin X T"
and "disjnt S T"
then have cloST: "closedin X (S ∪ T)"
by (simp add: closedin_Un)
moreover
have "continuous_map (subtopology X (S ∪ T)) euclideanreal (λx. if x ∈ S then 0 else 1)"
unfolding continuous_map_closedin
proof (intro conjI strip)
fix C :: "real set"
define D where "D ≡ {x ∈ topspace X. if x ∈ S then 0 ∈ C else x ∈ T ∧ 1 ∈ C}"
have "D ∈ {{}, S, T, S ∪ T}"
unfolding D_def
using closedin_subset [OF ‹closedin X S›] closedin_subset [OF ‹closedin X T›] ‹disjnt S T›
by (auto simp: disjnt_iff)
then have "closedin X D"
using ‹closedin X S› ‹closedin X T› closedin_empty by blast
with closedin_subset_topspace
show "closedin (subtopology X (S ∪ T)) {x ∈ topspace (subtopology X (S ∪ T)). (if x ∈ S then 0::real else 1) ∈ C}"
apply (simp add: D_def)
by (smt (verit, best) Collect_cong Collect_mono_iff Un_def closedin_subset_topspace)
qed auto
ultimately obtain g :: "'a ⇒ real" where
contg: "continuous_map X euclidean g" and gf: "∀x ∈ S ∪ T. g x = (if x ∈ S then 0 else 1)"
using R by blast
then show "∃f. continuous_map X euclideanreal f ∧ f ` S ⊆ {0} ∧ f ` T ⊆ {1}"
by (smt (verit) Un_iff ‹disjnt S T› disjnt_iff image_subset_iff insert_iff)
qed
qed
subsection ‹Random metric space stuff›
lemma metrizable_imp_k_space:
assumes "metrizable_space X"
shows "k_space X"
proof -
obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d"
using assms unfolding metrizable_space_def by metis
then interpret Metric_space M d
by blast
show ?thesis
unfolding k_space Xeq
proof clarsimp
fix S
assume "S ⊆ M" and S: "∀K. compactin mtopology K ⟶ closedin (subtopology mtopology K) (K ∩ S)"
have "l ∈ S"
if σ: "range σ ⊆ S" and l: "limitin mtopology σ l sequentially" for σ l
proof -
define K where "K ≡ insert l (range σ)"
interpret Submetric M d K
proof
show "K ⊆ M"
unfolding K_def using l limitin_mspace ‹S ⊆ M› σ by blast
qed
have "compactin mtopology K"
unfolding K_def using ‹S ⊆ M› σ
by (force intro: compactin_sequence_with_limit [OF l])
then have *: "closedin sub.mtopology (K ∩ S)"
by (simp add: S mtopology_submetric)
have "σ n ∈ K ∩ S" for n
by (simp add: K_def range_subsetD σ)
moreover have "limitin sub.mtopology σ l sequentially"
using l
unfolding sub.limit_metric_sequentially limit_metric_sequentially
by (force simp: K_def)
ultimately have "l ∈ K ∩ S"
by (meson * σ image_subsetI sub.metric_closedin_iff_sequentially_closed)
then show ?thesis
by simp
qed
then show "closedin mtopology S"
unfolding metric_closedin_iff_sequentially_closed
using ‹S ⊆ M› by blast
qed
qed
lemma (in Metric_space) k_space_mtopology: "k_space mtopology"
by (simp add: metrizable_imp_k_space metrizable_space_mtopology)
lemma k_space_euclideanreal: "k_space (euclidean :: 'a::metric_space topology)"
using metrizable_imp_k_space metrizable_space_euclidean by auto
subsection‹Hereditarily normal spaces›
lemma hereditarily_B:
assumes "⋀S T. separatedin X S T
⟹ ∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
shows "hereditarily normal_space X"
unfolding hereditarily_def
proof (intro strip)
fix W
assume "W ⊆ topspace X"
show "normal_space (subtopology X W)"
unfolding normal_space_def
proof clarify
fix S T
assume clo: "closedin (subtopology X W) S" "closedin (subtopology X W) T"
and "disjnt S T"
then have "separatedin (subtopology X W) S T"
by (simp add: separatedin_closed_sets)
then obtain U V where "openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
using assms [of S T] by (meson separatedin_subtopology)
then show "∃U V. openin (subtopology X W) U ∧ openin (subtopology X W) V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
apply (simp add: openin_subtopology_alt)
by (meson clo closedin_imp_subset disjnt_subset1 disjnt_subset2 inf_le2)
qed
qed
lemma hereditarily_C:
assumes "separatedin X S T" and norm: "⋀U. openin X U ⟹ normal_space (subtopology X U)"
shows "∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
proof -
define ST where "ST ≡ X closure_of S ∩ X closure_of T"
have subX: "S ⊆ topspace X" "T ⊆ topspace X"
by (meson ‹separatedin X S T› separation_closedin_Un_gen)+
have sub: "S ⊆ topspace X - ST" "T ⊆ topspace X - ST"
unfolding ST_def
by (metis Diff_mono Diff_triv ‹separatedin X S T› Int_lower1 Int_lower2 separatedin_def)+
have "normal_space (subtopology X (topspace X - ST))"
by (simp add: ST_def norm closedin_Int openin_diff)
moreover have " disjnt (subtopology X (topspace X - ST) closure_of S)
(subtopology X (topspace X - ST) closure_of T)"
using Int_absorb1 ST_def sub by (fastforce simp: disjnt_iff closure_of_subtopology)
ultimately show ?thesis
using sub subX
apply (simp add: normal_space_closures)
by (metis ST_def closedin_Int closedin_closure_of closedin_def openin_trans_full)
qed
lemma hereditarily_normal_space:
"hereditarily normal_space X ⟷ (∀U. openin X U ⟶ normal_space(subtopology X U))"
by (metis hereditarily_B hereditarily_C hereditarily)
lemma hereditarily_normal_separation:
"hereditarily normal_space X ⟷
(∀S T. separatedin X S T
⟶ (∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V))"
by (metis hereditarily_B hereditarily_C hereditarily)
lemma metrizable_imp_hereditarily_normal_space:
"metrizable_space X ⟹ hereditarily normal_space X"
by (simp add: hereditarily metrizable_imp_normal_space metrizable_space_subtopology)
lemma metrizable_space_separation:
"⟦metrizable_space X; separatedin X S T⟧
⟹ ∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
by (metis hereditarily hereditarily_C metrizable_imp_hereditarily_normal_space)
lemma hereditarily_normal_separation_pairwise:
"hereditarily normal_space X ⟷
(∀𝒰. finite 𝒰 ∧ (∀S ∈ 𝒰. S ⊆ topspace X) ∧ pairwise (separatedin X) 𝒰
⟶ (∃ℱ. (∀S ∈ 𝒰. openin X (ℱ S) ∧ S ⊆ ℱ S) ∧
pairwise (λS T. disjnt (ℱ S) (ℱ T)) 𝒰))"
(is "?lhs ⟷ ?rhs")
proof
assume L: ?lhs
show ?rhs
proof clarify
fix 𝒰
assume "finite 𝒰" and 𝒰: "∀S∈𝒰. S ⊆ topspace X"
and pw𝒰: "pairwise (separatedin X) 𝒰"
have "∃V W. openin X V ∧ openin X W ∧ S ⊆ V ∧
(∀T. T ∈ 𝒰 ∧ T ≠ S ⟶ T ⊆ W) ∧ disjnt V W"
if "S ∈ 𝒰" for S
proof -
have "separatedin X S (⋃(𝒰 - {S}))"
by (metis 𝒰 ‹finite 𝒰› pw𝒰 finite_Diff pairwise_alt separatedin_Union(1) that)
with L show ?thesis
unfolding hereditarily_normal_separation
by (smt (verit) Diff_iff UnionI empty_iff insert_iff subset_iff)
qed
then obtain ℱ 𝒢
where *: "⋀S. S ∈ 𝒰 ⟹ S ⊆ ℱ S ∧ (∀T. T ∈ 𝒰 ∧ T ≠ S ⟶ T ⊆ 𝒢 S)"
and ope: "⋀S. S ∈ 𝒰 ⟹ openin X (ℱ S) ∧ openin X (𝒢 S)"
and dis: "⋀S. S ∈ 𝒰 ⟹ disjnt (ℱ S) (𝒢 S)"
by metis
define ℋ where "ℋ ≡ λS. ℱ S ∩ (⋂T ∈ 𝒰 - {S}. 𝒢 T)"
show "∃ℱ. (∀S∈𝒰. openin X (ℱ S) ∧ S ⊆ ℱ S) ∧ pairwise (λS T. disjnt (ℱ S) (ℱ T)) 𝒰"
proof (intro exI conjI strip)
show "openin X (ℋ S)" if "S ∈ 𝒰" for S
unfolding ℋ_def
by (smt (verit) ope that DiffD1 ‹finite 𝒰› finite_Diff finite_imageI imageE openin_Int_Inter)
show "S ⊆ ℋ S" if "S ∈ 𝒰" for S
unfolding ℋ_def using "*" that by auto
show "pairwise (λS T. disjnt (ℋ S) (ℋ T)) 𝒰"
using dis by (fastforce simp: disjnt_iff pairwise_alt ℋ_def)
qed
qed
next
assume R: ?rhs
show ?lhs
unfolding hereditarily_normal_separation
proof (intro strip)
fix S T
assume "separatedin X S T"
show "∃U V. openin X U ∧ openin X V ∧ S ⊆ U ∧ T ⊆ V ∧ disjnt U V"
proof (cases "T=S")
case True
then show ?thesis
using ‹separatedin X S T› by force
next
case False
have "pairwise (separatedin X) {S, T}"
by (simp add: ‹separatedin X S T› pairwise_insert separatedin_sym)
moreover have "∀S∈{S, T}. S ⊆ topspace X"
by (metis ‹separatedin X S T› insertE separatedin_def singletonD)
ultimately show ?thesis
using R by (smt (verit) False finite.emptyI finite.insertI insertCI pairwiseD)
qed
qed
qed
lemma hereditarily_normal_space_perfect_map_image:
"⟦hereditarily normal_space X; perfect_map X Y f⟧ ⟹ hereditarily normal_space Y"
unfolding perfect_map_def proper_map_def
by (meson hereditarily_normal_space_continuous_closed_map_image)
lemma regular_second_countable_imp_hereditarily_normal_space:
assumes "regular_space X ∧ second_countable X"
shows "hereditarily normal_space X"
unfolding hereditarily
proof (intro regular_Lindelof_imp_normal_space strip)
show "regular_space (subtopology X S)" for S
by (simp add: assms regular_space_subtopology)
show "Lindelof_space (subtopology X S)" for S
using assms by (simp add: second_countable_imp_Lindelof_space second_countable_subtopology)
qed
subsection‹Completely regular spaces›
definition completely_regular_space where
"completely_regular_space X ≡
∀S x. closedin X S ∧ x ∈ topspace X - S
⟶ (∃f::'a⇒real. continuous_map X (top_of_set {0..1}) f ∧
f x = 0 ∧ (f ` S ⊆ {1}))"
lemma homeomorphic_completely_regular_space_aux:
assumes X: "completely_regular_space X" and hom: "X homeomorphic_space Y"
shows "completely_regular_space Y"
proof -
obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g"
and fg: "(∀x ∈ topspace X. g(f x) = x) ∧ (∀y ∈ topspace Y. f(g y) = y)"
using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce
show ?thesis
unfolding completely_regular_space_def
proof clarify
fix S x
assume A: "closedin Y S" and x: "x ∈ topspace Y" and "x ∉ S"
then have "closedin X (g`S)"
using hmg homeomorphic_map_closedness_eq by blast
moreover have "g x ∉ g`S"
by (meson A x ‹x ∉ S› closedin_subset hmg homeomorphic_imp_injective_map inj_on_image_mem_iff)
ultimately obtain φ where φ: "continuous_map X (top_of_set {0..1::real}) φ ∧ φ (g x) = 0 ∧ φ ` g`S ⊆ {1}"
by (metis DiffI X completely_regular_space_def hmg homeomorphic_imp_surjective_map image_eqI x)
then have "continuous_map Y (top_of_set {0..1::real}) (φ ∘ g)"
by (meson continuous_map_compose hmg homeomorphic_imp_continuous_map)
then show "∃ψ. continuous_map Y (top_of_set {0..1::real}) ψ ∧ ψ x = 0 ∧ ψ ` S ⊆ {1}"
by (metis φ comp_apply image_comp)
qed
qed
lemma homeomorphic_completely_regular_space:
assumes "X homeomorphic_space Y"
shows "completely_regular_space X ⟷ completely_regular_space Y"
by (meson assms homeomorphic_completely_regular_space_aux homeomorphic_space_sym)
lemma completely_regular_space_alt:
"completely_regular_space X ⟷
(∀S x. closedin X S ⟶ x ∈ topspace X - S
⟶ (∃f. continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` S ⊆ {1}))"
proof -
have "∃f. continuous_map X (top_of_set {0..1::real}) f ∧ f x = 0 ∧ f ` S ⊆ {1}"
if "closedin X S" "x ∈ topspace X - S" and f: "continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` S ⊆ {1}"
for S x f
proof (intro exI conjI)
show "continuous_map X (top_of_set {0..1}) (λx. max 0 (min (f x) 1))"
using that
by (auto simp: continuous_map_in_subtopology intro!: continuous_map_real_max continuous_map_real_min)
qed (use that in auto)
with continuous_map_in_subtopology show ?thesis
unfolding completely_regular_space_def by metis
qed
text ‹As above, but with @{term openin}›
lemma completely_regular_space_alt':
"completely_regular_space X ⟷
(∀S x. openin X S ⟶ x ∈ S
⟶ (∃f. continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` (topspace X - S) ⊆ {1}))"
apply (simp add: completely_regular_space_alt all_closedin)
by (meson openin_subset subsetD)
lemma completely_regular_space_gen_alt:
fixes a b::real
assumes "a ≠ b"
shows "completely_regular_space X ⟷
(∀S x. closedin X S ⟶ x ∈ topspace X - S
⟶ (∃f. continuous_map X euclidean f ∧ f x = a ∧ (f ` S ⊆ {b})))"
proof -
have "∃f. continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` S ⊆ {1}"
if "closedin X S" "x ∈ topspace X - S"
and f: "continuous_map X euclidean f ∧ f x = a ∧ f ` S ⊆ {b}"
for S x f
proof (intro exI conjI)
show "continuous_map X euclideanreal ((λx. inverse(b - a) * (x - a)) ∘ f)"
using that by (intro continuous_intros) auto
qed (use that assms in auto)
moreover
have "∃f. continuous_map X euclidean f ∧ f x = a ∧ f ` S ⊆ {b}"
if "closedin X S" "x ∈ topspace X - S"
and f: "continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` S ⊆ {1}"
for S x f
proof (intro exI conjI)
show "continuous_map X euclideanreal ((λx. a + (b - a) * x) ∘ f)"
using that by (intro continuous_intros) auto
qed (use that in auto)
ultimately show ?thesis
unfolding completely_regular_space_alt by meson
qed
text ‹As above, but with @{term openin}›
lemma completely_regular_space_gen_alt':
fixes a b::real
assumes "a ≠ b"
shows "completely_regular_space X ⟷
(∀S x. openin X S ⟶ x ∈ S
⟶ (∃f. continuous_map X euclidean f ∧ f x = a ∧ (f ` (topspace X - S) ⊆ {b})))"
apply (simp add: completely_regular_space_gen_alt[OF assms] all_closedin)
by (meson openin_subset subsetD)
lemma completely_regular_space_gen:
fixes a b::real
assumes "a < b"
shows "completely_regular_space X ⟷
(∀S x. closedin X S ∧ x ∈ topspace X - S
⟶ (∃f. continuous_map X (top_of_set {a..b}) f ∧ f x = a ∧ f ` S ⊆ {b}))"
proof -
have "∃f. continuous_map X (top_of_set {a..b}) f ∧ f x = a ∧ f ` S ⊆ {b}"
if "closedin X S" "x ∈ topspace X - S"
and f: "continuous_map X euclidean f ∧ f x = a ∧ f ` S ⊆ {b}"
for S x f
proof (intro exI conjI)
show "continuous_map X (top_of_set {a..b}) (λx. max a (min (f x) b))"
using that assms
by (auto simp: continuous_map_in_subtopology intro!: continuous_map_real_max continuous_map_real_min)
qed (use that assms in auto)
with continuous_map_in_subtopology assms show ?thesis
using completely_regular_space_gen_alt [of a b]
by (smt (verit) atLeastAtMost_singleton atLeastatMost_empty singletonI)
qed
lemma normal_imp_completely_regular_space_A:
assumes "normal_space X" "t1_space X"
shows "completely_regular_space X"
unfolding completely_regular_space_alt
proof clarify
fix x S
assume A: "closedin X S" "x ∉ S"
assume "x ∈ topspace X"
then have "closedin X {x}"
by (simp add: ‹t1_space X› closedin_t1_singleton)
with A ‹normal_space X› have "∃f. continuous_map X euclideanreal f ∧ f ` {x} ⊆ {0} ∧ f ` S ⊆ {1}"
using assms unfolding normal_space_iff_Urysohn_alt disjnt_iff by blast
then show "∃f. continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` S ⊆ {1}"
by auto
qed
lemma normal_imp_completely_regular_space_B:
assumes "normal_space X" "regular_space X"
shows "completely_regular_space X"
unfolding completely_regular_space_alt
proof clarify
fix x S
assume "closedin X S" "x ∉ S" "x ∈ topspace X"
then obtain U C where "openin X U" "closedin X C" "x ∈ U" "U ⊆ C" "C ⊆ topspace X - S"
using assms
unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of closedin_def by (metis Diff_iff)
then obtain f where "continuous_map X euclideanreal f ∧ f ` C ⊆ {0} ∧ f ` S ⊆ {1}"
using assms unfolding normal_space_iff_Urysohn_alt
by (metis Diff_iff ‹closedin X S› disjnt_iff subsetD)
then show "∃f. continuous_map X euclideanreal f ∧ f x = 0 ∧ f ` S ⊆ {1}"
by (meson ‹U ⊆ C› ‹x ∈ U› image_subset_iff singletonD subsetD)
qed
lemma normal_imp_completely_regular_space_gen:
"⟦normal_space X; t1_space X ∨ Hausdorff_space X ∨ regular_space X⟧ ⟹ completely_regular_space X"
using normal_imp_completely_regular_space_A normal_imp_completely_regular_space_B t1_or_Hausdorff_space by blast
lemma normal_imp_completely_regular_space:
"⟦normal_space X; Hausdorff_space X ∨ regular_space X⟧ ⟹ completely_regular_space X"
by (simp add: normal_imp_completely_regular_space_gen)
lemma (in Metric_space) completely_regular_space_mtopology:
"completely_regular_space mtopology"
by (simp add: normal_imp_completely_regular_space normal_space_mtopology regular_space_mtopology)
lemma metrizable_imp_completely_regular_space:
"metrizable_space X ⟹ completely_regular_space X"
by (simp add: metrizable_imp_normal_space metrizable_imp_regular_space normal_imp_completely_regular_space)
lemma completely_regular_space_discrete_topology:
"completely_regular_space(discrete_topology U)"
by (simp add: normal_imp_completely_regular_space normal_space_discrete_topology)
lemma completely_regular_space_subtopology:
assumes "completely_regular_space X"
shows "completely_regular_space (subtopology X S)"
unfolding completely_regular_space_def
proof clarify
fix A x
assume "closedin (subtopology X S) A" and x: "x ∈ topspace (subtopology X S)" and "x ∉ A"
then obtain T where "closedin X T" "A = S ∩ T" "x ∈ topspace X" "x ∈ S"
by (force simp: closedin_subtopology_alt image_iff)
then show "∃f. continuous_map (subtopology X S) (top_of_set {0::real..1}) f ∧ f x = 0 ∧ f ` A ⊆ {1}"
using assms ‹x ∉ A›
apply (simp add: completely_regular_space_def continuous_map_from_subtopology)
using continuous_map_from_subtopology by fastforce
qed
lemma completely_regular_space_retraction_map_image:
" ⟦retraction_map X Y r; completely_regular_space X⟧ ⟹ completely_regular_space Y"
using completely_regular_space_subtopology hereditary_imp_retractive_property homeomorphic_completely_regular_space by blast
lemma completely_regular_imp_regular_space:
assumes "completely_regular_space X"
shows "regular_space X"
proof -
have *: "∃U V. openin X U ∧ openin X V ∧ a ∈ U ∧ C ⊆ V ∧ disjnt U V"
if contf: "continuous_map X euclideanreal f" and a: "a ∈ topspace X - C" and "closedin X C"
and fim: "f ` topspace X ⊆ {0..1}" and f0: "f a = 0" and f1: "f ` C ⊆ {1}"
for C a f
proof (intro exI conjI)
show "openin X {x ∈ topspace X. f x ∈ {..<1 / 2}}" "openin X {x ∈ topspace X. f x ∈ {1 / 2<..}}"
using openin_continuous_map_preimage [OF contf]
by (meson open_lessThan open_greaterThan open_openin)+
show "a ∈ {x ∈ topspace X. f x ∈ {..<1 / 2}}"
using a f0 by auto
show "C ⊆ {x ∈ topspace X. f x ∈ {1 / 2<..}}"
using ‹closedin X C› f1 closedin_subset by auto
qed (auto simp: disjnt_iff)
show ?thesis
using assms
unfolding completely_regular_space_def regular_space_def continuous_map_in_subtopology
by (meson "*")
qed
proposition locally_compact_regular_imp_completely_regular_space:
assumes "locally_compact_space X" "Hausdorff_space X ∨ regular_space X"
shows "completely_regular_space X"
unfolding completely_regular_space_def
proof clarify
fix S x
assume "closedin X S" and "x ∈ topspace X" and "x ∉ S"
have "regular_space X"
using assms locally_compact_Hausdorff_imp_regular_space by blast
then have nbase: "neighbourhood_base_of (λC. compactin X C ∧ closedin X C) X"
using assms(1) locally_compact_regular_space_neighbourhood_base by blast
then obtain U M where "openin X U" "compactin X M" "closedin X M" "x ∈ U" "U ⊆ M" "M ⊆ topspace X - S"
unfolding neighbourhood_base_of by (metis (no_types, lifting) Diff_iff ‹closedin X S› ‹x ∈ topspace X› ‹x ∉ S› closedin_def)
then have "M ⊆ topspace X"
by blast
obtain V K where "openin X V" "closedin X K" "x ∈ V" "V ⊆ K" "K ⊆ U"
by (metis (no_types, lifting) ‹openin X U› ‹x ∈ U› neighbourhood_base_of nbase)
have "compact_space (subtopology X M)"
by (simp add: ‹compactin X M› compact_space_subtopology)
then have "normal_space (subtopology X M)"
by (simp add: ‹regular_space X› compact_Hausdorff_or_regular_imp_normal_space regular_space_subtopology)
moreover have "closedin (subtopology X M) K"
using ‹K ⊆ U› ‹U ⊆ M› ‹closedin X K› closedin_subset_topspace by fastforce
moreover have "closedin (subtopology X M) (M - U)"
by (simp add: ‹closedin X M› ‹openin X U› closedin_diff closedin_subset_topspace)
moreover have "disjnt K (M - U)"
by (meson DiffD2 ‹K ⊆ U› disjnt_iff subsetD)
ultimately obtain f::"'a⇒real" where contf: "continuous_map (subtopology X M) (top_of_set {0..1}) f"
and f0: "f ` K ⊆ {0}" and f1: "f ` (M - U) ⊆ {1}"
using Urysohn_lemma [of "subtopology X M" K "M-U" 0 1] by auto
then obtain g::"'a⇒real" where contg: "continuous_map (subtopology X M) euclidean g" and gim: "g ` M ⊆ {0..1}"
and g0: "⋀x. x ∈ K ⟹ g x = 0" and g1: "⋀x. ⟦x ∈ M; x ∉ U⟧ ⟹ g x = 1"
using ‹M ⊆ topspace X› by (force simp: continuous_map_in_subtopology image_subset_iff)
show "∃f::'a⇒real. continuous_map X (top_of_set {0..1}) f ∧ f x = 0 ∧ f ` S ⊆ {1}"
proof (intro exI conjI)
show "continuous_map X (top_of_set {0..1}) (λx. if x ∈ M then g x else 1)"
unfolding continuous_map_closedin
proof (intro strip conjI)
fix C
assume C: "closedin (top_of_set {0::real..1}) C"
have eq: "{x ∈ topspace X. (if x ∈ M then g x else 1) ∈ C} = {x ∈ M. g x ∈ C} ∪ (if 1 ∈ C then topspace X - U else {})"
using ‹U ⊆ M› ‹M ⊆ topspace X› g1 by auto
show "closedin X {x ∈ topspace X. (if x ∈ M then g x else 1) ∈ C}"
unfolding eq
proof (intro closedin_Un)
have "closedin euclidean C"
using C closed_closedin closedin_closed_trans by blast
then have "closedin (subtopology X M) {x ∈ M. g x ∈ C}"
using closedin_continuous_map_preimage_gen [OF contg] ‹M ⊆ topspace X› by auto
then show "closedin X {x ∈ M. g x ∈ C}"
using ‹closedin X M› closedin_trans_full by blast
qed (use ‹openin X U› in force)
qed (use gim in force)
show "(if x ∈ M then g x else 1) = 0"
using ‹U ⊆ M› ‹V ⊆ K› g0 ‹x ∈ U› ‹x ∈ V› by auto
show "(λx. if x ∈ M then g x else 1) ` S ⊆ {1}"
using ‹M ⊆ topspace X - S› by auto
qed
qed
lemma completely_regular_eq_regular_space:
"locally_compact_space X
⟹ (completely_regular_space X ⟷ regular_space X)"
using completely_regular_imp_regular_space locally_compact_regular_imp_completely_regular_space
by blast
lemma completely_regular_space_prod_topology:
"completely_regular_space (prod_topology X Y) ⟷
(prod_topology X Y) = trivial_topology ∨
completely_regular_space X ∧ completely_regular_space Y" (is "?lhs=?rhs")
proof
assume ?lhs then show ?rhs
by (rule topological_property_of_prod_component)
(auto simp: completely_regular_space_subtopology homeomorphic_completely_regular_space)
next
assume R: ?rhs
show ?lhs
proof (cases "(prod_topology X Y) = trivial_topology")
case False
then have X: "completely_regular_space X" and Y: "completely_regular_space Y"
using R by blast+
show ?thesis
unfolding completely_regular_space_alt'
proof clarify
fix W x y
assume "openin (prod_topology X Y) W" and "(x, y) ∈ W"
then obtain U V where "openin X U" "openin Y V" "x ∈ U" "y ∈ V" "U×V ⊆ W"
by (force simp: openin_prod_topology_alt)
then have "x ∈ topspace X" "y ∈ topspace Y"
using openin_subset by fastforce+
obtain f where contf: "continuous_map X euclideanreal f" and "f x = 0"
and f1: "⋀x. x ∈ topspace X ⟹ x ∉ U ⟹ f x = 1"
using X ‹openin X U› ‹x ∈ U› unfolding completely_regular_space_alt'
by (smt (verit, best) Diff_iff image_subset_iff singletonD)
obtain g where contg: "continuous_map Y euclideanreal g" and "g y = 0"
and g1: "⋀y. y ∈ topspace Y ⟹ y ∉ V ⟹ g y = 1"
using Y ‹openin Y V› ‹y ∈ V› unfolding completely_regular_space_alt'
by (smt (verit, best) Diff_iff image_subset_iff singletonD)
define h where "h ≡ λ(x,y). 1 - (1 - f x) * (1 - g y)"
show "∃h. continuous_map (prod_topology X Y) euclideanreal h ∧ h (x,y) = 0 ∧ h ` (topspace (prod_topology X Y) - W) ⊆ {1}"
proof (intro exI conjI)
have "continuous_map (prod_topology X Y) euclideanreal (f ∘ fst)"
using contf continuous_map_of_fst by blast
moreover
have "continuous_map (prod_topology X Y) euclideanreal (g ∘ snd)"
using contg continuous_map_of_snd by blast
ultimately
show "continuous_map (prod_topology X Y) euclideanreal h"
unfolding o_def h_def case_prod_unfold
by (intro continuous_intros) auto
show "h (x, y) = 0"
by (simp add: h_def ‹f x = 0› ‹g y = 0›)
show "h ` (topspace (prod_topology X Y) - W) ⊆ {1}"
using ‹U × V ⊆ W› f1 g1 by (force simp: h_def)
qed
qed
qed (force simp: completely_regular_space_def)
qed
proposition completely_regular_space_product_topology:
"completely_regular_space (product_topology X I) ⟷
(∃i∈I. X i = trivial_topology) ∨ (∀i ∈ I. completely_regular_space (X i))"
(is "?lhs ⟷ ?rhs")
proof
assume ?lhs then show ?rhs
by (rule topological_property_of_product_component)
(auto simp: completely_regular_space_subtopology homeomorphic_completely_regular_space)
next
assume R: ?rhs
show ?lhs
proof (cases "∃i∈I. X i = trivial_topology")
case False
show ?thesis
unfolding completely_regular_space_alt'
proof clarify
fix W x
assume W: "openin (product_topology X I) W" and "x ∈ W"
then obtain U where finU: "finite {i ∈ I. U i ≠ topspace (X i)}"
and ope: "⋀i. i∈I ⟹ openin (X i) (U i)"
and x: "x ∈ Pi⇩E I U" and "Pi⇩E I U ⊆ W"
by (auto simp: openin_product_topology_alt)
have "∀i ∈ I. openin (X i) (U i) ∧ x i ∈ U i
⟶ (∃f. continuous_map (X i) euclideanreal f ∧
f (x i) = 0 ∧ (∀x ∈ topspace(X i). x ∉ U i ⟶ f x = 1))"
using R unfolding completely_regular_space_alt'
by (smt (verit) DiffI False image_subset_iff singletonD)
with ope x have "⋀i. ∃f. i ∈ I ⟶ continuous_map (X i) euclideanreal f ∧
f (x i) = 0 ∧ (∀x ∈ topspace (X i). x ∉ U i ⟶ f x = 1)"
by auto
then obtain f where f: "⋀i. i ∈ I ⟹ continuous_map (X i) euclideanreal (f i) ∧
f i (x i) = 0 ∧ (∀x ∈ topspace (X i). x ∉ U i ⟶ f i x = 1)"
by metis
define h where "h ≡ λz. 1 - prod (λi. 1 - f i (z i)) {i∈I. U i ≠ topspace(X i)}"
show "∃h. continuous_map (product_topology X I) euclideanreal h ∧ h x = 0 ∧
h ` (topspace (product_topology X I) - W) ⊆ {1}"
proof (intro conjI exI)
have "continuous_map (product_topology X I) euclidean (f i ∘ (λx. x i))" if "i∈I" for i
using f that
by (blast intro: continuous_intros continuous_map_product_projection)
then show "continuous_map (product_topology X I) euclideanreal h"
unfolding h_def o_def by (intro continuous_intros) (auto simp: finU)
show "h x = 0"
by (simp add: h_def f)
show "h ` (topspace (product_topology X I) - W) ⊆ {1}"
proof -
have "∃i. i ∈ I ∧ U i ≠ topspace (X i) ∧ f i (x' i) = 1"
if "x' ∈ (Π⇩E i∈I. topspace (X i))" "x' ∉ W" for x'
using that ‹Pi⇩E I U ⊆ W› by (smt (verit, best) PiE_iff f in_mono)
then show ?thesis
by (auto simp: f h_def finU)
qed
qed
qed
qed (force simp: completely_regular_space_def)
qed
lemma zero_dimensional_imp_completely_regular_space:
assumes "X dim_le 0"
shows "completely_regular_space X"
proof (clarsimp simp: completely_regular_space_def)
fix C a
assume "closedin X C" and "a ∈ topspace X" and "a ∉ C"
then obtain U where "closedin X U" "openin X U" "a ∈ U" and U: "U ⊆ topspace X - C"
using assms by (force simp add: closedin_def dimension_le_0_neighbourhood_base_of_clopen open_neighbourhood_base_of)
show "∃f. continuous_map X (top_of_set {0::real..1}) f ∧ f a = 0 ∧ f ` C ⊆ {1}"
proof (intro exI conjI)
have "continuous_map X euclideanreal (λx. if x ∈ U then 0 else 1)"
using ‹closedin X U› ‹openin X U› apply (clarsimp simp: continuous_map_def closedin_def)
by (smt (verit) Diff_iff mem_Collect_eq openin_subopen subset_eq)
then show "continuous_map X (top_of_set {0::real..1}) (λx. if x ∈ U then 0 else 1)"
by (auto simp: continuous_map_in_subtopology)
qed (use U ‹a ∈ U› in auto)
qed
lemma zero_dimensional_imp_regular_space: "X dim_le 0 ⟹ regular_space X"
by (simp add: completely_regular_imp_regular_space zero_dimensional_imp_completely_regular_space)
lemma (in Metric_space) t1_space_mtopology:
"t1_space mtopology"
using Hausdorff_space_mtopology t1_or_Hausdorff_space by blast
subsection‹More generally, the k-ification functor›
definition kification_open
where "kification_open ≡
λX S. S ⊆ topspace X ∧ (∀K. compactin X K ⟶ openin (subtopology X K) (K ∩ S))"
definition kification
where "kification X ≡ topology (kification_open X)"
lemma istopology_kification_open: "istopology (kification_open X)"
unfolding istopology_def
proof (intro conjI strip)
show "kification_open X (S ∩ T)"
if "kification_open X S" and "kification_open X T" for S T
using that unfolding kification_open_def
by (smt (verit, best) inf.idem inf_commute inf_left_commute le_infI2 openin_Int)
show "kification_open X (⋃ 𝒦)" if "∀K∈𝒦. kification_open X K" for 𝒦
using that unfolding kification_open_def Int_Union by blast
qed
lemma openin_kification:
"openin (kification X) U ⟷
U ⊆ topspace X ∧
(∀K. compactin X K ⟶ openin (subtopology X K) (K ∩ U))"
by (metis topology_inverse' kification_def istopology_kification_open kification_open_def)
lemma openin_kification_finer:
"openin X S ⟹ openin (kification X) S"
by (simp add: openin_kification openin_subset openin_subtopology_Int2)
lemma topspace_kification [simp]:
"topspace(kification X) = topspace X"
by (meson openin_kification openin_kification_finer openin_topspace subset_antisym)
lemma closedin_kification:
"closedin (kification X) U ⟷
U ⊆ topspace X ∧
(∀K. compactin X K ⟶ closedin (subtopology X K) (K ∩ U))"
proof (cases "U ⊆ topspace X")
case True
then show ?thesis
by (simp add: closedin_def Diff_Int_distrib inf_commute le_infI2 openin_kification)
qed (simp add: closedin_def)
lemma closedin_kification_finer: "closedin X S ⟹ closedin (kification X) S"
by (simp add: closedin_def openin_kification_finer)
lemma kification_eq_self: "kification X = X ⟷ k_space X"
unfolding fun_eq_iff openin_kification k_space_alt openin_inject [symmetric]
by (metis openin_closedin_eq)
lemma compactin_kification [simp]:
"compactin (kification X) K ⟷ compactin X K" (is "?lhs ⟷ ?rhs")
proof
assume ?lhs then show ?rhs
by (simp add: compactin_contractive openin_kification_finer)
next
assume R: ?rhs
show ?lhs
unfolding compactin_def
proof (intro conjI strip)
show "K ⊆ topspace (kification X)"
by (simp add: R compactin_subset_topspace)
fix 𝒰
assume 𝒰: "Ball 𝒰 (openin (kification X)) ∧ K ⊆ ⋃ 𝒰"
then have *: "⋀U. U ∈ 𝒰 ⟹ U ⊆ topspace X ∧ openin (subtopology X K) (K ∩ U)"
by (simp add: R openin_kification)
have "K ⊆ topspace X" "compact_space (subtopology X K)"
using R compactin_subspace by force+
then have "∃V. finite V ∧ V ⊆ (λU. K ∩ U) ` 𝒰 ∧ ⋃ V = topspace (subtopology X K)"
unfolding compact_space
by (smt (verit, del_insts) Int_Union 𝒰 * image_iff inf.order_iff inf_commute topspace_subtopology)
then show "∃ℱ. finite ℱ ∧ ℱ ⊆ 𝒰 ∧ K ⊆ ⋃ ℱ"
by (metis Int_Union ‹K ⊆ topspace X› finite_subset_image inf.orderI topspace_subtopology_subset)
qed
qed
lemma compact_space_kification [simp]:
"compact_space(kification X) ⟷ compact_space X"
by (simp add: compact_space_def)
lemma kification_kification [simp]:
"kification(kification X) = kification X"
unfolding openin_inject [symmetric]
proof
fix U
show "openin (kification (kification X)) U = openin (kification X) U"
proof
show "openin (kification (kification X)) U ⟹ openin (kification X) U"
by (metis compactin_kification inf_commute openin_kification openin_subtopology topspace_kification)
qed (simp add: openin_kification_finer)
qed
lemma k_space_kification [iff]: "k_space(kification X)"
using kification_eq_self by fastforce
lemma continuous_map_into_kification:
assumes "k_space X"
shows "continuous_map X (kification Y) f ⟷ continuous_map X Y f" (is "?lhs ⟷ ?rhs")
proof
assume ?lhs then show ?rhs
by (simp add: continuous_map_def openin_kification_finer)
next
assume R: ?rhs
have "openin X {x ∈ topspace X. f x ∈ V}" if V: "openin (kification Y) V" for V
proof -
have "openin (subtopology X K) (K ∩ {x ∈ topspace X. f x ∈ V})"
if "compactin X K" for K
proof -
have "compactin Y (f ` K)"
using R image_compactin that by blast
then have "openin (subtopology Y (f ` K)) (f ` K ∩ V)"
by (meson V openin_kification)
then obtain U where U: "openin Y U" "f`K ∩ V = U ∩ f`K"
by (meson openin_subtopology)
show ?thesis
unfolding openin_subtopology
proof (intro conjI exI)
show "openin X {x ∈ topspace X. f x ∈ U}"
using R U openin_continuous_map_preimage_gen by (simp add: continuous_map_def)
qed (use U in auto)
qed
then show ?thesis
by (metis (full_types) Collect_subset assms k_space_open)
qed
with R show ?lhs
by (simp add: continuous_map_def)
qed
lemma continuous_map_from_kification:
"continuous_map X Y f ⟹ continuous_map (kification X) Y f"
by (simp add: continuous_map_openin_preimage_eq openin_kification_finer)
lemma continuous_map_kification:
"continuous_map X Y f ⟹ continuous_map (kification X) (kification Y) f"
by (simp add: continuous_map_from_kification continuous_map_into_kification)
lemma subtopology_kification_compact:
assumes "compactin X K"
shows "subtopology (kification X) K = subtopology X K"
unfolding openin_inject [symmetric]
proof
fix U
show "openin (subtopology (kification X) K) U = openin (subtopology X K) U"
by (metis assms inf_commute openin_kification openin_subset openin_subtopology)
qed
lemma subtopology_kification_finer:
assumes "openin (subtopology (kification X) S) U"
shows "openin (kification (subtopology X S)) U"
using assms
by (fastforce simp: openin_subtopology_alt image_iff openin_kification subtopology_subtopology compactin_subtopology)
lemma proper_map_from_kification:
assumes "k_space Y"
shows "proper_map (kification X) Y f ⟷ proper_map X Y f" (is "?lhs ⟷ ?rhs")
proof
assume ?lhs then show ?rhs
by (simp add: closed_map_def closedin_kification_finer proper_map_alt)
next
assume R: ?rhs
have "compactin Y K ⟹ compactin X {x ∈ topspace X. f x ∈ K}" for K
using R proper_map_alt by auto
with R show ?lhs
by (simp add: assms proper_map_into_k_space_eq subtopology_kification_compact)
qed
lemma perfect_map_from_kification:
"⟦k_space Y; perfect_map X Y f⟧ ⟹ perfect_map(kification X) Y f"
by (simp add: continuous_map_from_kification perfect_map_def proper_map_from_kification)
lemma k_space_perfect_map_image_eq:
assumes "Hausdorff_space X" "perfect_map X Y f"
shows "k_space X ⟷ k_space Y"
proof
show "k_space X ⟹ k_space Y"
using k_space_perfect_map_image assms by blast
assume "k_space Y"
have "homeomorphic_map (kification X) X id"
unfolding homeomorphic_eq_injective_perfect_map
proof (intro conjI perfect_map_from_composition_right [where f = id])
show "perfect_map (kification X) Y (f ∘ id)"
by (simp add: ‹k_space Y› assms(2) perfect_map_from_kification)
show "continuous_map (kification X) X id"
by (simp add: continuous_map_from_kification)
qed (use assms perfect_map_def in auto)
then show "k_space X"
using homeomorphic_k_space homeomorphic_space by blast
qed
subsection‹One-point compactifications and the Alexandroff extension construction›
lemma one_point_compactification_dense:
"⟦compact_space X; ¬ compactin X (topspace X - {a})⟧ ⟹ X closure_of (topspace X - {a}) = topspace X"
unfolding closure_of_complement
by (metis Diff_empty closedin_compact_space interior_of_eq_empty openin_closedin_eq subset_singletonD)
lemma one_point_compactification_interior:
"⟦compact_space X; ¬ compactin X (topspace X - {a})⟧ ⟹ X interior_of {a} = {}"
by (simp add: interior_of_eq_empty_complement one_point_compactification_dense)
proposition kc_space_one_point_compactification_gen:
assumes "compact_space X"
shows "kc_space X ⟷
openin X (topspace X - {a}) ∧ (∀K. compactin X K ∧ a∉K ⟶ closedin X K) ∧
k_space (subtopology X (topspace X - {a})) ∧ kc_space (subtopology X (topspace X - {a}))"
(is "?lhs ⟷ ?rhs")
proof
assume L: ?lhs show ?rhs
proof (intro conjI strip)
show "openin X (topspace X - {a})"
using L kc_imp_t1_space t1_space_openin_delete_alt by auto
then show "k_space (subtopology X (topspace X - {a}))"
by (simp add: L assms k_space_open_subtopology_aux)
show "closedin X k" if "compactin X k ∧ a ∉ k" for k :: "'a set"
using L kc_space_def that by blast
show "kc_space (subtopology X (topspace X - {a}))"
by (simp add: L kc_space_subtopology)
qed
next
assume R: ?rhs
show ?lhs
unfolding kc_space_def
proof (intro strip)
fix S
assume "compactin X S"
then have "S ⊆topspace X"
by (simp add: compactin_subset_topspace)
show "closedin X S"
proof (cases "a ∈ S")
case True
then have "topspace X - S = topspace X - {a} - (S - {a})"
by auto
moreover have "openin X (topspace X - {a} - (S - {a}))"
proof (rule openin_trans_full)
show "openin (subtopology X (topspace X - {a})) (topspace X - {a} - (S - {a}))"
proof
show "openin (subtopology X (topspace X - {a})) (topspace X - {a})"
using R openin_open_subtopology by blast
have "closedin (subtopology X ((topspace X - {a}) ∩ K)) (K ∩ (S - {a}))"
if "compactin X K" "K ⊆ topspace X - {a}" for K
proof (intro closedin_subset_topspace)
show "closedin X (K ∩ (S - {a}))"
using that
by (metis IntD1 Int_insert_right_if0 R True ‹compactin X S› closed_Int_compactin insert_Diff subset_Diff_insert)
qed (use that in auto)
moreover have "k_space (subtopology X (topspace X - {a}))"
using R by blast
moreover have "S-{a} ⊆ topspace X ∧ S-{a} ⊆ topspace X - {a}"
using ‹S ⊆ topspace X› by auto
ultimately show "closedin (subtopology X (topspace X - {a})) (S - {a})"
using ‹S ⊆ topspace X› True
by (simp add: k_space_def compactin_subtopology subtopology_subtopology)
qed
show "openin X (topspace X - {a})"
by (simp add: R)
qed
ultimately show ?thesis
by (simp add: ‹S ⊆ topspace X› closedin_def)
next
case False
then show ?thesis
by (simp add: R ‹compactin X S›)
qed
qed
qed
inductive Alexandroff_open for X where
base: "openin X U ⟹ Alexandroff_open X (Some ` U)"
| ext: "⟦compactin X C; closedin X C⟧ ⟹ Alexandroff_open X (insert None (Some ` (topspace X - C)))"
hide_fact (open) base ext
lemma Alexandroff_open_iff: "Alexandroff_open X S ⟷
(∃U. (S = Some ` U ∧ openin X U) ∨ (S = insert None (Some ` (topspace X - U)) ∧ compactin X U ∧ closedin X U))"
by (meson Alexandroff_open.cases Alexandroff_open.ext Alexandroff_open.base)
lemma Alexandroff_open_Un_aux:
assumes U: "openin X U" and "Alexandroff_open X T"
shows "Alexandroff_open X (Some ` U ∪ T)"
using ‹Alexandroff_open X T›
proof (induction rule: Alexandroff_open.induct)
case (base V)
then show ?case
by (metis Alexandroff_open.base U image_Un openin_Un)
next
case (ext C)
have "U ⊆ topspace X"
by (simp add: U openin_subset)
then have eq: "Some ` U ∪ insert None (Some ` (topspace X - C)) = insert None (Some ` (topspace X - (C ∩ (topspace X - U))))"
by force
have "closedin X (C ∩ (topspace X - U))"
using U ext.hyps(2) by blast
moreover
have "compactin X (C ∩ (topspace X - U))"
using U compact_Int_closedin ext.hyps(1) by blast
ultimately show ?case
unfolding eq using Alexandroff_open.ext by blast
qed
lemma Alexandroff_open_Un:
assumes "Alexandroff_open X S" and "Alexandroff_open X T"
shows "Alexandroff_open X (S ∪ T)"
using assms
proof (induction rule: Alexandroff_open.induct)
case (base U)
then show ?case
by (simp add: Alexandroff_open_Un_aux)
next
case (ext C)
then show ?case
by (smt (verit, best) Alexandroff_open_Un_aux Alexandroff_open_iff Un_commute Un_insert_left closedin_def insert_absorb2)
qed
lemma Alexandroff_open_Int_aux:
assumes U: "openin X U" and "Alexandroff_open X T"
shows "Alexandroff_open X (Some ` U ∩ T)"
using ‹Alexandroff_open X T›
proof (induction rule: Alexandroff_open.induct)
case (base V)
then show ?case
by (metis Alexandroff_open.base U image_Int inj_Some openin_Int)
next
case (ext C)
have eq: "Some ` U ∩ insert None (Some ` (topspace X - C)) = Some ` (topspace X - (C ∪ (topspace X - U)))"
by force
have "openin X (topspace X - (C ∪ (topspace X - U)))"
using U ext.hyps(2) by blast
then show ?case
unfolding eq using Alexandroff_open.base by blast
qed
proposition istopology_Alexandroff_open: "istopology (Alexandroff_open X)"
unfolding istopology_def
proof (intro conjI strip)
fix S T
assume "Alexandroff_open X S" and "Alexandroff_open X T"
then show "Alexandroff_open X (S ∩ T)"
proof (induction rule: Alexandroff_open.induct)
case (base U)
then show ?case
using Alexandroff_open_Int_aux by blast
next
case EC: (ext C)
show ?case
using ‹Alexandroff_open X T›
proof (induction rule: Alexandroff_open.induct)
case (base V)
then show ?case
by (metis Alexandroff_open.ext Alexandroff_open_Int_aux EC.hyps inf_commute)
next
case (ext D)
have eq: "insert None (Some ` (topspace X - C)) ∩ insert None (Some ` (topspace X - D))
= insert None (Some ` (topspace X - (C ∪ D)))"
by auto
show ?case
unfolding eq
by (simp add: Alexandroff_open.ext EC.hyps closedin_Un compactin_Un ext.hyps)
qed
qed
next
fix 𝒦
assume §: "∀K∈𝒦. Alexandroff_open X K"
show "Alexandroff_open X (⋃𝒦)"
proof (cases "None ∈ ⋃𝒦")
case True
have "∀K∈𝒦. ∃U. (openin X U ∧ K = Some ` U) ∨ (K = insert None (Some ` (topspace X - U)) ∧ compactin X U ∧ closedin X U)"
by (metis § Alexandroff_open_iff)
then obtain U where U:
"⋀K. K ∈ 𝒦 ⟹ openin X (U K) ∧ K = Some ` (U K)
∨ (K = insert None (Some ` (topspace X - U K)) ∧ compactin X (U K) ∧ closedin X (U K))"
by metis
define 𝒦N where "𝒦N ≡ {K ∈ 𝒦. None ∈ K}"
define A where "A ≡ ⋃K∈𝒦-𝒦N. U K"
define B where "B ≡ ⋂K∈𝒦N. U K"
have U1: "⋀K. K ∈ 𝒦-𝒦N ⟹ openin X (U K) ∧ K = Some ` (U K)"
using U 𝒦N_def by auto
have U2: "⋀K. K ∈ 𝒦N ⟹ K = insert None (Some ` (topspace X - U K)) ∧ compactin X (U K) ∧ closedin X (U K)"
using U 𝒦N_def by auto
have eqA: "⋃(𝒦-𝒦N) = Some ` A"
proof
show "⋃ (𝒦 - 𝒦N) ⊆ Some ` A"
by (metis A_def Sup_le_iff U1 UN_upper subset_image_iff)
show "Some ` A ⊆ ⋃ (𝒦 - 𝒦N)"
using A_def U1 by blast
qed
have eqB: "⋃𝒦N = insert None (Some ` (topspace X - B))"
using U2 True
by (auto simp: B_def image_iff 𝒦N_def)
have "⋃𝒦 = ⋃𝒦N ∪ ⋃(𝒦-𝒦N)"
by (auto simp: 𝒦N_def)
then have eq: "⋃𝒦 = (Some ` A) ∪ (insert None (Some ` (topspace X - B)))"
by (simp add: eqA eqB Un_commute)
show ?thesis
unfolding eq
proof (intro Alexandroff_open_Un Alexandroff_open.intros)
show "openin X A"
using A_def U1 by blast
show "closedin X B"
unfolding B_def using U2 True 𝒦N_def by auto
show "compactin X B"
by (metis B_def U2 eqB Inf_lower Union_iff ‹closedin X B› closed_compactin imageI insertI1)
qed
next
case False
then have "∀K∈𝒦. ∃U. openin X U ∧ K = Some ` U"
by (metis Alexandroff_open.simps UnionI § insertCI)
then obtain U where U: "∀K∈𝒦. openin X (U K) ∧ K = Some ` (U K)"
by metis
then have eq: "⋃𝒦 = Some ` (⋃ K∈𝒦. U K)"
using image_iff by fastforce
show ?thesis
unfolding eq by (simp add: U Alexandroff_open.base openin_clauses(3))
qed
qed
definition Alexandroff_compactification where
"Alexandroff_compactification X ≡ topology (Alexandroff_open X)"
lemma openin_Alexandroff_compactification:
"openin(Alexandroff_compactification X) V ⟷
(∃U. openin X U ∧ V = Some ` U) ∨
(∃C. compactin X C ∧ closedin X C ∧ V = insert None (Some ` (topspace X - C)))"
by (auto simp: Alexandroff_compactification_def istopology_Alexandroff_open Alexandroff_open.simps)
lemma topspace_Alexandroff_compactification [simp]:
"topspace(Alexandroff_compactification X) = insert None (Some ` topspace X)"
(is "?lhs = ?rhs")
proof
show "?lhs ⊆ ?rhs"
by (force simp: topspace_def openin_Alexandroff_compactification)
have "None ∈ topspace (Alexandroff_compactification X)"
by (meson closedin_empty compactin_empty insert_subset openin_Alexandroff_compactification openin_subset)
moreover have "Some x ∈ topspace (Alexandroff_compactification X)"
if "x ∈ topspace X" for x
by (meson that imageI openin_Alexandroff_compactification openin_subset openin_topspace subsetD)
ultimately show "?rhs ⊆ ?lhs"
by (auto simp: image_subset_iff)
qed
lemma closedin_Alexandroff_compactification:
"closedin (Alexandroff_compactification X) C ⟷
(∃K. compactin X K ∧ closedin X K ∧ C = Some ` K) ∨
(∃U. openin X U ∧ C = topspace(Alexandroff_compactification X) - Some ` U)"
(is "?lhs ⟷ ?rhs")
proof
show "?lhs ⟹ ?rhs"
apply (clarsimp simp: closedin_def openin_Alexandroff_compactification)
by (smt (verit) Diff_Diff_Int None_notin_image_Some image_set_diff inf.absorb_iff2 inj_Some insert_Diff_if subset_insert)
show "?rhs ⟹ ?lhs"
using openin_subset
by (fastforce simp: closedin_def openin_Alexandroff_compactification)
qed
lemma openin_Alexandroff_compactification_image_Some [simp]:
"openin(Alexandroff_compactification X) (Some ` U) ⟷ openin X U"
by (auto simp: openin_Alexandroff_compactification inj_image_eq_iff)
lemma closedin_Alexandroff_compactification_image_Some [simp]:
"closedin (Alexandroff_compactification X) (Some ` K) ⟷ compactin X K ∧ closedin X K"
by (auto simp: closedin_Alexandroff_compactification inj_image_eq_iff)
lemma open_map_Some: "open_map X (Alexandroff_compactification X) Some"
using open_map_def openin_Alexandroff_compactification by blast
lemma continuous_map_Some: "continuous_map X (Alexandroff_compactification X) Some"
unfolding continuous_map_def
proof (intro conjI strip)
fix U
assume "openin (Alexandroff_compactification X) U"
then consider V where "openin X V" "U = Some ` V"
| C where "compactin X C" "closedin X C" "U = insert None (Some ` (topspace X - C))"
by (auto simp: openin_Alexandroff_compactification)
then show "openin X {x ∈ topspace X. Some x ∈ U}"
proof cases
case 1
then show ?thesis
using openin_subopen openin_subset by fastforce
next
case 2
then show ?thesis
by (simp add: closedin_def image_iff set_diff_eq)
qed
qed auto
lemma embedding_map_Some: "embedding_map X (Alexandroff_compactification X) Some"
by (simp add: continuous_map_Some injective_open_imp_embedding_map open_map_Some)
lemma compact_space_Alexandroff_compactification [simp]:
"compact_space(Alexandroff_compactification X)"
proof (clarsimp simp: compact_space_alt image_subset_iff)
fix 𝒰 U
assume ope [rule_format]: "∀U∈𝒰. openin (Alexandroff_compactification X) U"
and cover: "∀x∈topspace X. ∃X∈𝒰. Some x ∈ X"
and "U ∈ 𝒰" "None ∈ U"
then have Usub: "U ⊆ insert None (Some ` topspace X)"
by (metis openin_subset topspace_Alexandroff_compactification)
with ope [OF ‹U ∈ 𝒰›] ‹None ∈ U›
obtain C where C: "compactin X C ∧ closedin X C ∧
insert None (Some ` topspace X) - U = Some ` C"
by (auto simp: openin_closedin closedin_Alexandroff_compactification)
then have D: "compactin (Alexandroff_compactification X) (insert None (Some ` topspace X) - U)"
by (metis continuous_map_Some image_compactin)
consider V where "openin X V" "U = Some ` V"
| C where "compactin X C" "closedin X C" "U = insert None (Some ` (topspace X - C))"
using ope [OF ‹U ∈ 𝒰›] by (auto simp: openin_Alexandroff_compactification)
then show "∃ℱ. finite ℱ ∧ ℱ ⊆ 𝒰 ∧ (∃X∈ℱ. None ∈ X) ∧ (∀x∈topspace X. ∃X∈ℱ. Some x ∈ X)"
proof cases
case 1
then show ?thesis
using ‹None ∈ U› by blast
next
case 2
obtain ℱ where "finite ℱ" "ℱ ⊆ 𝒰" "insert None (Some ` topspace X) - U ⊆ ⋃ℱ"
by (smt (verit, del_insts) C D Union_iff compactinD compactin_subset_topspace cover image_subset_iff ope subsetD)
with ‹U ∈ 𝒰› show ?thesis
by (rule_tac x="insert U ℱ" in exI) auto
qed
qed
lemma topspace_Alexandroff_compactification_delete:
"topspace(Alexandroff_compactification X) - {None} = Some ` topspace X"
by simp
lemma Alexandroff_compactification_dense:
assumes "¬ compact_space X"
shows "(Alexandroff_compactification X) closure_of (Some ` topspace X) =
topspace(Alexandroff_compactification X)"
unfolding topspace_Alexandroff_compactification_delete [symmetric]
proof (intro one_point_compactification_dense)
show "¬ compactin (Alexandroff_compactification X) (topspace (Alexandroff_compactification X) - {None})"
using assms compact_space_proper_map_preimage compact_space_subtopology embedding_map_Some embedding_map_def homeomorphic_imp_proper_map by fastforce
qed auto
lemma t0_space_one_point_compactification:
assumes "compact_space X ∧ openin X (topspace X - {a})"
shows "t0_space X ⟷ t0_space (subtopology X (topspace X - {a}))"
(is "?lhs ⟷ ?rhs")
proof
show "?lhs ⟹ ?rhs"
using t0_space_subtopology by blast
show "?rhs ⟹ ?lhs"
using assms
unfolding t0_space_def by (bestsimp simp flip: Int_Diff dest: openin_trans_full)
qed
lemma t0_space_Alexandroff_compactification [simp]:
"t0_space (Alexandroff_compactification X) ⟷ t0_space X"
using t0_space_one_point_compactification [of "Alexandroff_compactification X" None]
using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t0_space by fastforce
lemma t1_space_one_point_compactification:
assumes Xa: "openin X (topspace X - {a})"
and §: "⋀K. ⟦compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K⟧ ⟹ closedin X K"
shows "t1_space X ⟷ t1_space (subtopology X (topspace X - {a}))" (is "?lhs ⟷ ?rhs")
proof
show "?lhs ⟹ ?rhs"
using t1_space_subtopology by blast
assume R: ?rhs
show ?lhs
unfolding t1_space_closedin_singleton
proof (intro strip)
fix x
assume "x ∈ topspace X"
show "closedin X {x}"
proof (cases "x=a")
case True
then show ?thesis
using ‹x ∈ topspace X› Xa closedin_def by blast
next
case False
show ?thesis
by (simp add: "§" False R ‹x ∈ topspace X› closedin_t1_singleton)
qed
qed
qed
lemma closedin_Alexandroff_I:
assumes "compactin (Alexandroff_compactification X) K" "K ⊆ Some ` topspace X"
"closedin (Alexandroff_compactification X) T" "K = T ∩ Some ` topspace X"
shows "closedin (Alexandroff_compactification X) K"
proof -
obtain S where S: "S ⊆ topspace X" "K = Some ` S"
by (meson ‹K ⊆ Some ` topspace X› subset_imageE)
with assms have "compactin X S"
by (metis compactin_subtopology embedding_map_Some embedding_map_def homeomorphic_map_compactness)
moreover have "closedin X S"
using assms S
by (metis closedin_subtopology embedding_map_Some embedding_map_def homeomorphic_map_closedness)
ultimately show ?thesis
by (simp add: S)
qed
lemma t1_space_Alexandroff_compactification [simp]:
"t1_space(Alexandroff_compactification X) ⟷ t1_space X"
proof -
have "openin (Alexandroff_compactification X) (topspace (Alexandroff_compactification X) - {None})"
by auto
then show ?thesis
using t1_space_one_point_compactification [of "Alexandroff_compactification X" None]
using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t1_space
by (fastforce simp: compactin_subtopology closedin_Alexandroff_I closedin_subtopology)
qed
lemma kc_space_one_point_compactification:
assumes "compact_space X"
and ope: "openin X (topspace X - {a})"
and §: "⋀K. ⟦compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K⟧
⟹ closedin X K"
shows "kc_space X ⟷
k_space (subtopology X (topspace X - {a})) ∧ kc_space (subtopology X (topspace X - {a}))"
proof -
have "closedin X K"
if "kc_space (subtopology X (topspace X - {a}))" and "compactin X K" "a ∉ K" for K
using that unfolding kc_space_def
by (metis "§" Diff_empty compactin_subspace compactin_subtopology subset_Diff_insert)
then show ?thesis
by (metis ‹compact_space X› kc_space_one_point_compactification_gen ope)
qed
lemma kc_space_Alexandroff_compactification:
"kc_space(Alexandroff_compactification X) ⟷ (k_space X ∧ kc_space X)" (is "kc_space ?Y = _")
proof -
have "kc_space (Alexandroff_compactification X) ⟷
(k_space (subtopology ?Y (topspace ?Y - {None})) ∧ kc_space (subtopology ?Y (topspace ?Y - {None})))"
by (rule kc_space_one_point_compactification) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I)
also have "... ⟷ k_space X ∧ kc_space X"
using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_k_space homeomorphic_kc_space by simp blast
finally show ?thesis .
qed
proposition regular_space_one_point_compactification:
assumes "compact_space X" and ope: "openin X (topspace X - {a})"
and §: "⋀K. ⟦compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K⟧ ⟹ closedin X K"
shows "regular_space X ⟷
regular_space (subtopology X (topspace X - {a})) ∧ locally_compact_space (subtopology X (topspace X - {a}))"
(is "?lhs ⟷ ?rhs")
proof
show "?lhs ⟹ ?rhs"
using assms(1) compact_imp_locally_compact_space locally_compact_space_open_subset ope regular_space_subtopology by blast
assume R: ?rhs
let ?Xa = "subtopology X (topspace X - {a})"
show ?lhs
unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of imp_conjL
proof (intro strip)
fix W x
assume "openin X W" and "x ∈ W"
show "∃U V. openin X U ∧ closedin X V ∧ x ∈ U ∧ U ⊆ V ∧ V ⊆ W"
proof (cases "x=a")
case True
have "compactin ?Xa (topspace X - W)"
using ‹openin X W› assms(1) closedin_compact_space
by (metis Diff_mono True ‹x ∈ W› compactin_subtopology empty_subsetI insert_subset openin_closedin_eq order_refl)
then obtain V K where V: "openin ?Xa V" and K: "compactin ?Xa K" "closedin ?Xa K" and "topspace X - W ⊆ V" "V ⊆ K"
by (metis locally_compact_space_compact_closed_compact R)
show ?thesis
proof (intro exI conjI)
show "openin X (topspace X - K)"
using "§" K by blast
show "closedin X (topspace X - V)"
using V ope openin_trans_full by blast
show "x ∈ topspace X - K"
proof (rule)
show "x ∈ topspace X"
using ‹openin X W› ‹x ∈ W› openin_subset by blast
show "x ∉ K"
using K True closedin_imp_subset by blast
qed
show "topspace X - K ⊆ topspace X - V"
by (simp add: Diff_mono ‹V ⊆ K›)
show "topspace X - V ⊆ W"
using ‹topspace X - W ⊆ V› by auto
qed
next
case False
have "openin ?Xa ((topspace X - {a}) ∩ W)"
using ‹openin X W› openin_subtopology_Int2 by blast
moreover have "x ∈ (topspace X - {a}) ∩ W"
using ‹openin X W› ‹x ∈ W› openin_subset False by blast
ultimately obtain U V where "openin ?Xa U" "compactin ?Xa V" "closedin ?Xa V"
"x ∈ U" "U ⊆ V" "V ⊆ (topspace X - {a}) ∩ W"
using R locally_compact_regular_space_neighbourhood_base neighbourhood_base_of
by (metis (no_types, lifting))
then show ?thesis
by (meson "§" le_infE ope openin_trans_full)
qed
qed
qed
lemma regular_space_Alexandroff_compactification:
"regular_space(Alexandroff_compactification X) ⟷ regular_space X ∧ locally_compact_space X"
(is "regular_space ?Y = ?rhs")
proof -
have "regular_space ?Y ⟷
regular_space (subtopology ?Y (topspace ?Y - {None})) ∧ locally_compact_space (subtopology ?Y (topspace ?Y - {None}))"
by (rule regular_space_one_point_compactification) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I)
also have "... ⟷ regular_space X ∧ locally_compact_space X"
by (metis embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_locally_compact_space
homeomorphic_regular_space topspace_Alexandroff_compactification_delete)
finally show ?thesis .
qed
lemma Hausdorff_space_one_point_compactification:
assumes "compact_space X" and "openin X (topspace X - {a})"
and "⋀K. ⟦compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K⟧ ⟹ closedin X K"
shows "Hausdorff_space X ⟷
Hausdorff_space (subtopology X (topspace X - {a})) ∧ locally_compact_space (subtopology X (topspace X - {a}))"
(is "?lhs ⟷ ?rhs")
proof
show ?rhs if ?lhs
proof -
have "locally_compact_space (subtopology X (topspace X - {a}))"
using assms that compact_imp_locally_compact_space locally_compact_space_open_subset
by blast
with that show ?rhs
by (simp add: Hausdorff_space_subtopology)
qed
next
show "?rhs ⟹ ?lhs"
by (metis assms locally_compact_Hausdorff_or_regular regular_space_one_point_compactification
regular_t1_eq_Hausdorff_space t1_space_one_point_compactification)
qed
lemma Hausdorff_space_Alexandroff_compactification:
"Hausdorff_space(Alexandroff_compactification X) ⟷ Hausdorff_space X ∧ locally_compact_space X"
by (meson compact_Hausdorff_imp_regular_space compact_space_Alexandroff_compactification
locally_compact_Hausdorff_or_regular regular_space_Alexandroff_compactification
regular_t1_eq_Hausdorff_space t1_space_Alexandroff_compactification)
lemma completely_regular_space_Alexandroff_compactification:
"completely_regular_space(Alexandroff_compactification X) ⟷
completely_regular_space X ∧ locally_compact_space X"
by (metis regular_space_Alexandroff_compactification completely_regular_eq_regular_space
compact_imp_locally_compact_space compact_space_Alexandroff_compactification)
proposition Hausdorff_space_one_point_compactification_asymmetric_prod:
assumes "compact_space X"
shows "Hausdorff_space X ⟷
kc_space (prod_topology X (subtopology X (topspace X - {a}))) ∧
k_space (prod_topology X (subtopology X (topspace X - {a})))" (is "?lhs ⟷ ?rhs")
proof (cases "a ∈ topspace X")
case True
show ?thesis
proof
show ?rhs if ?lhs
proof
show "kc_space (prod_topology X (subtopology X (topspace X - {a})))"
using Hausdorff_imp_kc_space kc_space_prod_topology_right kc_space_subtopology that by blast
show "k_space (prod_topology X (subtopology X (topspace X - {a})))"
by (meson Hausdorff_imp_kc_space assms compact_imp_locally_compact_space k_space_prod_topology_left
kc_space_one_point_compactification_gen that)
qed
next
assume R: ?rhs
define D where "D ≡ (λx. (x,x)) ` (topspace X - {a})"
show ?lhs
proof (cases "topspace X = {a}")
case True
then show ?thesis
by (simp add: Hausdorff_space_def)
next
case False
with R True have "kc_space X"
using kc_space_retraction_map_image [of "prod_topology X (subtopology X (topspace X - {a}))" X fst]
by (metis Diff_subset insert_Diff retraction_map_fst topspace_discrete_topology topspace_subtopology_subset)
have "closedin (subtopology (prod_topology X (subtopology X (topspace X - {a}))) K) (K ∩ D)"
if "compactin (prod_topology X (subtopology X (topspace X - {a}))) K" for K
proof (intro closedin_subtopology_Int_subset[where V=K] closedin_subset_topspace)
show "fst ` K × snd ` K ∩ D ⊆ fst ` K × snd ` K" "K ⊆ fst ` K × snd ` K"
by force+
have eq: "(fst ` K × snd ` K ∩ D) = ((λx. (x,x)) ` (fst ` K ∩ snd ` K))"
using compactin_subset_topspace that by (force simp: D_def image_iff)
have "compactin (prod_topology X (subtopology X (topspace X - {a}))) (fst ` K × snd ` K ∩ D)"
unfolding eq
proof (rule image_compactin [of "subtopology X (topspace X - {a})"])
have "compactin X (fst ` K)" "compactin X (snd ` K)"
by (meson compactin_subtopology continuous_map_fst continuous_map_snd image_compactin that)+
moreover have "fst ` K ∩ snd ` K ⊆ topspace X - {a}"
using compactin_subset_topspace that by force
ultimately
show "compactin (subtopology X (topspace X - {a})) (fst ` K ∩ snd ` K)"
unfolding compactin_subtopology
by (meson ‹kc_space X› closed_Int_compactin kc_space_def)
show "continuous_map (subtopology X (topspace X - {a})) (prod_topology X (subtopology X (topspace X - {a}))) (λx. (x,x))"
by (simp add: continuous_map_paired)
qed
then show "closedin (prod_topology X (subtopology X (topspace X - {a}))) (fst ` K × snd ` K ∩ D)"
using R compactin_imp_closedin_gen by blast
qed
moreover have "D ⊆ topspace X × (topspace X ∩ (topspace X - {a}))"
by (auto simp: D_def)
ultimately have *: "closedin (prod_topology X (subtopology X (topspace X - {a}))) D"
using R by (auto simp: k_space)
have "x=y"
if "x ∈ topspace X" "y ∈ topspace X"
and §: "⋀T. ⟦(x,y) ∈ T; openin (prod_topology X X) T⟧ ⟹ ∃z ∈ topspace X. (z,z) ∈ T" for x y
proof (cases "x=a ∧ y=a")
case False
then consider "x≠a" | "y≠a"
by blast
then show ?thesis
proof cases
case 1
have "∃z ∈ topspace X - {a}. (z,z) ∈ T"
if "(y,x) ∈ T" "openin (prod_topology X (subtopology X (topspace X - {a}))) T" for T
proof -
have "(x,y) ∈ {z ∈ topspace (prod_topology X X). (snd z,fst z) ∈ T ∩ topspace X × (topspace X - {a})}"
by (simp add: "1" ‹x ∈ topspace X› ‹y ∈ topspace X› that)
moreover have "openin (prod_topology X X) {z ∈ topspace (prod_topology X X). (snd z,fst z) ∈ T ∩ topspace X × (topspace X - {a})}"
proof (rule openin_continuous_map_preimage)
show "continuous_map (prod_topology X X) (prod_topology X X) (λx. (snd x, fst x))"
by (simp add: continuous_map_fst continuous_map_pairedI continuous_map_snd)
have "openin (prod_topology X X) (topspace X × (topspace X - {a}))"
using ‹kc_space X› assms kc_space_one_point_compactification_gen openin_prod_Times_iff by fastforce
moreover have "openin (prod_topology X X) T"
using kc_space_one_point_compactification_gen [OF ‹compact_space X›] ‹kc_space X› that
by (metis openin_prod_Times_iff openin_topspace openin_trans_full prod_topology_subtopology(2))
ultimately show "openin (prod_topology X X) (T ∩ topspace X × (topspace X - {a}))"
by blast
qed
ultimately show ?thesis
by (smt (verit) § Int_iff fst_conv mem_Collect_eq mem_Sigma_iff snd_conv)
qed
then have "(y,x) ∈ prod_topology X (subtopology X (topspace X - {a})) closure_of D"
by (simp add: "1" D_def in_closure_of that)
then show ?thesis
using that "*" D_def closure_of_closedin by fastforce
next
case 2
have "∃z ∈ topspace X - {a}. (z,z) ∈ T"
if "(x,y) ∈ T" "openin (prod_topology X (subtopology X (topspace X - {a}))) T" for T
proof -
have "openin (prod_topology X X) (topspace X × (topspace X - {a}))"
using ‹kc_space X› assms kc_space_one_point_compactification_gen openin_prod_Times_iff by fastforce
moreover have XXT: "openin (prod_topology X X) T"
using kc_space_one_point_compactification_gen [OF ‹compact_space X›] ‹kc_space X› that
by (metis openin_prod_Times_iff openin_topspace openin_trans_full prod_topology_subtopology(2))
ultimately have "openin (prod_topology X X) (T ∩ topspace X × (topspace X - {a}))"
by blast
then show ?thesis
by (smt (verit) "§" Diff_subset XXT mem_Sigma_iff openin_subset subsetD that topspace_prod_topology topspace_subtopology_subset)
qed
then have "(x,y) ∈ prod_topology X (subtopology X (topspace X - {a})) closure_of D"
by (simp add: "2" D_def in_closure_of that)
then show ?thesis
using that "*" D_def closure_of_closedin by fastforce
qed
qed auto
then show ?thesis
unfolding Hausdorff_space_closedin_diagonal closure_of_subset_eq [symmetric]
by (force simp: closure_of_def)
qed
qed
next
case False
then show ?thesis
by (simp add: assms compact_imp_k_space compact_space_prod_topology kc_space_compact_prod_topology)
qed
lemma Hausdorff_space_Alexandroff_compactification_asymmetric_prod:
"Hausdorff_space(Alexandroff_compactification X) ⟷
kc_space(prod_topology (Alexandroff_compactification X) X) ∧
k_space(prod_topology (Alexandroff_compactification X) X)"
(is "Hausdorff_space ?Y = ?rhs")
proof -
have *: "subtopology (Alexandroff_compactification X)
(topspace (Alexandroff_compactification X) -
{None}) homeomorphic_space X"
using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_space_sym by fastforce
have "Hausdorff_space (Alexandroff_compactification X) ⟷
(kc_space (prod_topology ?Y (subtopology ?Y (topspace ?Y - {None}))) ∧
k_space (prod_topology ?Y (subtopology ?Y (topspace ?Y - {None}))))"
by (rule Hausdorff_space_one_point_compactification_asymmetric_prod) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I)
also have "... ⟷ ?rhs"
using homeomorphic_k_space homeomorphic_kc_space homeomorphic_space_prod_topology
homeomorphic_space_refl * by blast
finally show ?thesis .
qed
lemma kc_space_as_compactification_unique:
assumes "kc_space X" "compact_space X"
shows "openin X U ⟷
(if a ∈ U then U ⊆ topspace X ∧ compactin X (topspace X - U)
else openin (subtopology X (topspace X - {a})) U)"
proof (cases "a ∈ U")
case True
then show ?thesis
by (meson assms closedin_compact_space compactin_imp_closedin_gen openin_closedin_eq)
next
case False
then show ?thesis
by (metis Diff_empty kc_space_one_point_compactification_gen openin_open_subtopology openin_subset subset_Diff_insert assms)
qed
lemma kc_space_as_compactification_unique_explicit:
assumes "kc_space X" "compact_space X"
shows "openin X U ⟷
(if a ∈ U then U ⊆ topspace X ∧
compactin (subtopology X (topspace X - {a})) (topspace X - U) ∧
closedin (subtopology X (topspace X - {a})) (topspace X - U)
else openin (subtopology X (topspace X - {a})) U)"
apply (simp add: kc_space_subtopology compactin_imp_closedin_gen assms compactin_subtopology cong: conj_cong)
by (metis Diff_mono assms bot.extremum insert_subset kc_space_as_compactification_unique subset_refl)
lemma Alexandroff_compactification_unique:
assumes "kc_space X" "compact_space X" and a: "a ∈ topspace X"
shows "Alexandroff_compactification (subtopology X (topspace X - {a})) homeomorphic_space X"
(is "?Y homeomorphic_space X")
proof -
have [simp]: "topspace X ∩ (topspace X - {a}) = topspace X - {a}"
by auto
have [simp]: "insert None (Some ` A) = insert None (Some ` B) ⟷ A = B"
"insert None (Some ` A) ≠ Some ` B" for A B
by auto
have "quotient_map X ?Y (λx. if x = a then None else Some x)"
unfolding quotient_map_def
proof (intro conjI strip)
show "(λx. if x = a then None else Some x) ` topspace X = topspace ?Y"
using ‹a ∈ topspace X› by force
show "openin X {x ∈ topspace X. (if x = a then None else Some x) ∈ U} = openin ?Y U" (is "?L = ?R")
if "U ⊆ topspace ?Y" for U
proof (cases "None ∈ U")
case True
then obtain T where T[simp]: "U = insert None (Some ` T)"
by (metis Int_insert_right UNIV_I UNIV_option_conv inf.orderE inf_le2 subsetI subset_imageE)
have Tsub: "T ⊆ topspace X - {a}"
using in_these_eq that by auto
then have "{x ∈ topspace X. (if x = a then None else Some x) ∈ U} = insert a T"
by (auto simp: a image_iff cong: conj_cong)
then have "?L ⟷ openin X (insert a T)"
by metis
also have "… ⟷ ?R"
using Tsub assms
apply (simp add: openin_Alexandroff_compactification kc_space_as_compactification_unique_explicit [where a=a])
by (smt (verit, best) Diff_insert2 Diff_subset closedin_imp_subset double_diff)
finally show ?thesis .
next
case False
then obtain T where [simp]: "U = Some ` T"
by (metis Int_insert_right UNIV_I UNIV_option_conv inf.orderE inf_le2 subsetI subset_imageE)
have **: "⋀V. openin X V ⟹ openin X (V - {a})"
by (simp add: assms compactin_imp_closedin_gen openin_diff)
have Tsub: "T ⊆ topspace X - {a}"
using in_these_eq that by auto
then have "{x ∈ topspace X. (if x = a then None else Some x) ∈ U} = T"
by (auto simp: image_iff cong: conj_cong)
then show ?thesis
by (simp add: "**" Tsub openin_open_subtopology)
qed
qed
moreover have "inj_on (λx. if x = a then None else Some x) (topspace X)"
by (auto simp: inj_on_def)
ultimately show ?thesis
using homeomorphic_space_sym homeomorphic_space homeomorphic_map_def by blast
qed
subsection‹Extending continuous maps "pointwise" in a regular space›
lemma continuous_map_on_intermediate_closure_of:
assumes Y: "regular_space Y"
and T: "T ⊆ X closure_of S"
and f: "⋀t. t ∈ T ⟹ limitin Y f (f t) (atin_within X t S)"
shows "continuous_map (subtopology X T) Y f"
proof (clarsimp simp add: continuous_map_atin)
fix a
assume "a ∈ topspace X" and "a ∈ T"
have "f ` T ⊆ topspace Y"
by (metis f image_subsetI limitin_topspace)
have "∀⇩F x in atin_within X a T. f x ∈ W"
if W: "openin Y W" "f a ∈ W" for W
proof -
obtain V C where "openin Y V" "closedin Y C" "f a ∈ V" "V ⊆ C" "C ⊆ W"
by (metis Y W neighbourhood_base_of neighbourhood_base_of_closedin)
have "∀⇩F x in atin_within X a S. f x ∈ V"
by (metis ‹a ∈ T› ‹f a ∈ V› ‹openin Y V› f limitin_def)
then obtain U where "openin X U" "a ∈ U" and U: "∀x ∈ U - {a}. x ∈ S ⟶ f x ∈ V"
by (smt (verit) Diff_iff ‹a ∈ topspace X› eventually_atin_within insert_iff)
moreover have "f z ∈ W" if "z ∈ U" "z ≠ a" "z ∈ T" for z
proof -
have "z ∈ topspace X"
using ‹openin X U› openin_subset ‹z ∈ U› by blast
then have "f z ∈ topspace Y"
using ‹f ` T ⊆ topspace Y› ‹z ∈ T› by blast
{ assume "f z ∈ topspace Y" "f z ∉ C"
then have "∀⇩F x in atin_within X z S. f x ∈ topspace Y - C"
by (metis Diff_iff ‹closedin Y C› closedin_def f limitinD ‹z ∈ T›)
then obtain U' where U': "openin X U'" "z ∈ U'"
"⋀x. x ∈ U' - {z} ⟹ x ∈ S ⟹ f x ∉ C"
by (smt (verit) Diff_iff ‹z ∈ topspace X› eventually_atin_within insertCI)
then have *: "⋀D. z ∈ D ∧ openin X D ⟹ ∃y. y ∈ S ∧ y ∈ D"
by (meson T in_closure_of subsetD ‹z ∈ T›)
have False
using * [of "U ∩ U'"] U' U ‹V ⊆ C› ‹f a ∈ V› ‹f z ∉ C› ‹openin X U› that
by blast
}
then show ?thesis
using ‹C ⊆ W› ‹f z ∈ topspace Y› by auto
qed
ultimately have "∃U. openin X U ∧ a ∈ U ∧ (∀x ∈ U - {a}. x ∈ T ⟶ f x ∈ W)"
by blast
then show ?thesis
using eventually_atin_within by fastforce
qed
then show "limitin Y f (f a) (atin (subtopology X T) a)"
by (metis ‹a ∈ T› atin_subtopology_within f limitin_def)
qed
lemma continuous_map_on_intermediate_closure_of_eq:
assumes "regular_space Y" "S ⊆ T" and Tsub: "T ⊆ X closure_of S"
shows "continuous_map (subtopology X T) Y f ⟷ (∀t ∈ T. limitin Y f (f t) (atin_within X t S))"
(is "?lhs ⟷ ?rhs")
proof
assume L: ?lhs
show ?rhs
proof (clarsimp simp add: continuous_map_atin)
fix x
assume "x ∈ T"
with L Tsub closure_of_subset_topspace
have "limitin Y f (f x) (atin (subtopology X T) x)"
by (fastforce simp: continuous_map_atin)
then show "limitin Y f (f x) (atin_within X x S)"
using ‹x ∈ T› ‹S ⊆ T›
by (force simp: limitin_def atin_subtopology_within eventually_atin_within)
qed
next
show "?rhs ⟹ ?lhs"
using assms continuous_map_on_intermediate_closure_of by blast
qed
lemma continuous_map_extension_pointwise_alt:
assumes §: "regular_space Y" "S ⊆ T" "T ⊆ X closure_of S"
and f: "continuous_map (subtopology X S) Y f"
and lim: "⋀t. t ∈ T-S ⟹ ∃l. limitin Y f l (atin_within X t S)"
obtains g where "continuous_map (subtopology X T) Y g" "⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain g where g: "⋀t. t ∈ T ∧ t ∉ S ⟹ limitin Y f (g t) (atin_within X t S)"
by (metis Diff_iff lim)
let ?h = "λx. if x ∈ S then f x else g x"
show thesis
proof
have T: "T ⊆ topspace X"
using § closure_of_subset_topspace by fastforce
have "limitin Y ?h (f t) (atin_within X t S)" if "t ∈ T" "t ∈ S" for t
proof -
have "limitin Y f (f t) (atin_within X t S)"
by (meson T f limit_continuous_map_within subset_eq that)
then show ?thesis
by (simp add: eventually_atin_within limitin_def)
qed
moreover have "limitin Y ?h (g t) (atin_within X t S)" if "t ∈ T" "t ∉ S" for t
by (smt (verit, del_insts) eventually_atin_within g limitin_def that)
ultimately show "continuous_map (subtopology X T) Y ?h"
unfolding continuous_map_on_intermediate_closure_of_eq [OF §]
by (auto simp: § atin_subtopology_within)
qed auto
qed
lemma continuous_map_extension_pointwise:
assumes "regular_space Y" "S ⊆ T" and Tsub: "T ⊆ X closure_of S"
and ex: " ⋀x. x ∈ T ⟹ ∃g. continuous_map (subtopology X (insert x S)) Y g ∧
(∀x ∈ S. g x = f x)"
obtains g where "continuous_map (subtopology X T) Y g" "⋀x. x ∈ S ⟹ g x = f x"
proof (rule continuous_map_extension_pointwise_alt)
show "continuous_map (subtopology X S) Y f"
proof (clarsimp simp add: continuous_map_atin)
fix t
assume "t ∈ topspace X" and "t ∈ S"
then obtain g where g: "limitin Y g (g t) (atin (subtopology X (insert t S)) t)" and gf: "∀x ∈ S. g x = f x"
by (metis Int_iff ‹S ⊆ T› continuous_map_atin ex inf.orderE insert_absorb topspace_subtopology)
with ‹t ∈ S› show "limitin Y f (f t) (atin (subtopology X S) t)"
by (simp add: limitin_def atin_subtopology_within_if eventually_atin_within gf insert_absorb)
qed
show "∃l. limitin Y f l (atin_within X t S)" if "t ∈ T - S" for t
proof -
obtain g where g: "continuous_map (subtopology X (insert t S)) Y g" and gf: "∀x ∈ S. g x = f x"
using ‹S ⊆ T› ex ‹t ∈ T - S› by force
then have "limitin Y g (g t) (atin_within X t (insert t S))"
using Tsub in_closure_of limit_continuous_map_within that by fastforce
then show ?thesis
unfolding limitin_def
by (smt (verit) eventually_atin_within gf subsetD subset_insertI)
qed
qed (use assms in auto)
subsection‹Extending Cauchy continuous functions to the closure›
lemma Cauchy_continuous_map_extends_to_continuous_closure_of_aux:
assumes m2: "mcomplete_of m2" and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
and "S ⊆ mspace m1"
obtains g
where "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S))
(mtopology_of m2) g" "⋀x. x ∈ S ⟹ g x = f x"
proof (rule continuous_map_extension_pointwise_alt)
interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
by (simp add: Metric_space12_mspace_mdist)
interpret S: Metric_space "S ∩ mspace m1" "mdist m1"
by (simp add: L.M1.subspace)
show "regular_space (mtopology_of m2)"
by (simp add: Metric_space.regular_space_mtopology mtopology_of_def)
show "S ⊆ mtopology_of m1 closure_of S"
by (simp add: assms(3) closure_of_subset)
show "continuous_map (subtopology (mtopology_of m1) S) (mtopology_of m2) f"
by (metis Cauchy_continuous_imp_continuous_map f mtopology_of_submetric)
fix a
assume a: "a ∈ mtopology_of m1 closure_of S - S"
then obtain σ where ranσ: "range σ ⊆ S" "range σ ⊆ mspace m1"
and limσ: "limitin L.M1.mtopology σ a sequentially"
by (force simp: mtopology_of_def L.M1.closure_of_sequentially)
then have "L.M1.MCauchy σ"
by (simp add: L.M1.convergent_imp_MCauchy mtopology_of_def)
then have "L.M2.MCauchy (f ∘ σ)"
using f ranσ by (simp add: Cauchy_continuous_map_def L.M1.subspace Metric_space.MCauchy_def)
then obtain l where l: "limitin L.M2.mtopology (f ∘ σ) l sequentially"
by (meson L.M2.mcomplete_def m2 mcomplete_of_def)
have "limitin L.M2.mtopology f l (atin_within L.M1.mtopology a S)"
unfolding L.limit_atin_sequentially_within imp_conjL
proof (intro conjI strip)
show "l ∈ mspace m2"
using L.M2.limitin_mspace l by blast
fix ρ
assume "range ρ ⊆ S ∩ mspace m1 - {a}" and limρ: "limitin L.M1.mtopology ρ a sequentially"
then have ranρ: "range ρ ⊆ S" "range ρ ⊆ mspace m1" "⋀n. ρ n ≠ a"
by auto
have "a ∈ mspace m1"
using L.M1.limitin_mspace limρ by auto
have "S.MCauchy σ" "S.MCauchy ρ"
using L.M1.convergent_imp_MCauchy L.M1.MCauchy_def S.MCauchy_def limσ ranσ limρ ranρ by force+
then have "L.M2.MCauchy (f ∘ ρ)" "L.M2.MCauchy (f ∘ σ)"
using f by (auto simp: Cauchy_continuous_map_def)
then have ran_f: "range (λx. f (ρ x)) ⊆ mspace m2" "range (λx. f (σ x)) ⊆ mspace m2"
by (auto simp: L.M2.MCauchy_def)
have "(λn. mdist m2 (f (ρ n)) l) ⇢ 0"
proof (rule Lim_null_comparison)
have "mdist m2 (f (ρ n)) l ≤ mdist m2 (f (σ n)) l + mdist m2 (f (σ n)) (f (ρ n))" for n
using ‹l ∈ mspace m2› ran_f L.M2.triangle'' by (smt (verit, best) range_subsetD)
then show "∀⇩F n in sequentially. norm (mdist m2 (f (ρ n)) l) ≤ mdist m2 (f (σ n)) l + mdist m2 (f (σ n)) (f (ρ n))"
by force
define ψ where "ψ ≡ λn. if even n then σ (n div 2) else ρ (n div 2)"
have "(λn. mdist m1 (σ n) (ρ n)) ⇢ 0"
proof (rule Lim_null_comparison)
show "∀⇩F n in sequentially. norm (mdist m1 (σ n) (ρ n)) ≤ mdist m1 (σ n) a + mdist m1 (ρ n) a"
using L.M1.triangle' [of _ a] ranσ ranρ ‹a ∈ mspace m1› by (simp add: range_subsetD)
have "(λn. mdist m1 (σ n) a) ⇢ 0"
using L.M1.limitin_metric_dist_null limσ by blast
moreover have "(λn. mdist m1 (ρ n) a) ⇢ 0"
using L.M1.limitin_metric_dist_null limρ by blast
ultimately show "(λn. mdist m1 (σ n) a + mdist m1 (ρ n) a) ⇢ 0"
by (simp add: tendsto_add_zero)
qed
with ‹S.MCauchy σ› ‹S.MCauchy ρ› have "S.MCauchy ψ"
by (simp add: S.MCauchy_interleaving_gen ψ_def)
then have "L.M2.MCauchy (f ∘ ψ)"
by (metis Cauchy_continuous_map_def f mdist_submetric mspace_submetric)
then have "(λn. mdist m2 (f (σ n)) (f (ρ n))) ⇢ 0"
using L.M2.MCauchy_interleaving_gen [of "f ∘ σ" "f ∘ ρ"]
by (simp add: if_distrib ψ_def o_def cong: if_cong)
moreover have "∀⇩F n in sequentially. f (σ n) ∈ mspace m2 ∧ (λx. mdist m2 (f (σ x)) l) ⇢ 0"
using l by (auto simp: L.M2.limitin_metric_dist_null ‹l ∈ mspace m2›)
ultimately show "(λn. mdist m2 (f (σ n)) l + mdist m2 (f (σ n)) (f (ρ n))) ⇢ 0"
by (metis (mono_tags) tendsto_add_zero eventually_sequentially order_refl)
qed
with ran_f show "limitin L.M2.mtopology (f ∘ ρ) l sequentially"
by (auto simp: L.M2.limitin_metric_dist_null eventually_sequentially ‹l ∈ mspace m2›)
qed
then show "∃l. limitin (mtopology_of m2) f l (atin_within (mtopology_of m1) a S)"
by (force simp: mtopology_of_def)
qed auto
lemma Cauchy_continuous_map_extends_to_continuous_closure_of:
assumes "mcomplete_of m2"
and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
obtains g
where "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S))
(mtopology_of m2) g" "⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain g where cmg:
"continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of (mspace m1 ∩ S)))
(mtopology_of m2) g"
and gf: "(∀x ∈ mspace m1 ∩ S. g x = f x)"
using Cauchy_continuous_map_extends_to_continuous_closure_of_aux assms
by (metis inf_commute inf_le2 submetric_restrict)
define h where "h ≡ λx. if x ∈ topspace(mtopology_of m1) then g x else f x"
show thesis
proof
show "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S))
(mtopology_of m2) h"
unfolding h_def
proof (rule continuous_map_eq)
show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
by (metis closure_of_restrict cmg topspace_mtopology_of)
qed auto
qed (auto simp: gf h_def)
qed
lemma Cauchy_continuous_map_extends_to_continuous_intermediate_closure_of:
assumes "mcomplete_of m2"
and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
and T: "T ⊆ mtopology_of m1 closure_of S"
obtains g
where "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) g"
"(∀x ∈ S. g x = f x)"
by (metis Cauchy_continuous_map_extends_to_continuous_closure_of T assms(1) continuous_map_from_subtopology_mono f)
text ‹Technical lemma helpful for porting particularly ugly HOL Light proofs›
lemma all_in_closure_of:
assumes P: "∀x ∈ S. P x" and clo: "closedin X {x ∈ topspace X. P x}"
shows "∀x ∈ X closure_of S. P x"
proof -
have *: "topspace X ∩ S ⊆ {x ∈ topspace X. P x}"
using P by auto
show ?thesis
using closure_of_minimal [OF * clo] closure_of_restrict by fastforce
qed
lemma Lipschitz_continuous_map_on_intermediate_closure_aux:
assumes lcf: "Lipschitz_continuous_map (submetric m1 S) m2 f"
and "S ⊆ T" and Tsub: "T ⊆ (mtopology_of m1) closure_of S"
and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
and "S ⊆ mspace m1"
shows "Lipschitz_continuous_map (submetric m1 T) m2 f"
proof -
interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
by (simp add: Metric_space12_mspace_mdist)
interpret S: Metric_space "S ∩ mspace m1" "mdist m1"
by (simp add: L.M1.subspace)
have "T ⊆ mspace m1"
using Tsub by (auto simp: mtopology_of_def closure_of_def)
show ?thesis
unfolding Lipschitz_continuous_map_pos
proof
show "f ∈ mspace (submetric m1 T) → mspace m2"
by (metis cmf Metric_space.metric_continuous_map Metric_space_mspace_mdist mtopology_of_def
mtopology_of_submetric image_subset_iff_funcset)
define X where "X ≡ prod_topology (subtopology L.M1.mtopology T) (subtopology L.M1.mtopology T)"
obtain B::real where "B > 0" and B: "∀(x,y) ∈ S×S. mdist m2 (f x) (f y) ≤ B * mdist m1 x y"
using lcf ‹S ⊆ mspace m1› by (force simp: Lipschitz_continuous_map_pos)
have eq: "{z ∈ A. case z of (x,y) ⇒ p x y ≤ B * q x y} = {z ∈ A. ((λ(x,y). B * q x y - p x y)z) ∈ {0..}}"
for p q and A::"('a*'a)set"
by auto
have clo: "closedin X {z ∈ topspace X. case z of (x, y) ⇒ mdist m2 (f x) (f y) ≤ B * mdist m1 x y}"
unfolding eq
proof (rule closedin_continuous_map_preimage)
have *: "continuous_map X L.M2.mtopology (f ∘ fst)" "continuous_map X L.M2.mtopology (f ∘ snd)"
using cmf by (auto simp: mtopology_of_def X_def intro: continuous_map_compose continuous_map_fst continuous_map_snd)
then show "continuous_map X euclidean (λx. case x of (x, y) ⇒ B * mdist m1 x y - mdist m2 (f x) (f y))"
unfolding case_prod_unfold
proof (intro continuous_intros; simp add: mtopology_of_def o_def)
show "continuous_map X L.M1.mtopology fst" "continuous_map X L.M1.mtopology snd"
by (simp_all add: X_def continuous_map_subtopology_fst continuous_map_subtopology_snd flip: subtopology_Times)
qed
qed auto
have "mdist m2 (f x) (f y) ≤ B * mdist m1 x y" if "x ∈ T" "y ∈ T" for x y
using all_in_closure_of [OF B clo] ‹S ⊆ T› Tsub
by (fastforce simp: X_def subset_iff closure_of_Times closure_of_subtopology inf.absorb2
mtopology_of_def that)
then show "∃B>0. ∀x∈mspace (submetric m1 T).
∀y∈mspace (submetric m1 T).
mdist m2 (f x) (f y) ≤ B * mdist (submetric m1 T) x y"
using ‹0 < B› by auto
qed
qed
lemma Lipschitz_continuous_map_on_intermediate_closure:
assumes "Lipschitz_continuous_map (submetric m1 S) m2 f"
and "S ⊆ T" "T ⊆ (mtopology_of m1) closure_of S"
and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
shows "Lipschitz_continuous_map (submetric m1 T) m2 f"
by (metis Lipschitz_continuous_map_on_intermediate_closure_aux assms closure_of_subset_topspace subset_trans topspace_mtopology_of)
lemma Lipschitz_continuous_map_extends_to_closure_of:
assumes m2: "mcomplete_of m2"
and f: "Lipschitz_continuous_map (submetric m1 S) m2 f"
obtains g
where "Lipschitz_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
"⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain g
where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S))
(mtopology_of m2) g" "(∀x ∈ S. g x = f x)"
by (metis Cauchy_continuous_map_extends_to_continuous_closure_of Lipschitz_imp_Cauchy_continuous_map f m2)
have "Lipschitz_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
proof (rule Lipschitz_continuous_map_on_intermediate_closure)
show "Lipschitz_continuous_map (submetric m1 (mspace m1 ∩ S)) m2 g"
by (smt (verit, best) IntD2 Lipschitz_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict)
show "mspace m1 ∩ S ⊆ mtopology_of m1 closure_of S"
using closure_of_subset_Int by force
show "mtopology_of m1 closure_of S ⊆ mtopology_of m1 closure_of (mspace m1 ∩ S)"
by (metis closure_of_restrict subset_refl topspace_mtopology_of)
show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
by (simp add: g)
qed
with g that show thesis
by metis
qed
lemma Lipschitz_continuous_map_extends_to_intermediate_closure_of:
assumes "mcomplete_of m2"
and "Lipschitz_continuous_map (submetric m1 S) m2 f"
and "T ⊆ mtopology_of m1 closure_of S"
obtains g
where "Lipschitz_continuous_map (submetric m1 T) m2 g" "⋀x. x ∈ S ⟹ g x = f x"
by (metis Lipschitz_continuous_map_extends_to_closure_of Lipschitz_continuous_map_from_submetric_mono assms)
text ‹This proof uses the same trick to extend the function's domain to its closure›
lemma uniformly_continuous_map_on_intermediate_closure_aux:
assumes ucf: "uniformly_continuous_map (submetric m1 S) m2 f"
and "S ⊆ T" and Tsub: "T ⊆ (mtopology_of m1) closure_of S"
and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
and "S ⊆ mspace m1"
shows "uniformly_continuous_map (submetric m1 T) m2 f"
proof -
interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
by (simp add: Metric_space12_mspace_mdist)
interpret S: Metric_space "S ∩ mspace m1" "mdist m1"
by (simp add: L.M1.subspace)
have "T ⊆ mspace m1"
using Tsub by (auto simp: mtopology_of_def closure_of_def)
show ?thesis
unfolding uniformly_continuous_map_def
proof (intro conjI strip)
show "f ∈ mspace (submetric m1 T) → mspace m2"
by (metis cmf Metric_space.metric_continuous_map Metric_space_mspace_mdist
mtopology_of_def mtopology_of_submetric image_subset_iff_funcset)
fix ε::real
assume "ε > 0"
then obtain δ where "δ>0" and δ: "∀(x,y) ∈ S×S. mdist m1 x y < δ ⟶ mdist m2 (f x) (f y) ≤ ε/2"
using ucf ‹S ⊆ mspace m1› unfolding uniformly_continuous_map_def mspace_submetric
apply (simp add: Ball_def del: divide_const_simps)
by (metis IntD2 half_gt_zero inf.orderE less_eq_real_def)
define X where "X ≡ prod_topology (subtopology L.M1.mtopology T) (subtopology L.M1.mtopology T)"
text ‹A clever construction involving the union of two closed sets›
have eq: "{z ∈ A. case z of (x,y) ⇒ p x y < d ⟶ q x y ≤ e}
= {z ∈ A. ((λ(x,y). p x y - d)z) ∈ {0..}} ∪ {z ∈ A. ((λ(x,y). e - q x y)z) ∈ {0..}}"
for p q and d e::real and A::"('a*'a)set"
by auto
have clo: "closedin X {z ∈ topspace X. case z of (x, y) ⇒ mdist m1 x y < δ ⟶ mdist m2 (f x) (f y) ≤ ε/2}"
unfolding eq
proof (intro closedin_Un closedin_continuous_map_preimage)
have *: "continuous_map X L.M1.mtopology fst" "continuous_map X L.M1.mtopology snd"
by (metis X_def continuous_map_subtopology_fst subtopology_Times continuous_map_subtopology_snd)+
show "continuous_map X euclidean (λx. case x of (x, y) ⇒ mdist m1 x y - δ)"
unfolding case_prod_unfold
by (intro continuous_intros; simp add: mtopology_of_def *)
have *: "continuous_map X L.M2.mtopology (f ∘ fst)" "continuous_map X L.M2.mtopology (f ∘ snd)"
using cmf by (auto simp: mtopology_of_def X_def intro: continuous_map_compose continuous_map_fst continuous_map_snd)
then show "continuous_map X euclidean (λx. case x of (x, y) ⇒ ε / 2 - mdist m2 (f x) (f y))"
unfolding case_prod_unfold
by (intro continuous_intros; simp add: mtopology_of_def o_def)
qed auto
have "mdist m2 (f x) (f y) ≤ ε/2" if "x ∈ T" "y ∈ T" "mdist m1 x y < δ" for x y
using all_in_closure_of [OF δ clo] ‹S ⊆ T› Tsub
by (fastforce simp: X_def subset_iff closure_of_Times closure_of_subtopology inf.absorb2
mtopology_of_def that)
then show "∃δ>0. ∀x∈mspace (submetric m1 T). ∀y∈mspace (submetric m1 T). mdist (submetric m1 T) y x < δ ⟶ mdist m2 (f y) (f x) < ε"
using ‹0 < δ› ‹0 < ε› by fastforce
qed
qed
lemma uniformly_continuous_map_on_intermediate_closure:
assumes "uniformly_continuous_map (submetric m1 S) m2 f"
and "S ⊆ T" and"T ⊆ (mtopology_of m1) closure_of S"
and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
shows "uniformly_continuous_map (submetric m1 T) m2 f"
by (metis assms closure_of_subset_topspace subset_trans topspace_mtopology_of
uniformly_continuous_map_on_intermediate_closure_aux)
lemma uniformly_continuous_map_extends_to_closure_of:
assumes m2: "mcomplete_of m2"
and f: "uniformly_continuous_map (submetric m1 S) m2 f"
obtains g
where "uniformly_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
"⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain g
where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S))
(mtopology_of m2) g" "(∀x ∈ S. g x = f x)"
by (metis Cauchy_continuous_map_extends_to_continuous_closure_of uniformly_imp_Cauchy_continuous_map f m2)
have "uniformly_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
proof (rule uniformly_continuous_map_on_intermediate_closure)
show "uniformly_continuous_map (submetric m1 (mspace m1 ∩ S)) m2 g"
by (smt (verit, best) IntD2 uniformly_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict)
show "mspace m1 ∩ S ⊆ mtopology_of m1 closure_of S"
using closure_of_subset_Int by force
show "mtopology_of m1 closure_of S ⊆ mtopology_of m1 closure_of (mspace m1 ∩ S)"
by (metis closure_of_restrict subset_refl topspace_mtopology_of)
show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
by (simp add: g)
qed
with g that show thesis
by metis
qed
lemma uniformly_continuous_map_extends_to_intermediate_closure_of:
assumes "mcomplete_of m2"
and "uniformly_continuous_map (submetric m1 S) m2 f"
and "T ⊆ mtopology_of m1 closure_of S"
obtains g
where "uniformly_continuous_map (submetric m1 T) m2 g" "⋀x. x ∈ S ⟹ g x = f x"
by (metis uniformly_continuous_map_extends_to_closure_of uniformly_continuous_map_from_submetric_mono assms)
lemma Cauchy_continuous_map_on_intermediate_closure_aux:
assumes ucf: "Cauchy_continuous_map (submetric m1 S) m2 f"
and "S ⊆ T" and Tsub: "T ⊆ (mtopology_of m1) closure_of S"
and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
and "S ⊆ mspace m1"
shows "Cauchy_continuous_map (submetric m1 T) m2 f"
proof -
interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
by (simp add: Metric_space12_mspace_mdist)
interpret S: Metric_space "S ∩ mspace m1" "mdist m1"
by (simp add: L.M1.subspace)
interpret T: Metric_space T "mdist m1"
by (metis L.M1.subspace Tsub closure_of_subset_topspace dual_order.trans topspace_mtopology_of)
have "T ⊆ mspace m1"
using Tsub by (auto simp: mtopology_of_def closure_of_def)
then show ?thesis
proof (clarsimp simp: Cauchy_continuous_map_def Int_absorb2)
fix σ
assume σ: "T.MCauchy σ"
have "∃y∈S. mdist m1 (σ n) y < inverse (Suc n) ∧ mdist m2 (f (σ n)) (f y) < inverse (Suc n)" for n
proof -
have "σ n ∈ T"
using σ by (force simp: T.MCauchy_def)
moreover have "continuous_map (mtopology_of (submetric m1 T)) L.M2.mtopology f"
by (metis cmf mtopology_of_def mtopology_of_submetric)
ultimately obtain δ where "δ>0" and δ: "∀x ∈ T. mdist m1 (σ n) x < δ ⟶ mdist m2 (f(σ n)) (f x) < inverse (Suc n)"
using ‹T ⊆ mspace m1›
apply (simp add: mtopology_of_def Metric_space.metric_continuous_map L.M1.subspace Int_absorb2)
by (metis inverse_Suc of_nat_Suc)
have "∃y ∈ S. mdist m1 (σ n) y < min δ (inverse (Suc n))"
using ‹σ n ∈ T› Tsub ‹δ>0›
unfolding mtopology_of_def L.M1.metric_closure_of subset_iff mem_Collect_eq L.M1.in_mball
by (smt (verit, del_insts) inverse_Suc )
with δ ‹S ⊆ T› show ?thesis
by auto
qed
then obtain ρ where ρS: "⋀n. ρ n ∈ S" and ρ1: "⋀n. mdist m1 (σ n) (ρ n) < inverse (Suc n)"
and ρ2: "⋀n. mdist m2 (f (σ n)) (f (ρ n)) < inverse (Suc n)"
by metis
have "S.MCauchy ρ"
unfolding S.MCauchy_def
proof (intro conjI strip)
show "range ρ ⊆ S ∩ mspace m1"
using ‹S ⊆ mspace m1› by (auto simp: ρS)
fix ε :: real
assume "ε>0"
then obtain M where M: "⋀n n'. M ≤ n ⟹ M ≤ n' ⟹ mdist m1 (σ n) (σ n') < ε/2"
using σ unfolding T.MCauchy_def by (meson half_gt_zero)
have "∀⇩F n in sequentially. inverse (Suc n) < ε/4"
using Archimedean_eventually_inverse ‹0 < ε› divide_pos_pos zero_less_numeral by blast
then obtain N where N: "⋀n. N ≤ n ⟹ inverse (Suc n) < ε/4"
by (meson eventually_sequentially)
have "mdist m1 (ρ n) (ρ n') < ε" if "n ≥ max M N" "n' ≥ max M N" for n n'
proof -
have "mdist m1 (ρ n) (ρ n') ≤ mdist m1 (ρ n) (σ n) + mdist m1 (σ n) (ρ n')"
by (meson T.MCauchy_def T.triangle ρS σ ‹S ⊆ T› rangeI subset_iff)
also have "… ≤ mdist m1 (ρ n) (σ n) + mdist m1 (σ n) (σ n') + mdist m1 (σ n') (ρ n')"
by (smt (verit, best) T.MCauchy_def T.triangle ρS σ ‹S ⊆ T› in_mono rangeI)
also have "… < ε/4 + ε/2 + ε/4"
using ρ1[of n] ρ1[of n'] N[of n] N[of n'] that M[of n n'] by (simp add: T.commute)
also have "... ≤ ε"
by simp
finally show ?thesis .
qed
then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ mdist m1 (ρ n) (ρ n') < ε"
by blast
qed
then have fρ: "L.M2.MCauchy (f ∘ ρ)"
using ucf by (simp add: Cauchy_continuous_map_def)
show "L.M2.MCauchy (f ∘ σ)"
unfolding L.M2.MCauchy_def
proof (intro conjI strip)
show "range (f ∘ σ) ⊆ mspace m2"
using ‹T ⊆ mspace m1› σ cmf
apply (auto simp: )
by (metis Metric_space.metric_continuous_map Metric_space_mspace_mdist T.MCauchy_def image_eqI inf.absorb1 mspace_submetric mtopology_of_def mtopology_of_submetric range_subsetD subset_iff)
fix ε :: real
assume "ε>0"
then obtain M where M: "⋀n n'. M ≤ n ⟹ M ≤ n' ⟹ mdist m2 ((f ∘ ρ) n) ((f ∘ ρ) n') < ε/2"
using fρ unfolding L.M2.MCauchy_def by (meson half_gt_zero)
have "∀⇩F n in sequentially. inverse (Suc n) < ε/4"
using Archimedean_eventually_inverse ‹0 < ε› divide_pos_pos zero_less_numeral by blast
then obtain N where N: "⋀n. N ≤ n ⟹ inverse (Suc n) < ε/4"
by (meson eventually_sequentially)
have "mdist m2 ((f ∘ σ) n) ((f ∘ σ) n') < ε" if "n ≥ max M N" "n' ≥ max M N" for n n'
proof -
have "mdist m2 ((f ∘ σ) n) ((f ∘ σ) n') ≤ mdist m2 ((f ∘ σ) n) ((f ∘ ρ) n) + mdist m2 ((f ∘ ρ) n) ((f ∘ σ) n')"
by (meson L.M2.MCauchy_def ‹range (f ∘ σ) ⊆ mspace m2› fρ mdist_triangle rangeI subset_eq)
also have "… ≤ mdist m2 ((f ∘ σ) n) ((f ∘ ρ) n) + mdist m2 ((f ∘ ρ) n) ((f ∘ ρ) n') + mdist m2 ((f ∘ ρ) n') ((f ∘ σ) n')"
by (smt (verit) L.M2.MCauchy_def L.M2.triangle ‹range (f ∘ σ) ⊆ mspace m2› fρ range_subsetD)
also have "… < ε/4 + ε/2 + ε/4"
using ρ2[of n] ρ2[of n'] N[of n] N[of n'] that M[of n n'] by (simp add: L.M2.commute)
also have "... ≤ ε"
by simp
finally show ?thesis .
qed
then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ mdist m2 ((f ∘ σ) n) ((f ∘ σ) n') < ε"
by blast
qed
qed
qed
lemma Cauchy_continuous_map_on_intermediate_closure:
assumes "Cauchy_continuous_map (submetric m1 S) m2 f"
and "S ⊆ T" and "T ⊆ (mtopology_of m1) closure_of S"
and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
shows "Cauchy_continuous_map (submetric m1 T) m2 f"
by (metis Cauchy_continuous_map_on_intermediate_closure_aux assms closure_of_subset_topspace order.trans topspace_mtopology_of)
lemma Cauchy_continuous_map_extends_to_closure_of:
assumes m2: "mcomplete_of m2"
and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
obtains g
where "Cauchy_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
"⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain g
where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S))
(mtopology_of m2) g" "(∀x ∈ S. g x = f x)"
by (metis Cauchy_continuous_map_extends_to_continuous_closure_of f m2)
have "Cauchy_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
proof (rule Cauchy_continuous_map_on_intermediate_closure)
show "Cauchy_continuous_map (submetric m1 (mspace m1 ∩ S)) m2 g"
by (smt (verit, best) IntD2 Cauchy_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict)
show "mspace m1 ∩ S ⊆ mtopology_of m1 closure_of S"
using closure_of_subset_Int by force
show "mtopology_of m1 closure_of S ⊆ mtopology_of m1 closure_of (mspace m1 ∩ S)"
by (metis closure_of_restrict subset_refl topspace_mtopology_of)
show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
by (simp add: g)
qed
with g that show thesis
by metis
qed
lemma Cauchy_continuous_map_extends_to_intermediate_closure_of:
assumes "mcomplete_of m2"
and "Cauchy_continuous_map (submetric m1 S) m2 f"
and "T ⊆ mtopology_of m1 closure_of S"
obtains g
where "Cauchy_continuous_map (submetric m1 T) m2 g" "⋀x. x ∈ S ⟹ g x = f x"
by (metis Cauchy_continuous_map_extends_to_closure_of Cauchy_continuous_map_from_submetric_mono assms)
subsection‹Metric space of bounded functions›
context Metric_space
begin
definition fspace :: "'b set ⇒ ('b ⇒ 'a) set" where
"fspace ≡ λS. {f. f`S ⊆ M ∧ f ∈ extensional S ∧ mbounded (f`S)}"
definition fdist :: "['b set, 'b ⇒ 'a, 'b ⇒ 'a] ⇒ real" where
"fdist ≡ λS f g. if f ∈ fspace S ∧ g ∈ fspace S ∧ S ≠ {}
then Sup ((λx. d (f x) (g x)) ` S) else 0"
lemma fspace_empty [simp]: "fspace {} = {λx. undefined}"
by (auto simp: fspace_def)
lemma fdist_empty [simp]: "fdist {} = (λx y. 0)"
by (auto simp: fdist_def)
lemma fspace_in_M: "⟦f ∈ fspace S; x ∈ S⟧ ⟹ f x ∈ M"
by (auto simp: fspace_def)
lemma bdd_above_dist:
assumes f: "f ∈ fspace S" and g: "g ∈ fspace S" and "S ≠ {}"
shows "bdd_above ((λu. d (f u) (g u)) ` S)"
proof -
obtain a where "a ∈ S"
using ‹S ≠ {}› by blast
obtain B x C y where "B>0" and B: "f`S ⊆ mcball x B"
and "C>0" and C: "g`S ⊆ mcball y C"
using f g mbounded_pos by (auto simp: fspace_def)
have "d (f u) (g u) ≤ B + d x y + C" if "u∈S" for u
proof -
have "f u ∈ M"
by (meson B image_eqI mbounded_mcball mbounded_subset_mspace subsetD that)
have "g u ∈ M"
by (meson C image_eqI mbounded_mcball mbounded_subset_mspace subsetD that)
have "x ∈ M" "y ∈ M"
using B C that by auto
have "d (f u) (g u) ≤ d (f u) x + d x (g u)"
by (simp add: ‹f u ∈ M› ‹g u ∈ M› ‹x ∈ M› triangle)
also have "… ≤ d (f u) x + d x y + d y (g u)"
by (simp add: ‹f u ∈ M› ‹g u ∈ M› ‹x ∈ M› ‹y ∈ M› triangle)
also have "… ≤ B + d x y + C"
using B C commute that by fastforce
finally show ?thesis .
qed
then show ?thesis
by (meson bdd_above.I2)
qed
lemma Metric_space_funspace: "Metric_space (fspace S) (fdist S)"
proof
show *: "0 ≤ fdist S f g" for f g
by (auto simp: fdist_def intro: cSUP_upper2 [OF bdd_above_dist])
show "fdist S f g = fdist S g f" for f g
by (auto simp: fdist_def commute)
show "(fdist S f g = 0) = (f = g)"
if fg: "f ∈ fspace S" "g ∈ fspace S" for f g
proof
assume 0: "fdist S f g = 0"
show "f = g"
proof (cases "S={}")
case True
with 0 that show ?thesis
by (simp add: fdist_def fspace_def)
next
case False
with 0 fg have Sup0: "(SUP x∈S. d (f x) (g x)) = 0"
by (simp add: fdist_def)
have "d (f x) (g x) = 0" if "x∈S" for x
by (smt (verit) False Sup0 ‹x∈S› bdd_above_dist [OF fg] less_cSUP_iff nonneg)
with fg show "f=g"
by (simp add: fspace_def extensionalityI image_subset_iff)
qed
next
show "f = g ⟹ fdist S f g = 0"
using fspace_in_M [OF ‹g ∈ fspace S›] by (auto simp: fdist_def)
qed
show "fdist S f h ≤ fdist S f g + fdist S g h"
if fgh: "f ∈ fspace S" "g ∈ fspace S" "h ∈ fspace S" for f g h
proof (clarsimp simp add: fdist_def that)
assume "S ≠ {}"
have dfh: "d (f x) (h x) ≤ d (f x) (g x) + d (g x) (h x)" if "x∈S" for x
by (meson fgh fspace_in_M that triangle)
have bdd_fgh: "bdd_above ((λx. d (f x) (g x)) ` S)" "bdd_above ((λx. d (g x) (h x)) ` S)"
by (simp_all add: ‹S ≠ {}› bdd_above_dist that)
then obtain B C where B: "⋀x. x∈S ⟹ d (f x) (g x) ≤ B" and C: "⋀x. x∈S ⟹ d (g x) (h x) ≤ C"
by (auto simp: bdd_above_def)
then have "⋀x. x∈S ⟹ d (f x) (g x) + d (g x) (h x) ≤ B+C"
by force
then have bdd: "bdd_above ((λx. d (f x) (g x) + d (g x) (h x)) ` S)"
by (auto simp: bdd_above_def)
then have "(SUP x∈S. d (f x) (h x)) ≤ (SUP x∈S. d (f x) (g x) + d (g x) (h x))"
by (metis (mono_tags, lifting) cSUP_mono ‹S ≠ {}› dfh)
also have "… ≤ (SUP x∈S. d (f x) (g x)) + (SUP x∈S. d (g x) (h x))"
by (simp add: ‹S ≠ {}› bdd cSUP_le_iff bdd_fgh add_mono cSup_upper)
finally show "(SUP x∈S. d (f x) (h x)) ≤ (SUP x∈S. d (f x) (g x)) + (SUP x∈S. d (g x) (h x))" .
qed
qed
end
definition funspace where
"funspace S m ≡ metric (Metric_space.fspace (mspace m) (mdist m) S,
Metric_space.fdist (mspace m) (mdist m) S)"
lemma mspace_funspace [simp]:
"mspace (funspace S m) = Metric_space.fspace (mspace m) (mdist m) S"
by (simp add: Metric_space.Metric_space_funspace Metric_space.mspace_metric funspace_def)
lemma mdist_funspace [simp]:
"mdist (funspace S m) = Metric_space.fdist (mspace m) (mdist m) S"
by (simp add: Metric_space.Metric_space_funspace Metric_space.mdist_metric funspace_def)
lemma funspace_imp_welldefined:
"⟦f ∈ mspace (funspace S m); x ∈ S⟧ ⟹ f x ∈ mspace m"
by (simp add: Metric_space.fspace_def subset_iff)
lemma funspace_imp_extensional:
"f ∈ mspace (funspace S m) ⟹ f ∈ extensional S"
by (simp add: Metric_space.fspace_def)
lemma funspace_imp_bounded_image:
"f ∈ mspace (funspace S m) ⟹ Metric_space.mbounded (mspace m) (mdist m) (f ` S)"
by (simp add: Metric_space.fspace_def)
lemma funspace_imp_bounded:
"f ∈ mspace (funspace S m) ⟹ S = {} ∨ (∃c B. ∀x ∈ S. mdist m c (f x) ≤ B)"
by (auto simp: Metric_space.fspace_def Metric_space.mbounded)
lemma (in Metric_space) funspace_imp_bounded2:
assumes "f ∈ fspace S" "g ∈ fspace S"
obtains B where "⋀x. x ∈ S ⟹ d (f x) (g x) ≤ B"
proof -
have "mbounded (f ` S ∪ g ` S)"
using mbounded_Un assms by (force simp: fspace_def)
then show thesis
by (metis UnCI imageI mbounded_alt that)
qed
lemma funspace_imp_bounded2:
assumes "f ∈ mspace (funspace S m)" "g ∈ mspace (funspace S m)"
obtains B where "⋀x. x ∈ S ⟹ mdist m (f x) (g x) ≤ B"
by (metis Metric_space_mspace_mdist assms mspace_funspace Metric_space.funspace_imp_bounded2)
lemma (in Metric_space) funspace_mdist_le:
assumes fg: "f ∈ fspace S" "g ∈ fspace S" and "S ≠ {}"
shows "fdist S f g ≤ a ⟷ (∀x ∈ S. d (f x) (g x) ≤ a)"
using assms bdd_above_dist [OF fg] by (simp add: fdist_def cSUP_le_iff)
lemma funspace_mdist_le:
assumes "f ∈ mspace (funspace S m)" "g ∈ mspace (funspace S m)" and "S ≠ {}"
shows "mdist (funspace S m) f g ≤ a ⟷ (∀x ∈ S. mdist m (f x) (g x) ≤ a)"
using assms by (simp add: Metric_space.funspace_mdist_le)
lemma (in Metric_space) mcomplete_funspace:
assumes "mcomplete"
shows "mcomplete_of (funspace S Self)"
proof -
interpret F: Metric_space "fspace S" "fdist S"
by (simp add: Metric_space_funspace)
show ?thesis
proof (cases "S={}")
case True
then show ?thesis
by (simp add: mcomplete_of_def mcomplete_trivial_singleton)
next
case False
show ?thesis
proof (clarsimp simp: mcomplete_of_def Metric_space.mcomplete_def)
fix σ
assume σ: "F.MCauchy σ"
then have σM: "⋀n x. x ∈ S ⟹ σ n x ∈ M"
by (auto simp: F.MCauchy_def intro: fspace_in_M)
have fdist_less: "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ fdist S (σ n) (σ n') < ε" if "ε>0" for ε
using σ that by (auto simp: F.MCauchy_def)
have σext: "⋀n. σ n ∈ extensional S"
using σ unfolding F.MCauchy_def by (auto simp: fspace_def)
have σbd: "⋀n. mbounded (σ n ` S)"
using σ unfolding F.MCauchy_def by (simp add: fspace_def image_subset_iff)
have σin[simp]: "σ n ∈ fspace S" for n
using F.MCauchy_def σ by blast
have bd2: "⋀n n'. ∃B. ∀x ∈ S. d (σ n x) (σ n' x) ≤ B"
using σ unfolding F.MCauchy_def by (metis range_subsetD funspace_imp_bounded2)
have sup: "⋀n n' x0. x0 ∈ S ⟹ d (σ n x0) (σ n' x0) ≤ Sup ((λx. d (σ n x) (σ n' x)) ` S)"
proof (rule cSup_upper)
show "bdd_above ((λx. d (σ n x) (σ n' x)) ` S)" if "x0 ∈ S" for n n' x0
using that bd2 by (meson bdd_above.I2)
qed auto
have pcy: "MCauchy (λn. σ n x)" if "x ∈ S" for x
unfolding MCauchy_def
proof (intro conjI strip)
show "range (λn. σ n x) ⊆ M"
using σM that by blast
fix ε :: real
assume "ε > 0"
then obtain N where N: "⋀n n'. N ≤ n ⟶ N ≤ n' ⟶ fdist S (σ n) (σ n') < ε"
using σ by (force simp: F.MCauchy_def)
{ fix n n'
assume n: "N ≤ n" "N ≤ n'"
have "d (σ n x) (σ n' x) ≤ (SUP x∈S. d (σ n x) (σ n' x))"
using that sup by presburger
then have "d (σ n x) (σ n' x) ≤ fdist S (σ n) (σ n')"
by (simp add: fdist_def ‹S ≠ {}›)
with N n have "d (σ n x) (σ n' x) < ε"
by fastforce
} then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n x) (σ n' x) < ε"
by blast
qed
have "∃l. limitin mtopology (λn. σ n x) l sequentially" if "x ∈ S" for x
using assms mcomplete_def pcy ‹x ∈ S› by presburger
then obtain g0 where g0: "⋀x. x ∈ S ⟹ limitin mtopology (λn. σ n x) (g0 x) sequentially"
by metis
define g where "g ≡ restrict g0 S"
have gext: "g ∈ extensional S"
and glim: "⋀x. x ∈ S ⟹ limitin mtopology (λn. σ n x) (g x) sequentially"
by (auto simp: g_def g0)
have gwd: "g x ∈ M" if "x ∈ S" for x
using glim limitin_metric that by blast
have unif: "∃N. ∀x n. x ∈ S ⟶ N ≤ n ⟶ d (σ n x) (g x) < ε" if "ε>0" for ε
proof -
obtain N where N: "⋀n n'. N ≤ n ∧ N ≤ n' ⟹ Sup ((λx. d (σ n x) (σ n' x)) ` S) < ε/2"
using ‹S≠{}› ‹ε>0› fdist_less [of "ε/2"]
by (metis (mono_tags) σin fdist_def half_gt_zero)
show ?thesis
proof (intro exI strip)
fix x n
assume "x ∈ S" and "N ≤ n"
obtain N' where N': "⋀n. N' ≤ n ⟹ σ n x ∈ M ∧ d (σ n x) (g x) < ε/2"
by (metis ‹0 < ε› ‹x ∈ S› glim half_gt_zero limit_metric_sequentially)
have "d (σ n x) (g x) ≤ d (σ n x) (σ (max N N') x) + d (σ (max N N') x) (g x)"
using ‹x ∈ S› σM gwd triangle by presburger
also have "… < ε/2 + ε/2"
by (smt (verit) N N' ‹N ≤ n› ‹x ∈ S› max.cobounded1 max.cobounded2 sup)
finally show "d (σ n x) (g x) < ε" by simp
qed
qed
have "limitin F.mtopology σ g sequentially"
unfolding F.limit_metric_sequentially
proof (intro conjI strip)
obtain N where N: "⋀n n'. N ≤ n ∧ N ≤ n' ⟹ Sup ((λx. d (σ n x) (σ n' x)) ` S) < 1"
using fdist_less [of 1] ‹S≠{}› by (auto simp: fdist_def)
have "⋀x. x ∈ σ N ` S ⟹ x ∈ M"
using σM by blast
obtain a B where "a ∈ M" and B: "⋀x. x ∈ (σ N) ` S ⟹ d a x ≤ B"
by (metis False σM σbd ex_in_conv imageI mbounded_alt_pos)
have "d a (g x) ≤ B+1" if "x∈S" for x
proof -
have "d a (g x) ≤ d a (σ N x) + d (σ N x) (g x)"
by (simp add: ‹a ∈ M› σM gwd that triangle)
also have "… ≤ B+1"
proof -
have "d a (σ N x) ≤ B"
by (simp add: B that)
moreover
have False if 1: "d (σ N x) (g x) > 1"
proof -
obtain r where "1 < r" and r: "r < d (σ N x) (g x)"
using 1 dense by blast
then obtain N' where N': "⋀n. N' ≤ n ⟹ σ n x ∈ M ∧ d (σ n x) (g x) < r-1"
using glim [OF ‹x∈S›] by (fastforce simp: limit_metric_sequentially)
have "d (σ N x) (g x) ≤ d (σ N x) (σ (max N N') x) + d (σ (max N N') x) (g x)"
by (metis ‹x ∈ S› σM commute gwd triangle')
also have "… < 1 + (r-1)"
by (smt (verit) N N' ‹x ∈ S› max.cobounded1 max.cobounded2 max.idem sup)
finally have "d (σ N x) (g x) < r"
by simp
with r show False
by linarith
qed
ultimately show ?thesis
by force
qed
finally show ?thesis .
qed
with gwd ‹a ∈ M› have "mbounded (g ` S)"
unfolding mbounded by blast
with gwd gext show "g ∈ fspace S"
by (auto simp: fspace_def)
fix ε::real
assume "ε>0"
then obtain N where "⋀x n. x ∈ S ⟹ N ≤ n ⟹ d (σ n x) (g x) < ε/2"
by (meson unif half_gt_zero)
then have "fdist S (σ n) g ≤ ε/2" if "N ≤ n" for n
using ‹g ∈ fspace S› False that
by (force simp: funspace_mdist_le simp del: divide_const_simps)
then show "∃N. ∀n≥N. σ n ∈ fspace S ∧ fdist S (σ n) g < ε"
by (metis ‹0 < ε› σin add_strict_increasing field_sum_of_halves half_gt_zero)
qed
then show "∃x. limitin F.mtopology σ x sequentially"
by blast
qed
qed
qed
subsection‹Metric space of continuous bounded functions›
definition cfunspace where
"cfunspace X m ≡ submetric (funspace (topspace X) m) {f. continuous_map X (mtopology_of m) f}"
lemma mspace_cfunspace [simp]:
"mspace (cfunspace X m) =
{f. f ∈ topspace X → mspace m ∧ f ∈ extensional (topspace X) ∧
Metric_space.mbounded (mspace m) (mdist m) (f ` (topspace X)) ∧
continuous_map X (mtopology_of m) f}"
by (auto simp: cfunspace_def Metric_space.fspace_def)
lemma mdist_cfunspace_eq_mdist_funspace:
"mdist (cfunspace X m) = mdist (funspace (topspace X) m)"
by (auto simp: cfunspace_def)
lemma cfunspace_subset_funspace:
"mspace (cfunspace X m) ⊆ mspace (funspace (topspace X) m)"
by (simp add: cfunspace_def)
lemma cfunspace_mdist_le:
"⟦f ∈ mspace (cfunspace X m); g ∈ mspace (cfunspace X m); topspace X ≠ {}⟧
⟹ mdist (cfunspace X m) f g ≤ a ⟷ (∀x ∈ topspace X. mdist m (f x) (g x) ≤ a)"
by (simp add: cfunspace_def Metric_space.funspace_mdist_le)
lemma cfunspace_imp_bounded2:
assumes "f ∈ mspace (cfunspace X m)" "g ∈ mspace (cfunspace X m)"
obtains B where "⋀x. x ∈ topspace X ⟹ mdist m (f x) (g x) ≤ B"
by (metis assms all_not_in_conv cfunspace_mdist_le nle_le)
lemma cfunspace_mdist_lt:
"⟦compactin X (topspace X); f ∈ mspace (cfunspace X m);
g ∈ mspace (cfunspace X m); mdist (cfunspace X m) f g < a;
x ∈ topspace X⟧
⟹ mdist m (f x) (g x) < a"
by (metis (full_types) cfunspace_mdist_le empty_iff less_eq_real_def less_le_not_le)
lemma mdist_cfunspace_le:
assumes "0 ≤ B" and B: "⋀x. x ∈ topspace X ⟹ mdist m (f x) (g x) ≤ B"
shows "mdist (cfunspace X m) f g ≤ B"
proof (cases "X = trivial_topology")
case True
then show ?thesis
by (simp add: Metric_space.fdist_empty ‹B≥0› cfunspace_def)
next
case False
have bdd: "bdd_above ((λu. mdist m (f u) (g u)) ` topspace X)"
by (meson B bdd_above.I2)
with assms bdd show ?thesis
by (simp add: mdist_cfunspace_eq_mdist_funspace Metric_space.fdist_def cSUP_le_iff)
qed
lemma mdist_cfunspace_imp_mdist_le:
"⟦f ∈ mspace (cfunspace X m); g ∈ mspace (cfunspace X m);
mdist (cfunspace X m) f g ≤ a; x ∈ topspace X⟧ ⟹ mdist m (f x) (g x) ≤ a"
using cfunspace_mdist_le by blast
lemma compactin_mspace_cfunspace:
"compactin X (topspace X)
⟹ mspace (cfunspace X m) =
{f. (∀x ∈ topspace X. f x ∈ mspace m) ∧
f ∈ extensional (topspace X) ∧
continuous_map X (mtopology_of m) f}"
by (auto simp: Metric_space.compactin_imp_mbounded image_compactin mtopology_of_def)
lemma (in Metric_space) mcomplete_cfunspace:
assumes "mcomplete"
shows "mcomplete_of (cfunspace X Self)"
proof -
interpret F: Metric_space "fspace (topspace X)" "fdist (topspace X)"
by (simp add: Metric_space_funspace)
interpret S: Submetric "fspace (topspace X)" "fdist (topspace X)" "mspace (cfunspace X Self)"
proof
show "mspace (cfunspace X Self) ⊆ fspace (topspace X)"
by (metis cfunspace_subset_funspace mdist_Self mspace_Self mspace_funspace)
qed
show ?thesis
proof (cases "X = trivial_topology")
case True
then show ?thesis
by (simp add: mcomplete_of_def mcomplete_trivial_singleton mdist_cfunspace_eq_mdist_funspace cong: conj_cong)
next
case False
have *: "continuous_map X mtopology g"
if "range σ ⊆ mspace (cfunspace X Self)"
and g: "limitin F.mtopology σ g sequentially" for σ g
unfolding continuous_map_to_metric
proof (intro strip)
have σ: "⋀n. continuous_map X mtopology (σ n)"
using that by (auto simp: mtopology_of_def)
fix x and ε::real
assume "x ∈ topspace X" and "0 < ε"
then obtain N where N: "⋀n. N ≤ n ⟹ σ n ∈ fspace (topspace X) ∧ fdist (topspace X) (σ n) g < ε/3"
unfolding mtopology_of_def F.limitin_metric
by (metis F.limit_metric_sequentially divide_pos_pos g zero_less_numeral)
then obtain U where "openin X U" "x ∈ U"
and U: "⋀y. y ∈ U ⟹ σ N y ∈ mball (σ N x) (ε/3)"
by (metis Metric_space.continuous_map_to_metric Metric_space_axioms ‹0 < ε› ‹x ∈ topspace X› σ divide_pos_pos zero_less_numeral)
moreover
have "g y ∈ mball (g x) ε" if "y∈U" for y
proof -
have "U ⊆ topspace X"
using ‹openin X U› by (simp add: openin_subset)
have gx: "g x ∈ M"
by (meson F.limitin_mspace ‹x ∈ topspace X› fspace_in_M g)
have "y ∈ topspace X"
using ‹U ⊆ topspace X› that by auto
have gy: "g y ∈ M"
by (meson F.limitin_mspace[OF g] ‹U ⊆ topspace X› fspace_in_M subsetD that)
have "d (g x) (g y) < ε"
proof -
have *: "d (σ N x0) (g x0) ≤ ε/3" if "x0 ∈ topspace X" for x0
proof -
have "g ∈ fspace (topspace X)"
using F.limit_metric_sequentially g by blast
with N that have "bdd_above ((λx. d (σ N x) (g x)) ` topspace X)"
by (force intro: bdd_above_dist)
then have "d (σ N x0) (g x0) ≤ Sup ((λx. d (σ N x) (g x)) ` topspace X)"
by (simp add: cSup_upper that)
also have "… ≤ ε/3"
using g False N ‹g ∈ fspace (topspace X)›
by (fastforce simp: F.limit_metric_sequentially fdist_def)
finally show ?thesis .
qed
have "d (g x) (g y) ≤ d (g x) (σ N x) + d (σ N x) (g y)"
using U gx gy that triangle by force
also have "… < ε/3 + ε/3 + ε/3"
by (smt (verit) "*" U gy ‹x ∈ topspace X› ‹y ∈ topspace X› commute in_mball that triangle)
finally show ?thesis by simp
qed
with gx gy show ?thesis by simp
qed
ultimately show "∃U. openin X U ∧ x ∈ U ∧ (∀y∈U. g y ∈ mball (g x) ε)"
by blast
qed
have "S.sub.mcomplete"
proof (rule S.sequentially_closedin_mcomplete_imp_mcomplete)
show "F.mcomplete"
by (metis assms mcomplete_funspace mcomplete_of_def mdist_Self mdist_funspace mspace_Self mspace_funspace)
fix σ g
assume g: "range σ ⊆ mspace (cfunspace X Self) ∧ limitin F.mtopology σ g sequentially"
show "g ∈ mspace (cfunspace X Self)"
proof (simp add: mtopology_of_def, intro conjI)
show "g ∈ topspace X → M" "g ∈ extensional (topspace X)" "mbounded (g ` topspace X)"
using g F.limitin_mspace by (force simp: fspace_def)+
show "continuous_map X mtopology g"
using * g by blast
qed
qed
then show ?thesis
by (simp add: mcomplete_of_def mdist_cfunspace_eq_mdist_funspace)
qed
qed
subsection‹Existence of completion for any metric space M as a subspace of M=>R›
lemma (in Metric_space) metric_completion_explicit:
obtains f :: "['a,'a] ⇒ real" and S where
"S ⊆ mspace(funspace M euclidean_metric)"
"mcomplete_of (submetric (funspace M euclidean_metric) S)"
"f ∈ M → S"
"mtopology_of(funspace M euclidean_metric) closure_of f ` M = S"
"⋀x y. ⟦x ∈ M; y ∈ M⟧
⟹ mdist (funspace M euclidean_metric) (f x) (f y) = d x y"
proof -
define m':: "('a⇒real) metric" where "m' ≡ funspace M euclidean_metric"
show thesis
proof (cases "M = {}")
case True
then show ?thesis
using that by (simp add: mcomplete_of_def mcomplete_trivial)
next
case False
then obtain a where "a ∈ M"
by auto
define f where "f ≡ λx. (λu ∈ M. d x u - d a u)"
define S where "S ≡ mtopology_of(funspace M euclidean_metric) closure_of (f ` M)"
interpret S: Submetric "Met_TC.fspace M" "Met_TC.fdist M" "S ∩ Met_TC.fspace M"
by (simp add: Met_TC.Metric_space_funspace Submetric.intro Submetric_axioms_def)
have fim: "f ` M ⊆ mspace m'"
proof (clarsimp simp: m'_def Met_TC.fspace_def)
fix b
assume "b ∈ M"
then have "⋀c. ⟦c ∈ M⟧ ⟹ ¦d b c - d a c¦ ≤ d a b"
by (smt (verit, best) ‹a ∈ M› commute triangle'')
then have "(λx. d b x - d a x) ` M ⊆ cball 0 (d a b)"
by force
then show "f b ∈ extensional M ∧ bounded (f b ` M)"
by (metis bounded_cball bounded_subset f_def image_restrict_eq restrict_extensional_sub set_eq_subset)
qed
show thesis
proof
show "S ⊆ mspace (funspace M euclidean_metric)"
by (simp add: S_def in_closure_of subset_iff)
have "closedin S.mtopology (S ∩ Met_TC.fspace M)"
by (simp add: S_def closedin_Int funspace_def)
moreover have "S.mcomplete"
using Metric_space.mcomplete_funspace Met_TC.Metric_space_axioms by (fastforce simp: mcomplete_of_def)
ultimately show "mcomplete_of (submetric (funspace M euclidean_metric) S)"
by (simp add: S.closedin_eq_mcomplete mcomplete_of_def)
show "f ∈ M → S"
using S_def fim in_closure_of m'_def by fastforce
show "mtopology_of (funspace M euclidean_metric) closure_of f ` M = S"
by (auto simp: f_def S_def mtopology_of_def)
show "mdist (funspace M euclidean_metric) (f x) (f y) = d x y"
if "x ∈ M" "y ∈ M" for x y
proof -
have "∀c∈M. dist (f x c) (f y c) ≤ r ⟹ d x y ≤ r" for r
using that by (auto simp: f_def dist_real_def)
moreover have "dist (f x z) (f y z) ≤ r" if "d x y ≤ r" and "z ∈ M" for r z
using that ‹x ∈ M› ‹y ∈ M›
apply (simp add: f_def Met_TC.fdist_def dist_real_def)
by (smt (verit, best) commute triangle')
ultimately have "(SUP c ∈ M. dist (f x c) (f y c)) = d x y"
by (intro cSup_unique) auto
with that fim show ?thesis
using that fim by (simp add: Met_TC.fdist_def False m'_def image_subset_iff)
qed
qed
qed
qed
lemma (in Metric_space) metric_completion:
obtains f :: "['a,'a] ⇒ real" and m' where
"mcomplete_of m'"
"f ∈ M → mspace m' "
"mtopology_of m' closure_of f ` M = mspace m'"
"⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ mdist m' (f x) (f y) = d x y"
proof -
obtain f :: "['a,'a] ⇒ real" and S where
Ssub: "S ⊆ mspace(funspace M euclidean_metric)"
and mcom: "mcomplete_of (submetric (funspace M euclidean_metric) S)"
and fim: "f ∈ M → S"
and eqS: "mtopology_of(funspace M euclidean_metric) closure_of f ` M = S"
and eqd: "⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ mdist (funspace M euclidean_metric) (f x) (f y) = d x y"
using metric_completion_explicit by metis
define m' where "m' ≡ submetric (funspace M euclidean_metric) S"
show thesis
proof
show "mcomplete_of m'"
by (simp add: mcom m'_def)
show "f ∈ M → mspace m'"
using Ssub fim m'_def by auto
show "mtopology_of m' closure_of f ` M = mspace m'"
using eqS fim Ssub
by (force simp: m'_def mtopology_of_submetric closure_of_subtopology Int_absorb1 image_subset_iff_funcset)
show "mdist m' (f x) (f y) = d x y" if "x ∈ M" and "y ∈ M" for x y
using that eqd m'_def by force
qed
qed
lemma metrizable_space_completion:
assumes "metrizable_space X"
obtains f :: "['a,'a] ⇒ real" and Y where
"completely_metrizable_space Y" "embedding_map X Y f"
"Y closure_of (f ` (topspace X)) = topspace Y"
proof -
obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d"
using assms metrizable_space_def by blast
then interpret Metric_space M d by simp
obtain f :: "['a,'a] ⇒ real" and m' where
"mcomplete_of m'"
and fim: "f ∈ M → mspace m' "
and m': "mtopology_of m' closure_of f ` M = mspace m'"
and eqd: "⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ mdist m' (f x) (f y) = d x y"
by (metis metric_completion)
show thesis
proof
show "completely_metrizable_space (mtopology_of m')"
using ‹mcomplete_of m'›
unfolding completely_metrizable_space_def mcomplete_of_def mtopology_of_def
by (metis Metric_space_mspace_mdist)
show "embedding_map X (mtopology_of m') f"
using Metric_space12.isometry_imp_embedding_map
by (metis Metric_space12_def Metric_space_axioms Metric_space_mspace_mdist Xeq eqd fim
mtopology_of_def)
show "(mtopology_of m') closure_of f ` topspace X = topspace (mtopology_of m')"
by (simp add: Xeq m')
qed
qed
subsection‹Contractions›
lemma (in Metric_space) contraction_imp_unique_fixpoint:
assumes "f x = x" "f y = y"
and "f ∈ M → M"
and "k < 1"
and "⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ d (f x) (f y) ≤ k * d x y"
and "x ∈ M" "y ∈ M"
shows "x = y"
by (smt (verit, ccfv_SIG) mdist_pos_less mult_le_cancel_right1 assms)
text ‹Banach Fixed-Point Theorem (aka, Contraction Mapping Principle)›
lemma (in Metric_space) Banach_fixedpoint_thm:
assumes mcomplete and "M ≠ {}" and fim: "f ∈ M → M"
and "k < 1"
and con: "⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ d (f x) (f y) ≤ k * d x y"
obtains x where "x ∈ M" "f x = x"
proof -
obtain a where "a ∈ M"
using ‹M ≠ {}› by blast
show thesis
proof (cases "∀x ∈ M. f x = f a")
case True
then show ?thesis
by (metis ‹a ∈ M› fim image_subset_iff image_subset_iff_funcset that)
next
case False
then obtain b where "b ∈ M" and b: "f b ≠ f a"
by blast
have "k>0"
using Lipschitz_coefficient_pos [where f=f]
by (metis False ‹a ∈ M› con fim mdist_Self mspace_Self)
define σ where "σ ≡ λn. (f^^n) a"
have f_iter: "σ n ∈ M" for n
unfolding σ_def by (induction n) (use ‹a ∈ M› fim in auto)
show ?thesis
proof (cases "f a = a")
case True
then show ?thesis
using ‹a ∈ M› that by blast
next
case False
have "MCauchy σ"
proof -
show ?thesis
unfolding MCauchy_def
proof (intro conjI strip)
show "range σ ⊆ M"
using f_iter by blast
fix ε::real
assume "ε>0"
with ‹k < 1› ‹f a ≠ a› ‹a ∈ M› fim have gt0: "((1 - k) * ε) / d a (f a) > 0"
by (fastforce simp: divide_simps Pi_iff)
obtain N where "k^N < ((1-k) * ε) / d a (f a)"
using real_arch_pow_inv [OF gt0 ‹k < 1›] by blast
then have N: "⋀n. n ≥ N ⟹ k^n < ((1-k) * ε) / d a (f a)"
by (smt (verit) ‹0 < k› assms(4) power_decreasing)
have "∀n n'. n<n' ⟶ N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
proof (intro exI strip)
fix n n'
assume "n<n'" "N ≤ n" "N ≤ n'"
have "d (σ n) (σ n') ≤ (∑i=n..<n'. d (σ i) (σ (Suc i)))"
proof -
have "n < m ⟹ d (σ n) (σ m) ≤ (∑i=n..<m. d (σ i) (σ (Suc i)))" for m
proof (induction m)
case 0
then show ?case
by simp
next
case (Suc m)
then consider "n<m" | "m=n"
by linarith
then show ?case
proof cases
case 1
have "d (σ n) (σ (Suc m)) ≤ d (σ n) (σ m) + d (σ m) (σ (Suc m))"
by (simp add: f_iter triangle)
also have "… ≤ (∑i=n..<m. d (σ i) (σ (Suc i))) + d (σ m) (σ (Suc m))"
using Suc 1 by linarith
also have "… = (∑i = n..<Suc m. d (σ i) (σ (Suc i)))"
using "1" by force
finally show ?thesis .
qed auto
qed
with ‹n < n'› show ?thesis by blast
qed
also have "… ≤ (∑i=n..<n'. d a (f a) * k^i)"
proof (rule sum_mono)
fix i
assume "i ∈ {n..<n'}"
show "d (σ i) (σ (Suc i)) ≤ d a (f a) * k ^ i"
proof (induction i)
case 0
then show ?case
by (auto simp: σ_def)
next
case (Suc i)
have "d (σ (Suc i)) (σ (Suc (Suc i))) ≤ k * d (σ i) (σ (Suc i))"
using con σ_def f_iter fim by fastforce
also have "… ≤ d a (f a) * k ^ Suc i"
using Suc ‹0 < k› by auto
finally show ?case .
qed
qed
also have "… = d a (f a) * (∑i=n..<n'. k^i)"
by (simp add: sum_distrib_left)
also have "… = d a (f a) * (∑i=0..<n'-n. k^(i+n))"
using sum.shift_bounds_nat_ivl [of "power k" 0 n "n'-n"] ‹n < n'› by simp
also have "… = d a (f a) * k^n * (∑i<n'-n. k^i)"
by (simp add: power_add lessThan_atLeast0 flip: sum_distrib_right)
also have "… = d a (f a) * (k ^ n - k ^ n') / (1 - k)"
using ‹k < 1› ‹n < n'› apply (simp add: sum_gp_strict)
by (simp add: algebra_simps flip: power_add)
also have "… < ε"
using N ‹k < 1› ‹0 < ε› ‹0 < k› ‹N ≤ n›
apply (simp add: field_simps)
by (smt (verit) nonneg pos_less_divide_eq zero_less_divide_iff zero_less_power)
finally show "d (σ n) (σ n') < ε" .
qed
then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (σ n) (σ n') < ε"
by (metis ‹0 < ε› commute f_iter linorder_not_le local.mdist_zero nat_less_le)
qed
qed
then obtain l where l: "limitin mtopology σ l sequentially"
using ‹mcomplete› mcomplete_def by blast
show ?thesis
proof
show "l ∈ M"
using l limitin_mspace by blast
show "f l = l"
proof (rule limitin_metric_unique)
have "limitin mtopology (f ∘ σ) (f l) sequentially"
proof (rule continuous_map_limit)
have "Lipschitz_continuous_map Self Self f"
using con by (auto simp: Lipschitz_continuous_map_def fim)
then show "continuous_map mtopology mtopology f"
using Lipschitz_continuous_imp_continuous_map Self_def by force
qed (use l in auto)
moreover have "(f ∘ σ) = (λi. σ(i+1))"
by (auto simp: σ_def)
ultimately show "limitin mtopology (λn. (f^^n)a) (f l) sequentially"
using limitin_sequentially_offset_rev [of mtopology σ 1]
by (simp add: σ_def)
qed (use l in ‹auto simp: σ_def›)
qed
qed
qed
qed
subsection‹ The Baire Category Theorem›
text ‹Possibly relevant to the theorem "Baire" in Elementary Normed Spaces›
lemma (in Metric_space) metric_Baire_category:
assumes "mcomplete" "countable 𝒢"
and "⋀T. T ∈ 𝒢 ⟹ openin mtopology T ∧ mtopology closure_of T = M"
shows "mtopology closure_of ⋂𝒢 = M"
proof (cases "𝒢={}")
case False
then obtain U :: "nat ⇒ 'a set" where U: "range U = 𝒢"
by (metis ‹countable 𝒢› uncountable_def)
with assms have u_open: "⋀n. openin mtopology (U n)" and u_dense: "⋀n. mtopology closure_of (U n) = M"
by auto
have "⋂(range U) ∩ W ≠ {}" if W: "openin mtopology W" "W ≠ {}" for W
proof -
have "W ⊆ M"
using openin_mtopology W by blast
have "∃r' x'. 0 < r' ∧ r' < r/2 ∧ x' ∈ M ∧ mcball x' r' ⊆ mball x r ∩ U n"
if "r>0" "x ∈ M" for x r n
proof -
obtain z where z: "z ∈ U n ∩ mball x r"
using u_dense [of n] ‹r>0› ‹x ∈ M›
by (metis dense_intersects_open centre_in_mball_iff empty_iff openin_mball topspace_mtopology equals0I)
then have "z ∈ M" by auto
have "openin mtopology (U n ∩ mball x r)"
by (simp add: openin_Int u_open)
with ‹z ∈ M› z obtain e where "e>0" and e: "mcball z e ⊆ U n ∩ mball x r"
by (meson openin_mtopology_mcball)
define r' where "r' ≡ min e (r/4)"
show ?thesis
proof (intro exI conjI)
show "0 < r'" "r' < r / 2" "z ∈ M"
using ‹e>0› ‹r>0› ‹z ∈ M› by (auto simp: r'_def)
show "mcball z r' ⊆ mball x r ∩ U n"
using Metric_space.mcball_subset_concentric e r'_def by auto
qed
qed
then obtain nextr nextx
where nextr: "⋀r x n. ⟦r>0; x∈M⟧ ⟹ 0 < nextr r x n ∧ nextr r x n < r/2"
and nextx: "⋀r x n. ⟦r>0; x∈M⟧ ⟹ nextx r x n ∈ M"
and nextsub: "⋀r x n. ⟦r>0; x∈M⟧ ⟹ mcball (nextx r x n) (nextr r x n) ⊆ mball x r ∩ U n"
by metis
obtain x0 where x0: "x0 ∈ U 0 ∩ W"
by (metis W dense_intersects_open topspace_mtopology all_not_in_conv u_dense)
then have "x0 ∈ M"
using ‹W ⊆ M› by fastforce
obtain r0 where "0 < r0" "r0 < 1" and sub: "mcball x0 r0 ⊆ U 0 ∩ W"
proof -
have "openin mtopology (U 0 ∩ W)"
using W u_open by blast
then obtain r where "r>0" and r: "mball x0 r ⊆ U 0" "mball x0 r ⊆ W"
by (meson Int_subset_iff openin_mtopology x0)
define r0 where "r0 ≡ (min r 1) / 2"
show thesis
proof
show "0 < r0" "r0 < 1"
using ‹r>0› by (auto simp: r0_def)
show "mcball x0 r0 ⊆ U 0 ∩ W"
using r ‹0 < r0› r0_def by auto
qed
qed
define b where "b ≡ rec_nat (x0,r0) (λn (x,r). (nextx r x n, nextr r x n))"
have b0[simp]: "b 0 = (x0,r0)"
by (simp add: b_def)
have bSuc[simp]: "b (Suc n) = (let (x,r) = b n in (nextx r x n, nextr r x n))" for n
by (simp add: b_def)
define xf where "xf ≡ fst ∘ b"
define rf where "rf ≡ snd ∘ b"
have rfxf: "0 < rf n ∧ xf n ∈ M" for n
proof (induction n)
case 0
with ‹0 < r0› ‹x0 ∈ M› show ?case
by (auto simp: rf_def xf_def)
next
case (Suc n)
then show ?case
by (auto simp: rf_def xf_def case_prod_unfold nextr nextx Let_def)
qed
have mcball_sub: "mcball (xf (Suc n)) (rf (Suc n)) ⊆ mball (xf n) (rf n) ∩ U n" for n
using rfxf nextsub by (auto simp: xf_def rf_def case_prod_unfold Let_def)
have half: "rf (Suc n) < rf n / 2" for n
using rfxf nextr by (auto simp: xf_def rf_def case_prod_unfold Let_def)
then have "decseq rf"
using rfxf by (smt (verit, ccfv_threshold) decseq_SucI field_sum_of_halves)
have nested: "mball (xf n) (rf n) ⊆ mball (xf m) (rf m)" if "m ≤ n" for m n
using that
proof (induction n)
case (Suc n)
then show ?case
by (metis mcball_sub order.trans inf.boundedE le_Suc_eq mball_subset_mcball order.refl)
qed auto
have "MCauchy xf"
unfolding MCauchy_def
proof (intro conjI strip)
show "range xf ⊆ M"
using rfxf by blast
fix ε :: real
assume "ε>0"
then obtain N where N: "inverse (2^N) < ε"
using real_arch_pow_inv by (force simp flip: power_inverse)
have "d (xf n) (xf n') < ε" if "n ≤ n'" "N ≤ n" "N ≤ n'" for n n'
proof -
have *: "rf n < inverse (2 ^ n)" for n
proof (induction n)
case 0
then show ?case
by (simp add: ‹r0 < 1› rf_def)
next
case (Suc n)
with half show ?case
by simp (smt (verit))
qed
have "rf n ≤ rf N"
using ‹decseq rf› ‹N ≤ n› by (simp add: decseqD)
moreover
have "xf n' ∈ mball (xf n) (rf n)"
using nested rfxf ‹n ≤ n'› by blast
ultimately have "d (xf n) (xf n') < rf N"
by auto
also have "… < ε"
using "*" N order.strict_trans by blast
finally show ?thesis .
qed
then show "∃N. ∀n n'. N ≤ n ⟶ N ≤ n' ⟶ d (xf n) (xf n') < ε"
by (metis commute linorder_le_cases)
qed
then obtain l where l: "limitin mtopology xf l sequentially"
using ‹mcomplete› mcomplete_alt by blast
have l_in: "l ∈ mcball (xf n) (rf n)" for n
proof -
have "∀⇩F m in sequentially. xf m ∈ mcball (xf n) (rf n)"
unfolding eventually_sequentially
by (meson nested rfxf centre_in_mball_iff mball_subset_mcball subset_iff)
with l limitin_closedin show ?thesis
by (metis closedin_mcball trivial_limit_sequentially)
qed
then have "⋀n. l ∈ U n"
using mcball_sub by blast
moreover have "l ∈ W"
using l_in[of 0] sub by (auto simp: xf_def rf_def)
ultimately show ?thesis by auto
qed
with U show ?thesis
by (metis dense_intersects_open topspace_mtopology)
qed auto
lemma (in Metric_space) metric_Baire_category_alt:
assumes "mcomplete" "countable 𝒢"
and empty: "⋀T. T ∈ 𝒢 ⟹ closedin mtopology T ∧ mtopology interior_of T = {}"
shows "mtopology interior_of ⋃𝒢 = {}"
proof -
have *: "mtopology closure_of ⋂((-)M ` 𝒢) = M"
proof (intro metric_Baire_category conjI ‹mcomplete›)
show "countable ((-) M ` 𝒢)"
using ‹countable 𝒢› by blast
fix T
assume "T ∈ (-) M ` 𝒢"
then obtain U where U: "U ∈ 𝒢" "T = M-U" "U ⊆ M"
using empty metric_closedin_iff_sequentially_closed by force
with empty show "openin mtopology T" by blast
show "mtopology closure_of T = M"
using U by (simp add: closure_of_interior_of double_diff empty)
qed
with closure_of_eq show ?thesis
by (fastforce simp: interior_of_closure_of split: if_split_asm)
qed
text ‹Since all locally compact Hausdorff spaces are regular, the disjunction in the HOL Light version is redundant.›
lemma Baire_category_aux:
assumes "locally_compact_space X" "regular_space X"
and "countable 𝒢"
and empty: "⋀G. G ∈ 𝒢 ⟹ closedin X G ∧ X interior_of G = {}"
shows "X interior_of ⋃𝒢 = {}"
proof (cases "𝒢 = {}")
case True
then show ?thesis
by simp
next
case False
then obtain T :: "nat ⇒ 'a set" where T: "𝒢 = range T"
by (metis ‹countable 𝒢› uncountable_def)
with empty have Tempty: "⋀n. X interior_of (T n) = {}"
by auto
show ?thesis
proof (clarsimp simp: T interior_of_def)
fix z U
assume "z ∈ U" and opeA: "openin X U" and Asub: "U ⊆ ⋃ (range T)"
with openin_subset have "z ∈ topspace X"
by blast
have "neighbourhood_base_of (λC. compactin X C ∧ closedin X C) X"
using assms locally_compact_regular_space_neighbourhood_base by auto
then obtain V K where "openin X V" "compactin X K" "closedin X K" "z ∈ V" "V ⊆ K" "K ⊆ U"
by (metis (no_types, lifting) ‹z ∈ U› neighbourhood_base_of opeA)
have nb_closedin: "neighbourhood_base_of (closedin X) X"
using ‹regular_space X› neighbourhood_base_of_closedin by auto
have "∃Φ. ∀n. (Φ n ⊆ K ∧ closedin X (Φ n) ∧ X interior_of Φ n ≠ {} ∧ disjnt (Φ n) (T n)) ∧
Φ (Suc n) ⊆ Φ n"
proof (rule dependent_nat_choice)
show "∃x⊆K. closedin X x ∧ X interior_of x ≠ {} ∧ disjnt x (T 0)"
proof -
have False if "V ⊆ T 0"
using Tempty ‹openin X V› ‹z ∈ V› interior_of_maximal that by fastforce
then obtain x where "openin X (V - T 0) ∧ x ∈ V - T 0"
using T ‹openin X V› empty by blast
with nb_closedin
obtain N C where "openin X N" "closedin X C" "x ∈ N" "N ⊆ C" "C ⊆ V - T 0"
unfolding neighbourhood_base_of by metis
show ?thesis
proof (intro exI conjI)
show "C ⊆ K"
using ‹C ⊆ V - T 0› ‹V ⊆ K› by auto
show "X interior_of C ≠ {}"
by (metis ‹N ⊆ C› ‹openin X N› ‹x ∈ N› empty_iff interior_of_eq_empty)
show "disjnt C (T 0)"
using ‹C ⊆ V - T 0› disjnt_iff by fastforce
qed (use ‹closedin X C› in auto)
qed
show "∃L. (L ⊆ K ∧ closedin X L ∧ X interior_of L ≠ {} ∧ disjnt L (T (Suc n))) ∧ L ⊆ C"
if §: "C ⊆ K ∧ closedin X C ∧ X interior_of C ≠ {} ∧ disjnt C (T n)"
for C n
proof -
have False if "X interior_of C ⊆ T (Suc n)"
by (metis Tempty interior_of_eq_empty § openin_interior_of that)
then obtain x where "openin X (X interior_of C - T (Suc n)) ∧ x ∈ X interior_of C - T (Suc n)"
using T empty by fastforce
with nb_closedin
obtain N D where "openin X N" "closedin X D" "x ∈ N" "N ⊆ D" and D: "D ⊆ X interior_of C - T(Suc n)"
unfolding neighbourhood_base_of by metis
show ?thesis
proof (intro conjI exI)
show "D ⊆ K"
using D interior_of_subset § by fastforce
show "X interior_of D ≠ {}"
by (metis ‹N ⊆ D› ‹openin X N› ‹x ∈ N› empty_iff interior_of_eq_empty)
show "disjnt D (T (Suc n))"
using D disjnt_iff by fastforce
show "D ⊆ C"
using interior_of_subset [of X C] D by blast
qed (use ‹closedin X D› in auto)
qed
qed
then obtain Φ where Φ: "⋀n. Φ n ⊆ K ∧ closedin X (Φ n) ∧ X interior_of Φ n ≠ {} ∧ disjnt (Φ n) (T n)"
and "⋀n. Φ (Suc n) ⊆ Φ n" by metis
then have "decseq Φ"
by (simp add: decseq_SucI)
moreover have "⋀n. Φ n ≠ {}"
by (metis Φ bot.extremum_uniqueI interior_of_subset)
ultimately have "⋂(range Φ) ≠ {}"
by (metis Φ compact_space_imp_nest ‹compactin X K› compactin_subspace closedin_subset_topspace)
moreover have "U ⊆ {y. ∃x. y ∈ T x}"
using Asub by auto
with T have "{a. ∀n. a ∈ Φ n} ⊆ {}"
by (smt (verit) Asub Φ Collect_empty_eq UN_iff ‹K ⊆ U› disjnt_iff subset_iff)
ultimately show False
by blast
qed
qed
lemma Baire_category_alt:
assumes "completely_metrizable_space X ∨ locally_compact_space X ∧ regular_space X"
and "countable 𝒢"
and "⋀T. T ∈ 𝒢 ⟹ closedin X T ∧ X interior_of T = {}"
shows "X interior_of ⋃𝒢 = {}"
using Baire_category_aux [of X 𝒢] Metric_space.metric_Baire_category_alt
by (metis assms completely_metrizable_space_def)
lemma Baire_category:
assumes "completely_metrizable_space X ∨ locally_compact_space X ∧ regular_space X"
and "countable 𝒢"
and top: "⋀T. T ∈ 𝒢 ⟹ openin X T ∧ X closure_of T = topspace X"
shows "X closure_of ⋂𝒢 = topspace X"
proof (cases "𝒢={}")
case False
have *: "X interior_of ⋃((-)(topspace X) ` 𝒢) = {}"
proof (intro Baire_category_alt conjI assms)
show "countable ((-) (topspace X) ` 𝒢)"
using assms by blast
fix T
assume "T ∈ (-) (topspace X) ` 𝒢"
then obtain U where U: "U ∈ 𝒢" "T = (topspace X) - U" "U ⊆ (topspace X)"
by (meson top image_iff openin_subset)
then show "closedin X T"
by (simp add: closedin_diff top)
show "X interior_of T = {}"
using U top by (simp add: interior_of_closure_of double_diff)
qed
then show ?thesis
by (simp add: closure_of_eq_topspace interior_of_complement)
qed auto
subsection‹Sierpinski-Hausdorff type results about countable closed unions›
lemma locally_connected_not_countable_closed_union:
assumes "topspace X ≠ {}" and csX: "connected_space X"
and lcX: "locally_connected_space X"
and X: "completely_metrizable_space X ∨ locally_compact_space X ∧ Hausdorff_space X"
and "countable 𝒰" and pwU: "pairwise disjnt 𝒰"
and clo: "⋀C. C ∈ 𝒰 ⟹ closedin X C ∧ C ≠ {}"
and UU_eq: "⋃𝒰 = topspace X"
shows "𝒰 = {topspace X}"
proof -
define 𝒱 where "𝒱 ≡ (frontier_of) X ` 𝒰"
define B where "B ≡ ⋃𝒱"
then have Bsub: "B ⊆ topspace X"
by (simp add: Sup_le_iff 𝒱_def closedin_frontier_of closedin_subset)
have allsub: "A ⊆ topspace X" if "A ∈ 𝒰" for A
by (meson clo closedin_def that)
show ?thesis
proof (rule ccontr)
assume "𝒰 ≠ {topspace X}"
with assms have "∃A∈𝒰. ¬ (closedin X A ∧ openin X A)"
by (metis Union_empty connected_space_clopen_in singletonI subsetI subset_singleton_iff)
then have "B ≠ {}"
by (auto simp: B_def 𝒱_def frontier_of_eq_empty allsub)
moreover
have "subtopology X B interior_of B = B"
by (simp add: Bsub interior_of_openin openin_subtopology_refl)
ultimately have int_B_nonempty: "subtopology X B interior_of B ≠ {}"
by auto
have "subtopology X B interior_of ⋃𝒱 = {}"
proof (intro Baire_category_alt conjI)
have "⋃𝒰 ⊆ B ∪ ⋃((interior_of) X ` 𝒰)"
using clo closure_of_closedin by (fastforce simp: B_def 𝒱_def frontier_of_def)
moreover have "B ∪ ⋃((interior_of) X ` 𝒰) ⊆ ⋃𝒰"
using allsub clo frontier_of_subset_eq interior_of_subset by (fastforce simp: B_def 𝒱_def )
moreover have "disjnt B (⋃((interior_of) X ` 𝒰))"
using pwU
apply (clarsimp simp: B_def 𝒱_def frontier_of_def pairwise_def disjnt_iff)
by (metis clo closure_of_eq interior_of_subset subsetD)
ultimately have "B = topspace X - ⋃ ((interior_of) X ` 𝒰)"
by (auto simp: UU_eq disjnt_iff)
then have "closedin X B"
by fastforce
with X show "completely_metrizable_space (subtopology X B) ∨ locally_compact_space (subtopology X B) ∧ regular_space (subtopology X B)"
by (metis completely_metrizable_space_closedin locally_compact_Hausdorff_or_regular
locally_compact_space_closed_subset regular_space_subtopology)
show "countable 𝒱"
by (simp add: 𝒱_def ‹countable 𝒰›)
fix V
assume "V ∈ 𝒱"
then obtain S where S: "S ∈ 𝒰" "V = X frontier_of S"
by (auto simp: 𝒱_def)
show "closedin (subtopology X B) V"
by (metis B_def Sup_upper 𝒱_def ‹V ∈ 𝒱› closedin_frontier_of closedin_subset_topspace image_iff)
have "subtopology X B interior_of (X frontier_of S) = {}"
proof (clarsimp simp: interior_of_def openin_subtopology_alt)
fix a U
assume "a ∈ B" "a ∈ U" and opeU: "openin X U" and BUsub: "B ∩ U ⊆ X frontier_of S"
then have "a ∈ S"
by (meson IntI ‹S ∈ 𝒰› clo frontier_of_subset_closedin subsetD)
then obtain W C where "openin X W" "connectedin X C" "a ∈ W" "W ⊆ C" "C ⊆ U"
by (metis ‹a ∈ U› lcX locally_connected_space opeU)
have "W ∩ X frontier_of S ≠ {}"
using ‹B ∩ U ⊆ X frontier_of S› ‹a ∈ B› ‹a ∈ U› ‹a ∈ W› by auto
with frontier_of_openin_straddle_Int
obtain "W ∩ S ≠ {}" "W - S ≠ {}" "W ⊆ topspace X"
using ‹openin X W› by (metis openin_subset)
then obtain b where "b ∈ topspace X" "b ∈ W-S"
by blast
with UU_eq obtain T where "T ∈ 𝒰" "T ≠ S" "W ∩ T ≠ {}"
by auto
then have "disjnt S T"
by (metis ‹S ∈ 𝒰› pairwise_def pwU)
then have "C - T ≠ {}"
by (meson Diff_eq_empty_iff ‹W ⊆ C› ‹a ∈ S› ‹a ∈ W› disjnt_iff subsetD)
then have "C ∩ X frontier_of T ≠ {}"
using ‹W ∩ T ≠ {}› ‹W ⊆ C› ‹connectedin X C› connectedin_Int_frontier_of by blast
moreover have "C ∩ X frontier_of T = {}"
proof -
have "X frontier_of S ⊆ S" "X frontier_of T ⊆ T"
using frontier_of_subset_closedin ‹S ∈ 𝒰› ‹T ∈ 𝒰› clo by blast+
moreover have "X frontier_of T ∪ B = B"
using B_def 𝒱_def ‹T ∈ 𝒰› by blast
ultimately show ?thesis
using BUsub ‹C ⊆ U› ‹disjnt S T› unfolding disjnt_def by blast
qed
ultimately show False
by simp
qed
with S show "subtopology X B interior_of V = {}"
by meson
qed
then show False
using B_def int_B_nonempty by blast
qed
qed
lemma real_Sierpinski_lemma:
fixes a b::real
assumes "a ≤ b"
and "countable 𝒰" and pwU: "pairwise disjnt 𝒰"
and clo: "⋀C. C ∈ 𝒰 ⟹ closed C ∧ C ≠ {}"
and "⋃𝒰 = {a..b}"
shows "𝒰 = {{a..b}}"
proof -
have "locally_connected_space (top_of_set {a..b})"
by (simp add: locally_connected_real_interval)
moreover
have "completely_metrizable_space (top_of_set {a..b})"
by (metis box_real(2) completely_metrizable_space_cbox)
ultimately
show ?thesis
using locally_connected_not_countable_closed_union [of "subtopology euclidean {a..b}"] assms
apply (simp add: closedin_subtopology)
by (metis Union_upper inf.orderE)
qed
subsection‹The Tychonoff embedding›
lemma completely_regular_space_cube_embedding_explicit:
assumes "completely_regular_space X" "Hausdorff_space X"
shows "embedding_map X
(product_topology (λf. top_of_set {0..1::real})
(mspace (submetric (cfunspace X euclidean_metric)
{f. f ∈ topspace X → {0..1}})) )
(λx. λf ∈ mspace (submetric (cfunspace X euclidean_metric) {f. f ∈ topspace X → {0..1}}).
f x)"
proof -
define K where "K ≡ mspace(submetric (cfunspace X euclidean_metric) {f. f ∈ topspace X → {0..1::real}})"
define e where "e ≡ λx. λf∈K. f x"
have "e x ≠ e y" if xy: "x ≠ y" "x ∈ topspace X" "y ∈ topspace X" for x y
proof -
have "closedin X {x}"
by (simp add: ‹Hausdorff_space X› closedin_Hausdorff_singleton ‹x ∈ topspace X›)
then obtain f where contf: "continuous_map X euclideanreal f"
and f01: "⋀x. x ∈ topspace X ⟹ f x ∈ {0..1}" and fxy: "f y = 0" "f x = 1"
using ‹completely_regular_space X› xy unfolding completely_regular_space_def
by (smt (verit, ccfv_threshold) Diff_iff continuous_map_in_subtopology image_subset_iff singleton_iff)
then have "bounded (f ` topspace X)"
by (meson bounded_closed_interval bounded_subset image_subset_iff)
with contf f01 have "restrict f (topspace X) ∈ K"
by (auto simp: K_def)
with fxy xy show ?thesis
unfolding e_def by (metis restrict_apply' zero_neq_one)
qed
then have "inj_on e (topspace X)"
by (meson inj_onI)
then obtain e' where e': "⋀x. x ∈ topspace X ⟹ e' (e x) = x"
by (metis inv_into_f_f)
have "continuous_map (subtopology (product_topology (λf. top_of_set {0..1}) K) (e ` topspace X)) X e'"
proof (clarsimp simp add: continuous_map_atin limitin_atin openin_subtopology_alt e')
fix x U
assume "e x ∈ K →⇩E {0..1}" and "x ∈ topspace X" and "openin X U" and "x ∈ U"
then obtain g where contg: "continuous_map X (top_of_set {0..1}) g" and "g x = 0"
and gim: "g ` (topspace X - U) ⊆ {1::real}"
using ‹completely_regular_space X› unfolding completely_regular_space_def
by (metis Diff_iff openin_closedin_eq)
then have "bounded (g ` topspace X)"
by (meson bounded_closed_interval bounded_subset continuous_map_in_subtopology)
moreover have "g ∈ topspace X → {0..1}"
using contg by (simp add: continuous_map_def)
ultimately have g_in_K: "restrict g (topspace X) ∈ K"
using contg by (force simp add: K_def continuous_map_in_subtopology)
have "openin (top_of_set {0..1}) {0..<1::real}"
using open_real_greaterThanLessThan[of "-1" 1] by (force simp: openin_open)
moreover have "e x ∈ (Π⇩E f∈K. if f = restrict g (topspace X) then {0..<1} else {0..1})"
using ‹e x ∈ K →⇩E {0..1}› by (simp add: e_def ‹g x = 0› ‹x ∈ topspace X› PiE_iff)
moreover have "e y = e x"
if "y ∉ U" and ey: "e y ∈ (Π⇩E f∈K. if f = restrict g (topspace X) then {0..<1} else {0..1})"
and y: "y ∈ topspace X" for y
proof -
have "e y (restrict g (topspace X)) ∈ {0..<1}"
using ey by (smt (verit, ccfv_SIG) PiE_mem g_in_K)
with gim g_in_K y ‹y ∉ U› show ?thesis
by (fastforce simp: e_def)
qed
ultimately
show "∃W. openin (product_topology (λf. top_of_set {0..1}) K) W ∧ e x ∈ W ∧ e' ` (e ` topspace X ∩ W - {e x}) ⊆ U"
apply (rule_tac x="PiE K (λf. if f = restrict g (topspace X) then {0..<1} else {0..1})" in exI)
by (auto simp: openin_PiE_gen e')
qed
with e' have "embedding_map X (product_topology (λf. top_of_set {0..1}) K) e"
unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def
by (fastforce simp: e_def K_def continuous_map_in_subtopology continuous_map_componentwise)
then show ?thesis
by (simp add: K_def e_def)
qed
lemma completely_regular_space_cube_embedding:
fixes X :: "'a topology"
assumes "completely_regular_space X" "Hausdorff_space X"
obtains K:: "('a⇒real)set" and e
where "embedding_map X (product_topology (λf. top_of_set {0..1::real}) K) e"
using completely_regular_space_cube_embedding_explicit [OF assms] by metis
subsection ‹Urysohn and Tietze analogs for completely regular spaces›
text‹"Urysohn and Tietze analogs for completely regular spaces if (()) set is
assumed compact instead of closed. Note that Hausdorffness is *not* required: inside () proof
we factor through the Kolmogorov quotient." -- John Harrison›
lemma Urysohn_completely_regular_closed_compact:
fixes a b::real
assumes "a ≤ b" "completely_regular_space X" "closedin X S" "compactin X T" "disjnt S T"
obtains f where "continuous_map X (subtopology euclidean {a..b}) f" "f ` T ⊆ {a}" "f ` S ⊆ {b}"
proof -
obtain f where contf: "continuous_map X (subtopology euclideanreal {0..1}) f"
and f0: "f ` T ⊆ {0}" and f1: "f ` S ⊆ {1}"
proof (cases "T={}")
case True
show thesis
proof
show "continuous_map X (top_of_set {0..1}) (λx. 1::real)" "(λx. 1::real) ` T ⊆ {0}" "(λx. 1::real) ` S ⊆ {1}"
using True by auto
qed
next
case False
have "⋀t. t ∈ T ⟹ ∃f. continuous_map X (subtopology euclideanreal ({0..1})) f ∧ f t = 0 ∧ f ` S ⊆ {1}"
using assms unfolding completely_regular_space_def
by (meson DiffI compactin_subset_topspace disjnt_iff subset_eq)
then obtain g where contg: "⋀t. t ∈ T ⟹ continuous_map X (subtopology euclideanreal {0..1}) (g t)"
and g0: "⋀t. t ∈ T ⟹ g t t = 0"
and g1: "⋀t. t ∈ T ⟹ g t ` S ⊆ {1}"
by metis
then have g01: "⋀t. t ∈ T ⟹ g t ` topspace X ⊆ {0..1}"
by (meson continuous_map_in_subtopology)
define G where "G ≡ λt. {x ∈ topspace X. g t x ∈ {..<1/2}}"
have "Ball (G`T) (openin X)"
using contg unfolding G_def continuous_map_in_subtopology
by (smt (verit, best) Collect_cong openin_continuous_map_preimage image_iff open_lessThan open_openin)
moreover have "T ⊆ ⋃(G`T)"
using ‹compactin X T› g0 compactin_subset_topspace by (force simp: G_def)
ultimately have "∃ℱ. finite ℱ ∧ ℱ ⊆ G`T ∧ T ⊆ ⋃ ℱ"
using ‹compactin X T› unfolding compactin_def by blast
then obtain K where K: "finite K" "K ⊆ T" "T ⊆ ⋃(G`K)"
by (metis finite_subset_image)
with False have "K ≠ {}"
by fastforce
define f where "f ≡ λx. 2 * max 0 (Inf ((λt. g t x) ` K) - 1/2)"
have [simp]: "max 0 (x - 1/2) = 0 ⟷ x ≤ 1/2" for x::real
by force
have [simp]: "2 * max 0 (x - 1/2) = 1 ⟷ x = 1" for x::real
by (simp add: max_def_raw)
show thesis
proof
have "g t s = 1" if "s ∈ S" "t ∈ K" for s t
using ‹K ⊆ T› g1 that by auto
then show "f ` S ⊆ {1}"
using ‹K ≠ {}› by (simp add: f_def image_subset_iff)
have "(INF t∈K. g t x) ≤ 1/2" if "x ∈ T" for x
proof -
obtain k where "k ∈ K" "g k x < 1/2"
using K ‹x ∈ T› by (auto simp: G_def)
then show ?thesis
by (meson ‹finite K› cInf_le_finite dual_order.trans finite_imageI imageI less_le_not_le)
qed
then show "f ` T ⊆ {0}"
by (force simp: f_def)
have "⋀t. t ∈ K ⟹ continuous_map X euclideanreal (g t)"
using ‹K ⊆ T› contg continuous_map_in_subtopology by blast
moreover have "2 * max 0 ((INF t∈K. g t x) - 1/2) ≤ 1" if "x ∈ topspace X" for x
proof -
obtain k where "k ∈ K" "g k x ≤ 1"
using K ‹x ∈ topspace X› ‹K ≠ {}› g01 by (fastforce simp: G_def)
then have "(INF t∈K. g t x) ≤ 1"
by (meson ‹finite K› cInf_le_finite dual_order.trans finite_imageI imageI)
then show ?thesis
by (simp add: max_def_raw)
qed
ultimately show "continuous_map X (top_of_set {0..1}) f"
by (force simp: f_def continuous_map_in_subtopology intro!: ‹finite K› continuous_intros)
qed
qed
define g where "g ≡ λx. a + (b - a) * f x"
show thesis
proof
have "∀x∈topspace X. a + (b - a) * f x ≤ b"
using contf ‹a ≤ b› apply (simp add: continuous_map_in_subtopology image_subset_iff)
by (smt (verit, best) mult_right_le_one_le)
then show "continuous_map X (top_of_set {a..b}) g"
using contf ‹a ≤ b› unfolding g_def continuous_map_in_subtopology image_subset_iff
by (intro conjI continuous_intros; simp)
show "g ` T ⊆ {a}" "g ` S ⊆ {b}"
using f0 f1 by (auto simp: g_def)
qed
qed
lemma Urysohn_completely_regular_compact_closed:
fixes a b::real
assumes "a ≤ b" "completely_regular_space X" "compactin X S" "closedin X T" "disjnt S T"
obtains f where "continuous_map X (subtopology euclidean {a..b}) f" "f ` T ⊆ {a}" "f ` S ⊆ {b}"
proof -
obtain f where contf: "continuous_map X (subtopology euclidean {-b..-a}) f" and fim: "f ` T ⊆ {-a}" "f ` S ⊆ {-b}"
by (meson Urysohn_completely_regular_closed_compact assms disjnt_sym neg_le_iff_le)
show thesis
proof
show "continuous_map X (top_of_set {a..b}) (uminus ∘ f)"
using contf by (auto simp: continuous_map_in_subtopology o_def)
show "(uminus o f) ` T ⊆ {a}" "(uminus o f) ` S ⊆ {b}"
using fim by fastforce+
qed
qed
lemma Urysohn_completely_regular_compact_closed_alt:
fixes a b::real
assumes "completely_regular_space X" "compactin X S" "closedin X T" "disjnt S T"
obtains f where "continuous_map X euclideanreal f" "f ` T ⊆ {a}" "f ` S ⊆ {b}"
proof (cases a b rule: le_cases)
case le
then show ?thesis
by (meson Urysohn_completely_regular_compact_closed assms continuous_map_into_fulltopology that)
next
case ge
then show ?thesis
using Urysohn_completely_regular_closed_compact assms
by (metis Urysohn_completely_regular_closed_compact assms continuous_map_into_fulltopology disjnt_sym that)
qed
lemma Tietze_extension_comp_reg_aux:
fixes T :: "real set"
assumes "completely_regular_space X" "Hausdorff_space X" "compactin X S"
and T: "is_interval T" "T≠{}"
and contf: "continuous_map (subtopology X S) euclidean f" and fim: "f`S ⊆ T"
obtains g where "continuous_map X euclidean g" "g ` topspace X ⊆ T" "⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain K:: "('a⇒real)set" and e
where e0: "embedding_map X (product_topology (λf. top_of_set {0..1::real}) K) e"
using assms completely_regular_space_cube_embedding by blast
define cube where "cube ≡ product_topology (λf. top_of_set {0..1::real}) K"
have e: "embedding_map X cube e"
using e0 by (simp add: cube_def)
obtain e' where e': "homeomorphic_maps X (subtopology cube (e ` topspace X)) e e'"
using e by (force simp: cube_def embedding_map_def homeomorphic_map_maps)
then have conte: "continuous_map X (subtopology cube (e ` topspace X)) e"
and conte': "continuous_map (subtopology cube (e ` topspace X)) X e'"
and e'e: "∀x∈topspace X. e' (e x) = x"
by (auto simp: homeomorphic_maps_def)
have "Hausdorff_space cube"
unfolding cube_def
using Hausdorff_space_euclidean Hausdorff_space_product_topology Hausdorff_space_subtopology by blast
have "normal_space cube"
proof (rule compact_Hausdorff_or_regular_imp_normal_space)
show "compact_space cube"
unfolding cube_def
using compact_space_product_topology compact_space_subtopology compactin_euclidean_iff by blast
qed (use ‹Hausdorff_space cube› in auto)
moreover
have comp: "compactin cube (e ` S)"
by (meson ‹compactin X S› conte continuous_map_in_subtopology image_compactin)
then have "closedin cube (e ` S)"
by (intro compactin_imp_closedin ‹Hausdorff_space cube›)
moreover
have "continuous_map (subtopology cube (e ` S)) euclideanreal (f ∘ e')"
proof (intro continuous_map_compose)
show "continuous_map (subtopology cube (e ` S)) (subtopology X S) e'"
unfolding continuous_map_in_subtopology
proof
show "continuous_map (subtopology cube (e ` S)) X e'"
by (meson ‹compactin X S› compactin_subset_topspace conte' continuous_map_from_subtopology_mono image_mono)
show "e' ` topspace (subtopology cube (e ` S)) ⊆ S"
using ‹compactin X S› compactin_subset_topspace e'e by fastforce
qed
qed (simp add: contf)
moreover
have "(f ∘ e') ` e ` S ⊆ T"
using ‹compactin X S› compactin_subset_topspace e'e fim by fastforce
ultimately
obtain g where contg: "continuous_map cube euclidean g" and gsub: "g ` topspace cube ⊆ T"
and gf: "⋀x. x ∈ e`S ⟹ g x = (f ∘ e') x"
using Tietze_extension_realinterval T by metis
show thesis
proof
show "continuous_map X euclideanreal (g ∘ e)"
by (meson contg conte continuous_map_compose continuous_map_in_subtopology)
show "(g ∘ e) ` topspace X ⊆ T"
using gsub conte continuous_map_image_subset_topspace by fastforce
fix x
assume "x ∈ S"
then show "(g ∘ e) x = f x"
using gf ‹compactin X S› compactin_subset_topspace e'e by fastforce
qed
qed
lemma Tietze_extension_completely_regular:
assumes "completely_regular_space X" "compactin X S" "is_interval T" "T ≠ {}"
and contf: "continuous_map (subtopology X S) euclidean f" and fim: "f`S ⊆ T"
obtains g where "continuous_map X euclideanreal g" "g ` topspace X ⊆ T"
"⋀x. x ∈ S ⟹ g x = f x"
proof -
define Q where "Q ≡ Kolmogorov_quotient X ` (topspace X)"
obtain g where contg: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) euclidean g"
and gf: "⋀x. x ∈ S ⟹ g(Kolmogorov_quotient X x) = f x"
using Kolmogorov_quotient_lift_exists
by (metis ‹compactin X S› contf compactin_subset_topspace open_openin t0_space_def t1_space)
have "S ⊆ topspace X"
by (simp add: ‹compactin X S› compactin_subset_topspace)
then have [simp]: "Q ∩ Kolmogorov_quotient X ` S = Kolmogorov_quotient X ` S"
using Q_def by blast
have creg: "completely_regular_space (subtopology X Q)"
by (simp add: ‹completely_regular_space X› completely_regular_space_subtopology)
then have "regular_space (subtopology X Q)"
by (simp add: completely_regular_imp_regular_space)
then have "Hausdorff_space (subtopology X Q)"
using Q_def regular_t0_eq_Hausdorff_space t0_space_Kolmogorov_quotient by blast
moreover
have "compactin (subtopology X Q) (Kolmogorov_quotient X ` S)"
by (metis Q_def ‹compactin X S› image_compactin quotient_imp_continuous_map quotient_map_Kolmogorov_quotient)
ultimately obtain h where conth: "continuous_map (subtopology X Q) euclidean h"
and him: "h ` topspace (subtopology X Q) ⊆ T"
and hg: "⋀x. x ∈ Kolmogorov_quotient X ` S ⟹ h x = g x"
using Tietze_extension_comp_reg_aux [of "subtopology X Q" "Kolmogorov_quotient X ` S" T g]
apply (simp add: subtopology_subtopology creg contg assms)
using fim gf by blast
show thesis
proof
show "continuous_map X euclideanreal (h ∘ Kolmogorov_quotient X)"
by (metis Q_def conth continuous_map_compose quotient_imp_continuous_map quotient_map_Kolmogorov_quotient)
show "(h ∘ Kolmogorov_quotient X) ` topspace X ⊆ T"
using Q_def continuous_map_Kolmogorov_quotient continuous_map_image_subset_topspace him by fastforce
fix x
assume "x ∈ S" then show "(h ∘ Kolmogorov_quotient X) x = f x"
by (simp add: gf hg)
qed
qed
subsection‹Size bounds on connected or path-connected spaces›
lemma connected_space_imp_card_ge_alt:
assumes "connected_space X" "completely_regular_space X" "closedin X S" "S ≠ {}" "S ≠ topspace X"
shows "(UNIV::real set) ≲ topspace X"
proof -
have "S ⊆ topspace X"
using ‹closedin X S› closedin_subset by blast
then obtain a where "a ∈ topspace X" "a ∉ S"
using ‹S ≠ topspace X› by blast
have "(UNIV::real set) ≲ {0..1::real}"
using eqpoll_real_subset
by (meson atLeastAtMost_iff eqpoll_imp_lepoll eqpoll_sym less_eq_real_def zero_less_one)
also have "… ≲ topspace X"
proof -
obtain f where contf: "continuous_map X euclidean f"
and fim: "f ∈ (topspace X) → {0..1::real}"
and f0: "f a = 0" and f1: "f ` S ⊆ {1}"
using ‹completely_regular_space X›
unfolding completely_regular_space_def
by (metis Diff_iff ‹a ∈ topspace X› ‹a ∉ S› ‹closedin X S› continuous_map_in_subtopology image_subset_iff_funcset)
have "∃y∈topspace X. x = f y" if "0 ≤ x" and "x ≤ 1" for x
proof -
have "connectedin euclidean (f ` topspace X)"
using ‹connected_space X› connectedin_continuous_map_image connectedin_topspace contf by blast
moreover have "∃y. 0 = f y ∧ y ∈ topspace X"
using ‹a ∈ topspace X› f0 by auto
moreover have "∃y. 1 = f y ∧ y ∈ topspace X"
using ‹S ⊆ topspace X› ‹S ≠ {}› f1 by fastforce
ultimately show ?thesis
using that by (fastforce simp: is_interval_1 simp flip: is_interval_connected_1)
qed
then show ?thesis
unfolding lepoll_iff using atLeastAtMost_iff by blast
qed
finally show ?thesis .
qed
lemma connected_space_imp_card_ge_gen:
assumes "connected_space X" "normal_space X" "closedin X S" "closedin X T" "S ≠ {}" "T ≠ {}" "disjnt S T"
shows "(UNIV::real set) ≲ topspace X"
proof -
have "(UNIV::real set) ≲ {0..1::real}"
by (metis atLeastAtMost_iff eqpoll_real_subset eqpoll_imp_lepoll eqpoll_sym less_le_not_le zero_less_one)
also have "…≲ topspace X"
proof -
obtain f where contf: "continuous_map X euclidean f"
and fim: "f ∈ (topspace X) → {0..1::real}"
and f0: "f ` S ⊆ {0}" and f1: "f ` T ⊆ {1}"
using assms by (metis continuous_map_in_subtopology normal_space_iff_Urysohn image_subset_iff_funcset)
have "∃y∈topspace X. x = f y" if "0 ≤ x" and "x ≤ 1" for x
proof -
have "connectedin euclidean (f ` topspace X)"
using ‹connected_space X› connectedin_continuous_map_image connectedin_topspace contf by blast
moreover have "∃y. 0 = f y ∧ y ∈ topspace X"
using ‹closedin X S› ‹S ≠ {}› closedin_subset f0 by fastforce
moreover have "∃y. 1 = f y ∧ y ∈ topspace X"
using ‹closedin X T› ‹T ≠ {}› closedin_subset f1 by fastforce
ultimately show ?thesis
using that by (fastforce simp: is_interval_1 simp flip: is_interval_connected_1)
qed
then show ?thesis
unfolding lepoll_iff using atLeastAtMost_iff by blast
qed
finally show ?thesis .
qed
lemma connected_space_imp_card_ge:
assumes "connected_space X" "normal_space X" "t1_space X" and nosing: "¬ (∃a. topspace X ⊆ {a})"
shows "(UNIV::real set) ≲ topspace X"
proof -
obtain a b where "a ∈ topspace X" "b ∈ topspace X" "a ≠ b"
by (metis nosing singletonI subset_iff)
then have "{a} ≠ topspace X"
by force
with connected_space_imp_card_ge_alt assms show ?thesis
by (metis ‹a ∈ topspace X› closedin_t1_singleton insert_not_empty normal_imp_completely_regular_space_A)
qed
lemma connected_space_imp_infinite_gen:
"⟦connected_space X; t1_space X; ∄a. topspace X ⊆ {a}⟧ ⟹ infinite(topspace X)"
by (metis connected_space_discrete_topology finite_t1_space_imp_discrete_topology)
lemma connected_space_imp_infinite:
"⟦connected_space X; Hausdorff_space X; ∄a. topspace X ⊆ {a}⟧ ⟹ infinite(topspace X)"
by (simp add: Hausdorff_imp_t1_space connected_space_imp_infinite_gen)
lemma connected_space_imp_infinite_alt:
assumes "connected_space X" "regular_space X" "closedin X S" "S ≠ {}" "S ≠ topspace X"
shows "infinite(topspace X)"
proof -
have "S ⊆ topspace X"
using ‹closedin X S› closedin_subset by blast
then obtain a where a: "a ∈ topspace X" "a ∉ S"
using ‹S ≠ topspace X› by blast
have "∃Φ. ∀n. (disjnt (Φ n) S ∧ a ∈ Φ n ∧ openin X (Φ n)) ∧ Φ(Suc n) ⊂ Φ n"
proof (rule dependent_nat_choice)
show "∃T. disjnt T S ∧ a ∈ T ∧ openin X T"
by (metis Diff_iff a ‹closedin X S› closedin_def disjnt_iff)
fix V n
assume §: "disjnt V S ∧ a ∈ V ∧ openin X V"
then obtain U C where U: "openin X U" "closedin X C" "a ∈ U" "U ⊆ C" "C ⊆ V"
using ‹regular_space X› by (metis neighbourhood_base_of neighbourhood_base_of_closedin)
with assms have "U ⊂ V"
by (metis "§" ‹S ⊆ topspace X› connected_space_clopen_in disjnt_def empty_iff inf.absorb_iff2 inf.orderE psubsetI subset_trans)
with U show "∃U. (disjnt U S ∧ a ∈ U ∧ openin X U) ∧ U ⊂ V"
using "§" disjnt_subset1 by blast
qed
then obtain Φ where Φ: "⋀n. disjnt (Φ n) S ∧ a ∈ Φ n ∧ openin X (Φ n)"
and Φsub: "⋀n. Φ (Suc n) ⊂ Φ n" by metis
then have "decseq Φ"
by (simp add: decseq_SucI psubset_eq)
have "∀n. ∃x. x ∈ Φ n ∧ x ∉ Φ(Suc n)"
by (meson Φsub psubsetE subsetI)
then obtain f where fin: "⋀n. f n ∈ Φ n" and fout: "⋀n. f n ∉ Φ(Suc n)"
by metis
have "range f ⊆ topspace X"
by (meson Φ fin image_subset_iff openin_subset subset_iff)
moreover have "inj f"
by (metis Suc_le_eq ‹decseq Φ› decseq_def fin fout linorder_injI subsetD)
ultimately show ?thesis
using infinite_iff_countable_subset by blast
qed
lemma path_connected_space_imp_card_ge:
assumes "path_connected_space X" "Hausdorff_space X" and nosing: "¬ (∃x. topspace X ⊆ {x})"
shows "(UNIV::real set) ≲ topspace X"
proof -
obtain a b where "a ∈ topspace X" "b ∈ topspace X" "a ≠ b"
by (metis nosing singletonI subset_iff)
then obtain γ where γ: "pathin X γ" "γ 0 = a" "γ 1 = b"
by (meson ‹a ∈ topspace X› ‹b ∈ topspace X› ‹path_connected_space X› path_connected_space_def)
let ?Y = "subtopology X (γ ` (topspace (subtopology euclidean {0..1})))"
have "(UNIV::real set) ≲ topspace ?Y"
proof (intro compact_Hausdorff_or_regular_imp_normal_space connected_space_imp_card_ge)
show "connected_space ?Y"
using ‹pathin X γ› connectedin_def connectedin_path_image by auto
show "Hausdorff_space ?Y ∨ regular_space ?Y"
using Hausdorff_space_subtopology ‹Hausdorff_space X› by blast
show "t1_space ?Y"
using Hausdorff_imp_t1_space ‹Hausdorff_space X› t1_space_subtopology by blast
show "compact_space ?Y"
by (simp add: ‹pathin X γ› compact_space_subtopology compactin_path_image)
have "a ∈ topspace ?Y" "b ∈ topspace ?Y"
using γ pathin_subtopology by fastforce+
with ‹a ≠ b› show "∄x. topspace ?Y ⊆ {x}"
by blast
qed
also have "… ≲ γ ` {0..1}"
by (simp add: subset_imp_lepoll)
also have "… ≲ topspace X"
by (meson γ path_image_subset_topspace subset_imp_lepoll image_subset_iff_funcset)
finally show ?thesis .
qed
lemma connected_space_imp_uncountable:
assumes "connected_space X" "regular_space X" "Hausdorff_space X" "¬ (∃a. topspace X ⊆ {a})"
shows "¬ countable(topspace X)"
proof
assume coX: "countable (topspace X)"
with ‹regular_space X› have "normal_space X"
using countable_imp_Lindelof_space regular_Lindelof_imp_normal_space by blast
then have "(UNIV::real set) ≲ topspace X"
by (simp add: Hausdorff_imp_t1_space assms connected_space_imp_card_ge)
with coX show False
using countable_lepoll uncountable_UNIV_real by blast
qed
lemma path_connected_space_imp_uncountable:
assumes "path_connected_space X" "t1_space X" and nosing: "¬ (∃a. topspace X ⊆ {a})"
shows "¬ countable(topspace X)"
proof
assume coX: "countable (topspace X)"
obtain a b where "a ∈ topspace X" "b ∈ topspace X" "a ≠ b"
by (metis nosing singletonI subset_iff)
then obtain γ where "pathin X γ" "γ 0 = a" "γ 1 = b"
by (meson ‹a ∈ topspace X› ‹b ∈ topspace X› ‹path_connected_space X› path_connected_space_def)
then have "γ ` {0..1} ≲ topspace X"
by (meson path_image_subset_topspace subset_imp_lepoll image_subset_iff_funcset)
define 𝒜 where "𝒜 ≡ ((λa. {x ∈ {0..1}. γ x ∈ {a}}) ` topspace X) - {{}}"
have 𝒜01: "𝒜 = {{0..1}}"
proof (rule real_Sierpinski_lemma)
show "countable 𝒜"
using 𝒜_def coX by blast
show "disjoint 𝒜"
by (auto simp: 𝒜_def disjnt_iff pairwise_def)
show "⋃𝒜 = {0..1}"
using ‹pathin X γ› path_image_subset_topspace by (fastforce simp: 𝒜_def Bex_def)
fix C
assume "C ∈ 𝒜"
then obtain a where "a ∈ topspace X" and C: "C = {x ∈ {0..1}. γ x ∈ {a}}" "C ≠ {}"
by (auto simp: 𝒜_def)
then have "closedin X {a}"
by (meson ‹t1_space X› closedin_t1_singleton)
then have "closedin (top_of_set {0..1}) C"
using C ‹pathin X γ› closedin_continuous_map_preimage pathin_def by fastforce
then show "closed C ∧ C ≠ {}"
using C closedin_closed_trans by blast
qed auto
then have "{0..1} ∈ 𝒜"
by blast
then have "∃a ∈ topspace X. {0..1} ⊆ {x. γ x = a}"
using 𝒜_def image_iff by auto
then show False
using ‹γ 0 = a› ‹γ 1 = b› ‹a ≠ b› atLeastAtMost_iff zero_less_one_class.zero_le_one by blast
qed
subsection‹Lavrentiev extension etc›
lemma (in Metric_space) convergent_eq_zero_oscillation_gen:
assumes "mcomplete" and fim: "f ∈ (topspace X ∩ S) → M"
shows "(∃l. limitin mtopology f l (atin_within X a S)) ⟷
M ≠ {} ∧
(a ∈ topspace X
⟶ (∀ε>0. ∃U. openin X U ∧ a ∈ U ∧
(∀x ∈ (S ∩ U) - {a}. ∀y ∈ (S ∩ U) - {a}. d (f x) (f y) < ε)))"
proof (cases "M = {}")
case True
with limitin_mspace show ?thesis
by blast
next
case False
show ?thesis
proof (cases "a ∈ topspace X")
case True
let ?R = "∀ε>0. ∃U. openin X U ∧ a ∈ U ∧ (∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < ε)"
show ?thesis
proof (cases "a ∈ X derived_set_of S")
case True
have ?R
if "limitin mtopology f l (atin_within X a S)" for l
proof (intro strip)
fix ε::real
assume "ε>0"
with that ‹a ∈ topspace X›
obtain U where U: "openin X U" "a ∈ U" "l ∈ M"
and Uless: "∀x∈U - {a}. x ∈ S ⟶ f x ∈ M ∧ d (f x) l < ε/2"
unfolding limitin_metric eventually_atin_within by (metis Diff_iff insertI1 half_gt_zero [OF ‹ε>0›])
show "∃U. openin X U ∧ a ∈ U ∧ (∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < ε)"
proof (intro exI strip conjI)
fix x y
assume x: "x ∈ S ∩ U - {a}" and y: "y ∈ S ∩ U - {a}"
then have "d (f x) l < ε/2" "d (f y) l < ε/2" "f x ∈ M" "f y ∈ M"
using Uless by auto
then show "d (f x) (f y) < ε"
using triangle' ‹l ∈ M› by fastforce
qed (auto simp add: U)
qed
moreover have "∃l. limitin mtopology f l (atin_within X a S)"
if R [rule_format]: ?R
proof -
define F where "F ≡ λU. mtopology closure_of f ` (S ∩ U - {a})"
define 𝒞 where "𝒞 ≡ F ` {U. openin X U ∧ a ∈ U}"
have 𝒞_clo: "∀C ∈ 𝒞. closedin mtopology C"
by (force simp add: 𝒞_def F_def)
moreover have sub_mcball: "∃C a. C ∈ 𝒞 ∧ C ⊆ mcball a ε" if "ε>0" for ε
proof -
obtain U where U: "openin X U" "a ∈ U"
and Uless: "∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < ε"
using R [OF ‹ε>0›] by blast
then obtain b where b: "b ≠ a" "b ∈ S" "b ∈ U"
using True by (auto simp add: in_derived_set_of)
have "U ⊆ topspace X"
by (simp add: U(1) openin_subset)
have "f b ∈ M"
using b ‹openin X U› by (metis image_subset_iff_funcset Int_iff fim image_eqI openin_subset subsetD)
moreover
have "mtopology closure_of f ` ((S ∩ U) - {a}) ⊆ mcball (f b) ε"
proof (rule closure_of_minimal)
have "f y ∈ M" if "y ∈ S" and "y ∈ U" for y
using ‹U ⊆ topspace X› fim that by (auto simp: Pi_iff)
moreover
have "d (f b) (f y) ≤ ε" if "y ∈ S" "y ∈ U" "y ≠ a" for y
using that Uless b by force
ultimately show "f ` (S ∩ U - {a}) ⊆ mcball (f b) ε"
by (force simp: ‹f b ∈ M›)
qed auto
ultimately show ?thesis
using U by (auto simp add: 𝒞_def F_def)
qed
moreover have "⋂ℱ ≠ {}" if "finite ℱ" "ℱ ⊆ 𝒞" for ℱ
proof -
obtain 𝒢 where sub: "𝒢 ⊆ {U. openin X U ∧ a ∈ U}" and eq: "ℱ = F ` 𝒢" and "finite 𝒢"
by (metis (no_types, lifting) 𝒞_def ‹ℱ ⊆ 𝒞› ‹finite ℱ› finite_subset_image)
then have "U ⊆ topspace X" if "U ∈ 𝒢" for U
using openin_subset that by auto
then have "T ⊆ mtopology closure_of T"
if "T ∈ (λU. f ` (S ∩ U - {a})) ` 𝒢" for T
using that fim by (fastforce simp add: intro!: closure_of_subset)
moreover
have ain: "a ∈ ⋂ (insert (topspace X) 𝒢)" "openin X (⋂ (insert (topspace X) 𝒢))"
using True in_derived_set_of sub ‹finite 𝒢› by (fastforce intro!: openin_Inter)+
then obtain y where "y ≠ a" "y ∈ S" and y: "y ∈ ⋂ (insert (topspace X) 𝒢)"
by (meson ‹a ∈ X derived_set_of S› sub in_derived_set_of)
then have "f y ∈ ⋂ℱ"
using eq that ain fim by (auto simp add: F_def image_subset_iff in_closure_of)
then show ?thesis by blast
qed
ultimately have "⋂𝒞 ≠ {}"
using ‹mcomplete› mcomplete_fip by metis
then obtain b where "b ∈ ⋂𝒞"
by auto
then have "b ∈ M"
using sub_mcball 𝒞_clo mbounded_alt_pos mbounded_empty metric_closedin_iff_sequentially_closed by force
have "limitin mtopology f b (atin_within X a S)"
proof (clarsimp simp: limitin_metric ‹b ∈ M›)
fix ε :: real
assume "ε > 0"
then obtain U where U: "openin X U" "a ∈ U" and subU: "U ⊆ topspace X"
and Uless: "∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < ε/2"
by (metis R half_gt_zero openin_subset)
then obtain x where x: "x ∈ S" "x ∈ U" "x ≠ a" and fx: "f x ∈ mball b (ε/2)"
using ‹b ∈ ⋂𝒞›
apply (simp add: 𝒞_def F_def closure_of_def del: divide_const_simps)
by (metis Diff_iff Int_iff centre_in_mball_iff in_mball openin_mball singletonI zero_less_numeral)
moreover
have "d (f y) b < ε" if "y ∈ U" "y ≠ a" "y ∈ S" for y
proof -
have "d (f x) (f y) < ε/2"
using Uless that x by force
moreover have "d b (f x) < ε/2"
using fx by simp
ultimately show ?thesis
using triangle [of b "f x" "f y"] subU that ‹b ∈ M› commute fim fx by fastforce
qed
ultimately show "∀⇩F x in atin_within X a S. f x ∈ M ∧ d (f x) b < ε"
using fim U
apply (simp add: eventually_atin_within del: divide_const_simps flip: image_subset_iff_funcset)
by (smt (verit, del_insts) Diff_iff Int_iff imageI insertI1 openin_subset subsetD)
qed
then show ?thesis ..
qed
ultimately
show ?thesis
by (meson True ‹M ≠ {}› in_derived_set_of)
next
case False
have "(∃l. limitin mtopology f l (atin_within X a S))"
by (metis ‹M ≠ {}› False derived_set_of_trivial_limit equals0I limitin_trivial topspace_mtopology)
moreover have "∀e>0. ∃U. openin X U ∧ a ∈ U ∧ (∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < e)"
by (metis Diff_iff False IntE True in_derived_set_of insert_iff)
ultimately show ?thesis
using limitin_mspace by blast
qed
next
case False
then show ?thesis
by (metis derived_set_of_trivial_limit ex_in_conv in_derived_set_of limitin_mspace limitin_trivial topspace_mtopology)
qed
qed
text ‹The HOL Light proof uses some ugly tricks to share common parts of what are two separate proofs for the two cases›
lemma (in Metric_space) gdelta_in_points_of_convergence_within:
assumes "mcomplete"
and f: "continuous_map (subtopology X S) mtopology f ∨ t1_space X ∧ f ∈ S → M"
shows "gdelta_in X {x ∈ topspace X. ∃l. limitin mtopology f l (atin_within X x S)}"
proof -
have fim: "f ∈ (topspace X ∩ S) → M"
using continuous_map_image_subset_topspace f by force
show ?thesis
proof (cases "M={}")
case True
then show ?thesis
by (smt (verit) Collect_cong empty_def empty_iff gdelta_in_empty limitin_mspace)
next
case False
define A where "A ≡ {a ∈ topspace X. ∀ε>0. ∃U. openin X U ∧ a ∈ U ∧ (∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < ε)}"
have "gdelta_in X A"
using f
proof (elim disjE conjE)
assume cm: "continuous_map (subtopology X S) mtopology f"
define C where "C ≡ λr. ⋃{U. openin X U ∧ (∀x ∈ S∩U. ∀y ∈ S∩U. d (f x) (f y) < r)}"
define B where "B ≡ (⋂n. C(inverse(Suc n)))"
define D where "D ≡ (⋂ (C ` {0<..}))"
have "D=B"
unfolding B_def C_def D_def
apply (intro Inter_eq_Inter_inverse_Suc Sup_subset_mono)
by (smt (verit, ccfv_threshold) Collect_mono_iff)
have "B ⊆ topspace X"
using openin_subset by (force simp add: B_def C_def)
have "(countable intersection_of openin X) B"
unfolding B_def C_def
by (intro relative_to_inc countable_intersection_of_Inter countable_intersection_of_inc) auto
then have "gdelta_in X B"
unfolding gdelta_in_def by (intro relative_to_subset_inc ‹B ⊆ topspace X›)
moreover have "A=D"
proof (intro equalityI subsetI)
fix a
assume x: "a ∈ A"
then have "a ∈ topspace X"
using A_def by blast
show "a ∈ D"
proof (clarsimp simp: D_def C_def ‹a ∈ topspace X›)
fix ε::real assume "ε > 0"
then obtain U where "openin X U" "a ∈ U" and U: "(∀x∈S ∩ U - {a}. ∀y∈S ∩ U - {a}. d (f x) (f y) < ε)"
using x by (force simp: A_def)
show "∃T. openin X T ∧ (∀x∈S ∩ T. ∀y∈S ∩ T. d (f x) (f y) < ε) ∧ a ∈ T"
proof (cases "a ∈ S")
case True
then obtain V where "openin X V" "a ∈ V" and V: "∀x. x ∈ S ∧ x ∈ V ⟶ f a ∈ M ∧ f x ∈ M ∧ d (f a) (f x) < ε"
using ‹a ∈ topspace X› ‹ε > 0› cm
by (force simp add: continuous_map_to_metric openin_subtopology_alt Ball_def)
show ?thesis
proof (intro exI conjI strip)
show "openin X (U ∩ V)"
using ‹openin X U› ‹openin X V› by blast
show "a ∈ U ∩ V"
using ‹a ∈ U› ‹a ∈ V› by blast
show "⋀x y. ⟦x ∈ S ∩ (U ∩ V); y ∈ S ∩ (U ∩ V)⟧ ⟹ d (f x) (f y) < ε"
by (metis DiffI Int_iff U V commute singletonD)
qed
next
case False then show ?thesis
using U ‹a ∈ U› ‹openin X U› by auto
qed
qed
next
fix x
assume x: "x ∈ D"
then have "x ∈ topspace X"
using ‹B ⊆ topspace X› ‹D=B› by blast
with x show "x ∈ A"
apply (clarsimp simp: D_def C_def A_def)
by (meson DiffD1 greaterThan_iff)
qed
ultimately show ?thesis
by (simp add: ‹D=B›)
next
assume "t1_space X" "f ∈ S → M"
define C where "C ≡ λr. ⋃{U. openin X U ∧
(∃b ∈ topspace X. ∀x ∈ S∩U - {b}. ∀y ∈ S∩U - {b}. d (f x) (f y) < r)}"
define B where "B ≡ (⋂n. C(inverse(Suc n)))"
define D where "D ≡ (⋂ (C ` {0<..}))"
have "D=B"
unfolding B_def C_def D_def
apply (intro Inter_eq_Inter_inverse_Suc Sup_subset_mono)
by (smt (verit, ccfv_threshold) Collect_mono_iff)
have "B ⊆ topspace X"
using openin_subset by (force simp add: B_def C_def)
have "(countable intersection_of openin X) B"
unfolding B_def C_def
by (intro relative_to_inc countable_intersection_of_Inter countable_intersection_of_inc) auto
then have "gdelta_in X B"
unfolding gdelta_in_def by (intro relative_to_subset_inc ‹B ⊆ topspace X›)
moreover have "A=D"
proof (intro equalityI subsetI)
fix x
assume x: "x ∈ D"
then have "x ∈ topspace X"
using ‹B ⊆ topspace X› ‹D=B› by blast
show "x ∈ A"
proof (clarsimp simp: A_def ‹x ∈ topspace X›)
fix ε :: real
assume "ε>0"
then obtain U b where "openin X U" "b ∈ topspace X"
and U: "∀x∈S ∩ U - {b}. ∀y∈S ∩ U - {b}. d (f x) (f y) < ε" and "x ∈ U"
using x by (auto simp: D_def C_def A_def Ball_def)
then have "openin X (U-{b})"
by (meson ‹t1_space X› t1_space_openin_delete_alt)
then show "∃U. openin X U ∧ x ∈ U ∧ (∀xa∈S ∩ U - {x}. ∀y∈S ∩ U - {x}. d (f xa) (f y) < ε)"
using U ‹openin X U› ‹x ∈ U› by auto
qed
next
show "⋀x. x ∈ A ⟹ x ∈ D"
unfolding A_def D_def C_def
by clarsimp meson
qed
ultimately show ?thesis
by (simp add: ‹D=B›)
qed
then show ?thesis
by (simp add: A_def convergent_eq_zero_oscillation_gen False fim ‹mcomplete› cong: conj_cong)
qed
qed
lemma gdelta_in_points_of_convergence_within:
assumes Y: "completely_metrizable_space Y"
and f: "continuous_map (subtopology X S) Y f ∨ t1_space X ∧ f ∈ S → topspace Y"
shows "gdelta_in X {x ∈ topspace X. ∃l. limitin Y f l (atin_within X x S)}"
using assms
unfolding completely_metrizable_space_def
using Metric_space.gdelta_in_points_of_convergence_within Metric_space.topspace_mtopology by fastforce
lemma Lavrentiev_extension_gen:
assumes "S ⊆ topspace X" and Y: "completely_metrizable_space Y"
and contf: "continuous_map (subtopology X S) Y f"
obtains U g where "gdelta_in X U" "S ⊆ U"
"continuous_map (subtopology X (X closure_of S ∩ U)) Y g"
"⋀x. x ∈ S ⟹ g x = f x"
proof -
define U where "U ≡ {x ∈ topspace X. ∃l. limitin Y f l (atin_within X x S)}"
have "S ⊆ U"
using that contf limit_continuous_map_within subsetD [OF ‹S ⊆ topspace X›]
by (fastforce simp: U_def)
then have "S ⊆ X closure_of S ∩ U"
by (simp add: ‹S ⊆ topspace X› closure_of_subset)
moreover
have "⋀t. t ∈ X closure_of S ∩ U - S ⟹ ∃l. limitin Y f l (atin_within X t S)"
using U_def by blast
moreover have "regular_space Y"
by (simp add: Y completely_metrizable_imp_metrizable_space metrizable_imp_regular_space)
ultimately
obtain g where g: "continuous_map (subtopology X (X closure_of S ∩ U)) Y g"
and gf: "⋀x. x ∈ S ⟹ g x = f x"
using continuous_map_extension_pointwise_alt assms by blast
show thesis
proof
show "gdelta_in X U"
by (simp add: U_def Y contf gdelta_in_points_of_convergence_within)
show "continuous_map (subtopology X (X closure_of S ∩ U)) Y g"
by (simp add: g)
qed (use ‹S ⊆ U› gf in auto)
qed
lemma Lavrentiev_extension:
assumes "S ⊆ topspace X"
and X: "metrizable_space X ∨ topspace X ⊆ X closure_of S"
and Y: "completely_metrizable_space Y"
and contf: "continuous_map (subtopology X S) Y f"
obtains U g where "gdelta_in X U" "S ⊆ U" "U ⊆ X closure_of S"
"continuous_map (subtopology X U) Y g" "⋀x. x ∈ S ⟹ g x = f x"
proof -
obtain U g where "gdelta_in X U" "S ⊆ U"
and contg: "continuous_map (subtopology X (X closure_of S ∩ U)) Y g"
and gf: "⋀x. x ∈ S ⟹ g x = f x"
using Lavrentiev_extension_gen Y assms(1) contf by blast
define V where "V ≡ X closure_of S ∩ U"
show thesis
proof
show "gdelta_in X V"
by (metis V_def X ‹gdelta_in X U› closed_imp_gdelta_in closedin_closure_of closure_of_subset_topspace gdelta_in_Int gdelta_in_topspace subset_antisym)
show "S ⊆ V"
by (simp add: V_def ‹S ⊆ U› assms(1) closure_of_subset)
show "continuous_map (subtopology X V) Y g"
by (simp add: V_def contg)
qed (auto simp: gf V_def)
qed
subsection‹Embedding in products and hence more about completely metrizable spaces›
lemma (in Metric_space) gdelta_homeomorphic_space_closedin_product:
assumes S: "⋀i. i ∈ I ⟹ openin mtopology (S i)"
obtains T where "closedin (prod_topology mtopology (powertop_real I)) T"
"subtopology mtopology (⋂i ∈ I. S i) homeomorphic_space
subtopology (prod_topology mtopology (powertop_real I)) T"
proof (cases "I={}")
case True
then have top: "topspace (prod_topology mtopology (powertop_real I)) = (M × {(λx. undefined)})"
by simp
show ?thesis
proof
show "closedin (prod_topology mtopology (powertop_real I)) (M × {(λx. undefined)})"
by (metis top closedin_topspace)
have "subtopology mtopology (⋂(S ` I)) homeomorphic_space mtopology"
by (simp add: True product_topology_empty_discrete)
also have "… homeomorphic_space (prod_topology mtopology (discrete_topology {λx. undefined}))"
by (meson homeomorphic_space_sym prod_topology_homeomorphic_space_left)
finally
show "subtopology mtopology (⋂(S ` I)) homeomorphic_space subtopology (prod_topology mtopology (powertop_real I)) (M × {(λx. undefined)})"
by (smt (verit, ccfv_SIG) True product_topology_empty_discrete subtopology_topspace top)
qed
next
case False
have SM: "⋀i. i ∈ I ⟹ S i ⊆ M"
using assms openin_mtopology by blast
then have "(⋂i ∈ I. S i) ⊆ M"
using False by blast
define dd where "dd ≡ λi. if i∉I ∨ S i = M then λu. 1 else (λu. INF x ∈ M - S i. d u x)"
have [simp]: "bdd_below (d u ` A)" for u A
by (meson bdd_belowI2 nonneg)
have cont_dd: "continuous_map (subtopology mtopology (S i)) euclidean (dd i)" if "i ∈ I" for i
proof -
have "dist (Inf (d x ` (M - S i))) (Inf (d y ` (M - S i))) ≤ d x y"
if "x ∈ S i" "x ∈ M" "y ∈ S i" "y ∈ M" "S i ≠ M" for x y
proof -
have [simp]: "¬ M ⊆ S i"
using SM ‹S i ≠ M› ‹i ∈ I› by auto
have "⋀u. ⟦u ∈ M; u ∉ S i⟧ ⟹ Inf (d x ` (M - S i)) ≤ d x y + d y u"
apply (clarsimp simp add: cInf_le_iff_less)
by (smt (verit) DiffI triangle ‹x ∈ M› ‹y ∈ M›)
then have "Inf (d x ` (M - S i)) - d x y ≤ Inf (d y ` (M - S i))"
by (force simp add: le_cInf_iff)
moreover
have "⋀u. ⟦u ∈ M; u ∉ S i⟧ ⟹ Inf (d y ` (M - S i)) ≤ d x u + d x y"
apply (clarsimp simp add: cInf_le_iff_less)
by (smt (verit) DiffI triangle'' ‹x ∈ M› ‹y ∈ M›)
then have "Inf (d y ` (M - S i)) - d x y ≤ Inf (d x ` (M - S i))"
by (force simp add: le_cInf_iff)
ultimately show ?thesis
by (simp add: dist_real_def abs_le_iff)
qed
then have *: "Lipschitz_continuous_map (submetric Self (S i)) euclidean_metric (λu. Inf (d u ` (M - S i)))"
unfolding Lipschitz_continuous_map_def by (force intro!: exI [where x=1])
then show ?thesis
using Lipschitz_continuous_imp_continuous_map [OF *]
by (simp add: dd_def Self_def mtopology_of_submetric )
qed
have dd_pos: "0 < dd i x" if "x ∈ S i" for i x
proof (clarsimp simp add: dd_def)
assume "i ∈ I" and "S i ≠ M"
have opeS: "openin mtopology (S i)"
by (simp add: ‹i ∈ I› assms)
then obtain r where "r>0" and r: "⋀y. ⟦y ∈ M; d x y < r⟧ ⟹ y ∈ S i"
by (meson ‹x ∈ S i› in_mball openin_mtopology subsetD)
then have "⋀y. y ∈ M - S i ⟹ d x y ≥ r"
by (meson Diff_iff linorder_not_le)
then have "Inf (d x ` (M - S i)) ≥ r"
by (meson Diff_eq_empty_iff SM ‹S i ≠ M› ‹i ∈ I› cINF_greatest set_eq_subset)
with ‹r>0› show "0 < Inf (d x ` (M - S i))" by simp
qed
define f where "f ≡ λx. (x, λi∈I. inverse(dd i x))"
define T where "T ≡ f ` (⋂i ∈ I. S i)"
show ?thesis
proof
show "closedin (prod_topology mtopology (powertop_real I)) T"
unfolding closure_of_subset_eq [symmetric]
proof
show "T ⊆ topspace (prod_topology mtopology (powertop_real I))"
using False SM by (auto simp: T_def f_def)
have "(x,ds) ∈ T"
if §: "⋀U. ⟦(x, ds) ∈ U; openin (prod_topology mtopology (powertop_real I)) U⟧ ⟹ ∃y∈T. y ∈ U"
and "x ∈ M" and ds: "ds ∈ I →⇩E UNIV" for x ds
proof -
have ope: "∃x. x ∈ ⋂(S ` I) ∧ f x ∈ U × V"
if "x ∈ U" and "ds ∈ V" and "openin mtopology U" and "openin (powertop_real I) V" for U V
using § [of "U×V"] that by (force simp add: T_def openin_prod_Times_iff)
have x_in_INT: "x ∈ ⋂(S ` I)"
proof clarify
fix i
assume "i ∈ I"
show "x ∈ S i"
proof (rule ccontr)
assume "x ∉ S i"
have "openin (powertop_real I) {z ∈ topspace (powertop_real I). z i ∈ {ds i - 1 <..< ds i + 1}}"
proof (rule openin_continuous_map_preimage)
show "continuous_map (powertop_real I) euclidean (λx. x i)"
by (metis ‹i ∈ I› continuous_map_product_projection)
qed auto
then obtain y where "y ∈ S i" "y ∈ M" and dxy: "d x y < inverse (¦ds i¦ + 1)"
and intvl: "inverse (dd i y) ∈ {ds i - 1 <..< ds i + 1}"
using ope [of "mball x (inverse(abs(ds i) + 1))" "{z ∈ topspace(powertop_real I). z i ∈ {ds i - 1<..<ds i + 1}}"]
‹x ∈ M› ‹ds ∈ I →⇩E UNIV› ‹i ∈ I›
by (fastforce simp add: f_def)
have "¬ M ⊆ S i"
using ‹x ∉ S i› ‹x ∈ M› by blast
have "inverse (¦ds i¦ + 1) ≤ dd i y"
using intvl ‹y ∈ S i› dd_pos [of y i]
by (smt (verit, ccfv_threshold) greaterThanLessThan_iff inverse_inverse_eq le_imp_inverse_le)
also have "… ≤ d x y"
using ‹i ∈ I› ‹¬ M ⊆ S i› ‹x ∉ S i› ‹x ∈ M›
apply (simp add: dd_def cInf_le_iff_less)
using commute by force
finally show False
using dxy by linarith
qed
qed
moreover have "ds = (λi∈I. inverse (dd i x))"
proof (rule PiE_ext [OF ds])
fix i
assume "i ∈ I"
define e where "e ≡ ¦ds i - inverse (dd i x)¦"
{ assume con: "e > 0"
have "continuous_map (subtopology mtopology (S i)) euclidean (λx. inverse (dd i x))"
using dd_pos cont_dd ‹i ∈ I›
by (fastforce simp: intro!: continuous_map_real_inverse)
then have "openin (subtopology mtopology (S i))
{z ∈ topspace (subtopology mtopology (S i)).
inverse (dd i z) ∈ {inverse(dd i x) - e/2<..<inverse(dd i x) + e/2}}"
using openin_continuous_map_preimage open_greaterThanLessThan open_openin by blast
then obtain U where "openin mtopology U"
and U: "{z ∈ S i. inverse (dd i x) - e/2 < inverse (dd i z) ∧
inverse (dd i z) < inverse (dd i x) + e/2}
= U ∩ S i"
using SM ‹i ∈ I› by (auto simp: openin_subtopology)
have "x ∈ U"
using U x_in_INT ‹i ∈ I› con by fastforce
have "ds ∈ {z ∈ topspace (powertop_real I). z i ∈ {ds i - e / 2<..<ds i + e/2}}"
by (simp add: con ds)
moreover
have "openin (powertop_real I) {z ∈ topspace (powertop_real I). z i ∈ {ds i - e / 2<..<ds i + e/2}}"
proof (rule openin_continuous_map_preimage)
show "continuous_map (powertop_real I) euclidean (λx. x i)"
by (metis ‹i ∈ I› continuous_map_product_projection)
qed auto
ultimately obtain y where "y ∈ ⋂(S ` I) ∧ f y ∈ U × {z ∈ topspace (powertop_real I). z i ∈ {ds i - e / 2<..<ds i + e/2}}"
using ope ‹x ∈ U› ‹openin mtopology U› ‹x ∈ U›
by presburger
with ‹i ∈ I› U
have False unfolding set_eq_iff f_def e_def by simp (smt (verit) field_sum_of_halves)
}
then show "ds i = (λi∈I. inverse (dd i x)) i"
using ‹i ∈ I› by (force simp: e_def)
qed auto
ultimately show ?thesis
by (auto simp: T_def f_def)
qed
then show "prod_topology mtopology (powertop_real I) closure_of T ⊆ T"
by (auto simp: closure_of_def)
qed
have eq: "(⋂(S ` I) × (I →⇩E UNIV) ∩ f ` (M ∩ ⋂(S ` I))) = (f ` ⋂(S ` I))"
using False SM by (force simp: f_def image_iff)
have "continuous_map (subtopology mtopology (⋂(S ` I))) euclidean (dd i)" if "i ∈ I" for i
by (meson INT_lower cont_dd continuous_map_from_subtopology_mono that)
then have "continuous_map (subtopology mtopology (⋂(S ` I))) (powertop_real I) (λx. λi∈I. inverse (dd i x))"
using dd_pos by (fastforce simp: continuous_map_componentwise intro!: continuous_map_real_inverse)
then have "embedding_map (subtopology mtopology (⋂(S ` I))) (prod_topology (subtopology mtopology (⋂(S ` I))) (powertop_real I)) f"
by (simp add: embedding_map_graph f_def)
moreover have "subtopology (prod_topology (subtopology mtopology (⋂(S ` I))) (powertop_real I))
(f ` topspace (subtopology mtopology (⋂(S ` I)))) =
subtopology (prod_topology mtopology (powertop_real I)) T"
by (simp add: prod_topology_subtopology subtopology_subtopology T_def eq)
ultimately
show "subtopology mtopology (⋂(S ` I)) homeomorphic_space subtopology (prod_topology mtopology (powertop_real I)) T"
by (metis embedding_map_imp_homeomorphic_space)
qed
qed
lemma gdelta_homeomorphic_space_closedin_product:
assumes "metrizable_space X" and "⋀i. i ∈ I ⟹ openin X (S i)"
obtains T where "closedin (prod_topology X (powertop_real I)) T"
"subtopology X (⋂i ∈ I. S i) homeomorphic_space
subtopology (prod_topology X (powertop_real I)) T"
using Metric_space.gdelta_homeomorphic_space_closedin_product
by (metis assms metrizable_space_def)
lemma open_homeomorphic_space_closedin_product:
assumes "metrizable_space X" and "openin X S"
obtains T where "closedin (prod_topology X euclideanreal) T"
"subtopology X S homeomorphic_space
subtopology (prod_topology X euclideanreal) T"
proof -
obtain T where cloT: "closedin (prod_topology X (powertop_real {()})) T"
and homT: "subtopology X S homeomorphic_space
subtopology (prod_topology X (powertop_real {()})) T"
using gdelta_homeomorphic_space_closedin_product [of X "{()}" "λi. S"] assms
by auto
have "prod_topology X (powertop_real {()}) homeomorphic_space prod_topology X euclideanreal"
by (meson homeomorphic_space_prod_topology homeomorphic_space_refl homeomorphic_space_singleton_product)
then obtain f where f: "homeomorphic_map (prod_topology X (powertop_real {()})) (prod_topology X euclideanreal) f"
unfolding homeomorphic_space by metis
show thesis
proof
show "closedin (prod_topology X euclideanreal) (f ` T)"
using cloT f homeomorphic_map_closedness_eq by blast
moreover have "T = topspace (subtopology (prod_topology X (powertop_real {()})) T)"
by (metis cloT closedin_subset topspace_subtopology_subset)
ultimately show "subtopology X S homeomorphic_space subtopology (prod_topology X euclideanreal) (f ` T)"
by (smt (verit, best) closedin_subset f homT homeomorphic_map_subtopologies homeomorphic_space
homeomorphic_space_trans topspace_subtopology topspace_subtopology_subset)
qed
qed
lemma completely_metrizable_space_gdelta_in_alt:
assumes X: "completely_metrizable_space X"
and S: "(countable intersection_of openin X) S"
shows "completely_metrizable_space (subtopology X S)"
proof -
obtain 𝒰 where "countable 𝒰" "S = ⋂𝒰" and ope: "⋀U. U ∈ 𝒰 ⟹ openin X U"
using S by (force simp add: intersection_of_def)
then have 𝒰: "completely_metrizable_space (powertop_real 𝒰)"
by (simp add: completely_metrizable_space_euclidean completely_metrizable_space_product_topology)
obtain C where "closedin (prod_topology X (powertop_real 𝒰)) C"
and sub: "subtopology X (⋂𝒰) homeomorphic_space
subtopology (prod_topology X (powertop_real 𝒰)) C"
by (metis gdelta_homeomorphic_space_closedin_product X completely_metrizable_imp_metrizable_space ope INF_identity_eq)
moreover have "completely_metrizable_space (prod_topology X (powertop_real 𝒰))"
by (simp add: completely_metrizable_space_prod_topology X 𝒰)
ultimately have "completely_metrizable_space (subtopology (prod_topology X (powertop_real 𝒰)) C)"
using completely_metrizable_space_closedin by blast
then show ?thesis
using ‹S = ⋂𝒰› sub homeomorphic_completely_metrizable_space by blast
qed
lemma completely_metrizable_space_gdelta_in:
"⟦completely_metrizable_space X; gdelta_in X S⟧
⟹ completely_metrizable_space (subtopology X S)"
by (simp add: completely_metrizable_space_gdelta_in_alt gdelta_in_alt)
lemma completely_metrizable_space_openin:
"⟦completely_metrizable_space X; openin X S⟧
⟹ completely_metrizable_space (subtopology X S)"
by (simp add: completely_metrizable_space_gdelta_in open_imp_gdelta_in)
lemma (in Metric_space) locally_compact_imp_completely_metrizable_space:
assumes "locally_compact_space mtopology"
shows "completely_metrizable_space mtopology"
proof -
obtain f :: "['a,'a] ⇒ real" and m' where
"mcomplete_of m'" and fim: "f ∈ M → mspace m'"
and clo: "mtopology_of m' closure_of f ` M = mspace m'"
and d: "⋀x y. ⟦x ∈ M; y ∈ M⟧ ⟹ mdist m' (f x) (f y) = d x y"
by (metis metric_completion)
then have "embedding_map mtopology (mtopology_of m') f"
unfolding mtopology_of_def
by (metis Metric_space12.isometry_imp_embedding_map Metric_space12_mspace_mdist mdist_metric mspace_metric)
then have hom: "mtopology homeomorphic_space subtopology (mtopology_of m') (f ` M)"
by (metis embedding_map_imp_homeomorphic_space topspace_mtopology)
have "locally_compact_space (subtopology (mtopology_of m') (f ` M))"
using assms hom homeomorphic_locally_compact_space by blast
moreover have "Hausdorff_space (mtopology_of m')"
by (simp add: Metric_space.Hausdorff_space_mtopology mtopology_of_def)
ultimately have "openin (mtopology_of m') (f ` M)"
by (simp add: clo dense_locally_compact_openin_Hausdorff_space fim image_subset_iff_funcset)
then
have "completely_metrizable_space (subtopology (mtopology_of m') (f ` M))"
using ‹mcomplete_of m'› unfolding mcomplete_of_def mtopology_of_def
by (metis Metric_space.completely_metrizable_space_mtopology Metric_space_mspace_mdist completely_metrizable_space_openin )
then show ?thesis
using hom homeomorphic_completely_metrizable_space by blast
qed
lemma locally_compact_imp_completely_metrizable_space:
assumes "metrizable_space X" and "locally_compact_space X"
shows "completely_metrizable_space X"
by (metis Metric_space.locally_compact_imp_completely_metrizable_space assms metrizable_space_def)
lemma completely_metrizable_space_imp_gdelta_in:
assumes X: "metrizable_space X" and "S ⊆ topspace X"
and XS: "completely_metrizable_space (subtopology X S)"
shows "gdelta_in X S"
proof -
obtain U f where "gdelta_in X U" "S ⊆ U" and U: "U ⊆ X closure_of S"
and contf: "continuous_map (subtopology X U) (subtopology X S) f"
and fid: "⋀x. x ∈ S ⟹ f x = x"
using Lavrentiev_extension[of S X "subtopology X S" id] assms by auto
then have "f ` topspace (subtopology X U) ⊆ topspace (subtopology X S)"
using continuous_map_image_subset_topspace by blast
then have "f`U ⊆ S"
by (metis ‹gdelta_in X U› ‹S ⊆ topspace X› gdelta_in_subset topspace_subtopology_subset)
moreover
have "Hausdorff_space (subtopology X U)"
by (simp add: Hausdorff_space_subtopology X metrizable_imp_Hausdorff_space)
then have "⋀x. x ∈ U ⟹ f x = x"
using U fid contf forall_in_closure_of_eq [of _ "subtopology X U" S "subtopology X U" f id]
by (metis ‹S ⊆ U› closure_of_subtopology_open continuous_map_id continuous_map_in_subtopology id_apply inf.orderE subtopology_subtopology)
ultimately have "U ⊆ S"
by auto
then show ?thesis
using ‹S ⊆ U› ‹gdelta_in X U› by auto
qed
lemma completely_metrizable_space_eq_gdelta_in:
"⟦completely_metrizable_space X; S ⊆ topspace X⟧
⟹ completely_metrizable_space (subtopology X S) ⟷ gdelta_in X S"
using completely_metrizable_imp_metrizable_space completely_metrizable_space_gdelta_in completely_metrizable_space_imp_gdelta_in by blast
lemma gdelta_in_eq_completely_metrizable_space:
"completely_metrizable_space X
⟹ gdelta_in X S ⟷ S ⊆ topspace X ∧ completely_metrizable_space (subtopology X S)"
by (metis completely_metrizable_space_eq_gdelta_in gdelta_in_alt)
subsection‹ Theorems from Kuratowski›
text‹Kuratowski, Remark on an Invariance Theorem, \emph{Fundamenta Mathematicae} \textbf{37} (1950), pp. 251-252.
The idea is that in suitable spaces, to show "number of components of the complement" (without
distinguishing orders of infinity) is a homeomorphic invariant, it
suffices to show it for closed subsets. Kuratowski states the main result
for a "locally connected continuum", and seems clearly to be implicitly
assuming that means metrizable. We call out the general topological
hypotheses more explicitly, which do not however include connectedness. ›
lemma separation_by_closed_intermediates_count:
assumes X: "hereditarily normal_space X"
and "finite 𝒰"
and pwU: "pairwise (separatedin X) 𝒰"
and nonempty: "{} ∉ 𝒰"
and UU: "⋃𝒰 = topspace X - S"
obtains C where "closedin X C" "C ⊆ S"
"⋀D. ⟦closedin X D; C ⊆ D; D ⊆ S⟧
⟹ ∃𝒱. 𝒱 ≈ 𝒰 ∧ pairwise (separatedin X) 𝒱 ∧ {} ∉ 𝒱 ∧ ⋃𝒱 = topspace X - D"
proof -
obtain F where F: "⋀S. S ∈ 𝒰 ⟹ openin X (F S) ∧ S ⊆ F S"
and pwF: "pairwise (λS T. disjnt (F S) (F T)) 𝒰"
using assms by (smt (verit, best) Diff_subset Sup_le_iff hereditarily_normal_separation_pairwise)
show thesis
proof
show "closedin X (topspace X - ⋃(F ` 𝒰))"
using F by blast
show "topspace X - ⋃(F ` 𝒰) ⊆ S"
using UU F by auto
show "∃𝒱. 𝒱 ≈ 𝒰 ∧ pairwise (separatedin X) 𝒱 ∧ {} ∉ 𝒱 ∧ ⋃𝒱 = topspace X - C"
if "closedin X C" "C ⊆ S" and C: "topspace X - ⋃(F ` 𝒰) ⊆ C" for C
proof (intro exI conjI strip)
have "inj_on (λS. F S - C) 𝒰"
using pwF F
unfolding inj_on_def pairwise_def disjnt_iff
by (metis Diff_iff UU UnionI nonempty subset_empty subset_eq ‹C ⊆ S›)
then show "(λS. F S - C) ` 𝒰 ≈ 𝒰"
by simp
show "pairwise (separatedin X) ((λS. F S - C) ` 𝒰)"
using ‹closedin X C› F pwF by (force simp: pairwise_def openin_diff separatedin_open_sets disjnt_iff)
show "{} ∉ (λS. F S - C) ` 𝒰"
using nonempty UU ‹C ⊆ S› F
by clarify (metis DiffD2 Diff_eq_empty_iff F UnionI subset_empty subset_eq)
show "(⋃S∈𝒰. F S - C) = topspace X - C"
using UU F C openin_subset by fastforce
qed
qed
qed
lemma separation_by_closed_intermediates_gen:
assumes X: "hereditarily normal_space X"
and discon: "¬ connectedin X (topspace X - S)"
obtains C where "closedin X C" "C ⊆ S"
"⋀D. ⟦closedin X D; C ⊆ D; D ⊆ S⟧ ⟹ ¬ connectedin X (topspace X - D)"
proof -
obtain C1 C2 where Ueq: "C1 ∪ C2 = topspace X - S" and "C1 ≠ {}" "C2 ≠ {}"
and sep: "separatedin X C1 C2" and "C1 ≠ C2"
by (metis Diff_subset connectedin_eq_not_separated discon separatedin_refl)
then obtain C where "closedin X C" "C ⊆ S"
and C: "⋀D. ⟦closedin X D; C ⊆ D; D ⊆ S⟧
⟹ ∃𝒱. 𝒱 ≈ {C1,C2} ∧ pairwise (separatedin X) 𝒱 ∧ {} ∉ 𝒱 ∧ ⋃𝒱 = topspace X - D"
using separation_by_closed_intermediates_count [of X "{C1,C2}" S] X
apply (simp add: pairwise_insert separatedin_sym)
by metis
have "¬ connectedin X (topspace X - D)"
if D: "closedin X D" "C ⊆ D" "D ⊆ S" for D
proof -
obtain V1 V2 where *: "pairwise (separatedin X) {V1,V2}" "{} ∉ {V1,V2}"
"⋃{V1,V2} = topspace X - D" "V1≠V2"
by (metis C [OF D] ‹C1 ≠ C2› eqpoll_doubleton_iff)
then have "disjnt V1 V2"
by (metis pairwise_insert separatedin_imp_disjoint singleton_iff)
with * show ?thesis
by (auto simp add: connectedin_eq_not_separated pairwise_insert)
qed
then show thesis
using ‹C ⊆ S› ‹closedin X C› that by auto
qed
lemma separation_by_closed_intermediates_eq_count:
fixes n::nat
assumes lcX: "locally_connected_space X" and hnX: "hereditarily normal_space X"
shows "(∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - S) ⟷
(∃C. closedin X C ∧ C ⊆ S ∧
(∀D. closedin X D ∧ C ⊆ D ∧ D ⊆ S
⟶ (∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - D)))"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
by (smt (verit, best) hnX separation_by_closed_intermediates_count eqpoll_iff_finite_card eqpoll_trans)
next
assume R: ?rhs
show ?lhs
proof (cases "n=0")
case True
with R show ?thesis
by fastforce
next
case False
obtain C where "closedin X C" "C ⊆ S"
and C: "⋀D. ⟦closedin X D; C ⊆ D; D ⊆ S⟧
⟹ ∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - D"
using R by force
then have "C ⊆ topspace X"
by (simp add: closedin_subset)
define 𝒰 where "𝒰 ≡ {D ∈ connected_components_of (subtopology X (topspace X - C)). D-S ≠ {}}"
have ope𝒰: "openin X U" if "U ∈ 𝒰" for U
using that ‹closedin X C› lcX locally_connected_space_open_connected_components
by (fastforce simp add: closedin_def 𝒰_def)
have "{} ∉ 𝒰"
by (auto simp: 𝒰_def)
have "pairwise disjnt 𝒰"
using connected_components_of_disjoint by (fastforce simp add: pairwise_def 𝒰_def)
show ?lhs
proof (rule ccontr)
assume con: "∄𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - S"
have card𝒰: "finite 𝒰 ∧ card 𝒰 < n"
proof (rule ccontr)
assume "¬ (finite 𝒰 ∧ card 𝒰 < n)"
then obtain 𝒱 where "𝒱 ⊆ 𝒰" "finite 𝒱" "card 𝒱 = n"
by (metis infinite_arbitrarily_large linorder_not_less obtain_subset_with_card_n)
then obtain T where "T ∈ 𝒱"
using False by force
define 𝒲 where "𝒲 ≡ insert (topspace X - S - ⋃(𝒱 - {T})) ((λD. D - S) ` (𝒱 - {T}))"
have "⋃𝒲 = topspace X - S"
using ‹⋀U. U ∈ 𝒰 ⟹ openin X U› ‹𝒱 ⊆ 𝒰› topspace_def by (fastforce simp: 𝒲_def)
moreover have "{} ∉ 𝒲"
proof -
obtain a where "a ∈ T" "a ∉ S"
using 𝒰_def ‹T ∈ 𝒱› ‹𝒱 ⊆ 𝒰› by blast
then have "a ∈ topspace X"
using ‹T ∈ 𝒱› ope𝒰 ‹𝒱 ⊆ 𝒰› openin_subset by blast
moreover have "a ∉ ⋃(𝒱 - {T})"
using diff_Union_pairwise_disjoint [of 𝒱 "{T}"] ‹disjoint 𝒰› pairwise_subset ‹T ∈ 𝒱› ‹𝒱 ⊆ 𝒰› ‹a ∈ T›
by auto
ultimately have "topspace X - S - ⋃(𝒱 - {T}) ≠ {}"
using ‹a ∉ S› by blast
moreover have "⋀V. V ∈ 𝒱 - {T} ⟹ V - S ≠ {}"
using 𝒰_def ‹𝒱 ⊆ 𝒰› by blast
ultimately show ?thesis
by (metis (no_types, lifting) 𝒲_def image_iff insert_iff)
qed
moreover have "disjoint 𝒱"
using ‹𝒱 ⊆ 𝒰› ‹disjoint 𝒰› pairwise_subset by blast
then have inj: "inj_on (λD. D - S) (𝒱 - {T})"
unfolding inj_on_def using ‹𝒱 ⊆ 𝒰› disjointD 𝒰_def inf_commute by blast
have "finite 𝒲" "card 𝒲 = n"
using ‹{} ∉ 𝒲› ‹n ≠ 0› ‹T ∈ 𝒱›
by (auto simp add: 𝒲_def ‹finite 𝒱› card_insert_if card_image inj ‹card 𝒱 = n›)
moreover have "pairwise (separatedin X) 𝒲"
proof -
have "disjoint 𝒲"
using ‹disjoint 𝒱› by (auto simp: 𝒲_def pairwise_def disjnt_iff)
have "pairwise (separatedin (subtopology X (topspace X - S))) 𝒲"
proof (intro pairwiseI)
fix A B
assume §: "A ∈ 𝒲" "B ∈ 𝒲" "A ≠ B"
then have "disjnt A B"
by (meson ‹disjoint 𝒲› pairwiseD)
have "closedin (subtopology X (topspace X - C)) (⋃(𝒱 - {T}))"
using 𝒰_def ‹𝒱 ⊆ 𝒰› closedin_connected_components_of ‹finite 𝒱›
by (force simp add: intro!: closedin_Union)
with ‹C ⊆ S› have "openin (subtopology X (topspace X - S)) (topspace X - S - ⋃(𝒱 - {T}))"
by (fastforce simp add: openin_closedin_eq closedin_subtopology Int_absorb1)
moreover have "⋀V. V ∈ 𝒱 ∧ V≠T ⟹ openin (subtopology X (topspace X - S)) (V - S)"
using ‹𝒱 ⊆ 𝒰› ope𝒰
by (metis IntD2 Int_Diff inf.orderE openin_subset openin_subtopology)
ultimately have "openin (subtopology X (topspace X - S)) A" "openin (subtopology X (topspace X - S)) B"
using § 𝒲_def by blast+
with ‹disjnt A B› show "separatedin (subtopology X (topspace X - S)) A B"
using separatedin_open_sets by blast
qed
then show ?thesis
by (simp add: pairwise_def separatedin_subtopology)
qed
ultimately show False
by (metis con eqpoll_iff_finite_card)
qed
obtain 𝒱 where "𝒱 ≈ {..<n} " "{} ∉ 𝒱"
and pw𝒱: "pairwise (separatedin X) 𝒱" and UV: "⋃𝒱 = topspace X - (topspace X - ⋃𝒰)"
proof -
have "closedin X (topspace X - ⋃𝒰)"
using ope𝒰 by blast
moreover have "C ⊆ topspace X - ⋃𝒰"
using ‹C ⊆ topspace X› connected_components_of_subset by (fastforce simp: 𝒰_def)
moreover have "topspace X - ⋃𝒰 ⊆ S"
using Union_connected_components_of [of "subtopology X (topspace X - C)"] ‹C ⊆ S›
by (auto simp: 𝒰_def)
ultimately show thesis
by (metis C that)
qed
have "𝒱 ≲ 𝒰"
proof (rule lepoll_relational_full)
have "⋃𝒱 = ⋃𝒰"
by (simp add: Sup_le_iff UV double_diff ope𝒰 openin_subset)
then show "∃U. U ∈ 𝒰 ∧ ¬ disjnt U V" if "V ∈ 𝒱" for V
using that
by (metis ‹{} ∉ 𝒱› disjnt_Union1 disjnt_self_iff_empty)
show "C1 = C2"
if "T ∈ 𝒰" and "C1 ∈ 𝒱" and "C2 ∈ 𝒱" and "¬ disjnt T C1" and "¬ disjnt T C2" for T C1 C2
proof (cases "C1=C2")
case False
then have "connectedin X T"
using 𝒰_def connectedin_connected_components_of connectedin_subtopology ‹T ∈ 𝒰› by blast
have "T ⊆ C1 ∪ ⋃(𝒱 - {C1})"
using ‹⋃𝒱 = ⋃𝒰› ‹T ∈ 𝒰› by auto
with ‹connectedin X T›
have "¬ separatedin X C1 (⋃(𝒱 - {C1}))"
unfolding connectedin_eq_not_separated_subset
by (smt (verit) that False disjnt_def UnionI disjnt_iff insertE insert_Diff)
with that show ?thesis
by (metis (no_types, lifting) ‹𝒱 ≈ {..<n}› eqpoll_iff_finite_card finite_Diff pairwiseD pairwise_alt pw𝒱 separatedin_Union(1) separatedin_def)
qed auto
qed
then show False
by (metis ‹𝒱 ≈ {..<n}› card𝒰 eqpoll_iff_finite_card leD lepoll_iff_card_le)
qed
qed
qed
lemma separation_by_closed_intermediates_eq_gen:
assumes "locally_connected_space X" "hereditarily normal_space X"
shows "¬ connectedin X (topspace X - S) ⟷
(∃C. closedin X C ∧ C ⊆ S ∧
(∀D. closedin X D ∧ C ⊆ D ∧ D ⊆ S ⟶ ¬ connectedin X (topspace X - D)))"
(is "?lhs = ?rhs")
proof -
have *: "(∃𝒰::'a set set. 𝒰 ≈ {..<Suc (Suc 0)} ∧ P 𝒰) ⟷ (∃A B. A≠B ∧ P{A,B})" for P
by (metis One_nat_def eqpoll_doubleton_iff lessThan_Suc lessThan_empty_iff zero_neq_one)
have *: "(∃C1 C2. separatedin X C1 C2 ∧ C1≠C2 ∧ C1≠{} ∧ C2≠{} ∧ C1 ∪ C2 = topspace X - S) ⟷
(∃C. closedin X C ∧ C ⊆ S ∧
(∀D. closedin X D ∧ C ⊆ D ∧ D ⊆ S
⟶ (∃C1 C2. separatedin X C1 C2 ∧ C1≠C2 ∧ C1≠{} ∧ C2≠{} ∧ C1 ∪ C2 = topspace X - D)))"
using separation_by_closed_intermediates_eq_count [OF assms, of "Suc(Suc 0)" S]
apply (simp add: * pairwise_insert separatedin_sym cong: conj_cong)
apply (simp add: eq_sym_conv conj_ac)
done
with separatedin_refl
show ?thesis
apply (simp add: connectedin_eq_not_separated)
by (smt (verit, best) separatedin_refl)
qed
lemma lepoll_connnected_components_connectedin:
assumes "⋀C. C ∈ 𝒰 ⟹ connectedin X C" "⋃𝒰 = topspace X"
shows "connected_components_of X ≲ 𝒰"
proof -
have "connected_components_of X ≲ 𝒰 - {{}}"
proof (rule lepoll_relational_full)
show "∃U. U ∈ 𝒰 - {{}} ∧ U ⊆ V"
if "V ∈ connected_components_of X" for V
using that unfolding connected_components_of_def image_iff
by (metis Union_iff assms connected_component_of_maximal empty_iff insert_Diff_single insert_iff)
show "V = V'"
if "U ∈ 𝒰 - {{}}" "V ∈ connected_components_of X" "V' ∈ connected_components_of X" "U ⊆ V" "U ⊆ V'"
for U V V'
by (metis DiffD2 disjointD insertCI le_inf_iff pairwise_disjoint_connected_components_of subset_empty that)
qed
also have "… ≲ 𝒰"
by (simp add: subset_imp_lepoll)
finally show ?thesis .
qed
lemma lepoll_connected_components_alt:
"{..<n::nat} ≲ connected_components_of X ⟷
n = 0 ∨ (∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X)"
(is "?lhs ⟷ ?rhs")
proof (cases "n=0")
next
case False
show ?thesis
proof
assume L: ?lhs
with False show ?rhs
proof (induction n rule: less_induct)
case (less n)
show ?case
proof (cases "n≤1")
case True
with less.prems have "topspace X ≠ {}" "n=1"
by (fastforce simp add: connected_components_of_def)+
then have "{} ∉ {topspace X}"
by blast
with ‹n=1› show ?thesis
by (simp add: eqpoll_iff_finite_card card_Suc_eq flip: ex_simps)
next
case False
then have "n-1 ≠ 0"
by linarith
have n1_lesspoll: "{..<n-1} ≺ {..<n}"
using False lesspoll_iff_finite_card by fastforce
also have "… ≲ connected_components_of X"
using less by blast
finally have "{..<n-1} ≲ connected_components_of X"
using lesspoll_imp_lepoll by blast
then obtain 𝒰 where Ueq: "𝒰 ≈ {..<n-1}" and "{} ∉ 𝒰"
and pwU: "pairwise (separatedin X) 𝒰" and UU: "⋃𝒰 = topspace X"
by (meson ‹n - 1 ≠ 0› diff_less gr0I less zero_less_one)
show ?thesis
proof (cases "∀C ∈ 𝒰. connectedin X C")
case True
then show ?thesis
using lepoll_connnected_components_connectedin [of 𝒰 X] less.prems
by (metis UU Ueq lepoll_antisym lepoll_trans lepoll_trans2 lesspoll_def n1_lesspoll)
next
case False
with UU obtain C A B where ABC: "C ∈ 𝒰" "A ∪ B = C" "A ≠ {}" "B ≠ {}" and sep: "separatedin X A B"
by (fastforce simp add: connectedin_eq_not_separated)
define 𝒱 where "𝒱 ≡ insert A (insert B (𝒰 - {C}))"
have "𝒱 ≈ {..<n}"
proof -
have "A ≠ B"
using ‹B ≠ {}› sep by auto
moreover obtain "A ∉ 𝒰" "B ∉ 𝒰"
using pwU unfolding pairwise_def
by (metis ABC sep separatedin_Un(1) separatedin_refl separatedin_sym)
moreover have "card 𝒰 = n-1" "finite 𝒰"
using Ueq eqpoll_iff_finite_card by blast+
ultimately
have "card (insert A (insert B (𝒰 - {C}))) = n"
using ‹C ∈ 𝒰› by (auto simp add: card_insert_if)
then show ?thesis
using 𝒱_def ‹finite 𝒰› eqpoll_iff_finite_card by blast
qed
moreover have "{} ∉ 𝒱"
using ABC 𝒱_def ‹{} ∉ 𝒰› by blast
moreover have "⋃𝒱 = topspace X"
using ABC UU 𝒱_def by auto
moreover have "pairwise (separatedin X) 𝒱"
using pwU sep ABC unfolding 𝒱_def
apply (simp add: separatedin_sym pairwise_def)
by (metis member_remove remove_def separatedin_Un(1))
ultimately show ?thesis
by blast
qed
qed
qed
next
assume ?rhs
then obtain 𝒰 where "𝒰 ≈ {..<n}" "{} ∉ 𝒰" and pwU: "pairwise (separatedin X) 𝒰" and UU: "⋃𝒰 = topspace X"
using False by force
have "card (connected_components_of X) ≥ n" if "finite (connected_components_of X)"
proof -
have "𝒰 ≲ connected_components_of X"
proof (rule lepoll_relational_full)
show "∃T. T ∈ connected_components_of X ∧ ¬ disjnt T C" if "C ∈ 𝒰" for C
by (metis that UU Union_connected_components_of Union_iff ‹{} ∉ 𝒰› disjnt_iff equals0I)
show "(C1::'a set) = C2"
if "T ∈ connected_components_of X" and "C1 ∈ 𝒰" "C2 ∈ 𝒰" "¬ disjnt T C1" "¬ disjnt T C2" for T C1 C2
proof (rule ccontr)
assume "C1 ≠ C2"
then have "connectedin X T"
by (simp add: connectedin_connected_components_of that(1))
moreover have "¬ separatedin X C1 (⋃(𝒰 - {C1}))"
using ‹connectedin X T› pwU unfolding pairwise_def
by (smt (verit) Sup_upper UU Union_connected_components_of ‹C1 ≠ C2› complete_lattice_class.Sup_insert connectedin_subset_separated_union disjnt_subset2 disjnt_sym insert_Diff separatedin_imp_disjoint that)
ultimately show False
using ‹𝒰 ≈ {..<n}›
apply (simp add: connectedin_eq_not_separated_subset eqpoll_iff_finite_card)
by (metis Sup_upper UU finite_Diff pairwise_alt pwU separatedin_Union(1) that(2))
qed
qed
then show ?thesis
by (metis ‹𝒰 ≈ {..<n}› eqpoll_iff_finite_card lepoll_iff_card_le that)
qed
then show ?lhs
by (metis card_lessThan finite_lepoll_infinite finite_lessThan lepoll_iff_card_le)
qed
qed auto
subsection‹A perfect set in common cases must have at least the cardinality of the continuum›
lemma (in Metric_space) lepoll_perfect_set:
assumes "mcomplete"
and "mtopology derived_set_of S = S" "S ≠ {}"
shows "(UNIV::real set) ≲ S"
proof -
have "S ⊆ M"
using assms(2) derived_set_of_infinite_mball by blast
have "(UNIV::real set) ≲ (UNIV::nat set set)"
using eqpoll_imp_lepoll eqpoll_sym nat_sets_eqpoll_reals by blast
also have "… ≲ S"
proof -
have "∃y z δ. y ∈ S ∧ z ∈ S ∧ 0 < δ ∧ δ < ε/2 ∧
mcball y δ ⊆ mcball x ε ∧ mcball z δ ⊆ mcball x ε ∧ disjnt (mcball y δ) (mcball z δ)"
if "x ∈ S" "0 < ε" for x ε
proof -
define S' where "S' ≡ S ∩ mball x (ε/4)"
have "infinite S'"
using derived_set_of_infinite_mball [of S] assms that S'_def
by (smt (verit, ccfv_SIG) mem_Collect_eq zero_less_divide_iff)
then have "⋀x y z. ¬ (S' ⊆ {x,y,z})"
using finite_subset by auto
then obtain l r where lr: "l ∈ S'" "r ∈ S'" "l≠r" "l≠x" "r≠x"
by (metis insert_iff subsetI)
show ?thesis
proof (intro exI conjI)
show "l ∈ S" "r ∈ S" "d l r / 3 > 0"
using lr by (auto simp: S'_def)
show "d l r / 3 < ε/2" "mcball l (d l r / 3) ⊆ mcball x ε" "mcball r (d l r / 3) ⊆ mcball x ε"
using lr by (clarsimp simp: S'_def, smt (verit) commute triangle'')+
show "disjnt (mcball l (d l r / 3)) (mcball r (d l r / 3))"
using lr by (simp add: S'_def disjnt_iff) (smt (verit, best) mdist_pos_less triangle')
qed
qed
then obtain l r δ
where lrS: "⋀x ε. ⟦x ∈ S; 0 < ε⟧ ⟹ l x ε ∈ S ∧ r x ε ∈ S"
and δ: "⋀x ε. ⟦x ∈ S; 0 < ε⟧ ⟹ 0 < δ x ε ∧ δ x ε < ε/2"
and "⋀x ε. ⟦x ∈ S; 0 < ε⟧ ⟹ mcball (l x ε) (δ x ε) ⊆ mcball x ε ∧ mcball (r x ε) (δ x ε) ⊆ mcball x ε ∧
disjnt (mcball (l x ε) (δ x ε)) (mcball (r x ε) (δ x ε))"
by metis
then have lr_mcball: "⋀x ε. ⟦x ∈ S; 0 < ε⟧ ⟹ mcball (l x ε) (δ x ε) ⊆ mcball x ε ∧ mcball (r x ε) (δ x ε) ⊆ mcball x ε "
and lr_disjnt: "⋀x ε. ⟦x ∈ S; 0 < ε⟧ ⟹ disjnt (mcball (l x ε) (δ x ε)) (mcball (r x ε) (δ x ε))"
by metis+
obtain a where "a ∈ S"
using ‹S ≠ {}› by blast
define xe where "xe ≡
λB. rec_nat (a,1) (λn (x,γ). ((if n∈B then r else l) x γ, δ x γ))"
have [simp]: "xe b 0 = (a,1)" for b
by (simp add: xe_def)
have "xe B (Suc n) = (let (x,γ) = xe B n in ((if n∈B then r else l) x γ, δ x γ))" for B n
by (simp add: xe_def)
define x where "x ≡ λB n. fst (xe B n)"
define γ where "γ ≡ λB n. snd (xe B n)"
have [simp]: "x B 0 = a" "γ B 0 = 1" for B
by (simp_all add: x_def γ_def xe_def)
have x_Suc[simp]: "x B (Suc n) = ((if n∈B then r else l) (x B n) (γ B n))"
and γ_Suc[simp]: "γ B (Suc n) = δ (x B n) (γ B n)" for B n
by (simp_all add: x_def γ_def xe_def split: prod.split)
interpret Submetric M d S
proof qed (use ‹S ⊆ M› in metis)
have "closedin mtopology S"
by (metis assms(2) closure_of closure_of_eq inf.absorb_iff2 subset subset_Un_eq subset_refl topspace_mtopology)
with ‹mcomplete›
have "sub.mcomplete"
by (metis closedin_mcomplete_imp_mcomplete)
have *: "x B n ∈ S ∧ γ B n > 0" for B n
by (induction n) (auto simp: ‹a ∈ S› lrS δ)
with subset have E: "x B n ∈ M" for B n
by blast
have γ_le: "γ B n ≤ (1/2)^n" for B n
proof(induction n)
case 0 then show ?case by auto
next
case (Suc n)
then show ?case
by simp (smt (verit) "*" δ field_sum_of_halves)
qed
{ fix B
have "⋀n. sub.mcball (x B (Suc n)) (γ B (Suc n)) ⊆ sub.mcball (x B n) (γ B n)"
by (smt (verit, best) "*" Int_iff γ_Suc x_Suc in_mono lr_mcball mcball_submetric_eq subsetI)
then have mon: "monotone (≤) (λx y. y ⊆ x) (λn. sub.mcball (x B n) (γ B n))"
by (simp add: decseq_SucI)
have "∃n a. sub.mcball (x B n) (γ B n) ⊆ sub.mcball a ε" if "ε>0" for ε
proof -
obtain n where "(1/2)^n < ε"
using ‹0 < ε› real_arch_pow_inv by force
with γ_le have ε: "γ B n ≤ ε"
by (smt (verit))
show ?thesis
proof (intro exI)
show "sub.mcball (x B n) (γ B n) ⊆ sub.mcball (x B n) ε"
by (simp add: ε sub.mcball_subset_concentric)
qed
qed
then have "∃l. l ∈ S ∧ (⋂n. sub.mcball (x B n) (γ B n)) = {l}"
using ‹sub.mcomplete› mon
unfolding sub.mcomplete_nest_sing
apply (drule_tac x="λn. sub.mcball (x B n) (γ B n)" in spec)
by (meson * order.asym sub.closedin_mcball sub.mcball_eq_empty)
}
then obtain z where z: "⋀B. z B ∈ S ∧ (⋂n. sub.mcball (x B n) (γ B n)) = {z B}"
by metis
show ?thesis
unfolding lepoll_def
proof (intro exI conjI)
show "inj z"
proof (rule inj_onCI)
fix B C
assume eq: "z B = z C" and "B ≠ C"
then have ne: "sym_diff B C ≠ {}"
by blast
define n where "n ≡ LEAST k. k ∈ (sym_diff B C)"
with ne have n: "n ∈ sym_diff B C"
by (metis Inf_nat_def1 LeastI)
then have non: "n ∈ B ⟷ n ∉ C"
by blast
have H: "z C ∈ sub.mcball (x B (Suc n)) (γ B (Suc n)) ∧ z C ∈ sub.mcball (x C (Suc n)) (γ C (Suc n))"
using z [of B] z [of C] apply (simp add: lrS set_eq_iff non *)
by (smt (verit, best) γ_Suc eq non x_Suc)
have "k ∈ B ⟷ k ∈ C" if "k<n" for k
using that unfolding n_def by (meson DiffI UnCI not_less_Least)
moreover have "(∀m. m < p ⟶ (m ∈ B ⟷ m ∈ C)) ⟹ x B p = x C p ∧ γ B p = γ C p" for p
by (induction p) auto
ultimately have "x B n = x C n" "γ B n = γ C n"
by blast+
then show False
using lr_disjnt * H non
by (smt (verit) IntD2 γ_Suc disjnt_iff mcball_submetric_eq x_Suc)
qed
show "range z ⊆ S"
using z by blast
qed
qed
finally show ?thesis .
qed
lemma lepoll_perfect_set_aux:
assumes lcX: "locally_compact_space X" and hsX: "Hausdorff_space X"
and eq: "X derived_set_of topspace X = topspace X" and "topspace X ≠ {}"
shows "(UNIV::real set) ≲ topspace X"
proof -
have "(UNIV::real set) ≲ (UNIV::nat set set)"
using eqpoll_imp_lepoll eqpoll_sym nat_sets_eqpoll_reals by blast
also have "… ≲ topspace X"
proof -
obtain z where z: "z ∈ topspace X"
using assms by blast
then obtain U K where "openin X U" "compactin X K" "U ≠ {}" "U ⊆ K"
by (metis emptyE lcX locally_compact_space_def)
then have "closedin X K"
by (simp add: compactin_imp_closedin hsX)
have intK_ne: "X interior_of K ≠ {}"
using ‹U ≠ {}› ‹U ⊆ K› ‹openin X U› interior_of_eq_empty by blast
have "∃D E. closedin X D ∧ D ⊆ K ∧ X interior_of D ≠ {} ∧
closedin X E ∧ E ⊆ K ∧ X interior_of E ≠ {} ∧
disjnt D E ∧ D ⊆ C ∧ E ⊆ C"
if "closedin X C" "C ⊆ K" and C: "X interior_of C ≠ {}" for C
proof -
obtain z where z: "z ∈ X interior_of C" "z ∈ topspace X"
using C interior_of_subset_topspace by fastforce
obtain x y where "x ∈ X interior_of C" "y ∈ X interior_of C" "x≠y"
by (metis z eq in_derived_set_of openin_interior_of)
then have "x ∈ topspace X" "y ∈ topspace X"
using interior_of_subset_topspace by force+
with hsX obtain V W where "openin X V" "openin X W" "x ∈ V" "y ∈ W" "disjnt V W"
by (metis Hausdorff_space_def ‹x ≠ y›)
have *: "⋀W x. openin X W ∧ x ∈ W
⟹ ∃U V. openin X U ∧ closedin X V ∧ x ∈ U ∧ U ⊆ V ∧ V ⊆ W"
using lcX hsX locally_compact_Hausdorff_imp_regular_space neighbourhood_base_of_closedin neighbourhood_base_of
by metis
obtain M D where MD: "openin X M" "closedin X D" "y ∈ M" "M ⊆ D" "D ⊆ X interior_of C ∩ W"
using * [of "X interior_of C ∩ W" y]
using ‹openin X W› ‹y ∈ W› ‹y ∈ X interior_of C› by fastforce
obtain N E where NE: "openin X N" "closedin X E" "x ∈ N" "N ⊆ E" "E ⊆ X interior_of C ∩ V"
using * [of "X interior_of C ∩ V" x]
using ‹openin X V› ‹x ∈ V› ‹x ∈ X interior_of C› by fastforce
show ?thesis
proof (intro exI conjI)
show "X interior_of D ≠ {}" "X interior_of E ≠ {}"
using MD NE by (fastforce simp: interior_of_def)+
show "disjnt D E"
by (meson MD(5) NE(5) ‹disjnt V W› disjnt_subset1 disjnt_sym le_inf_iff)
qed (use MD NE ‹C ⊆ K› interior_of_subset in force)+
qed
then obtain L R where
LR: "⋀C. ⟦closedin X C; C ⊆ K; X interior_of C ≠ {}⟧
⟹ closedin X (L C) ∧ (L C) ⊆ K ∧ X interior_of (L C) ≠ {} ∧
closedin X (R C) ∧ (R C) ⊆ K ∧ X interior_of (R C) ≠ {}"
and disjLR: "⋀C. ⟦closedin X C; C ⊆ K; X interior_of C ≠ {}⟧
⟹ disjnt (L C) (R C) ∧ (L C) ⊆ C ∧ (R C) ⊆ C"
by metis
define d where "d ≡ λB. rec_nat K (λn. if n ∈ B then R else L)"
have d0[simp]: "d B 0 = K" for B
by (simp add: d_def)
have [simp]: "d B (Suc n) = (if n ∈ B then R else L) (d B n)" for B n
by (simp add: d_def)
have d_correct: "closedin X (d B n) ∧ d B n ⊆ K ∧ X interior_of (d B n) ≠ {}" for B n
proof (induction n)
case 0
then show ?case by (auto simp: ‹closedin X K› intK_ne)
next
case (Suc n) with LR show ?case by auto
qed
have "(⋂n. d B n) ≠ {}" for B
proof (rule compact_space_imp_nest)
show "compact_space (subtopology X K)"
by (simp add: ‹compactin X K› compact_space_subtopology)
show "closedin (subtopology X K) (d B n)" for n :: nat
by (simp add: closedin_subset_topspace d_correct)
show "d B n ≠ {}" for n :: nat
by (metis d_correct interior_of_empty)
show "antimono (d B)"
proof (rule antimonoI [OF transitive_stepwise_le])
fix n
show "d B (Suc n) ⊆ d B n"
by (simp add: d_correct disjLR)
qed auto
qed
then obtain x where x: "⋀B. x B ∈ (⋂n. d B n)"
unfolding set_eq_iff by (metis empty_iff)
show ?thesis
unfolding lepoll_def
proof (intro exI conjI)
show "inj x"
proof (rule inj_onCI)
fix B C
assume eq: "x B = x C" and "B≠C"
then have ne: "sym_diff B C ≠ {}"
by blast
define n where "n ≡ LEAST k. k ∈ (sym_diff B C)"
with ne have n: "n ∈ sym_diff B C"
by (metis Inf_nat_def1 LeastI)
then have non: "n ∈ B ⟷ n ∉ C"
by blast
have "k ∈ B ⟷ k ∈ C" if "k<n" for k
using that unfolding n_def by (meson DiffI UnCI not_less_Least)
moreover have "(∀m. m < p ⟶ (m ∈ B ⟷ m ∈ C)) ⟹ d B p = d C p" for p
by (induction p) auto
ultimately have "d B n = d C n"
by blast
then have "disjnt (d B (Suc n)) (d C (Suc n))"
by (simp add: d_correct disjLR disjnt_sym non)
then show False
by (metis InterE disjnt_iff eq rangeI x)
qed
show "range x ⊆ topspace X"
using x d0 ‹compactin X K› compactin_subset_topspace d_correct by fastforce
qed
qed
finally show ?thesis .
qed
lemma lepoll_perfect_set:
assumes X: "completely_metrizable_space X ∨ locally_compact_space X ∧ Hausdorff_space X"
and S: "X derived_set_of S = S" "S ≠ {}"
shows "(UNIV::real set) ≲ S"
using X
proof
assume "completely_metrizable_space X"
with assms show "(UNIV::real set) ≲ S"
by (metis Metric_space.lepoll_perfect_set completely_metrizable_space_def)
next
assume "locally_compact_space X ∧ Hausdorff_space X"
then show "(UNIV::real set) ≲ S"
using lepoll_perfect_set_aux [of "subtopology X S"]
by (metis Hausdorff_space_subtopology S closedin_derived_set_of closedin_subset derived_set_of_subtopology
locally_compact_space_closed_subset subtopology_topspace topspace_subtopology topspace_subtopology_subset)
qed
lemma Kuratowski_aux1:
assumes "⋀S T. R S T ⟹ R T S"
shows "(∀S T n. R S T ⟶ (f S ≈ {..<n::nat} ⟷ f T ≈ {..<n::nat})) ⟷
(∀n S T. R S T ⟶ {..<n::nat} ≲ f S ⟶ {..<n::nat} ≲ f T)"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
by (meson eqpoll_iff_finite_card eqpoll_sym finite_lepoll_infinite finite_lessThan lepoll_trans2)
next
assume ?rhs then show ?lhs
by (smt (verit, best) lepoll_iff_finite_card assms eqpoll_iff_finite_card finite_lepoll_infinite
finite_lessThan le_Suc_eq lepoll_antisym lepoll_iff_card_le not_less_eq_eq)
qed
lemma Kuratowski_aux2:
"pairwise (separatedin (subtopology X (topspace X - S))) 𝒰 ∧ {} ∉ 𝒰 ∧
⋃𝒰 = topspace(subtopology X (topspace X - S)) ⟷
pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - S"
by (auto simp: pairwise_def separatedin_subtopology)
proposition Kuratowski_component_number_invariance_aux:
assumes "compact_space X" and HsX: "Hausdorff_space X"
and lcX: "locally_connected_space X" and hnX: "hereditarily normal_space X"
and hom: "(subtopology X S) homeomorphic_space (subtopology X T)"
and leXS: "{..<n::nat} ≲ connected_components_of (subtopology X (topspace X - S))"
assumes §: "⋀S T.
⟦closedin X S; closedin X T; (subtopology X S) homeomorphic_space (subtopology X T);
{..<n::nat} ≲ connected_components_of (subtopology X (topspace X - S))⟧
⟹ {..<n::nat} ≲ connected_components_of (subtopology X (topspace X - T))"
shows "{..<n::nat} ≲ connected_components_of (subtopology X (topspace X - T))"
proof (cases "n=0")
case False
obtain f g where homf: "homeomorphic_map (subtopology X S) (subtopology X T) f"
and homg: "homeomorphic_map (subtopology X T) (subtopology X S) g"
and gf: "⋀x. x ∈ topspace (subtopology X S) ⟹ g(f x) = x"
and fg: "⋀y. y ∈ topspace (subtopology X T) ⟹ f(g y) = y"
and f: "f ∈ topspace (subtopology X S) → topspace (subtopology X T)"
and g: "g ∈ topspace (subtopology X T) → topspace (subtopology X S)"
using homeomorphic_space_unfold hom by metis
obtain C where "closedin X C" "C ⊆ S"
and C: "⋀D. ⟦closedin X D; C ⊆ D; D ⊆ S⟧
⟹ ∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - D"
using separation_by_closed_intermediates_eq_count [of X n S] assms
by (smt (verit, ccfv_threshold) False Kuratowski_aux2 lepoll_connected_components_alt)
have "∃C. closedin X C ∧ C ⊆ T ∧
(∀D. closedin X D ∧ C ⊆ D ∧ D ⊆ T
⟶ (∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧
{} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - D))"
proof (intro exI, intro conjI strip)
have "compactin X (f ` C)"
by (meson ‹C ⊆ S› ‹closedin X C› assms(1) closedin_compact_space compactin_subtopology homeomorphic_map_compactness_eq homf)
then show "closedin X (f ` C)"
using ‹Hausdorff_space X› compactin_imp_closedin by blast
show "f ` C ⊆ T"
by (meson ‹C ⊆ S› ‹closedin X C› closedin_imp_subset closedin_subset_topspace homeomorphic_map_closedness_eq homf)
fix D'
assume D': "closedin X D' ∧ f ` C ⊆ D' ∧ D' ⊆ T"
define D where "D ≡ g ` D'"
have "compactin X D"
unfolding D_def
by (meson D' ‹compact_space X› closedin_compact_space compactin_subtopology homeomorphic_map_compactness_eq homg)
then have "closedin X D"
by (simp add: assms(2) compactin_imp_closedin)
moreover have "C ⊆ D"
using D' D_def ‹C ⊆ S› ‹closedin X C› closedin_subset gf image_iff by fastforce
moreover have "D ⊆ S"
by (metis D' D_def assms(1) closedin_compact_space compactin_subtopology homeomorphic_map_compactness_eq homg)
ultimately obtain 𝒰 where "𝒰 ≈ {..<n}" "pairwise (separatedin X) 𝒰" "{} ∉ 𝒰" "⋃𝒰 = topspace X - D"
using C by meson
moreover have "(subtopology X D) homeomorphic_space (subtopology X D')"
unfolding homeomorphic_space_def
proof (intro exI)
have "subtopology X D = subtopology (subtopology X S) D"
by (simp add: ‹D ⊆ S› inf.absorb2 subtopology_subtopology)
moreover have "subtopology X D' = subtopology (subtopology X T) D'"
by (simp add: D' inf.absorb2 subtopology_subtopology)
moreover have "homeomorphic_maps (subtopology X T) (subtopology X S) g f"
by (simp add: fg gf homeomorphic_maps_map homf homg)
ultimately
have "homeomorphic_maps (subtopology X D') (subtopology X D) g f"
by (metis D' D_def ‹closedin X D› closedin_subset homeomorphic_maps_subtopologies topspace_subtopology Int_absorb1)
then show "homeomorphic_maps (subtopology X D) (subtopology X D') f g"
using homeomorphic_maps_sym by blast
qed
ultimately show "∃𝒰. 𝒰 ≈ {..<n} ∧ pairwise (separatedin X) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃ 𝒰 = topspace X - D'"
by (smt (verit, ccfv_SIG) § D' False ‹closedin X D› Kuratowski_aux2 lepoll_connected_components_alt)
qed
then have "∃𝒰. 𝒰 ≈ {..<n} ∧
pairwise (separatedin (subtopology X (topspace X - T))) 𝒰 ∧ {} ∉ 𝒰 ∧ ⋃𝒰 = topspace X - T"
using separation_by_closed_intermediates_eq_count [of X n T] Kuratowski_aux2 lcX hnX by auto
with False show ?thesis
using lepoll_connected_components_alt by fastforce
qed auto
theorem Kuratowski_component_number_invariance:
assumes "compact_space X" "Hausdorff_space X" "locally_connected_space X" "hereditarily normal_space X"
shows "((∀S T n.
closedin X S ∧ closedin X T ∧
(subtopology X S) homeomorphic_space (subtopology X T)
⟶ (connected_components_of
(subtopology X (topspace X - S)) ≈ {..<n::nat} ⟷
connected_components_of
(subtopology X (topspace X - T)) ≈ {..<n::nat})) ⟷
(∀S T n.
(subtopology X S) homeomorphic_space (subtopology X T)
⟶ (connected_components_of
(subtopology X (topspace X - S)) ≈ {..<n::nat} ⟷
connected_components_of
(subtopology X (topspace X - T)) ≈ {..<n::nat})))"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
then show ?rhs
apply (subst (asm) Kuratowski_aux1, use homeomorphic_space_sym in blast)
apply (subst Kuratowski_aux1, use homeomorphic_space_sym in blast)
apply (blast intro: Kuratowski_component_number_invariance_aux assms)
done
qed blast
end