Documentation

Mathlib.Topology.Category.UniformSpace

The category of uniform spaces #

We construct the category of uniform spaces, show that the complete separated uniform spaces form a reflective subcategory, and hence possess all limits that uniform spaces do.

TODO: show that uniform spaces actually have all limits!

structure UniformSpaceCat :
Type (u + 1)

An object in the category of uniform spaces.

Instances For
    structure UniformSpaceCat.Hom (X : UniformSpaceCat) (Y : UniformSpaceCat) :
    Type (max u_1 u_2)

    A bundled uniform continuous map.

    Instances For
      theorem UniformSpaceCat.Hom.ext {X : UniformSpaceCat} {Y : UniformSpaceCat} {x y : X.Hom Y} (hom' : x.hom' = y.hom') :
      x = y
      Equations
      • One or more equations did not get rendered due to their size.
      Equations
      Equations
      • One or more equations did not get rendered due to their size.
      @[reducible, inline]

      Turn a morphism in UniformSpaceCat back into a function which is UniformContinuous.

      Equations
      Instances For
        @[reducible, inline]
        abbrev UniformSpaceCat.ofHom {X Y : Type u} [UniformSpace X] [UniformSpace Y] (f : { f : XY // UniformContinuous f }) :
        { carrier := X, str := inst✝ } { carrier := Y, str := inst✝¹ }

        Typecheck a function which is UniformContinuous as a morphism in UniformSpaceCat.

        Equations
        Instances For
          theorem UniformSpaceCat.coe_of (X : Type u) [UniformSpace X] :
          { carrier := X, str := inst✝ }.carrier = X
          @[simp]
          theorem UniformSpaceCat.hom_ofHom {X Y : Type u} [UniformSpace X] [UniformSpace Y] (f : { f : XY // UniformContinuous f }) :
          theorem UniformSpaceCat.coe_mk {X Y : UniformSpaceCat} (f : X.carrierY.carrier) (hf : UniformContinuous f) :
          { hom' := f, hf }.hom = f

          The forgetful functor from uniform spaces to topological spaces.

          Equations
          • One or more equations did not get rendered due to their size.
          structure CpltSepUniformSpace :
          Type (u + 1)

          A (bundled) complete separated uniform space.

          Instances For

            The function forgetting that a complete separated uniform spaces is complete and separated.

            Equations
            Instances For

              Construct a bundled UniformSpace from the underlying type and the appropriate typeclasses.

              Equations
              Instances For
                @[simp]
                Equations

                The functor turning uniform spaces into complete separated uniform spaces.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For

                  The completion functor is left adjoint to the forgetful functor.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    Equations
                    • One or more equations did not get rendered due to their size.