Documentation

Mathlib.AlgebraicTopology.RelativeCellComplex.Basic

Relative cell complexes #

In this file, we define a structure RelativeCellComplex which expresses that a morphism f : X ⟶ Y is a transfinite composition of morphisms, all of which consists in attaching cells. Here, we allow a different family of authorized cells at each step. For example, (relative) CW-complexes are defined in the file Mathlib.Topology.CWComplex.Abstract.Basic by requiring that at the nth step, we attach n-disks along their boundaries.

This structure RelativeCellComplex is also used in the formalization of the small object argument, see the file Mathlib.CategoryTheory.SmallObject.IsCardinalForSmallObjectArgument.

References #

structure HomotopicalAlgebra.RelativeCellComplex {C : Type u} [CategoryTheory.Category.{v, u} C] {J : Type w'} [LinearOrder J] [OrderBot J] [SuccOrder J] [WellFoundedLT J] {α : JType t} {A B : (j : J) → α jC} (basicCell : (j : J) → (i : α j) → A j i B j i) {X Y : C} (f : X Y) extends CategoryTheory.TransfiniteCompositionOfShape J f :
Type (max (max (max (max t u) v) (w + 1)) w')

Let J be a well ordered type. Assume that for each j : J, we have a family basicCell j of morphisms. A relative cell complex is a morphism f : X ⟶ Y which is a transfinite composition of morphisms in such a way that at the step j : J, we attach cells in the family basicCell j.

Instances For
    structure HomotopicalAlgebra.RelativeCellComplex.Cells {C : Type u} [CategoryTheory.Category.{v, u} C] {J : Type w'} [LinearOrder J] [OrderBot J] [SuccOrder J] [WellFoundedLT J] {α : JType t} {A B : (j : J) → α jC} {basicCell : (j : J) → (i : α j) → A j i B j i} {X Y : C} {f : X Y} (c : RelativeCellComplex basicCell f) :
    Type (max u_1 w')

    The index type of cells in a relative cell complex.

    • j : J

      the step where the cell is added

    • hj : ¬IsMax self.j
    • k : (c.attachCells self.j ).ι

      the index of the cell

    Instances For
      def HomotopicalAlgebra.RelativeCellComplex.Cells.i {C : Type u} [CategoryTheory.Category.{v, u} C] {J : Type w'} [LinearOrder J] [OrderBot J] [SuccOrder J] [WellFoundedLT J] {α : JType t} {A B : (j : J) → α jC} {basicCell : (j : J) → (i : α j) → A j i B j i} {X Y : C} {f : X Y} {c : RelativeCellComplex basicCell f} (γ : c.Cells) :
      α γ.j

      Given a cell γ in a relative cell complex, this is the corresponding index in the family of morphisms basicCell γ.j.

      Equations
      Instances For
        def HomotopicalAlgebra.RelativeCellComplex.Cells.ι {C : Type u} [CategoryTheory.Category.{v, u} C] {J : Type w'} [LinearOrder J] [OrderBot J] [SuccOrder J] [WellFoundedLT J] {α : JType t} {A B : (j : J) → α jC} {basicCell : (j : J) → (i : α j) → A j i B j i} {X Y : C} {f : X Y} {c : RelativeCellComplex basicCell f} (γ : c.Cells) :
        B γ.j γ.i Y

        The inclusion of a cell.

        Equations
        Instances For
          theorem HomotopicalAlgebra.RelativeCellComplex.hom_ext {C : Type u} [CategoryTheory.Category.{v, u} C] {J : Type w'} [LinearOrder J] [OrderBot J] [SuccOrder J] [WellFoundedLT J] {α : JType t} {A B : (j : J) → α jC} {basicCell : (j : J) → (i : α j) → A j i B j i} {X Y : C} {f : X Y} (c : RelativeCellComplex basicCell f) {Z : C} {φ₁ φ₂ : Y Z} (h₀ : CategoryTheory.CategoryStruct.comp f φ₁ = CategoryTheory.CategoryStruct.comp f φ₂) (h : ∀ (γ : c.Cells), CategoryTheory.CategoryStruct.comp γ.ι φ₁ = CategoryTheory.CategoryStruct.comp γ.ι φ₂) :
          φ₁ = φ₂

          If f is a relative cell complex with respect to a constant family of morphisms g, then f is a transfinite composition of pushouts of coproducts of morphisms in the family g.

          Equations
          Instances For