Edges and "triangles" in simplicial sets #
Given a simplicial set X, we introduce two types:
- Given
0-simplicesx₀andx₁, we defineEdge x₀ x₁which is the type of1-simplices with facesx₁andx₀respectively; - Given
0-simplicesx₀,x₁,x₂, edgese₀₁ : Edge x₀ x₁,e₁₂ : Edge x₁ x₂,e₀₂ : Edge x₀ x₂, a structureCompStruct e₀₁ e₁₂ e₀₂which records the data of a2-simplex with facese₁₂,e₀₂ande₀₁respectively. This data will allow to obtain relations in the homotopy category ofX. (This API parallels similar definitions for2-truncated simplicial sets. The definitions in this file are definitionally equal to their2-truncated counterparts.)
In a simplicial set, an edge from a vertex x₀ to x₁ is
a 1-simplex with prescribed 0-dimensional faces.
Equations
- SSet.Edge x₀ x₁ = SSet.Truncated.Edge x₀ x₁
Instances For
Constructor for SSet.Edge which takes as an input a term in the definitionally
equal type SSet.Truncated.Edge for the 2-truncation of the simplicial set.
(This definition is made to contain abuse of defeq in other definitions.)
Equations
Instances For
The edge of the 2-truncation of a simplicial set X that is induced
by an edge of X.
Equations
- e.toTruncated = e
Instances For
In a simplicial set, an edge from a vertex x₀ to x₁ is
a 1-simplex with prescribed 0-dimensional faces.
Equations
- e.edge = e.toTruncated.edge
Instances For
Constructor for edges in a simplicial set.
Equations
- SSet.Edge.mk edge src_eq tgt_eq = SSet.Edge.ofTruncated { edge := edge, src_eq := ⋯, tgt_eq := ⋯ }
Instances For
The constant edge on a 0-simplex.
Equations
Instances For
The image of an edge by a morphism of simplicial sets.
Equations
- e.map f = SSet.Edge.ofTruncated (e.toTruncated.map ((SSet.truncation 2).map f))
Instances For
The edge given by a 1-simplex.
Equations
- SSet.Edge.mk' s = SSet.Edge.mk s ⋯ ⋯
Instances For
Let x₀, x₁, x₂ be 0-simplices of a simplicial set X,
e₀₁ an edge from x₀ to x₁, e₁₂ an edge from x₁ to x₂,
e₀₂ an edge from x₀ to x₂. This is the data of a 2-simplex whose
faces are respectively e₀₂, e₁₂ and e₀₁. Such structures shall provide
relations in the homotopy category of arbitrary simplicial sets
(and specialized constructions for quasicategories and Kan complexes.).
Equations
- e₀₁.CompStruct e₁₂ e₀₂ = e₀₁.toTruncated.CompStruct e₁₂.toTruncated e₀₂.toTruncated
Instances For
Constructor for SSet.Edge.CompStruct which takes as an input a term in the
definitionally equal type SSet.Truncated.Edge.CompStruct for the 2-truncation of
the simplicial set. (This definition is made to contain abuse of defeq in
other definitions.)
Equations
Instances For
Conversion from SSet.Edge.CompStruct to SSet.Truncated.Edge.CompStruct.
Equations
- h.toTruncated = h
Instances For
The underlying 2-simplex in a structure SSet.Edge.CompStruct.
Equations
- h.simplex = h.toTruncated.simplex
Instances For
Constructor for SSet.Edge.CompStruct.
Equations
- SSet.Edge.CompStruct.mk simplex d₂ d₀ d₁ = { simplex := simplex, d₂ := ⋯, d₀ := ⋯, d₁ := ⋯ }
Instances For
e : Edge x₀ x₁ is a composition of Edge.id x₀ with e.
Equations
Instances For
e : Edge x₀ x₁ is a composition of e with Edge.id x₁
Equations
Instances For
The image of a Edge.CompStruct by a morphism of simplicial sets.
Equations
- h.map f = SSet.Edge.CompStruct.ofTruncated (h.toTruncated.map ((SSet.truncation 2).map f))