Theory Lindelof_Spaces
sectionβΉLindel\"of spacesβΊ
theory Lindelof_Spaces
imports T1_Spaces
begin
definition Lindelof_space where
  "Lindelof_space X β‘
        βπ°. (βU β π°. openin X U) β§ βπ° = topspace X
            βΆ (βπ±. countable π± β§ π± β π° β§ βπ± = topspace X)"
lemma Lindelof_spaceD:
  "β¦Lindelof_space X; βU. U β π° βΉ openin X U; βπ° = topspace Xβ§
  βΉ βπ±. countable π± β§ π± β π° β§ βπ± = topspace X"
  by (auto simp: Lindelof_space_def)
lemma Lindelof_space_alt:
   "Lindelof_space X β·
        (βπ°. (βU β π°. openin X U) β§ topspace X β βπ°
             βΆ (βπ±. countable π± β§ π± β π° β§ topspace X β βπ±))"
  unfolding Lindelof_space_def
  using openin_subset by fastforce
lemma compact_imp_Lindelof_space:
   "compact_space X βΉ Lindelof_space X"
  unfolding Lindelof_space_def compact_space
  by (meson uncountable_infinite)
lemma Lindelof_space_topspace_empty:
   "topspace X = {} βΉ Lindelof_space X"
  using compact_imp_Lindelof_space compact_space_topspace_empty by blast
lemma Lindelof_space_Union:
  assumes π°: "countable π°" and lin: "βU. U β π° βΉ Lindelof_space (subtopology X U)"
  shows "Lindelof_space (subtopology X (βπ°))"
proof -
  have "βπ±. countable π± β§ π± β β± β§ βπ° β© βπ± = topspace X β© βπ°"
    if β±: "β± β Collect (openin X)" and UF: "βπ° β© ββ± = topspace X β© βπ°"
    for β±
  proof -
    have "βU. β¦U β π°; U β© ββ± = topspace X β© Uβ§
               βΉ βπ±. countable π± β§ π± β β± β§ U β© βπ± = topspace X β© U"
      using lin β±
      unfolding Lindelof_space_def openin_subtopology_alt Ball_def subset_iff [symmetric]
      by (simp add: all_subset_image imp_conjL ex_countable_subset_image)
    then obtain g where g: "βU. β¦U β π°; U β© ββ± = topspace X β© Uβ§
                               βΉ countable (g U) β§ (g U) β β± β§ U β© β(g U) = topspace X β© U"
      by metis
    show ?thesis
    proof (intro exI conjI)
      show "countable (β(g ` π°))"
        using Int_commute UF g  by (fastforce intro: countable_UN [OF π°])
      show "β(g ` π°) β β±"
        using g UF by blast
      show "βπ° β© β(β(g ` π°)) = topspace X β© βπ°"
      proof
        show "βπ° β© β(β(g ` π°)) β topspace X β© βπ°"
          using g UF by blast
        show "topspace X β© βπ° β βπ° β© β(β(g ` π°))"
        proof clarsimp
          show "βyβπ°. βWβg y. x β W"
            if "x β topspace X" "x β V" "V β π°" for x V
          proof -
            have "V β© ββ± = topspace X β© V"
              using UF βΉV β π°βΊ by blast
            with that g [OF βΉV β π°βΊ]  show ?thesis by blast
          qed
        qed
      qed
    qed
  qed
  then show ?thesis
      unfolding Lindelof_space_def openin_subtopology_alt Ball_def subset_iff [symmetric]
      by (simp add: all_subset_image imp_conjL ex_countable_subset_image)
qed
lemma countable_imp_Lindelof_space:
  assumes "countable(topspace X)"
  shows "Lindelof_space X"
proof -
  have "Lindelof_space (subtopology X (βx β topspace X. {x}))"
  proof (rule Lindelof_space_Union)
    show "countable ((Ξ»x. {x}) ` topspace X)"
      using assms by blast
    show "Lindelof_space (subtopology X U)"
      if "U β (Ξ»x. {x}) ` topspace X" for U
    proof -
      have "compactin X U"
        using that by force
      then show ?thesis
        by (meson compact_imp_Lindelof_space compact_space_subtopology)
    qed
  qed
  then show ?thesis
    by simp
qed
lemma Lindelof_space_subtopology:
   "Lindelof_space(subtopology X S) β·
        (βπ°. (βU β π°. openin X U) β§ topspace X β© S β βπ°
            βΆ (βV. countable V β§ V β π° β§ topspace X β© S β βV))"
proof -
  have *: "(S β© βπ° = topspace X β© S) = (topspace X β© S β βπ°)"
    if "βx. x β π° βΉ openin X x" for π°
    by (blast dest: openin_subset [OF that])
  moreover have "(π± β π° β§ S β© βπ± = topspace X β© S) = (π± β π° β§ topspace X β© S β βπ±)"
    if "βx. x β π° βΆ openin X x" "topspace X β© S β βπ°" "countable π±" for π° π±
    using that * by blast
  ultimately show ?thesis
    unfolding Lindelof_space_def openin_subtopology_alt Ball_def
    apply (simp add: all_subset_image imp_conjL ex_countable_subset_image flip: subset_iff)
    apply (intro all_cong1 imp_cong ex_cong, auto)
    done
qed
lemma Lindelof_space_subtopology_subset:
   "S β topspace X
        βΉ (Lindelof_space(subtopology X S) β·
             (βπ°. (βU β π°. openin X U) β§ S β βπ°
                 βΆ (βV. countable V β§ V β π° β§ S β βV)))"
  by (metis Lindelof_space_subtopology topspace_subtopology topspace_subtopology_subset)
lemma Lindelof_space_closedin_subtopology:
  assumes X: "Lindelof_space X" and clo: "closedin X S"
  shows "Lindelof_space (subtopology X S)"
proof -
  have "S β topspace X"
    by (simp add: clo closedin_subset)
  then show ?thesis
  proof (clarsimp simp add: Lindelof_space_subtopology_subset)
    show "βV. countable V β§ V β β± β§ S β βV"
      if "βUββ±. openin X U" and "S β ββ±" for β±
    proof -
      have "βπ±. countable π± β§ π± β insert (topspace X - S) β± β§ βπ± = topspace X"
      proof (rule Lindelof_spaceD [OF X, of "insert (topspace X - S) β±"])
        show "openin X U"
          if "U β insert (topspace X - S) β±" for U
          using that βΉβUββ±. openin X UβΊ clo by blast
        show "β(insert (topspace X - S) β±) = topspace X"
          apply auto
          apply (meson in_mono openin_closedin_eq that(1))
          using UnionE βΉS β ββ±βΊ by auto
      qed
      then obtain π± where "countable π±" "π± β insert (topspace X - S) β±" "βπ± = topspace X"
        by metis
      with βΉS β topspace XβΊ
      show ?thesis
        by (rule_tac x="(π± - {topspace X - S})" in exI) auto
    qed
  qed
qed
lemma Lindelof_space_continuous_map_image:
  assumes X: "Lindelof_space X" and f: "continuous_map X Y f" and fim: "f ` (topspace X) = topspace Y"
  shows "Lindelof_space Y"
proof -
  have "βπ±. countable π± β§ π± β π° β§ βπ± = topspace Y"
    if π°: "βU. U β π° βΉ openin Y U" and UU: "βπ° = topspace Y" for π°
  proof -
    define π± where "π± β‘ (Ξ»U. {x β topspace X. f x β U}) ` π°"
    have "βV. V β π± βΉ openin X V"
      unfolding π±_def using π° continuous_map f by fastforce
    moreover have "βπ± = topspace X"
      unfolding π±_def using UU fim by fastforce
    ultimately have "βπ². countable π² β§ π² β π± β§ βπ² = topspace X"
      using X by (simp add: Lindelof_space_def)
    then obtain π where "countable π" "π β π°" and π: "(βUβπ. {x β topspace X. f x β U}) = topspace X"
      by (metis (no_types, lifting) π±_def countable_subset_image)
    moreover have "βπ = topspace Y"
    proof
      show "βπ β topspace Y"
        using UU π βΉπ β π°βΊ by fastforce
      have "y β βπ" if "y β topspace Y" for y
      proof -
        obtain x where "x β topspace X" "y = f x"
          using that fim by (metis βΉy β topspace YβΊ imageE)
        with π show ?thesis by auto
      qed
      then show "topspace Y β βπ" by blast
    qed
    ultimately show ?thesis
      by blast
  qed
  then show ?thesis
    unfolding Lindelof_space_def
    by auto
qed
lemma Lindelof_space_quotient_map_image:
   "β¦quotient_map X Y q; Lindelof_space Xβ§ βΉ Lindelof_space Y"
  by (meson Lindelof_space_continuous_map_image quotient_imp_continuous_map quotient_imp_surjective_map)
lemma Lindelof_space_retraction_map_image:
   "β¦retraction_map X Y r; Lindelof_space Xβ§ βΉ Lindelof_space Y"
  using Abstract_Topology.retraction_imp_quotient_map Lindelof_space_quotient_map_image by blast
lemma locally_finite_cover_of_Lindelof_space:
  assumes X: "Lindelof_space X" and UU: "topspace X β βπ°" and fin: "locally_finite_in X π°"
  shows "countable π°"
proof -
  have UU_eq: "βπ° = topspace X"
    by (meson UU fin locally_finite_in_def subset_antisym)
  obtain T where T: "βx. x β topspace X βΉ openin X (T x) β§ x β T x β§ finite {U β π°. U β© T x β  {}}"
    using fin unfolding locally_finite_in_def by metis
  then obtain I where "countable I" "I β topspace X" and I: "topspace X β β(T ` I)"
    using X unfolding Lindelof_space_alt
    by (drule_tac x="image T (topspace X)" in spec) (auto simp: ex_countable_subset_image)
  show ?thesis
  proof (rule countable_subset)
    have "βi. i β I βΉ countable {U β π°. U β© T i β  {}}"
      using T
      by (meson βΉI β topspace XβΊ in_mono uncountable_infinite)
    then show "countable (insert {} (βiβI. {U β π°. U β© T i β  {}}))"
      by (simp add: βΉcountable IβΊ)
  qed (use UU_eq I in auto)
qed
lemma Lindelof_space_proper_map_preimage:
  assumes f: "proper_map X Y f" and Y: "Lindelof_space Y"
  shows "Lindelof_space X"
proof (clarsimp simp: Lindelof_space_alt)
  show "βπ±. countable π± β§ π± β π° β§ topspace X β βπ±"
    if π°: "βUβπ°. openin X U" and sub_UU: "topspace X β βπ°" for π°
  proof -
    have "βπ±. finite π± β§ π± β π° β§ {x β topspace X. f x = y} β βπ±" if "y β topspace Y" for y
    proof (rule compactinD)
      show "compactin X {x β topspace X. f x = y}"
        using f proper_map_def that by fastforce
    qed (use sub_UU π° in auto)
    then obtain π± where π±: "βy. y β topspace Y βΉ finite (π± y) β§ π± y β π° β§ {x β topspace X. f x = y} β β(π± y)"
      by meson
    define π² where "π² β‘ (Ξ»y. topspace Y - image f (topspace X - β(π± y))) ` topspace Y"
    have "βU β π². openin Y U"
      using f π° π± unfolding π²_def proper_map_def closed_map_def
      by (simp add: closedin_diff openin_Union openin_diff subset_iff)
    moreover have "topspace Y β βπ²"
      using π± unfolding π²_def by clarsimp fastforce
    ultimately have "βπ±. countable π± β§ π± β π² β§ topspace Y β βπ±"
      using Y by (simp add: Lindelof_space_alt)
    then obtain I where "countable I" "I β topspace Y"
      and I: "topspace Y β (βiβI. topspace Y - f ` (topspace X - β(π± i)))"
      unfolding π²_def ex_countable_subset_image by metis
    show ?thesis
    proof (intro exI conjI)
      have "βi. i β I βΉ countable (π± i)"
        by (meson π± βΉI β topspace YβΊ in_mono uncountable_infinite)
      with βΉcountable IβΊ show "countable (β(π± ` I))"
        by auto
      show "β(π± ` I) β π°"
        using π± βΉI β topspace YβΊ by fastforce
      show "topspace X β β(β(π± ` I))"
      proof
        show "x β β (β (π± ` I))" if "x β topspace X" for x
        proof -
          have "f x β topspace Y"
            by (meson f image_subset_iff proper_map_imp_subset_topspace that)
          then show ?thesis
            using that I by auto
        qed
      qed
    qed
  qed
qed
lemma Lindelof_space_perfect_map_image:
   "β¦Lindelof_space X; perfect_map X Y fβ§ βΉ Lindelof_space Y"
  using Lindelof_space_quotient_map_image perfect_imp_quotient_map by blast
lemma Lindelof_space_perfect_map_image_eq:
   "perfect_map X Y f βΉ Lindelof_space X β· Lindelof_space Y"
  using Lindelof_space_perfect_map_image Lindelof_space_proper_map_preimage perfect_map_def by blast
end