Documentation

Mathlib.CategoryTheory.Idempotents.Karoubi

The Karoubi envelope of a category #

In this file, we define the Karoubi envelope Karoubi C of a category C.

Main constructions and definitions #

structure CategoryTheory.Idempotents.Karoubi (C : Type u_1) [Category.{u_2, u_1} C] :
Type (max u_1 u_2)

In a preadditive category C, when an object X decomposes as X ≅ P ⨿ Q, one may consider P as a direct factor of X and up to unique isomorphism, it is determined by the obvious idempotent X ⟶ P ⟶ X which is the projection onto P with kernel Q. More generally, one may define a formal direct factor of an object X : C : it consists of an idempotent p : XX which is thought as the "formal image" of p. The type Karoubi C shall be the type of the objects of the karoubi envelope of C. It makes sense for any category C.

  • X : C

    an object of the underlying category

  • p : self.X self.X

    an endomorphism of the object

  • idem : CategoryStruct.comp self.p self.p = self.p

    the condition that the given endomorphism is an idempotent

Instances For
    @[simp]
    theorem CategoryTheory.Idempotents.Karoubi.idem_assoc {C : Type u_1} [Category.{u_2, u_1} C] (self : Karoubi C) {Z : C} (h : self.X Z) :
    theorem CategoryTheory.Idempotents.Karoubi.ext {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (h_X : P.X = Q.X) (h_p : CategoryStruct.comp P.p (eqToHom h_X) = CategoryStruct.comp (eqToHom h_X) Q.p) :
    P = Q

    A morphism P ⟶ Q in the category Karoubi C is a morphism in the underlying category C which satisfies a relation, which in the preadditive case, expresses that it induces a map between the corresponding "formal direct factors" and that it vanishes on the complement formal direct factor.

    Instances For
      theorem CategoryTheory.Idempotents.Karoubi.Hom.ext {C : Type u_1} {inst✝ : Category.{u_2, u_1} C} {P Q : Karoubi C} {x y : P.Hom Q} (f : x.f = y.f) :
      x = y
      Equations
      • P.instInhabitedHomOfPreadditive Q = { default := { f := 0, comm := } }
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.p_comp {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f : P.Hom Q) :
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.p_comp_assoc {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f : P.Hom Q) {Z : C} (h : Q.X Z) :
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.comp_p {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f : P.Hom Q) :
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.comp_p_assoc {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f : P.Hom Q) {Z : C} (h : Q.X Z) :
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.hom_ext_iff {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} {f g : P Q} :
      f = g f.f = g.f
      theorem CategoryTheory.Idempotents.Karoubi.hom_ext {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f g : P Q) (h : f.f = g.f) :
      f = g
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.comp_f {C : Type u_1} [Category.{u_2, u_1} C] {P Q R : Karoubi C} (f : P Q) (g : Q R) :
      @[deprecated "No deprecation message was provided." (since := "2024-07-15")]
      theorem CategoryTheory.Idempotents.Karoubi.id_eq {C : Type u_1} [Category.{u_2, u_1} C] {P : Karoubi C} :
      CategoryStruct.id P = { f := P.p, comm := }

      It is possible to coerce an object of C into an object of Karoubi C. See also the functor toKaroubi.

      Equations
      theorem CategoryTheory.Idempotents.Karoubi.coe_X {C : Type u_1} [Category.{u_2, u_1} C] (X : C) :
      { X := X, p := CategoryStruct.id X, idem := }.X = X
      @[simp]
      theorem CategoryTheory.Idempotents.Karoubi.coe_p {C : Type u_1} [Category.{u_2, u_1} C] (X : C) :
      { X := X, p := CategoryStruct.id X, idem := }.p = CategoryStruct.id X

      The obvious fully faithful functor toKaroubi sends an object X : C to the obvious formal direct factor of X given by 𝟙 X.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem CategoryTheory.Idempotents.toKaroubi_map_f (C : Type u_1) [Category.{u_2, u_1} C] {X✝ Y✝ : C} (f : X✝ Y✝) :
        ((toKaroubi C).map f).f = f
        @[simp]
        theorem CategoryTheory.Idempotents.toKaroubi_obj_X (C : Type u_1) [Category.{u_2, u_1} C] (X : C) :
        ((toKaroubi C).obj X).X = X
        Equations
        @[simp]
        theorem CategoryTheory.Idempotents.instAdd_add {C : Type u_1} [Category.{u_2, u_1} C] [Preadditive C] {P Q : Karoubi C} (f g : P Q) :
        f + g = { f := f.f + g.f, comm := }
        Equations
        @[simp]
        theorem CategoryTheory.Idempotents.instNeg_neg {C : Type u_1} [Category.{u_2, u_1} C] [Preadditive C] {P Q : Karoubi C} (f : P Q) :
        -f = { f := -f.f, comm := }
        Equations
        @[simp]
        theorem CategoryTheory.Idempotents.instZero_zero {C : Type u_1} [Category.{u_2, u_1} C] [Preadditive C] {P Q : Karoubi C} :
        0 = { f := 0, comm := }

        The map sending f : P ⟶ Q to f.f : P.X ⟶ Q.X is additive.

        Equations
        • P.inclusionHom Q = { toFun := fun (f : P Q) => f.f, map_zero' := , map_add' := }
        Instances For
          @[simp]
          theorem CategoryTheory.Idempotents.Karoubi.inclusionHom_apply {C : Type u_1} [Category.{u_2, u_1} C] [Preadditive C] (P Q : Karoubi C) (f : P Q) :
          (P.inclusionHom Q) f = f.f
          @[simp]
          theorem CategoryTheory.Idempotents.Karoubi.sum_hom {C : Type u_1} [Category.{u_3, u_1} C] [Preadditive C] {P Q : Karoubi C} {α : Type u_2} (s : Finset α) (f : α(P Q)) :
          (∑ xs, f x).f = xs, (f x).f

          The category Karoubi C is preadditive if C is.

          Equations

          If C is idempotent complete, the functor toKaroubi : C ⥤ Karoubi C is an equivalence.

          def CategoryTheory.Idempotents.Karoubi.decompId_i {C : Type u_1} [Category.{u_2, u_1} C] (P : Karoubi C) :
          P { X := P.X, p := CategoryStruct.id P.X, idem := }

          The split mono which appears in the factorisation decompId P.

          Equations
          • P.decompId_i = { f := P.p, comm := }
          Instances For
            @[simp]
            theorem CategoryTheory.Idempotents.Karoubi.decompId_i_f {C : Type u_1} [Category.{u_2, u_1} C] (P : Karoubi C) :
            P.decompId_i.f = P.p
            def CategoryTheory.Idempotents.Karoubi.decompId_p {C : Type u_1} [Category.{u_2, u_1} C] (P : Karoubi C) :
            { X := P.X, p := CategoryStruct.id P.X, idem := } P

            The split epi which appears in the factorisation decompId P.

            Equations
            • P.decompId_p = { f := P.p, comm := }
            Instances For
              @[simp]
              theorem CategoryTheory.Idempotents.Karoubi.decompId_p_f {C : Type u_1} [Category.{u_2, u_1} C] (P : Karoubi C) :
              P.decompId_p.f = P.p

              The formal direct factor of P.X given by the idempotent P.p in the category C is actually a direct factor in the category Karoubi C.

              theorem CategoryTheory.Idempotents.Karoubi.decomp_p {C : Type u_1} [Category.{u_2, u_1} C] (P : Karoubi C) :
              (toKaroubi C).map P.p = CategoryStruct.comp P.decompId_p P.decompId_i
              theorem CategoryTheory.Idempotents.Karoubi.decompId_p_toKaroubi {C : Type u_1} [Category.{u_2, u_1} C] (X : C) :
              ((toKaroubi C).obj X).decompId_p = CategoryStruct.id { X := ((toKaroubi C).obj X).X, p := CategoryStruct.id ((toKaroubi C).obj X).X, idem := }
              theorem CategoryTheory.Idempotents.Karoubi.decompId_i_naturality {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f : P Q) :
              CategoryStruct.comp f Q.decompId_i = CategoryStruct.comp P.decompId_i { f := f.f, comm := }
              theorem CategoryTheory.Idempotents.Karoubi.decompId_p_naturality {C : Type u_1} [Category.{u_2, u_1} C] {P Q : Karoubi C} (f : P Q) :
              CategoryStruct.comp P.decompId_p f = CategoryStruct.comp { f := f.f, comm := } Q.decompId_p
              @[simp]
              theorem CategoryTheory.Idempotents.Karoubi.zsmul_hom {C : Type u_1} [Category.{u_2, u_1} C] [Preadditive C] {P Q : Karoubi C} (n : ) (f : P Q) :
              (n f).f = n f.f