Relative morphisms of simplicial sets #
Given two simplicial sets X and Y, and subcomplexes A of X, and B of Y,
we introduce a type RelativeMorphism A B φ of morphisms X ⟶ Y which induce
a given morphism of simplicial sets A ⟶ B. We define homotopies between
these relative morphisms and introduce the quotient type of homotopy classes.
Given a morphism of simplicial sets φ : A ⟶ B with A : X.Subcomplex and
B : Y.Subcomplex, this is the type of morphisms X ⟶ Y which induce φ.
The underlying morphism of simplicial sets.
Instances For
The constant relative morphism when the given subcomplex of the target
is generated by a 0-simplex.
Equations
- SSet.RelativeMorphism.const = { map := SSet.const y, comm := ⋯ }
Instances For
Given a morphism f : X ⟶ Y of simplicial set which sends a 0-simplex x to y,
this is the pointed morphism (X, x) ⟶ (Y, y).
Equations
- SSet.RelativeMorphism.ofSimplex₀ f x y h = { map := f, comm := ⋯ }
Instances For
The type of homotopies between morphisms of simplicial sets relatively to given subcomplexes.
The homotopy.
- rel : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategoryStruct.whiskerRight A.ι (stdSimplex.obj (SimplexCategory.mk 1))) self.h = CategoryTheory.CategoryStruct.comp (CategoryTheory.SemiCartesianMonoidalCategory.fst A.toSSet (stdSimplex.obj (SimplexCategory.mk 1))) (CategoryTheory.CategoryStruct.comp φ B.ι)
Instances For
Two equal relative morphisms are homotopic.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A relative morphism is homotopic to itself.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The type of homotopy classes of relative morphisms.
Equations
- SSet.RelativeMorphism.HomotopyClass A B φ = Quot fun (f g : SSet.RelativeMorphism A B φ) => Nonempty (f.Homotopy g)
Instances For
The homotopy class of a relative morphism.
Equations
- f.homotopyClass = Quot.mk (fun (f g : SSet.RelativeMorphism A B φ) => Nonempty (f.Homotopy g)) f
Instances For
The composition of relative morphisms.
Instances For
The postcomposition of an homotopy with a relative morphism.
Equations
Instances For
The precomposition of an homotopy with a relative morphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The postcomposition of an homotopy class by a relative morphism.
Equations
- h.postcomp f' fac = Quot.lift (fun (f : SSet.RelativeMorphism A B φ) => (f.comp f' fac).homotopyClass) ⋯ h
Instances For
The precomposition of an homotopy class with a relative morphism.
Equations
- h.precomp f' fac = Quot.lift (fun (f : SSet.RelativeMorphism B C ψ) => (f'.comp f fac).homotopyClass) ⋯ h