Documentation

Mathlib.CategoryTheory.Localization.Construction

Construction of the localized category #

This file constructs the localized category, obtained by formally inverting a class of maps W : MorphismProperty C in a category C.

We first construct a quiver LocQuiver W whose objects are the same as those of C and whose maps are the maps in C and placeholders for the formal inverses of the maps in W.

The localized category W.Localization is obtained by taking the quotient of the path category of LocQuiver W by the congruence generated by four types of relations.

The obvious functor Q W : C ⥤ W.Localization satisfies the universal property of the localization. Indeed, if G : C ⥤ D sends morphisms in W to isomorphisms in D (i.e. we have hG : W.IsInvertedBy G), then there exists a unique functor G' : W.Localization ⥤ D such that Q W ≫ G' = G. This G' is lift G hG. The expected property of lift G hG if expressed by the lemma fac and the uniqueness is expressed by uniq.

References #

If W : MorphismProperty C, LocQuiver W is a quiver with the same objects as C, and whose morphisms are those in C and placeholders for formal inverses of the morphisms in W.

  • obj : C

    underlying object

Instances For

    The morphism in the path category associated to a morphism in the original category.

    Equations
    Instances For

      The morphism in the path category corresponding to a formal inverse.

      Equations
      Instances For

        The relations by which we take the quotient in order to get the localized category.

        Instances For

          The localized category obtained by formally inverting the morphisms in W : MorphismProperty C

          Equations
          Instances For

            The obvious functor C ⥤ W.Localization

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              def CategoryTheory.Localization.Construction.wIso {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} {X : C} {Y : C} (w : X Y) (hw : W w) :
              W.Q.obj X W.Q.obj Y

              The isomorphism in W.Localization associated to a morphism w in W

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[reducible, inline]
                abbrev CategoryTheory.Localization.Construction.winv {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} {X : C} {Y : C} (w : X Y) (hw : W w) :
                W.Q.obj Y W.Q.obj X

                The formal inverse in W.Localization of a morphism w in W.

                Equations
                Instances For
                  @[simp]

                  The lifting of a functor to the path category of LocQuiver W

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

                    The lifting of a functor C ⥤ D inverting W as a functor W.Localization ⥤ D

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      theorem CategoryTheory.Localization.Construction.uniq {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} {D : Type uD} [CategoryTheory.Category.{uD', uD} D] (G₁ : CategoryTheory.Functor W.Localization D) (G₂ : CategoryTheory.Functor W.Localization D) (h : W.Q.comp G₁ = W.Q.comp G₂) :
                      G₁ = G₂

                      The canonical bijection between objects in a category and its localization with respect to a morphism_property W

                      Equations
                      Instances For
                        theorem CategoryTheory.Localization.Construction.morphismProperty_is_top {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} (P : CategoryTheory.MorphismProperty W.Localization) [P.IsStableUnderComposition] (hP₁ : ∀ ⦃X Y : C⦄ (f : X Y), P (W.Q.map f)) (hP₂ : ∀ ⦃X Y : C⦄ (w : X Y) (hw : W w), P (CategoryTheory.Localization.Construction.winv w hw)) :
                        P =

                        A MorphismProperty in W.Localization is satisfied by all morphisms in the localized category if it contains the image of the morphisms in the original category, the inverses of the morphisms in W and if it is stable under composition

                        theorem CategoryTheory.Localization.Construction.morphismProperty_is_top' {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} (P : CategoryTheory.MorphismProperty W.Localization) [P.IsStableUnderComposition] (hP₁ : ∀ ⦃X Y : C⦄ (f : X Y), P (W.Q.map f)) (hP₂ : ∀ ⦃X Y : W.Localization⦄ (e : X Y), P e.homP e.inv) :
                        P =

                        A MorphismProperty in W.Localization is satisfied by all morphisms in the localized category if it contains the image of the morphisms in the original category, if is stable under composition and if the property is stable by passing to inverses.

                        def CategoryTheory.Localization.Construction.NatTransExtension.app {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} {D : Type uD} [CategoryTheory.Category.{uD', uD} D] {F₁ : CategoryTheory.Functor W.Localization D} {F₂ : CategoryTheory.Functor W.Localization D} (τ : W.Q.comp F₁ W.Q.comp F₂) (X : W.Localization) :
                        F₁.obj X F₂.obj X

                        If F₁ and F₂ are functors W.Localization ⥤ D and if we have τ : W.Q ⋙ F₁ ⟶ W.Q ⋙ F₂, we shall define a natural transformation F₁ ⟶ F₂. This is the app field of this natural transformation.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def CategoryTheory.Localization.Construction.natTransExtension {C : Type uC} [CategoryTheory.Category.{uC', uC} C] {W : CategoryTheory.MorphismProperty C} {D : Type uD} [CategoryTheory.Category.{uD', uD} D] {F₁ : CategoryTheory.Functor W.Localization D} {F₂ : CategoryTheory.Functor W.Localization D} (τ : W.Q.comp F₁ W.Q.comp F₂) :
                          F₁ F₂

                          If F₁ and F₂ are functors W.Localization ⥤ D, a natural transformation F₁ ⟶ F₂ can be obtained from a natural transformation W.Q ⋙ F₁ ⟶ W.Q ⋙ F₂.

                          Equations
                          Instances For

                            The functor (W.Localization ⥤ D) ⥤ (W.FunctorsInverting D) induced by the composition with W.Q : C ⥤ W.Localization.

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

                              The function (W.FunctorsInverting D) ⥤ (W.Localization ⥤ D) induced by Construction.lift.

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

                                The equivalence of categories (W.localization ⥤ D) ≌ (W.FunctorsInverting D) induced by the composition with W.Q : C ⥤ W.localization.

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