Documentation

Mathlib.Algebra.Homology.ComplexShape

Shapes of homological complexes #

We define a structure ComplexShape ι for describing the shapes of homological complexes indexed by a type ι. This is intended to capture chain complexes and cochain complexes, indexed by either or , as well as more exotic examples.

Rather than insisting that the indexing type has a succ function specifying where differentials should go, inside c : ComplexShape we have c.Rel : ι → ι → Prop, and when we define HomologicalComplex we only allow nonzero differentials d i j from i to j if c.Rel i j. Further, we require that { j // c.Rel i j } and { i // c.Rel i j } are subsingletons. This means that the shape consists of some union of lines, rays, intervals, and circles.

Convenience functions c.next and c.prev provide these related elements when they exist, and return their input otherwise.

This design aims to avoid certain problems arising from dependent type theory. In particular we never have to ensure morphisms d i : X i ⟶ X (succ i) compose as expected (which would often require rewriting by equations in the indexing type). Instead such identities become separate proof obligations when verifying that a complex we've constructed is of the desired shape.

If α is an AddRightCancelSemigroup, then we define up α : ComplexShape α, the shape appropriate for cohomology, so d : X i ⟶ X j is nonzero only when j = i + 1, as well as down α : ComplexShape α, appropriate for homology, so d : X i ⟶ X j is nonzero only when i = j + 1. (Later we'll introduce CochainComplex and ChainComplex as abbreviations for HomologicalComplex with one of these shapes baked in.)

theorem ComplexShape.ext {ι : Type u_1} (x : ComplexShape ι) (y : ComplexShape ι) (Rel : x.Rel = y.Rel) :
x = y
theorem ComplexShape.ext_iff {ι : Type u_1} (x : ComplexShape ι) (y : ComplexShape ι) :
x = y x.Rel = y.Rel
structure ComplexShape (ι : Type u_1) :
Type u_1

A c : ComplexShape ι describes the shape of a chain complex, with chain groups indexed by ι. Typically ι will be , , or Fin n.

There is a relation Rel : ι → ι → Prop, and we will only allow a non-zero differential from i to j when Rel i j.

There are axioms which imply { j // c.Rel i j } and { i // c.Rel i j } are subsingletons. This means that the shape consists of some union of lines, rays, intervals, and circles.

Below we define c.next and c.prev which provide these related elements.

  • Rel : ιιProp

    Nonzero differentials X i ⟶ X j shall be allowed on homological complexes when Rel i j holds.

  • next_eq : ∀ {i j j' : ι}, self.Rel i jself.Rel i j'j = j'

    There is at most one nonzero differential from X i.

  • prev_eq : ∀ {i i' j : ι}, self.Rel i jself.Rel i' ji = i'

    There is at most one nonzero differential to X j.

Instances For
    @[simp]
    theorem ComplexShape.refl_Rel (ι : Type u_2) (i : ι) (j : ι) :
    (ComplexShape.refl ι).Rel i j = (i = j)

    The complex shape where only differentials from each X.i to itself are allowed.

    This is mostly only useful so we can describe the relation of "related in k steps" below.

    Equations
    Instances For
      @[simp]
      theorem ComplexShape.symm_Rel {ι : Type u_1} (c : ComplexShape ι) (i : ι) (j : ι) :
      (ComplexShape.symm c).Rel i j = c.Rel j i
      def ComplexShape.symm {ι : Type u_1} (c : ComplexShape ι) :

      The reverse of a ComplexShape.

      Equations
      • ComplexShape.symm c = { Rel := fun (i j : ι) => c.Rel j i, next_eq := , prev_eq := }
      Instances For
        theorem ComplexShape.symm_bijective {ι : Type u_1} :
        Function.Bijective ComplexShape.symm
        def ComplexShape.trans {ι : Type u_1} (c₁ : ComplexShape ι) (c₂ : ComplexShape ι) :

        The "composition" of two ComplexShapes.

        We need this to define "related in k steps" later.

        Equations
        Instances For
          instance ComplexShape.subsingleton_next {ι : Type u_1} (c : ComplexShape ι) (i : ι) :
          Subsingleton { j : ι // c.Rel i j }
          Equations
          • =
          instance ComplexShape.subsingleton_prev {ι : Type u_1} (c : ComplexShape ι) (j : ι) :
          Subsingleton { i : ι // c.Rel i j }
          Equations
          • =
          def ComplexShape.next {ι : Type u_1} (c : ComplexShape ι) (i : ι) :
          ι

          An arbitrary choice of index j such that Rel i j, if such exists. Returns i otherwise.

          Equations
          Instances For
            def ComplexShape.prev {ι : Type u_1} (c : ComplexShape ι) (j : ι) :
            ι

            An arbitrary choice of index i such that Rel i j, if such exists. Returns j otherwise.

            Equations
            Instances For
              theorem ComplexShape.next_eq' {ι : Type u_1} (c : ComplexShape ι) {i : ι} {j : ι} (h : c.Rel i j) :
              theorem ComplexShape.prev_eq' {ι : Type u_1} (c : ComplexShape ι) {i : ι} {j : ι} (h : c.Rel i j) :
              @[simp]
              theorem ComplexShape.up'_Rel {α : Type u_2} [AddRightCancelSemigroup α] (a : α) (i : α) (j : α) :
              (ComplexShape.up' a).Rel i j = (i + a = j)

              The ComplexShape allowing differentials from X i to X (i+a). (For example when a = 1, a cohomology theory indexed by or )

              Equations
              Instances For
                @[simp]
                theorem ComplexShape.down'_Rel {α : Type u_2} [AddRightCancelSemigroup α] (a : α) (i : α) (j : α) :
                (ComplexShape.down' a).Rel i j = (j + a = i)

                The ComplexShape allowing differentials from X (j+a) to X j. (For example when a = 1, a homology theory indexed by or )

                Equations
                Instances For
                  theorem ComplexShape.down'_mk {α : Type u_2} [AddRightCancelSemigroup α] (a : α) (i : α) (j : α) (h : j + a = i) :
                  (ComplexShape.down' a).Rel i j
                  @[simp]
                  theorem ComplexShape.up_Rel (α : Type u_2) [AddRightCancelSemigroup α] [One α] (i : α) (j : α) :
                  (ComplexShape.up α).Rel i j = (i + 1 = j)

                  The ComplexShape appropriate for cohomology, so d : X i ⟶ X j only when j = i + 1.

                  Equations
                  Instances For
                    @[simp]
                    theorem ComplexShape.down_Rel (α : Type u_2) [AddRightCancelSemigroup α] [One α] (i : α) (j : α) :
                    (ComplexShape.down α).Rel i j = (j + 1 = i)

                    The ComplexShape appropriate for homology, so d : X i ⟶ X j only when i = j + 1.

                    Equations
                    Instances For
                      theorem ComplexShape.down_mk {α : Type u_2} [AddRightCancelSemigroup α] [One α] (i : α) (j : α) (h : j + 1 = i) :
                      (ComplexShape.down α).Rel i j