Documentation

Mathlib.AlgebraicTopology.ModelCategory.LeftHomotopy

Left homotopies in model categories #

We introduce the types Precylinder.LeftHomotopy and Cylinder.LeftHomotopy of homotopies between morphisms X ⟶ Y relative to a (pre)cylinder of X. Given two morphisms f and g, we introduce the relation LeftHomotopyRel f g asserting the existence of a cylinder object P and a left homotopy P.LeftHomotopy f g, and we define the quotient type LeftHomotopyClass X Y. We show that if X is a cofibrant object in a model category, then LeftHomotopyRel is an equivalence relation on X ⟶ Y.

References #

structure HomotopicalAlgebra.Precylinder.LeftHomotopy {C : Type u} [CategoryTheory.Category.{v, u} C] {X : C} (P : Precylinder X) {Y : C} (f g : X Y) :

Given a precylinder P for X, two maps f and g in X ⟶ Y are homotopic relative to P when there is a morphism h : P.I ⟶ Y such that P.i₀ ≫ h = f and P.i₁ ≫ h = g.

Instances For

    f : X ⟶ Y is left homotopic to itself relative to any precylinder.

    Equations
    Instances For

      If f and g are homotopic relative to a precylinder P, then g and f are homotopic relative to P.symm

      Equations
      • h.symm = { h := h.h, h₀ := , h₁ := }
      Instances For
        @[simp]
        theorem HomotopicalAlgebra.Precylinder.LeftHomotopy.symm_h {C : Type u} [CategoryTheory.Category.{v, u} C] {X : C} {P : Precylinder X} {Y : C} {f g : X Y} (h : P.LeftHomotopy f g) :
        h.symm.h = h.h
        noncomputable def HomotopicalAlgebra.Precylinder.LeftHomotopy.trans {C : Type u} [CategoryTheory.Category.{v, u} C] {X : C} {P : Precylinder X} {Y : C} {f₀ f₁ f₂ : X Y} (h : P.LeftHomotopy f₀ f₁) {P' : Precylinder X} (h' : P'.LeftHomotopy f₁ f₂) [CategoryTheory.Limits.HasPushout P.i₁ P'.i₀] :
        (P.trans P').LeftHomotopy f₀ f₂

        If f₀ is homotopic to f₁ relative to a precylinder P, and f₁ is homotopic to f₂ relative to P', then f₀ is homotopic to f₂ relative to P.trans P'.

        Equations
        Instances For
          @[simp]
          theorem HomotopicalAlgebra.Precylinder.LeftHomotopy.trans_h {C : Type u} [CategoryTheory.Category.{v, u} C] {X : C} {P : Precylinder X} {Y : C} {f₀ f₁ f₂ : X Y} (h : P.LeftHomotopy f₀ f₁) {P' : Precylinder X} (h' : P'.LeftHomotopy f₁ f₂) [CategoryTheory.Limits.HasPushout P.i₁ P'.i₀] :

          Left homotopies are compatible with postcomposition.

          Equations
          Instances For
            @[simp]
            theorem HomotopicalAlgebra.Precylinder.LeftHomotopy.postcomp_h {C : Type u} [CategoryTheory.Category.{v, u} C] {X : C} {P : Precylinder X} {Y : C} {f g : X Y} (h : P.LeftHomotopy f g) {Z : C} (p : Y Z) :
            @[reducible, inline]

            Given a cylinder P for X, two maps f and g in X ⟶ Y are homotopic relative to P when there is a morphism h : P.I ⟶ Y such that P.i₀ ≫ h = f and P.i₁ ≫ h = g.

            Equations
            Instances For
              @[reducible, inline]

              f : X ⟶ Y is left homotopic to itself relative to any cylinder.

              Equations
              Instances For
                @[reducible, inline]

                If f and g are homotopic relative to a cylinder P, then g and f are homotopic relative to P.symm.

                Equations
                Instances For
                  @[reducible, inline]

                  Left homotopies are compatible with postcomposition.

                  Equations
                  Instances For
                    @[reducible, inline]
                    noncomputable abbrev HomotopicalAlgebra.Cylinder.LeftHomotopy.trans {C : Type u} [CategoryTheory.Category.{v, u} C] {X Y : C} [ModelCategory C] {P : Cylinder X} [IsCofibrant X] {f₀ f₁ f₂ : X Y} (h : P.LeftHomotopy f₀ f₁) {P' : Cylinder X} [P'.IsGood] (h' : P'.LeftHomotopy f₁ f₂) [CategoryTheory.Limits.HasPushout P.i₁ P'.i₀] :
                    (P.trans P').LeftHomotopy f₀ f₂

                    If f₀ : X ⟶ Y is homotopic to f₁ relative to a cylinder P, and f₁ is homotopic to f₂ relative to a good cylinder P', then f₀ is homotopic to f₂ relative to the cylinder P.trans P' when X is cofibrant.

                    Equations
                    Instances For
                      theorem HomotopicalAlgebra.Cylinder.LeftHomotopy.covering_homotopy {C : Type u} [CategoryTheory.Category.{v, u} C] [ModelCategory C] {A E B : C} {P : Cylinder A} {f₀ f₁ : A B} [IsCofibrant A] [P.IsGood] (h : P.LeftHomotopy f₀ f₁) (p : E B) [Fibration p] (l₀ : A E) (hl₀ : CategoryTheory.CategoryStruct.comp l₀ p = f₀ := by cat_disch) :
                      ∃ (l₁ : A E) (h' : P.LeftHomotopy l₀ l₁), CategoryTheory.CategoryStruct.comp h'.h p = h.h

                      The covering homotopy theorem: if p : E ⟶ B is a fibration, l₀ : A ⟶ E is a morphism, if there is a left homotopy h between the composition f₀ := l₀ ≫ p and a morphism f₁ : A ⟶ B, then there exists a morphism l₁ : A ⟶ E and a left homotopy h' from l₀ to l₁ which is compatible with h (in particular, l₁ ≫ p = f₁).

                      The left homotopy relation on morphisms in a category with weak equivalences.

                      Equations
                      Instances For
                        theorem HomotopicalAlgebra.LeftHomotopyRel.trans {C : Type u} [CategoryTheory.Category.{v, u} C] {X Y : C} [ModelCategory C] {f₀ f₁ f₂ : X Y} [IsCofibrant X] (h : LeftHomotopyRel f₀ f₁) (h' : LeftHomotopyRel f₁ f₂) :
                        LeftHomotopyRel f₀ f₂

                        In a category with weak equivalences, this is the quotient of the type of morphisms X ⟶ Y by the equivalence relation generated by left homotopies.

                        Equations
                        Instances For