|
| 1 | +/- |
| 2 | +Copyright (c) 2023 Christopher Hoskin. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Christopher Hoskin |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module topology.order.lower_topology |
| 7 | +! leanprover-community/mathlib commit 98e83c3d541c77cdb7da20d79611a780ff8e7d90 |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.Topology.Homeomorph |
| 12 | +import Mathlib.Topology.Order.Lattice |
| 13 | +import Mathlib.Order.Hom.CompleteLattice |
| 14 | + |
| 15 | +/-! |
| 16 | +# Lower topology |
| 17 | +
|
| 18 | +This file introduces the lower topology on a preorder as the topology generated by the complements |
| 19 | +of the closed intervals to infinity. |
| 20 | +
|
| 21 | +## Main statements |
| 22 | +
|
| 23 | +- `LowerTopology.t0Space` - the lower topology on a partial order is T₀ |
| 24 | +- `LowerTopology.isTopologicalBasis` - the complements of the upper closures of finite |
| 25 | + subsets form a basis for the lower topology |
| 26 | +- `LowerTopology.continuousInf` - the inf map is continuous with respect to the lower topology |
| 27 | +
|
| 28 | +## Implementation notes |
| 29 | +
|
| 30 | +A type synonym `WithLowerTopology` is introduced and for a preorder `α`, `WithLowerTopology α` |
| 31 | +is made an instance of `TopologicalSpace` by the topology generated by the complements of the |
| 32 | +closed intervals to infinity. |
| 33 | +
|
| 34 | +We define a mixin class `LowerTopology` for the class of types which are both a preorder and a |
| 35 | +topology and where the topology is generated by the complements of the closed intervals to infinity. |
| 36 | +It is shown that `WithLowerTopology α` is an instance of `LowerTopology`. |
| 37 | +
|
| 38 | +## Motivation |
| 39 | +
|
| 40 | +The lower topology is used with the `Scott` topology to define the Lawson topology. The restriction |
| 41 | +of the lower topology to the spectrum of a complete lattice coincides with the hull-kernel topology. |
| 42 | +
|
| 43 | +## References |
| 44 | +
|
| 45 | +* [Gierz et al, *A Compendium of Continuous Lattices*][GierzEtAl1980] |
| 46 | +
|
| 47 | +## Tags |
| 48 | +
|
| 49 | +lower topology, preorder |
| 50 | +-/ |
| 51 | + |
| 52 | + |
| 53 | +variable (α β : Type _) |
| 54 | + |
| 55 | +open Set TopologicalSpace |
| 56 | + |
| 57 | +/-- Type synonym for a preorder equipped with the lower topology. -/ |
| 58 | +def WithLowerTopology := α |
| 59 | +#align with_lower_topology WithLowerTopology |
| 60 | + |
| 61 | +variable {α β} |
| 62 | + |
| 63 | +namespace WithLowerTopology |
| 64 | + |
| 65 | +/-- `toLower` is the identity function to the `WithLowerTopology` of a type. -/ |
| 66 | +@[match_pattern] |
| 67 | +def toLower : α ≃ WithLowerTopology α := Equiv.refl _ |
| 68 | +#align with_lower_topology.to_lower WithLowerTopology.toLower |
| 69 | + |
| 70 | +/-- `ofLower` is the identity function from the `WithLowerTopology` of a type. -/ |
| 71 | +@[match_pattern] |
| 72 | +def ofLower : WithLowerTopology α ≃ α := Equiv.refl _ |
| 73 | +#align with_lower_topology.of_lower WithLowerTopology.ofLower |
| 74 | + |
| 75 | +@[simp] |
| 76 | +theorem to_withLowerTopology_symm_eq : (@toLower α).symm = ofLower := |
| 77 | + rfl |
| 78 | +#align with_lower_topology.to_with_lower_topology_symm_eq WithLowerTopology.to_withLowerTopology_symm_eq |
| 79 | + |
| 80 | +@[simp] |
| 81 | +theorem of_withLowerTopology_symm_eq : (@ofLower α).symm = toLower := |
| 82 | + rfl |
| 83 | +#align with_lower_topology.of_with_lower_topology_symm_eq WithLowerTopology.of_withLowerTopology_symm_eq |
| 84 | + |
| 85 | +@[simp] |
| 86 | +theorem toLower_ofLower (a : WithLowerTopology α) : toLower (ofLower a) = a := |
| 87 | + rfl |
| 88 | +#align with_lower_topology.to_lower_of_lower WithLowerTopology.toLower_ofLower |
| 89 | + |
| 90 | +@[simp] |
| 91 | +theorem ofLower_toLower (a : α) : ofLower (toLower a) = a := |
| 92 | + rfl |
| 93 | +#align with_lower_topology.of_lower_to_lower WithLowerTopology.ofLower_toLower |
| 94 | + |
| 95 | +-- porting note: removed @[simp] to make linter happy |
| 96 | +theorem toLower_inj {a b : α} : toLower a = toLower b ↔ a = b := |
| 97 | + Iff.rfl |
| 98 | +#align with_lower_topology.to_lower_inj WithLowerTopology.toLower_inj |
| 99 | + |
| 100 | +-- porting note: removed @[simp] to make linter happy |
| 101 | +theorem ofLower_inj {a b : WithLowerTopology α} : ofLower a = ofLower b ↔ a = b := |
| 102 | + Iff.rfl |
| 103 | +#align with_lower_topology.of_lower_inj WithLowerTopology.ofLower_inj |
| 104 | + |
| 105 | +/-- A recursor for `WithLowerTopology`. Use as `induction x using WithLowerTopology.rec`. -/ |
| 106 | +protected def rec {β : WithLowerTopology α → Sort _} (h : ∀ a, β (toLower a)) : ∀ a, β a := fun a => |
| 107 | + h (ofLower a) |
| 108 | +#align with_lower_topology.rec WithLowerTopology.rec |
| 109 | + |
| 110 | +instance [Nonempty α] : Nonempty (WithLowerTopology α) := |
| 111 | + ‹Nonempty α› |
| 112 | + |
| 113 | +instance [Inhabited α] : Inhabited (WithLowerTopology α) := |
| 114 | + ‹Inhabited α› |
| 115 | + |
| 116 | +variable [Preorder α] |
| 117 | + |
| 118 | +instance : Preorder (WithLowerTopology α) := |
| 119 | + ‹Preorder α› |
| 120 | + |
| 121 | +instance : TopologicalSpace (WithLowerTopology α) := |
| 122 | + generateFrom { s | ∃ a, Ici aᶜ = s } |
| 123 | + |
| 124 | +theorem isOpen_preimage_ofLower (S : Set α) : |
| 125 | + IsOpen (WithLowerTopology.ofLower ⁻¹' S) ↔ |
| 126 | + (generateFrom { s : Set α | ∃ a : α, Ici aᶜ = s }).IsOpen S := |
| 127 | + Iff.rfl |
| 128 | +#align with_lower_topology.is_open_preimage_of_lower WithLowerTopology.isOpen_preimage_ofLower |
| 129 | + |
| 130 | +theorem isOpen_def (T : Set (WithLowerTopology α)) : |
| 131 | + IsOpen T ↔ |
| 132 | + (generateFrom { s : Set α | ∃ a : α, Ici aᶜ = s }).IsOpen (WithLowerTopology.toLower ⁻¹' T) := |
| 133 | + Iff.rfl |
| 134 | +#align with_lower_topology.is_open_def WithLowerTopology.isOpen_def |
| 135 | + |
| 136 | +end WithLowerTopology |
| 137 | + |
| 138 | +/-- |
| 139 | +The lower topology is the topology generated by the complements of the closed intervals to infinity. |
| 140 | +-/ |
| 141 | +class LowerTopology (α : Type _) [t : TopologicalSpace α] [Preorder α] : Prop where |
| 142 | + topology_eq_lowerTopology : t = generateFrom { s | ∃ a, Ici aᶜ = s } |
| 143 | +#align lower_topology LowerTopology |
| 144 | + |
| 145 | +instance [Preorder α] : LowerTopology (WithLowerTopology α) := |
| 146 | + ⟨rfl⟩ |
| 147 | + |
| 148 | +namespace LowerTopology |
| 149 | + |
| 150 | +/-- The complements of the upper closures of finite sets are a collection of lower sets |
| 151 | +which form a basis for the lower topology. -/ |
| 152 | +def lowerBasis (α : Type _) [Preorder α] := |
| 153 | + { s : Set α | ∃ t : Set α, t.Finite ∧ (upperClosure t : Set α)ᶜ = s } |
| 154 | +#align lower_topology.lower_basis LowerTopology.lowerBasis |
| 155 | + |
| 156 | +section Preorder |
| 157 | + |
| 158 | +variable (α) |
| 159 | +variable [Preorder α] [TopologicalSpace α] [LowerTopology α] {s : Set α} |
| 160 | + |
| 161 | +lemma topology_eq : ‹_› = generateFrom { s | ∃ a : α, (Ici a)ᶜ = s } := topology_eq_lowerTopology |
| 162 | + |
| 163 | +variable {α} |
| 164 | + |
| 165 | +/-- If `α` is equipped with the lower topology, then it is homeomorphic to `WithLowerTopology α`. |
| 166 | +-/ |
| 167 | +def withLowerTopologyHomeomorph : WithLowerTopology α ≃ₜ α := |
| 168 | + WithLowerTopology.ofLower.toHomeomorphOfInducing ⟨by erw [topology_eq α, induced_id]; rfl⟩ |
| 169 | +#align lower_topology.with_lower_topology_homeomorph LowerTopology.withLowerTopologyHomeomorph |
| 170 | + |
| 171 | +theorem isOpen_iff_generate_Ici_compl : IsOpen s ↔ GenerateOpen { t | ∃ a, Ici aᶜ = t } s := by |
| 172 | + rw [topology_eq α]; rfl |
| 173 | +#align lower_topology.is_open_iff_generate_Ici_compl LowerTopology.isOpen_iff_generate_Ici_compl |
| 174 | + |
| 175 | +/-- Left-closed right-infinite intervals [a, ∞) are closed in the lower topology. -/ |
| 176 | +theorem isClosed_Ici (a : α) : IsClosed (Ici a) := |
| 177 | + isOpen_compl_iff.1 <| isOpen_iff_generate_Ici_compl.2 <| GenerateOpen.basic _ ⟨a, rfl⟩ |
| 178 | +#align lower_topology.is_closed_Ici LowerTopology.isClosed_Ici |
| 179 | + |
| 180 | +/-- The upper closure of a finite set is closed in the lower topology. -/ |
| 181 | +theorem isClosed_upperClosure (h : s.Finite) : IsClosed (upperClosure s : Set α) := by |
| 182 | + simp only [← UpperSet.infᵢ_Ici, UpperSet.coe_infᵢ] |
| 183 | + exact isClosed_bunionᵢ h fun a _ => isClosed_Ici a |
| 184 | +#align lower_topology.is_closed_upper_closure LowerTopology.isClosed_upperClosure |
| 185 | + |
| 186 | +/-- Every set open in the lower topology is a lower set. -/ |
| 187 | +theorem isLowerSet_of_isOpen (h : IsOpen s) : IsLowerSet s := by |
| 188 | + -- porting note: `rw` leaves a shadowed assumption |
| 189 | + replace h := isOpen_iff_generate_Ici_compl.1 h |
| 190 | + induction h |
| 191 | + case basic u h' => obtain ⟨a, rfl⟩ := h'; exact (isUpperSet_Ici a).compl |
| 192 | + case univ => exact isLowerSet_univ |
| 193 | + case inter u v _ _ hu2 hv2 => exact hu2.inter hv2 |
| 194 | + case unionₛ _ _ ih => exact isLowerSet_unionₛ ih |
| 195 | +#align lower_topology.is_lower_set_of_is_open LowerTopology.isLowerSet_of_isOpen |
| 196 | + |
| 197 | +theorem isUpperSet_of_isClosed (h : IsClosed s) : IsUpperSet s := |
| 198 | + isLowerSet_compl.1 <| isLowerSet_of_isOpen h.isOpen_compl |
| 199 | +#align lower_topology.is_upper_set_of_is_closed LowerTopology.isUpperSet_of_isClosed |
| 200 | + |
| 201 | +/-- |
| 202 | +The closure of a singleton `{a}` in the lower topology is the left-closed right-infinite interval |
| 203 | +[a, ∞). |
| 204 | +-/ |
| 205 | +@[simp] |
| 206 | +theorem closure_singleton (a : α) : closure {a} = Ici a := |
| 207 | + Subset.antisymm ((closure_minimal fun _ h => h.ge) <| isClosed_Ici a) <| |
| 208 | + (isUpperSet_of_isClosed isClosed_closure).Ici_subset <| subset_closure rfl |
| 209 | +#align lower_topology.closure_singleton LowerTopology.closure_singleton |
| 210 | + |
| 211 | +protected theorem isTopologicalBasis : IsTopologicalBasis (lowerBasis α) := by |
| 212 | + convert isTopologicalBasis_of_subbasis (topology_eq α) |
| 213 | + simp_rw [lowerBasis, coe_upperClosure, compl_unionᵢ] |
| 214 | + ext s |
| 215 | + constructor |
| 216 | + · rintro ⟨F, hF, rfl⟩ |
| 217 | + refine' ⟨(fun a => Ici aᶜ) '' F, ⟨hF.image _, image_subset_iff.2 fun _ _ => ⟨_, rfl⟩⟩, _⟩ |
| 218 | + simp only [interₛ_image] |
| 219 | + · rintro ⟨F, ⟨hF, hs⟩, rfl⟩ |
| 220 | + haveI := hF.to_subtype |
| 221 | + rw [subset_def, Subtype.forall'] at hs |
| 222 | + choose f hf using hs |
| 223 | + exact ⟨_, finite_range f, by simp_rw [binterᵢ_range, hf, interₛ_eq_interᵢ]⟩ |
| 224 | +#align lower_topology.is_topological_basis LowerTopology.isTopologicalBasis |
| 225 | + |
| 226 | +/-- A function `f : β → α` with lower topology in the codomain is continuous provided that the |
| 227 | +preimage of every interval `Set.Ici a` is a closed set. |
| 228 | +
|
| 229 | +TODO: upgrade to an `iff`. -/ |
| 230 | +lemma continuous_of_Ici [TopologicalSpace β] {f : β → α} (h : ∀ a, IsClosed (f ⁻¹' (Ici a))) : |
| 231 | + Continuous f := by |
| 232 | + obtain rfl := LowerTopology.topology_eq α |
| 233 | + refine continuous_generateFrom ?_ |
| 234 | + rintro _ ⟨a, rfl⟩ |
| 235 | + exact (h a).isOpen_compl |
| 236 | + |
| 237 | +end Preorder |
| 238 | + |
| 239 | +section PartialOrder |
| 240 | + |
| 241 | +variable [PartialOrder α] [TopologicalSpace α] [LowerTopology α] |
| 242 | + |
| 243 | +-- see Note [lower instance priority] |
| 244 | +/-- The lower topology on a partial order is T₀. -/ |
| 245 | +instance (priority := 90) t0Space : T0Space α := |
| 246 | + (t0Space_iff_inseparable α).2 fun x y h => |
| 247 | + Ici_injective <| by simpa only [inseparable_iff_closure_eq, closure_singleton] using h |
| 248 | + |
| 249 | +end PartialOrder |
| 250 | + |
| 251 | +end LowerTopology |
| 252 | + |
| 253 | +instance [Preorder α] [TopologicalSpace α] [LowerTopology α] [OrderBot α] [Preorder β] |
| 254 | + [TopologicalSpace β] [LowerTopology β] [OrderBot β] : LowerTopology (α × β) where |
| 255 | + topology_eq_lowerTopology := by |
| 256 | + refine' le_antisymm (le_generateFrom _) _ |
| 257 | + · rintro _ ⟨x, rfl⟩ |
| 258 | + exact ((LowerTopology.isClosed_Ici _).prod <| LowerTopology.isClosed_Ici _).isOpen_compl |
| 259 | + rw [(LowerTopology.isTopologicalBasis.prod LowerTopology.isTopologicalBasis).eq_generateFrom, |
| 260 | + le_generateFrom_iff_subset_isOpen, image2_subset_iff] |
| 261 | + rintro _ ⟨s, hs, rfl⟩ _ ⟨t, ht, rfl⟩ |
| 262 | + dsimp |
| 263 | + simp_rw [coe_upperClosure, compl_unionᵢ, prod_eq, preimage_interᵢ, preimage_compl] |
| 264 | + -- without `let`, `refine` tries to use the product topology and fails |
| 265 | + let _ : TopologicalSpace (α × β) := generateFrom { s | ∃ a, Ici aᶜ = s } |
| 266 | + refine (isOpen_binterᵢ hs fun a _ => ?_).inter (isOpen_binterᵢ ht fun b _ => ?_) |
| 267 | + · exact GenerateOpen.basic _ ⟨(a, ⊥), by simp [Ici_prod_eq, prod_univ]⟩ |
| 268 | + · exact GenerateOpen.basic _ ⟨(⊥, b), by simp [Ici_prod_eq, univ_prod]⟩ |
| 269 | + |
| 270 | +section CompleteLattice |
| 271 | + |
| 272 | +variable [CompleteLattice α] [CompleteLattice β] [TopologicalSpace α] [LowerTopology α] |
| 273 | + [TopologicalSpace β] [LowerTopology β] |
| 274 | + |
| 275 | +protected theorem InfₛHom.continuous (f : InfₛHom α β) : Continuous f := by |
| 276 | + refine LowerTopology.continuous_of_Ici fun b => ?_ |
| 277 | + convert LowerTopology.isClosed_Ici (infₛ <| f ⁻¹' Ici b) |
| 278 | + refine' Subset.antisymm (fun a => infₛ_le) fun a ha => le_trans _ <| |
| 279 | + OrderHomClass.mono (f : α →o β) ha |
| 280 | + refine' LE.le.trans _ (map_infₛ f _).ge |
| 281 | + simp |
| 282 | +#align Inf_hom.continuous InfₛHom.continuous |
| 283 | + |
| 284 | +-- see Note [lower instance priority] |
| 285 | +instance (priority := 90) LowerTopology.continuousInf : ContinuousInf α := |
| 286 | + ⟨(infInfₛHom : InfₛHom (α × α) α).continuous⟩ |
| 287 | +#align lower_topology.to_has_continuous_inf LowerTopology.continuousInf |
| 288 | + |
| 289 | +end CompleteLattice |
| 290 | + |
0 commit comments