Documentation

Mathlib.Topology.Category.TopCat.OpenNhds

The category of open neighborhoods of a point #

Given an object X of the category TopCat of topological spaces and a point x : X, this file builds the type OpenNhds x of open neighborhoods of x in X and endows it with the partial order given by inclusion and the corresponding category structure (as a full subcategory of the poset category Set X). This is used in Topology.Sheaves.Stalks to build the stalk of a sheaf at x as a limit over OpenNhds x.

Main declarations #

Besides OpenNhds, the main constructions here are:

def TopologicalSpace.OpenNhds {X : TopCat} (x : X) :

The type of open neighbourhoods of a point x in a (bundled) topological space.

Equations
Instances For
    Equations
    instance TopologicalSpace.OpenNhds.opensNhds.instFunLike {X : TopCat} {x : X} {U V : OpenNhds x} :
    FunLike (U V) U.obj V.obj
    Equations
    @[simp]
    theorem TopologicalSpace.OpenNhds.apply_mk {X : TopCat} {x : X} {U V : OpenNhds x} (f : U V) (y : X) (hy : y U.obj) :
    f y, hy = y,
    @[simp]
    theorem TopologicalSpace.OpenNhds.val_apply {X : TopCat} {x : X} {U V : OpenNhds x} (f : U V) (y : U.obj) :
    (f y) = y
    @[simp]
    theorem TopologicalSpace.OpenNhds.coe_id {X : TopCat} {x : X} {U : OpenNhds x} (f : U U) :
    f = id
    theorem TopologicalSpace.OpenNhds.id_apply {X : TopCat} {x : X} {U : OpenNhds x} (f : U U) (y : U.obj) :
    f y = y
    @[simp]
    theorem TopologicalSpace.OpenNhds.comp_apply {X : TopCat} {x : X} {U V W : OpenNhds x} (f : U V) (g : V W) (x✝ : U.obj) :
    def TopologicalSpace.OpenNhds.infLELeft {X : TopCat} {x : X} (U V : OpenNhds x) :
    U V U

    The inclusion U ⊓ V ⟶ U as a morphism in the category of open sets.

    Equations
    Instances For
      def TopologicalSpace.OpenNhds.infLERight {X : TopCat} {x : X} (U V : OpenNhds x) :
      U V V

      The inclusion U ⊓ V ⟶ V as a morphism in the category of open sets.

      Equations
      Instances For

        The inclusion functor from open neighbourhoods of x to open sets in the ambient topological space.

        Equations
        Instances For
          @[simp]
          theorem TopologicalSpace.OpenNhds.inclusion_obj {X : TopCat} (x : X) (U : Opens X) (p : x U) :
          (inclusion x).obj { obj := U, property := p } = U
          @[deprecated TopologicalSpace.OpenNhds.isOpenEmbedding (since := "2024-10-18")]

          Alias of TopologicalSpace.OpenNhds.isOpenEmbedding.

          The preimage functor from neighborhoods of f x to neighborhoods of x.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem TopologicalSpace.OpenNhds.map_obj {X Y : TopCat} (f : X Y) (x : X) (U : Opens Y) (q : (CategoryTheory.ConcreteCategory.hom f) x U) :
            (map f x).obj { obj := U, property := q } = { obj := (Opens.map f).obj U, property := q }
            @[simp]
            theorem TopologicalSpace.OpenNhds.map_id_obj' {X : TopCat} (x : X) (U : Set X) (p : IsOpen U) (q : (CategoryTheory.ConcreteCategory.hom (CategoryTheory.CategoryStruct.id X)) x { carrier := U, is_open' := p }) :
            (map (CategoryTheory.CategoryStruct.id X) x).obj { obj := { carrier := U, is_open' := p }, property := q } = { obj := { carrier := U, is_open' := p }, property := q }

            Opens.map f and OpenNhds.map f form a commuting square (up to natural isomorphism) with the inclusion functors into Opens X.

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

              An open map f : X ⟶ Y induces a functor OpenNhds x ⥤ OpenNhds (f x).

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem IsOpenMap.functorNhds_map {X Y : TopCat} {f : X Y} (h : IsOpenMap (CategoryTheory.ConcreteCategory.hom f)) (x : X) {X✝ Y✝ : TopologicalSpace.OpenNhds x} (i : X✝ Y✝) :

                An open map f : X ⟶ Y induces an adjunction between OpenNhds x and OpenNhds (f x).

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

                  An inducing map f : X ⟶ Y induces a functor open_nhds x ⥤ open_nhds (f x).

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem Topology.IsInducing.functorNhds_map {X Y : TopCat} {f : X Y} (h : IsInducing (CategoryTheory.ConcreteCategory.hom f)) (x : X) {X✝ Y✝ : TopologicalSpace.OpenNhds x} (a✝ : X✝.obj Y✝.obj) :
                    (h.functorNhds x).map a✝ = h.functor.map a✝

                    An inducing map f : X ⟶ Y induces an adjunction between open_nhds x and open_nhds (f x).

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