Documentation

Mathlib.AlgebraicTopology.SimplicialObject.II

A construction by Gabriel and Zisman #

In this file, we construct a cosimplicial object SimplexCategory.II in SimplexCategoryᵒᵖ, i.e. a functor SimplexCategory ⥤ SimplexCategoryᵒᵖ. If we identify SimplexCategory with the category of finite nonempty linearly ordered types, this functor could be interpreted as the contravariant functor which sends a finite nonempty linearly ordered type T to T →o Fin 2 (with f ≤ g ↔ ∀ i, g i ≤ f i, which turns out to be a linear order); in particular, it sends Fin (n + 1) to a linearly ordered type which is isomorphic to Fin (n + 2). As a result, we define SimplexCategory.II as a functor which sends ⦋n⦌ to ⦋n + 1⦌: on morphisms, it sends faces to degeneracies and vice versa. This construction appeared in Calculus of fractions and homotopy theory, chapter III, paragraph 1.1, by Gabriel and Zisman.

References #

def SimplexCategory.II.finset {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) :
Finset (Fin (n + 2))

Auxiliary definition for map'. Given f : Fin (n + 1) →o Fin (m + 1) and x : Fin (m + 2), map' f x shall be the smallest element in this finset f x : Finset (Fin (n + 2)).

Equations
Instances For
    theorem SimplexCategory.II.mem_finset_iff {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) (i : Fin (n + 2)) :
    i finset f x i = Fin.last (n + 1) ∃ (h : i Fin.last (n + 1)), x (f (i.castPred h)).castSucc
    @[simp]
    theorem SimplexCategory.II.last_mem_finset {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) :
    Fin.last (n + 1) finset f x
    @[simp]
    theorem SimplexCategory.II.castSucc_mem_finset_iff {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) (i : Fin (n + 1)) :
    theorem SimplexCategory.II.nonempty_finset {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) :
    def SimplexCategory.II.map' {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) :
    Fin (n + 2)

    Auxiliary definition for the definition of the action of the functor SimplexCategory.II on morphisms.

    Equations
    Instances For
      theorem SimplexCategory.II.map'_eq_last_iff {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) :
      map' f x = Fin.last (n + 1) ∀ (i : Fin (n + 1)), (f i).castSucc < x
      theorem SimplexCategory.II.map'_eq_castSucc_iff {n m : } (f : Fin (n + 1) →o Fin (m + 1)) (x : Fin (m + 2)) (y : Fin (n + 1)) :
      map' f x = y.castSucc x (f y).castSucc i < y, (f i).castSucc < x
      @[simp]
      theorem SimplexCategory.II.map'_last {n m : } (f : Fin (n + 1) →o Fin (m + 1)) :
      map' f (Fin.last (m + 1)) = Fin.last (n + 1)
      @[simp]
      theorem SimplexCategory.II.map'_zero {n m : } (f : Fin (n + 1) →o Fin (m + 1)) :
      map' f 0 = 0
      @[simp]
      theorem SimplexCategory.II.map'_id {n : } (x : Fin (n + 2)) :
      theorem SimplexCategory.II.map'_map' {n m p : } (f : Fin (n + 1) →o Fin (m + 1)) (g : Fin (m + 1) →o Fin (p + 1)) (x : Fin (p + 2)) :
      map' f (map' g x) = map' (g.comp f) x
      @[simp]
      theorem SimplexCategory.II.map'_predAbove {n : } (i : Fin (n + 1)) (x : Fin (n + 2)) :
      map' { toFun := i.predAbove, monotone' := } x = i.succ.castSucc.succAbove x
      theorem SimplexCategory.II.monotone_map' {n m : } (f : Fin (n + 1) →o Fin (m + 1)) :

      The functor SimplexCategory ⥤ SimplexCategoryᵒᵖ (i.e. a cosimplicial object in SimplexCategoryᵒᵖ) which sends ⦋n⦌ to the object in SimplexCategoryᵒᵖ that is associated to the linearly ordered type ⦋n + 1⦌ (which could be identified to the ordered type ⦋n⦌ →o ⦋1⦌).

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem SimplexCategory.II_δ {n : } (i : Fin (n + 2)) :
        II.δ i = (σ i).op
        @[simp]
        theorem SimplexCategory.II_σ {n : } (i : Fin (n + 1)) :