Documentation

Mathlib.AlgebraicTopology.ModelCategory.RightHomotopy

Right homotopies in model categories #

We introduce the types PrepathObject.RightHomotopy and PathObject.RightHomotopy of homotopies between morphisms X ⟶ Y relative to a (pre)path object of Y. Given two morphisms f and g, we introduce the relation RightHomotopyRel f g asserting the existence of a path object P and a right homotopy P.RightHomotopy f g, and we define the quotient type RightHomotopyClass X Y. We show that if Y is a fibrant object in a model category, then RightHomotopyRel is an equivalence relation on X ⟶ Y.

(This file dualizes the definitions in AlgebraicTopology.ModelCategory.LeftHomotopy.)

References #

Given a pre-path object P for Y, two maps f and g in X ⟶ Y are homotopic relative to P when there is a morphism h : X ⟶ P.P such that h ≫ P.p₀ = f and h ≫ P.p₁ = g.

Instances For

    f : X ⟶ Y is right homotopic to itself relative to any pre-path object.

    Equations
    Instances For

      If f and g are homotopic relative to a pre-path object P, then g and f are homotopic relative to P.symm

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

        If f₀ is homotopic to f₁ relative to a pre-path object 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.PrepathObject.RightHomotopy.trans_h {C : Type u} [CategoryTheory.Category.{v, u} C] {Y : C} {P : PrepathObject Y} {X : C} {f₀ f₁ f₂ : X Y} (h : P.RightHomotopy f₀ f₁) {P' : PrepathObject Y} (h' : P'.RightHomotopy f₁ f₂) [CategoryTheory.Limits.HasPullback P.p₁ P'.p₀] :

          Right homotopies are compatible with precomposition.

          Equations
          Instances For
            @[simp]
            @[reducible, inline]

            Given a path object 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 right homotopic to itself relative to any path object.

              Equations
              Instances For
                @[reducible, inline]

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

                Equations
                Instances For
                  @[reducible, inline]

                  Right homotopies are compatible with precomposition.

                  Equations
                  Instances For
                    @[reducible, inline]
                    noncomputable abbrev HomotopicalAlgebra.PathObject.RightHomotopy.trans {C : Type u} [CategoryTheory.Category.{v, u} C] {X Y : C} [ModelCategory C] {P : PathObject Y} [IsFibrant Y] {f₀ f₁ f₂ : X Y} (h : P.RightHomotopy f₀ f₁) {P' : PathObject Y} [P'.IsGood] (h' : P'.RightHomotopy f₁ f₂) [CategoryTheory.Limits.HasPullback P.p₁ P'.p₀] :
                    (P.trans P').RightHomotopy f₀ f₂

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

                    Equations
                    Instances For
                      theorem HomotopicalAlgebra.PathObject.RightHomotopy.homotopy_extension {C : Type u} [CategoryTheory.Category.{v, u} C] [ModelCategory C] {A B X : C} {P : PathObject B} {f₀ f₁ : A B} [IsFibrant B] [P.IsGood] (h : P.RightHomotopy f₀ f₁) (i : A X) [Cofibration i] (l₀ : X B) (hl₀ : CategoryTheory.CategoryStruct.comp i l₀ = f₀ := by cat_disch) :
                      ∃ (l₁ : X B) (h' : P.RightHomotopy l₀ l₁), CategoryTheory.CategoryStruct.comp i h'.h = h.h

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

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

                      Equations
                      Instances For
                        theorem HomotopicalAlgebra.RightHomotopyRel.trans {C : Type u} [CategoryTheory.Category.{v, u} C] {X Y : C} [ModelCategory C] {f₀ f₁ f₂ : X Y} [IsFibrant Y] (h : RightHomotopyRel f₀ f₁) (h' : RightHomotopyRel 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 right homotopies.

                        Equations
                        Instances For