Documentation

Mathlib.Topology.Order.LawsonTopology

Lawson topology #

This file introduces the Lawson topology on a preorder.

Main definitions #

Main statements #

Implementation notes #

A type synonym Topology.WithLawson is introduced and for a preorder α, Topology.WithLawson α is made an instance of TopologicalSpace by Topology.lawson.

We define a mixin class Topology.IsLawson for the class of types which are both a preorder and a topology and where the topology is Topology.lawson. It is shown that Topology.WithLawson α is an instance of Topology.IsLawson.

References #

Tags #

Lawson topology, preorder

Lawson topology #

The Lawson topology is defined as the meet of Topology.lower and the Topology.scott.

Equations
Instances For

    Predicate for an ordered topological space to be equipped with its Lawson topology.

    The Lawson topology is defined as the meet of Topology.lower and the Topology.scott.

    Instances

      The complements of the upper closures of finite sets intersected with Scott open sets form a basis for the lawson topology.

      Equations
      Instances For
        def Topology.WithLawson (α : Type u_3) :
        Type u_3

        Type synonym for a preorder equipped with the Lawson topology.

        Equations
        Instances For
          @[match_pattern]

          toLawson is the identity function to the WithLawson of a type.

          Equations
          Instances For
            @[match_pattern]

            ofLawson is the identity function from the WithLawson of a type.

            Equations
            Instances For
              @[simp]
              theorem Topology.WithLawson.to_Lawson_symm_eq {α : Type u_1} :
              Topology.WithLawson.toLawson.symm = Topology.WithLawson.ofLawson
              @[simp]
              theorem Topology.WithLawson.of_Lawson_symm_eq {α : Type u_1} :
              Topology.WithLawson.ofLawson.symm = Topology.WithLawson.toLawson
              @[simp]
              theorem Topology.WithLawson.toLawson_ofLawson {α : Type u_1} (a : Topology.WithLawson α) :
              Topology.WithLawson.toLawson (Topology.WithLawson.ofLawson a) = a
              @[simp]
              theorem Topology.WithLawson.ofLawson_toLawson {α : Type u_1} (a : α) :
              Topology.WithLawson.ofLawson (Topology.WithLawson.toLawson a) = a
              @[simp]
              theorem Topology.WithLawson.toLawson_inj {α : Type u_1} {a : α} {b : α} :
              Topology.WithLawson.toLawson a = Topology.WithLawson.toLawson b a = b
              @[simp]
              theorem Topology.WithLawson.ofLawson_inj {α : Type u_1} {a : Topology.WithLawson α} {b : Topology.WithLawson α} :
              Topology.WithLawson.ofLawson a = Topology.WithLawson.ofLawson b a = b
              def Topology.WithLawson.rec {α : Type u_1} {β : Topology.WithLawson αSort u_3} (h : (a : α) → β (Topology.WithLawson.toLawson a)) (a : Topology.WithLawson α) :
              β a

              A recursor for WithLawson. Use as induction' x.

              Equations
              Instances For
                Equations
                • = inst
                Equations
                • Topology.WithLawson.instInhabited = inst
                Equations
                • Topology.WithLawson.instPreorder = inst
                Equations

                If α is equipped with the Lawson topology, then it is homeomorphic to WithLawson α.

                Equations
                • Topology.WithLawson.homeomorph = Topology.WithLawson.ofLawson.toHomeomorphOfInducing
                Instances For
                  theorem Topology.WithLawson.isOpen_preimage_ofLawson {α : Type u_1} [Preorder α] {S : Set α} :
                  IsOpen (Topology.WithLawson.ofLawson ⁻¹' S) TopologicalSpace.IsOpen S
                  theorem Topology.WithLawson.isClosed_preimage_ofLawson {α : Type u_1} [Preorder α] {S : Set α} :
                  IsClosed (Topology.WithLawson.ofLawson ⁻¹' S) IsClosed S
                  theorem Topology.WithLawson.isOpen_def {α : Type u_1} [Preorder α] {T : Set (Topology.WithLawson α)} :
                  IsOpen T TopologicalSpace.IsOpen (Topology.WithLawson.toLawson ⁻¹' T)
                  theorem Topology.lawsonClosed_of_scottClosed {α : Type u_1} [Preorder α] (s : Set α) (h : IsClosed (Topology.WithScott.ofScott ⁻¹' s)) :
                  IsClosed (Topology.WithLawson.ofLawson ⁻¹' s)
                  theorem Topology.lawsonClosed_of_lowerClosed {α : Type u_1} [Preorder α] (s : Set α) (h : IsClosed (Topology.WithLower.ofLower ⁻¹' s)) :
                  IsClosed (Topology.WithLawson.ofLawson ⁻¹' s)
                  theorem Topology.lawsonOpen_iff_scottOpen_of_isUpperSet {α : Type u_1} [Preorder α] {s : Set α} (h : IsUpperSet s) :
                  IsOpen (Topology.WithLawson.ofLawson ⁻¹' s) IsOpen (Topology.WithScott.ofScott ⁻¹' s)

                  An upper set is Lawson open if and only if it is Scott open

                  An upper set is Lawson open if and only if it is Scott open

                  A lower set is Lawson closed if and only if it is closed under sups of directed sets

                  @[instance 90]

                  The Lawson topology on a partial order is T₀.

                  Equations
                  • =