Documentation

Mathlib.Topology.Category.Sequential

The category of sequential topological spaces #

We define the category Sequential of sequential topological spaces. We follow the usual template for defining categories of topological spaces, by giving it the induced category structure from TopCat.

structure Sequential :
Type (u + 1)

The type sequential topological spaces.

  • toTop : TopCat

    The underlying topological space of an object of Sequential.

  • is_sequential : SequentialSpace self.toTop

    The underlying topological space is sequential.

Instances For
    Equations
    Equations

    Constructor for objects of the category Sequential.

    Equations
    Instances For
      @[simp]
      def Sequential.isoOfHomeo {X Y : Sequential} (f : X.toTop ≃ₜ Y.toTop) :
      X Y

      Construct an isomorphism from a homeomorphism.

      Equations
      • Sequential.isoOfHomeo f = { hom := { toFun := f, continuous_toFun := }, inv := { toFun := f.symm, continuous_toFun := }, hom_inv_id := , inv_hom_id := }
      Instances For
        @[simp]
        theorem Sequential.isoOfHomeo_hom {X Y : Sequential} (f : X.toTop ≃ₜ Y.toTop) :
        (Sequential.isoOfHomeo f).hom = { toFun := f, continuous_toFun := }
        @[simp]
        theorem Sequential.isoOfHomeo_inv {X Y : Sequential} (f : X.toTop ≃ₜ Y.toTop) :
        (Sequential.isoOfHomeo f).inv = { toFun := f.symm, continuous_toFun := }
        def Sequential.homeoOfIso {X Y : Sequential} (f : X Y) :
        X.toTop ≃ₜ Y.toTop

        Construct a homeomorphism from an isomorphism.

        Equations
        • Sequential.homeoOfIso f = { toFun := f.hom, invFun := f.inv, left_inv := , right_inv := , continuous_toFun := , continuous_invFun := }
        Instances For
          @[simp]
          theorem Sequential.homeoOfIso_symm_apply {X Y : Sequential} (f : X Y) (a : (CategoryTheory.forget Sequential).obj Y) :
          (Sequential.homeoOfIso f).symm a = f.inv a
          @[simp]
          def Sequential.isoEquivHomeo {X Y : Sequential} :
          (X Y) (X.toTop ≃ₜ Y.toTop)

          The equivalence between isomorphisms in Sequential and homeomorphisms of topological spaces.

          Equations
          • Sequential.isoEquivHomeo = { toFun := Sequential.homeoOfIso, invFun := Sequential.isoOfHomeo, left_inv := , right_inv := }
          Instances For
            @[simp]
            theorem Sequential.isoEquivHomeo_apply {X Y : Sequential} (f : X Y) :
            Sequential.isoEquivHomeo f = Sequential.homeoOfIso f
            @[simp]
            theorem Sequential.isoEquivHomeo_symm_apply {X Y : Sequential} (f : X.toTop ≃ₜ Y.toTop) :
            Sequential.isoEquivHomeo.symm f = Sequential.isoOfHomeo f