Documentation

Mathlib.CategoryTheory.Functor.FullyFaithful

Full and faithful functors #

We define typeclasses Full and Faithful, decorating functors. These typeclasses carry no data. However, we also introduce a structure Functor.FullyFaithful which contains the data of the inverse map (F.obj X ⟶ F.obj Y) ⟶ (X ⟶ Y) of the map induced on morphisms by a functor F.

Main definitions and results #

See CategoryTheory.Equivalence.of_fullyFaithful_ess_surj for the fact that a functor is an equivalence if and only if it is fully faithful and essentially surjective.

A functor F : C ⥤ D is full if for each X Y : C, F.map is surjective.

Stacks Tag 001C

Instances

    A functor F : C ⥤ D is faithful if for each X Y : C, F.map is injective.

    Stacks Tag 001C

    Instances
      theorem CategoryTheory.Functor.map_injective_iff {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) [F.Faithful] {X Y : C} (f g : X Y) :
      F.map f = F.map g f = g
      noncomputable def CategoryTheory.Functor.preimage {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {X Y : C} (F : Functor C D) [F.Full] (f : F.obj X F.obj Y) :
      X Y

      The choice of a preimage of a morphism under a full functor.

      Equations
      Instances For
        @[simp]
        theorem CategoryTheory.Functor.map_preimage {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) [F.Full] {X Y : C} (f : F.obj X F.obj Y) :
        F.map (F.preimage f) = f
        @[simp]
        theorem CategoryTheory.Functor.preimage_comp {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} {X Y Z : C} [F.Full] [F.Faithful] (f : F.obj X F.obj Y) (g : F.obj Y F.obj Z) :
        @[simp]
        theorem CategoryTheory.Functor.preimage_map {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} {X Y : C} [F.Full] [F.Faithful] (f : X Y) :
        F.preimage (F.map f) = f
        noncomputable def CategoryTheory.Functor.preimageIso {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) {X Y : C} [F.Full] [F.Faithful] (f : F.obj X F.obj Y) :
        X Y

        If F : C ⥤ D is fully faithful, every isomorphism F.obj X ≅ F.obj Y has a preimage.

        Equations
        Instances For
          @[simp]
          theorem CategoryTheory.Functor.preimageIso_inv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) {X Y : C} [F.Full] [F.Faithful] (f : F.obj X F.obj Y) :
          @[simp]
          theorem CategoryTheory.Functor.preimageIso_hom {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) {X Y : C} [F.Full] [F.Faithful] (f : F.obj X F.obj Y) :
          @[simp]
          theorem CategoryTheory.Functor.preimageIso_mapIso {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) {X Y : C} [F.Full] [F.Faithful] (f : X Y) :
          F.preimageIso (F.mapIso f) = f
          structure CategoryTheory.Functor.FullyFaithful {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) :
          Type (max (max u₁ v₁) v₂)

          Structure containing the data of inverse map (F.obj X ⟶ F.obj Y) ⟶ (X ⟶ Y) of F.map in order to express that F is a fully faithful functor.

          • preimage {X Y : C} (f : F.obj X F.obj Y) : X Y

            The inverse map (F.obj X ⟶ F.obj Y) ⟶ (X ⟶ Y) of F.map.

          • map_preimage {X Y : C} (f : F.obj X F.obj Y) : F.map (self.preimage f) = f
          • preimage_map {X Y : C} (f : X Y) : self.preimage (F.map f) = f
          Instances For

            A FullyFaithful structure can be obtained from the assumption the F is both full and faithful.

            Equations
            Instances For

              The identity functor is fully faithful.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem CategoryTheory.Functor.FullyFaithful.id_preimage (C : Type u₁) [Category.{v₁, u₁} C] {X✝ Y✝ : C} (f : (Functor.id C).obj X✝ (Functor.id C).obj Y✝) :
                (id C).preimage f = f
                def CategoryTheory.Functor.FullyFaithful.homEquiv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} :
                (X Y) (F.obj X F.obj Y)

                The equivalence (X ⟶ Y) ≃ (F.obj X ⟶ F.obj Y) given by h : F.FullyFaithful.

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Functor.FullyFaithful.homEquiv_apply {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (a✝ : X Y) :
                  hF.homEquiv a✝ = F.map a✝
                  @[simp]
                  theorem CategoryTheory.Functor.FullyFaithful.homEquiv_symm_apply {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (f : F.obj X F.obj Y) :
                  theorem CategoryTheory.Functor.FullyFaithful.map_injective {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} {f g : X Y} (h : F.map f = F.map g) :
                  f = g
                  def CategoryTheory.Functor.FullyFaithful.preimageIso {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (e : F.obj X F.obj Y) :
                  X Y

                  The unique isomorphism X ≅ Y which induces an isomorphism F.obj X ≅ F.obj Y when hF : F.FullyFaithful.

                  Equations
                  Instances For
                    @[simp]
                    theorem CategoryTheory.Functor.FullyFaithful.preimageIso_hom {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (e : F.obj X F.obj Y) :
                    @[simp]
                    theorem CategoryTheory.Functor.FullyFaithful.preimageIso_inv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (e : F.obj X F.obj Y) :
                    theorem CategoryTheory.Functor.FullyFaithful.isIso_of_isIso_map {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (f : X Y) [IsIso (F.map f)] :
                    def CategoryTheory.Functor.FullyFaithful.isoEquiv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} :
                    (X Y) (F.obj X F.obj Y)

                    The equivalence (X ≅ Y) ≃ (F.obj X ≅ F.obj Y) given by h : F.FullyFaithful.

                    Equations
                    Instances For
                      @[simp]
                      @[simp]
                      theorem CategoryTheory.Functor.FullyFaithful.isoEquiv_apply {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F : Functor C D} (hF : F.FullyFaithful) {X Y : C} (i : X Y) :
                      hF.isoEquiv i = F.mapIso i

                      Fully faithful functors are stable by composition.

                      Equations
                      Instances For
                        @[simp]
                        theorem CategoryTheory.Functor.FullyFaithful.comp_preimage {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u_1} [Category.{u_2, u_1} E] {F : Functor C D} (hF : F.FullyFaithful) {G : Functor D E} (hG : G.FullyFaithful) {X✝ Y✝ : C} (f : (F.comp G).obj X✝ (F.comp G).obj Y✝) :
                        (hF.comp hG).preimage f = hF.preimage (hG.preimage f)

                        If F ⋙ G is fully faithful and G is faithful, then F is fully faithful.

                        Equations
                        Instances For
                          theorem CategoryTheory.isIso_of_fully_faithful {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) [F.Full] [F.Faithful] {X Y : C} (f : X Y) [IsIso (F.map f)] :

                          If the image of a morphism under a fully faithful functor in an isomorphism, then the original morphisms is also an isomorphism.

                          theorem CategoryTheory.Functor.Full.of_iso {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F F' : Functor C D} [F.Full] (α : F F') :
                          F'.Full

                          If F is full, and naturally isomorphic to some F', then F' is also full.

                          theorem CategoryTheory.Functor.Faithful.of_comp_iso {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] {F : Functor C D} {G : Functor D E} {H : Functor C E} [H.Faithful] (h : F.comp G H) :
                          theorem CategoryTheory.Functor.Faithful.of_comp_eq {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] {F : Functor C D} {G : Functor D E} {H : Functor C E} [ℋ : H.Faithful] (h : F.comp G = H) :
                          def CategoryTheory.Functor.Faithful.div {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] (F : Functor C E) (G : Functor D E) [G.Faithful] (obj : CD) (h_obj : ∀ (X : C), G.obj (obj X) = F.obj X) (map : {X Y : C} → (X Y) → (obj X obj Y)) (h_map : ∀ {X Y : C} {f : X Y}, HEq (G.map (map f)) (F.map f)) :

                          “Divide” a functor by a faithful functor.

                          Equations
                          Instances For
                            theorem CategoryTheory.Functor.Faithful.div_comp {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] (F : Functor C E) [F.Faithful] (G : Functor D E) [G.Faithful] (obj : CD) (h_obj : ∀ (X : C), G.obj (obj X) = F.obj X) (map : {X Y : C} → (X Y) → (obj X obj Y)) (h_map : ∀ {X Y : C} {f : X Y}, HEq (G.map (map f)) (F.map f)) :
                            (Faithful.div F G obj h_obj map h_map).comp G = F
                            theorem CategoryTheory.Functor.Faithful.div_faithful {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] (F : Functor C E) [F.Faithful] (G : Functor D E) [G.Faithful] (obj : CD) (h_obj : ∀ (X : C), G.obj (obj X) = F.obj X) (map : {X Y : C} → (X Y) → (obj X obj Y)) (h_map : ∀ {X Y : C} {f : X Y}, HEq (G.map (map f)) (F.map f)) :
                            (Faithful.div F G obj h_obj map h_map).Faithful
                            instance CategoryTheory.Functor.Full.comp {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] (F : Functor C D) (G : Functor D E) [F.Full] [G.Full] :
                            (F.comp G).Full

                            If F ⋙ G is full and G is faithful, then F is full.

                            theorem CategoryTheory.Functor.Full.of_comp_faithful_iso {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] {F : Functor C D} {G : Functor D E} {H : Functor C E} [H.Full] [G.Faithful] (h : F.comp G H) :

                            If F ⋙ G is full and G is faithful, then F is full.

                            noncomputable def CategoryTheory.Functor.fullyFaithfulCancelRight {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] {F G : Functor C D} (H : Functor D E) [H.Full] [H.Faithful] (comp_iso : F.comp H G.comp H) :
                            F G

                            Given a natural isomorphism between F ⋙ H and G ⋙ H for a fully faithful functor H, we can 'cancel' it to give a natural iso between F and G.

                            Equations
                            Instances For
                              @[simp]
                              theorem CategoryTheory.Functor.fullyFaithfulCancelRight_hom_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] {F G : Functor C D} {H : Functor D E} [H.Full] [H.Faithful] (comp_iso : F.comp H G.comp H) (X : C) :
                              (fullyFaithfulCancelRight H comp_iso).hom.app X = H.preimage (comp_iso.hom.app X)
                              @[simp]
                              theorem CategoryTheory.Functor.fullyFaithfulCancelRight_inv_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] {F G : Functor C D} {H : Functor D E} [H.Full] [H.Faithful] (comp_iso : F.comp H G.comp H) (X : C) :
                              (fullyFaithfulCancelRight H comp_iso).inv.app X = H.preimage (comp_iso.inv.app X)