Documentation

Mathlib.CategoryTheory.Shift.Twist

Twisting a shift #

Given a category C equipped with a shift by a monoid A, we introduce a structure t : TwistShiftData C A which consists of a collection of invertible elements in the center of the category C (typically, C will be preadditive, and these will be signs), which allow to introduce a type synonym category t.Category with identical shift functors as C but where the isomorphisms shiftFunctorAdd have been modified.

structure CategoryTheory.TwistShiftData (C : Type u) [Category.{v, u} C] (A : Type w) [AddMonoid A] [HasShift C A] :
Type (max (max u v) w)

Given a category C equipped with a shift by a monoid A

  • z (a b : A) : (CatCenter C)ˣ

    The invertible elements in the center of C which are used to modify the shiftFunctorAdd isomorphisms .

  • z_zero_zero : self.z 0 0 = 1
  • assoc (a b c : A) : self.z (a + b) c * self.z a b = self.z a (b + c) * self.z b c
  • commShift (a b : A) : NatTrans.CommShift (↑(self.z a b)) A
Instances For
    @[simp]
    theorem CategoryTheory.TwistShiftData.z_zero_right {C : Type u} [Category.{v, u} C] {A : Type w} [AddMonoid A] [HasShift C A] (t : TwistShiftData C A) (a : A) :
    t.z a 0 = 1
    @[simp]
    theorem CategoryTheory.TwistShiftData.z_zero_left {C : Type u} [Category.{v, u} C] {A : Type w} [AddMonoid A] [HasShift C A] (t : TwistShiftData C A) (b : A) :
    t.z 0 b = 1
    @[simp]
    theorem CategoryTheory.TwistShiftData.shift_z_app {C : Type u} [Category.{v, u} C] {A : Type w} [AddMonoid A] [HasShift C A] (t : TwistShiftData C A) (a b c : A) (X : C) :
    (shiftFunctor C c).map ((↑(t.z a b)).app X) = (↑(t.z a b)).app ((shiftFunctor C c).obj X)

    Given t : TwistShiftData C A, this is a type synonym for the category C, which the same shift functors as C but where the shiftFunctorAdd isomorphisms have been modified using t.

    Equations
    Instances For

      Given t : TwistShiftData C A, the shift on the category TwistShift t has the same shift functors as C, the same isomorphism shiftFunctorZero isomorphism, but the shiftFunctorAdd isomorphisms are modified using t.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        noncomputable def CategoryTheory.TwistShiftData.shiftIso {C : Type u} [Category.{v, u} C] {A : Type w} [AddMonoid A] [HasShift C A] (t : TwistShiftData C A) (m : A) :

        Given t : TwistShiftData C A, the shift functors on t.Category identify to the shift functors on C.

        Equations
        Instances For
          @[simp]
          theorem CategoryTheory.TwistShiftData.shiftFunctorAdd'_hom_app {C : Type u} [Category.{v, u} C] {A : Type w} [AddMonoid A] [HasShift C A] (t : TwistShiftData C A) (i j k : A) (h : i + j = k) (X : t.Category) :
          @[simp]