Documentation

Mathlib.CategoryTheory.Sites.Equivalence

Equivalences of sheaf categories #

Given a site (C, J) and a category D which is equivalent to C, with C and D possibly large and possibly in different universes, we transport the Grothendieck topology J on C to D and prove that the sheaf categories are equivalent.

We also prove that sheafification and the property HasSheafCompose transport nicely over this equivalence, and apply it to essentially small sites. We also provide instances for existence of sufficiently small limits in the sheaf category on the essentially small site.

Main definitions #

@[instance 900]
instance CategoryTheory.Equivalence.instIsCoverDenseOfIsEquivalence {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor D C) [G.IsEquivalence] :
G.IsCoverDense J

The functor in the equivalence of sheaf categories.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem CategoryTheory.Equivalence.sheafCongr.functor_obj_val_map {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (F : Sheaf J A) {X✝ Y✝ : Dᵒᵖ} (f : X✝ Y✝) :
    ((functor J K e A).obj F).val.map f = F.val.map (e.inverse.map f.unop).op
    @[simp]
    theorem CategoryTheory.Equivalence.sheafCongr.functor_map_val_app {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] {X✝ Y✝ : Sheaf J A} (f : X✝ Y✝) (X : Dᵒᵖ) :
    ((functor J K e A).map f).val.app X = f.val.app (Opposite.op (e.inverse.obj (Opposite.unop X)))
    @[simp]
    theorem CategoryTheory.Equivalence.sheafCongr.functor_obj_val_obj {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (F : Sheaf J A) (X : Dᵒᵖ) :
    ((functor J K e A).obj F).val.obj X = F.val.obj (Opposite.op (e.inverse.obj (Opposite.unop X)))

    The inverse in the equivalence of sheaf categories.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem CategoryTheory.Equivalence.sheafCongr.inverse_obj_val_obj {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (F : Sheaf K A) (X : Cᵒᵖ) :
      ((inverse J K e A).obj F).val.obj X = F.val.obj (Opposite.op (e.functor.obj (Opposite.unop X)))
      @[simp]
      theorem CategoryTheory.Equivalence.sheafCongr.inverse_obj_val_map {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (F : Sheaf K A) {X✝ Y✝ : Cᵒᵖ} (f : X✝ Y✝) :
      ((inverse J K e A).obj F).val.map f = F.val.map (e.functor.map f.unop).op
      @[simp]
      theorem CategoryTheory.Equivalence.sheafCongr.inverse_map_val_app {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] {X✝ Y✝ : Sheaf K A} (f : X✝ Y✝) (X : Cᵒᵖ) :
      ((inverse J K e A).map f).val.app X = f.val.app (Opposite.op (e.functor.obj (Opposite.unop X)))

      The unit iso in the equivalence of sheaf categories.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem CategoryTheory.Equivalence.sheafCongr.unitIso_hom_app_val_app {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (X : Sheaf J A) (X✝ : Cᵒᵖ) :
        ((unitIso J K e A).hom.app X).val.app X✝ = X.val.map (e.unitIso.inv.app (Opposite.unop X✝)).op
        @[simp]
        theorem CategoryTheory.Equivalence.sheafCongr.unitIso_inv_app_val_app {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (X : Sheaf J A) (X✝ : Cᵒᵖ) :
        ((unitIso J K e A).inv.app X).val.app X✝ = X.val.map (e.unitIso.hom.app (Opposite.unop X✝)).op

        The counit iso in the equivalence of sheaf categories.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          theorem CategoryTheory.Equivalence.sheafCongr.counitIso_inv_app_val_app {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (X : Sheaf K A) (X✝ : Dᵒᵖ) :
          ((counitIso J K e A).inv.app X).val.app X✝ = X.val.map (e.counitIso.hom.app (Opposite.unop X✝)).op
          @[simp]
          theorem CategoryTheory.Equivalence.sheafCongr.counitIso_hom_app_val_app {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] (X : Sheaf K A) (X✝ : Dᵒᵖ) :
          ((counitIso J K e A).hom.app X).val.app X✝ = X.val.map (e.counitIso.inv.app (Opposite.unop X✝)).op

          The equivalence of sheaf categories.

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

            Transport a presheaf to the equivalent category and sheafify there.

            Equations
            Instances For
              noncomputable def CategoryTheory.Equivalence.transportIsoSheafToPresheaf {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) (A : Type u₃) [Category.{v₃, u₃} A] [Functor.IsDenseSubsite K J e.inverse] :
              (sheafCongr J K e A).functor.comp ((sheafToPresheaf K A).comp e.op.congrLeft.inverse) sheafToPresheaf J A

              An auxiliary definition for the sheafification adjunction.

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

                Transporting and sheafifying is left adjoint to taking the underlying presheaf.

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

                  Transport HasSheafify along an equivalence of sites.

                  theorem CategoryTheory.Equivalence.hasSheafCompose {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (e : C D) [Functor.IsDenseSubsite K J e.inverse] {A : Type u_1} [Category.{u_3, u_1} A] {B : Type u_2} [Category.{u_4, u_2} B] (F : Functor A B) [K.HasSheafCompose F] :
                  J.HasSheafCompose F

                  Transport to a small model and sheafify there.

                  Equations
                  Instances For

                    Transporting to a small model and sheafifying there is left adjoint to the underlying presheaf functor

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      instance CategoryTheory.hasSheafComposeEssentiallySmallSite {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) (A : Type u₃) [Category.{v₃, u₃} A] (B : Type u₄) [Category.{v₄, u₄} B] (F : Functor A B) [EssentiallySmall.{w, v₁, u₁} C] [((equivSmallModel C).inverse.inducedTopology J).HasSheafCompose F] :
                      J.HasSheafCompose F
                      theorem CategoryTheory.GrothendieckTopology.W_inverseImage_whiskeringLeft {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (G : Functor D C) {A : Type u₃} [Category.{v₃, u₃} A] [G.IsCoverDense J] [G.Full] [G.IsContinuous K J] [(G.sheafPushforwardContinuous A K J).EssSurj] :
                      K.W.inverseImage ((whiskeringLeft Dᵒᵖ Cᵒᵖ A).obj G.op) = J.W
                      theorem CategoryTheory.GrothendieckTopology.W_whiskerLeft_iff {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (G : Functor D C) {A : Type u₃} [Category.{v₃, u₃} A] [G.IsCoverDense J] [G.Full] [G.IsContinuous K J] [(G.sheafPushforwardContinuous A K J).EssSurj] {P Q : Functor Cᵒᵖ A} (f : P Q) :
                      K.W (whiskerLeft G.op f) J.W f
                      theorem CategoryTheory.GrothendieckTopology.PreservesSheafification.transport {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (G : Functor D C) {A : Type u₃} [Category.{v₃, u₃} A] (B : Type u₄) [Category.{v₄, u₄} B] (F : Functor A B) [G.IsCoverDense J] [G.Full] [G.IsContinuous K J] [G.IsContinuous K J] [(G.sheafPushforwardContinuous B K J).EssSurj] [(G.sheafPushforwardContinuous A K J).EssSurj] [K.PreservesSheafification F] :
                      J.PreservesSheafification F
                      theorem CategoryTheory.GrothendieckTopology.WEqualsLocallyBijective.transport {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {D : Type u₂} [Category.{v₂, u₂} D] (K : GrothendieckTopology D) (G : Functor D C) {A : Type u₃} [Category.{v₃, u₃} A] [G.IsCoverDense J] [G.Full] [G.IsContinuous K J] [(G.sheafPushforwardContinuous A K J).EssSurj] [G.IsCocontinuous K J] [HasForget A] [K.WEqualsLocallyBijective A] (hG : CoverPreserving K J G) :
                      J.WEqualsLocallyBijective A
                      instance CategoryTheory.GrothendieckTopology.instPreservesSheafification_1 {C : Type u₁} [Category.{v₁, u₁} C] (J : GrothendieckTopology C) {A : Type u₃} [Category.{v₃, u₃} A] (B : Type u₄) [Category.{v₄, u₄} B] (F : Functor A B) [EssentiallySmall.{w, v₁, u₁} C] [∀ (X : Cᵒᵖ), Limits.HasLimitsOfShape (StructuredArrow X (equivSmallModel C).inverse.op) A] [∀ (X : Cᵒᵖ), Limits.HasLimitsOfShape (StructuredArrow X (equivSmallModel C).inverse.op) B] [((equivSmallModel C).inverse.inducedTopology J).PreservesSheafification F] :
                      J.PreservesSheafification F