Documentation

Mathlib.Algebra.Homology.ExactSequence

Exact sequences #

A sequence of n composable arrows S : ComposableArrows C (i.e. a functor S : Fin (n + 1) ⥤ C) is said to be exact (S.Exact) if the composition of two consecutive arrows are zero (S.IsComplex) and the diagram is exact at each i for 1 ≤ i < n.

Together with the inductive construction of composable arrows ComposableArrows.precomp, this is useful in order to state that certain finite sequences of morphisms are exact (e.g the snake lemma), even though in the applications it would usually be more convenient to use individual lemmas expressing the exactness at a particular object.

This implementation is a refactor of exact_seq with appeared in the Liquid Tensor Experiement as a property of lists in Arrow C.

F : ComposableArrows C n is a complex if all compositions of two consecutive arrows are zero.

Instances For
    @[reducible]

    The short complex consisting of maps S.map' i j and S.map' j k when we know that S : ComposableArrows C n satisfies S.IsComplex.

    Equations
    Instances For
      @[inline, reducible]

      The short complex consisting of maps S.map' i (i + 1) and S.map' (i + 1) (i + 2) when we know that S : ComposableArrows C n satisfies S.IsComplex.

      Equations
      Instances For

        F : ComposableArrows C n is exact if it is a complex and that all short complexes consisting of two consecutive arrows are exact.

        Instances For
          @[simp]
          theorem CategoryTheory.ComposableArrows.sc'Map_τ₃ {C : Type u_1} [CategoryTheory.Category.{u_2, u_1} C] [CategoryTheory.Limits.HasZeroMorphisms C] {n : } {S₁ : CategoryTheory.ComposableArrows C n} {S₂ : CategoryTheory.ComposableArrows C n} (φ : S₁ S₂) (h₁ : CategoryTheory.ComposableArrows.IsComplex S₁) (h₂ : CategoryTheory.ComposableArrows.IsComplex S₂) (i : ) (j : ) (k : ) (hij : autoParam (i + 1 = j) _auto✝) (hjk : autoParam (j + 1 = k) _auto✝) (hk : autoParam (k n) _auto✝) :
          (CategoryTheory.ComposableArrows.sc'Map φ h₁ h₂ i j k hij hjk hk).τ₃ = φ.app { val := k, isLt := }
          @[simp]
          theorem CategoryTheory.ComposableArrows.sc'Map_τ₁ {C : Type u_1} [CategoryTheory.Category.{u_2, u_1} C] [CategoryTheory.Limits.HasZeroMorphisms C] {n : } {S₁ : CategoryTheory.ComposableArrows C n} {S₂ : CategoryTheory.ComposableArrows C n} (φ : S₁ S₂) (h₁ : CategoryTheory.ComposableArrows.IsComplex S₁) (h₂ : CategoryTheory.ComposableArrows.IsComplex S₂) (i : ) (j : ) (k : ) (hij : autoParam (i + 1 = j) _auto✝) (hjk : autoParam (j + 1 = k) _auto✝) (hk : autoParam (k n) _auto✝) :
          (CategoryTheory.ComposableArrows.sc'Map φ h₁ h₂ i j k hij hjk hk).τ₁ = φ.app { val := i, isLt := }
          @[simp]
          theorem CategoryTheory.ComposableArrows.sc'Map_τ₂ {C : Type u_1} [CategoryTheory.Category.{u_2, u_1} C] [CategoryTheory.Limits.HasZeroMorphisms C] {n : } {S₁ : CategoryTheory.ComposableArrows C n} {S₂ : CategoryTheory.ComposableArrows C n} (φ : S₁ S₂) (h₁ : CategoryTheory.ComposableArrows.IsComplex S₁) (h₂ : CategoryTheory.ComposableArrows.IsComplex S₂) (i : ) (j : ) (k : ) (hij : autoParam (i + 1 = j) _auto✝) (hjk : autoParam (j + 1 = k) _auto✝) (hk : autoParam (k n) _auto✝) :
          (CategoryTheory.ComposableArrows.sc'Map φ h₁ h₂ i j k hij hjk hk).τ₂ = φ.app { val := j, isLt := }

          Functoriality maps for ComposableArrows.sc'.

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

            The isomorphism S₁.sc' _ i j k ≅ S₂.sc' _ i j k induced by an isomorphism S₁ ≅ S₂ in ComposableArrows C n.

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

              The isomorphism S₁.sc _ i ≅ S₂.sc _ i induced by an isomorphism S₁ ≅ S₂ in ComposableArrows C n.

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

                If S : ComposableArrows C (n + 2) is such that the first two arrows form an exact sequence and that the tail S.δ₀ is exact, then S is also exact. See ShortComplex.SnakeInput.snake_lemma in Algebra.Homology.ShortComplex.SnakeLemma for a use of this lemma.