Documentation

Mathlib.Topology.Sheaves.Skyscraper

Skyscraper (pre)sheaves #

A skyscraper (pre)sheaf 𝓕 : (Pre)Sheaf C X is the (pre)sheaf with value A at point p₀ that is supported only at open sets contain p₀, i.e. 𝓕(U) = A if p₀ ∈ U and 𝓕(U) = * if p₀ ∉ U where * is a terminal object of C. In terms of stalks, 𝓕 is supported at all specializations of p₀, i.e. if p₀ ⤳ x then 𝓕ₓ ≅ A and if ¬ p₀ ⤳ x then 𝓕ₓ ≅ *.

Main definitions #

Main statements #

TODO: generalize universe level when calculating stalks, after generalizing universe level of stalk.

@[simp]
theorem skyscraperPresheaf_map {X : TopCat} (p₀ : X) [(U : TopologicalSpace.Opens X) → Decidable (p₀ U)] {C : Type v} [CategoryTheory.Category.{w, v} C] [CategoryTheory.Limits.HasTerminal C] (A : C) {U : (TopologicalSpace.Opens X)ᵒᵖ} {V : (TopologicalSpace.Opens X)ᵒᵖ} (i : U V) :
(skyscraperPresheaf p₀ A).map i = if h : p₀ V.unop then CategoryTheory.eqToHom else CategoryTheory.Limits.IsTerminal.from (CategoryTheory.Limits.terminalIsTerminal) ((fun (U : (TopologicalSpace.Opens X)ᵒᵖ) => if p₀ U.unop then A else ⊤_ C) U)
@[simp]
theorem skyscraperPresheaf_obj {X : TopCat} (p₀ : X) [(U : TopologicalSpace.Opens X) → Decidable (p₀ U)] {C : Type v} [CategoryTheory.Category.{w, v} C] [CategoryTheory.Limits.HasTerminal C] (A : C) (U : (TopologicalSpace.Opens X)ᵒᵖ) :
(skyscraperPresheaf p₀ A).obj U = if p₀ U.unop then A else ⊤_ C

A skyscraper presheaf is a presheaf supported at a single point: if p₀ ∈ X is a specified point, then the skyscraper presheaf 𝓕 with value A is defined by U ↦ A if p₀ ∈ U and U ↦ * if p₀ ∉ A where * is some terminal object.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem SkyscraperPresheafFunctor.map'_app {X : TopCat} (p₀ : X) [(U : TopologicalSpace.Opens X) → Decidable (p₀ U)] {C : Type v} [CategoryTheory.Category.{w, v} C] [CategoryTheory.Limits.HasTerminal C] {a : C} {b : C} (f : a b) (U : (TopologicalSpace.Opens X)ᵒᵖ) :
    (SkyscraperPresheafFunctor.map' p₀ f).app U = if h : p₀ U.unop then CategoryTheory.CategoryStruct.comp (CategoryTheory.eqToHom ) (CategoryTheory.CategoryStruct.comp f (CategoryTheory.eqToHom )) else CategoryTheory.Limits.IsTerminal.from (CategoryTheory.Limits.terminalIsTerminal) ((skyscraperPresheaf p₀ a).obj U)

    Taking skyscraper presheaf at a point is functorial: c ↦ skyscraper p₀ c defines a functor by sending every f : a ⟶ b to the natural transformation α defined as: α(U) = f : a ⟶ b if p₀ ∈ U and the unique morphism to a terminal object in C if p₀ ∉ U.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem skyscraperPresheafFunctor_map {X : TopCat} (p₀ : X) [(U : TopologicalSpace.Opens X) → Decidable (p₀ U)] {C : Type v} [CategoryTheory.Category.{w, v} C] [CategoryTheory.Limits.HasTerminal C] :
      ∀ {X_1 Y : C} (f : X_1 Y), (skyscraperPresheafFunctor p₀).map f = SkyscraperPresheafFunctor.map' p₀ f

      Taking skyscraper presheaf at a point is functorial: c ↦ skyscraper p₀ c defines a functor by sending every f : a ⟶ b to the natural transformation α defined as: α(U) = f : a ⟶ b if p₀ ∈ U and the unique morphism to a terminal object in C if p₀ ∉ U.

      Equations
      Instances For
        @[simp]
        theorem skyscraperPresheafCoconeOfSpecializes_pt {X : TopCat} (p₀ : X) [(U : TopologicalSpace.Opens X) → Decidable (p₀ U)] {C : Type v} [CategoryTheory.Category.{u, v} C] (A : C) [CategoryTheory.Limits.HasTerminal C] {y : X} (h : p₀ y) :

        The cocone at A for the stalk functor of skyscraperPresheaf p₀ A when y ∈ closure {p₀}

        Equations
        Instances For

          The cocone at A for the stalk functor of skyscraperPresheaf p₀ A when y ∈ closure {p₀} is a colimit

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

            If y ∈ closure {p₀}, then the stalk of skyscraperPresheaf p₀ A at y is A.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem skyscraperPresheafCocone_pt {X : TopCat} (p₀ : X) [(U : TopologicalSpace.Opens X) → Decidable (p₀ U)] {C : Type v} [CategoryTheory.Category.{u, v} C] (A : C) [CategoryTheory.Limits.HasTerminal C] (y : X) :

              The cocone at * for the stalk functor of skyscraperPresheaf p₀ A when y ∉ closure {p₀}

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

                The cocone at * for the stalk functor of skyscraperPresheaf p₀ A when y ∉ closure {p₀} is a colimit

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

                  If y ∉ closure {p₀}, then the stalk of skyscraperPresheaf p₀ A at y is isomorphic to a terminal object.

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

                    If y ∉ closure {p₀}, then the stalk of skyscraperPresheaf p₀ A at y is a terminal object

                    Equations
                    Instances For

                      The skyscraper presheaf supported at p₀ with value A is the sheaf that assigns A to all opens U that contain p₀ and assigns * otherwise.

                      Equations
                      Instances For

                        Taking skyscraper sheaf at a point is functorial: c ↦ skyscraper p₀ c defines a functor by sending every f : a ⟶ b to the natural transformation α defined as: α(U) = f : a ⟶ b if p₀ ∈ U and the unique morphism to a terminal object in C if p₀ ∉ U.

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

                          If f : 𝓕.stalk p₀ ⟶ c, then a natural transformation 𝓕 ⟶ skyscraperPresheaf p₀ c can be defined by: 𝓕.germ p₀ ≫ f : 𝓕(U) ⟶ c if p₀ ∈ U and the unique morphism to a terminal object if p₀ ∉ U.

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

                            If f : 𝓕 ⟶ skyscraperPresheaf p₀ c is a natural transformation, then there is a morphism 𝓕.stalk p₀ ⟶ c defined as the morphism from colimit to cocone at c.

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

                              The unit in Presheaf.stalkFunctor ⊣ skyscraperPresheafFunctor

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

                                skyscraperPresheafFunctor is the right adjoint of Presheaf.stalkFunctor

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

                                  Taking stalks of a sheaf is the left adjoint functor to skyscraperSheafFunctor

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