Documentation

Mathlib.AlgebraicTopology.SimplicialSet.Simplices

The preordered type of simplices of a simplicial set #

In this file, we define the type X.S of simplices of a simplicial set X, where a simplex consists of the data of dim : ℕ and simplex : X _⦋dim⦌. We endow this type with a preorder defined by x ≤ y ↔ Subcomplex.ofSimplex x.simplex ≤ Subcomplex.ofSimplex y.simplex. In particular, as a preordered type, X.S is a category, but this is not what is called "the category of simplices of X" in the literature (and which is X.Elementsᵒᵖ in mathlib).

TODO (@joelriou) #

structure SSet.S (X : SSet) :

The type of simplices of a simplicial set X. This type X.S is in bijection with X.Elements (see SSet.S.equivElements), but X.S is not what the literature names "category of simplices of X", as the category on X.S comes from a preorder (see S.le_iff_nonempty_hom).

Instances For
    theorem SSet.S.mk_surjective {X : SSet} (s : X.S) :
    ∃ (n : ) (x : X.obj (Opposite.op (SimplexCategory.mk n))), s = { dim := n, simplex := x }
    def SSet.S.map {X Y : SSet} (f : X Y) (s : X.S) :
    Y.S

    The image of a simplex by a morphism of simplicial sets.

    Equations
    Instances For
      theorem SSet.S.dim_eq_of_eq {X : SSet} {s t : X.S} (h : s = t) :
      s.dim = t.dim
      theorem SSet.S.dim_eq_of_mk_eq {X : SSet} {n m : } {x : X.obj (Opposite.op (SimplexCategory.mk n))} {y : X.obj (Opposite.op (SimplexCategory.mk m))} (h : { dim := n, simplex := x } = { dim := m, simplex := y }) :
      n = m
      def SSet.S.cast {X : SSet} (s : X.S) {d : } (hd : s.dim = d) :
      X.S

      When s : X.S is such that s.dim = d, this is a term that is equal to s, but whose dimension if definitionally equal to d.

      Equations
      Instances For
        @[simp]
        theorem SSet.S.cast_dim {X : SSet} (s : X.S) {d : } (hd : s.dim = d) :
        (s.cast hd).dim = d
        theorem SSet.S.cast_eq_self {X : SSet} (s : X.S) {d : } (hd : s.dim = d) :
        s.cast hd = s
        @[simp]
        theorem SSet.S.cast_simplex_rfl {X : SSet} (s : X.S) :
        (s.cast ).simplex = s.simplex
        theorem SSet.S.ext_iff' {X : SSet} (s t : X.S) :
        s = t ∃ (h : s.dim = t.dim), (s.cast h).simplex = t.simplex
        theorem SSet.S.ext_iff {X : SSet} {n : } (x y : X.obj (Opposite.op (SimplexCategory.mk n))) :
        { dim := n, simplex := x } = { dim := n, simplex := y } x = y
        @[reducible, inline]
        abbrev SSet.S.subcomplex {X : SSet} (s : X.S) :

        The subcomplex generated by a simplex.

        Equations
        Instances For

          If s : X.S and t : X.S are simplices of a simplicial set, s ≤ t means that the subcomplex generated by s is contained in the subcomplex generated by t, see SSet.S.le_def and SSet.S.le_iff. Note that the category structure on X.S induced by this preorder is not the "category of simplices" of X (which is see X.Elementsᵒᵖ); see SSet.S.le_iff_nonempty_hom for the precise relation.

          Equations
          theorem SSet.S.le_def {X : SSet} {s t : X.S} :
          theorem SSet.S.le_iff {X : SSet} {s t : X.S} :
          theorem SSet.S.mk_map_le {X : SSet} {n m : } (x : X.obj (Opposite.op (SimplexCategory.mk n))) (f : SimplexCategory.mk m SimplexCategory.mk n) :
          { dim := m, simplex := X.map f.op x } { dim := n, simplex := x }
          theorem SSet.S.mk_map_eq_iff_of_mono {X : SSet} {n m : } (x : X.obj (Opposite.op (SimplexCategory.mk n))) (f : SimplexCategory.mk m SimplexCategory.mk n) [CategoryTheory.Mono f] :
          { dim := m, simplex := X.map f.op x } = { dim := n, simplex := x } CategoryTheory.IsIso f

          The type of simplices of X : SSet.{u} identifies to the type of elements of X considered as a functor SimplexCategoryᵒᵖ ⥤ Type u. (Note that this is not an (anti)equivalence of categories, see S.le_iff_nonempty_hom.)

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For