Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat(AlgebraicTopology/SimplexCategory/SimplicialObject): definitions of simplicial objects by generators and relation #21748

Open
wants to merge 38 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
38 commits
Select commit Hold shift + click to select a range
53a6f43
Initial commit
robin-carlier Feb 11, 2025
b5404f9
lake shake
robin-carlier Feb 11, 2025
de99346
Apply suggestions from code review
robin-carlier Feb 12, 2025
8307c3e
fix build
robin-carlier Feb 12, 2025
f36825f
change namespace
robin-carlier Feb 12, 2025
89e5264
mirror names
robin-carlier Feb 12, 2025
0a75491
better nat lemmas
robin-carlier Feb 12, 2025
a8ba882
Initial
robin-carlier Feb 11, 2025
be0bd7a
adapt to changes in #21741
robin-carlier Feb 12, 2025
789297a
rebase on #21742
robin-carlier Feb 12, 2025
75a9378
Initial
robin-carlier Feb 11, 2025
58f12ca
lake shake fix
robin-carlier Feb 11, 2025
25db868
fix build
robin-carlier Feb 12, 2025
92cc4ea
Normal forms for compositions of degeneracies
robin-carlier Feb 11, 2025
80f0501
whitespaces
robin-carlier Feb 12, 2025
c78261e
Initial
robin-carlier Feb 11, 2025
c08f0c4
Initial
robin-carlier Feb 11, 2025
ad3d1f4
fix build
robin-carlier Feb 12, 2025
a42c226
Initial
robin-carlier Feb 11, 2025
530752f
fix build
robin-carlier Feb 12, 2025
99c76f2
change namespace
robin-carlier Feb 13, 2025
92ac6ea
Initial
robin-carlier Feb 11, 2025
1e84b00
adapt to changes in #21741
robin-carlier Feb 12, 2025
7679d60
adapt namespace change
robin-carlier Feb 13, 2025
ee6df34
rebase on #21742
robin-carlier Feb 12, 2025
3b3e305
Initial
robin-carlier Feb 11, 2025
4721ff8
lake shake fix
robin-carlier Feb 11, 2025
89d037b
fix build
robin-carlier Feb 12, 2025
12bbc18
adapt namespace
robin-carlier Feb 13, 2025
0e442a8
Normal forms for compositions of degeneracies
robin-carlier Feb 11, 2025
fcb5478
Normal forms for compositions of degeneracies
robin-carlier Feb 11, 2025
f74e79a
Initial
robin-carlier Feb 11, 2025
96a23a8
Normal forms for compositions of degeneracies
robin-carlier Feb 11, 2025
9f8f1f1
Initial
robin-carlier Feb 11, 2025
2d0f4ed
fix build
robin-carlier Feb 12, 2025
9be0291
adapt namespace
robin-carlier Feb 13, 2025
592477f
Merge branch 'RC_SimplexCategoryGenRel7' into RC_SimplexCategoryGenRel8
robin-carlier Feb 13, 2025
e9d9f11
change namespaces
robin-carlier Feb 13, 2025
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
5 changes: 5 additions & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1128,11 +1128,16 @@ import Mathlib.AlgebraicTopology.Quasicategory.Basic
import Mathlib.AlgebraicTopology.Quasicategory.Nerve
import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal
import Mathlib.AlgebraicTopology.SimplexCategory
import Mathlib.AlgebraicTopology.SimplexCategory.GeneratorsRelations.Basic
import Mathlib.AlgebraicTopology.SimplexCategory.GeneratorsRelations.EpiMono
import Mathlib.AlgebraicTopology.SimplexCategory.GeneratorsRelations.Equivalence
import Mathlib.AlgebraicTopology.SimplexCategory.GeneratorsRelations.NormalForms
import Mathlib.AlgebraicTopology.SimplicialCategory.Basic
import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject
import Mathlib.AlgebraicTopology.SimplicialNerve
import Mathlib.AlgebraicTopology.SimplicialObject.Basic
import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal
import Mathlib.AlgebraicTopology.SimplicialObject.GeneratorsRelations
import Mathlib.AlgebraicTopology.SimplicialObject.Split
import Mathlib.AlgebraicTopology.SimplicialSet.Basic
import Mathlib.AlgebraicTopology.SimplicialSet.Boundary
Expand Down
Original file line number Diff line number Diff line change
@@ -0,0 +1,253 @@
/-
Copyright (c) 2025 Robin Carlier. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Robin Carlier
-/
import Mathlib.AlgebraicTopology.SimplexCategory
import Mathlib.CategoryTheory.PathCategory.Basic
/-! # Presentation of the simplex category by generator and relations.

We introduce `SimplexCategoryGenRel` as the category presented by generating
morphisms `δ i : [n] ⟶ [n + 1]` and `σ i : [n + 1] ⟶ [n]` and subject to the
simplicial identities, and we provide induction principles for reasoning about
objects and morphisms in this category.

This category admits a canonical functor `toSimplexCategory` to the usual simplex category.
The fact that this functor is an equivalence will be recorded in a separate file.
-/
open CategoryTheory

/-- The objects of the free simplex quiver are the natural numbers. -/
def FreeSimplexQuiver := ℕ

/-- Making an object of `FreeSimplexQuiver` out of a natural number. -/
def FreeSimplexQuiver.mk (n : ℕ) : FreeSimplexQuiver := n

/-- Getting back the natural number from the objects. -/
def FreeSimplexQuiver.len (x : FreeSimplexQuiver) : ℕ := x

namespace FreeSimplexQuiver

/-- A morphism in `FreeSimplexQuiver` is either a face map (`δ`) or a degeneracy map (`σ`). -/
inductive Hom : FreeSimplexQuiver → FreeSimplexQuiver → Type
| δ {n : ℕ} (i : Fin (n + 2)) : Hom (.mk n) (.mk (n + 1))
| σ {n : ℕ} (i : Fin (n + 1)) : Hom (.mk (n + 1)) (.mk n)

instance Quiv : Quiver FreeSimplexQuiver where
Hom := FreeSimplexQuiver.Hom

/-- `FreeSimplexQuiver.δ i` represents the `i`-th face map `.mk n ⟶ .mk (n + 1)`. -/
abbrev δ {n : ℕ} (i : Fin (n + 2)) : FreeSimplexQuiver.mk n ⟶ .mk (n + 1) :=
FreeSimplexQuiver.Hom.δ i

/-- `FreeSimplexQuiver.σ i` represents `i`-th degeneracy map `.mk (n + 1) ⟶ .mk n`. -/
abbrev σ {n : ℕ} (i : Fin (n + 1)) : FreeSimplexQuiver.mk (n + 1) ⟶ .mk n :=
FreeSimplexQuiver.Hom.σ i

/-- `FreeSimplexQuiver.homRel` is the relation on morphisms freely generated on the
five simplicial identities. -/
inductive homRel : HomRel (Paths FreeSimplexQuiver)
| δ_comp_δ {n : ℕ} {i j : Fin (n + 2)} (H : i ≤ j) : homRel
(Paths.of.map (δ i) ≫ Paths.of.map (δ j.succ))
(Paths.of.map (δ j) ≫ Paths.of.map (δ i.castSucc))
| δ_comp_σ_of_le {n : ℕ} {i : Fin (n + 2)} {j : Fin (n + 1)} (H : i ≤ j.castSucc) : homRel
(Paths.of.map (δ i.castSucc) ≫ Paths.of.map (σ j.succ))
(Paths.of.map (σ j) ≫ Paths.of.map (δ i))
| δ_comp_σ_self {n : ℕ} {i : Fin (n + 1)} : homRel
(Paths.of.map (δ i.castSucc) ≫ Paths.of.map (σ i)) (𝟙 _)
| δ_comp_σ_succ {n : ℕ} {i : Fin (n + 1)} : homRel
(Paths.of.map (δ i.succ) ≫ Paths.of.map (σ i)) (𝟙 _)
| δ_comp_σ_of_gt {n : ℕ} {i : Fin (n + 2)} {j : Fin (n + 1)} (H : j.castSucc < i) : homRel
(Paths.of.map (δ i.succ) ≫ Paths.of.map (σ j.castSucc))
(Paths.of.map (σ j) ≫ Paths.of.map (δ i))
| σ_comp_σ {n : ℕ} {i j : Fin (n + 1)} (H : i ≤ j) : homRel
(Paths.of.map (σ i.castSucc) ≫ Paths.of.map (σ j))
(Paths.of.map (σ j.succ) ≫ Paths.of.map (σ i))

end FreeSimplexQuiver

/-- SimplexCategory is the category presented by generators and relation by the simplicial
identities.-/
def SimplexCategoryGenRel := Quotient FreeSimplexQuiver.homRel
deriving Category

/-- `SimplexCategoryGenRel.mk` is the main constructor for objects of `SimplexCategoryGenRel`. -/
def SimplexCategoryGenRel.mk (n : ℕ) : SimplexCategoryGenRel where
as := Paths.of.obj n

namespace SimplexCategoryGenRel

/-- `SimplexCategoryGenRel.δ i` is the `i`-th face map `.mk n ⟶ .mk (n + 1)`. -/
abbrev δ {n : ℕ} (i : Fin (n + 2)) : (SimplexCategoryGenRel.mk n) ⟶ .mk (n + 1) :=
(Quotient.functor FreeSimplexQuiver.homRel).map <| Paths.of.map (.δ i)

/-- `SimplexCategoryGenRel.σ i` is the `i`-th degeneracy map `.mk (n + 1) ⟶ .mk n`. -/
abbrev σ {n : ℕ} (i : Fin (n + 1)) :
(SimplexCategoryGenRel.mk (n + 1)) ⟶ (SimplexCategoryGenRel.mk n) :=
(Quotient.functor FreeSimplexQuiver.homRel).map <| Paths.of.map (.σ i)

/-- The length of an object of `SimplexCategoryGenRel`. -/
def len (x : SimplexCategoryGenRel) : ℕ := by rcases x with ⟨n⟩; exact n

@[simp]
lemma mk_len (n : ℕ) : (len (mk n)) = n := rfl

section InductionPrinciples

/-- An induction principle for reasonning about morphisms properties in SimplexCategoryGenRel. -/
lemma hom_induction (P : MorphismProperty SimplexCategoryGenRel)
(hi : ∀ {n : ℕ}, P (𝟙 (.mk n)))
(hc₁ : ∀ {n m : ℕ} (u : .mk n ⟶ .mk m) (i : Fin (m + 2)), P u → P (u ≫ δ i))
(hc₂ : ∀ {n m : ℕ} (u : .mk n ⟶ .mk (m + 1)) (i : Fin (m + 1)), P u → P (u ≫ σ i))
{a b : SimplexCategoryGenRel} (f : a ⟶ b) :
P f := by
apply CategoryTheory.Quotient.induction (P := (fun f ↦ P f))
apply Paths.induction
· exact hi
· rintro _ _ _ _ ⟨⟩
· exact hc₁ _ _
· exact hc₂ _ _

/-- An induction principle for reasonning about morphisms in SimplexCategoryGenRel, where we compose
with generators on the right. -/
lemma hom_induction' (P : MorphismProperty SimplexCategoryGenRel)
(hi : ∀ {n : ℕ}, P (𝟙 (SimplexCategoryGenRel.mk n)))
(hc₁ : ∀ {n m : ℕ} (u : SimplexCategoryGenRel.mk (m + 1) ⟶ SimplexCategoryGenRel.mk n)
(i : Fin (m + 2)), P u → P (δ i ≫ u))
(hc₂ : ∀ {n m : ℕ} (u : SimplexCategoryGenRel.mk m ⟶ SimplexCategoryGenRel.mk n)
(i : Fin (m + 1)), P u → P (σ i ≫ u )) {a b : SimplexCategoryGenRel} (f : a ⟶ b) :
P f := by
apply CategoryTheory.Quotient.induction (P := (fun f ↦ P f))
apply Paths.induction'
· exact hi
· rintro _ _ _ ⟨⟩ _
· exact hc₁ _ _
· exact hc₂ _ _

/-- An induction principle for reasonning about morphisms properties in SimplexCategoryGenRel. -/
lemma morphismProperty_eq_top (P : MorphismProperty SimplexCategoryGenRel)
(hi : ∀ {n : ℕ}, P (𝟙 (.mk n)))
(hc₁ : ∀ {n m : ℕ} (u : .mk n ⟶ .mk m) (i : Fin (m + 2)),
P u → P (u ≫ δ i))
(hc₂ : ∀ {n m : ℕ} (u : .mk n ⟶ .mk (m + 1))
(i : Fin (m + 1)), P u → P (u ≫ σ i)) :
P = ⊤ := by
ext; constructor
· simp
· intro _
apply hom_induction <;> assumption

/-- An induction principle for reasonning about morphisms properties in SimplexCategoryGenRel,
where we compose with generators on the right. -/
lemma hom_induction_eq_top' (P : MorphismProperty SimplexCategoryGenRel)
(hi : ∀ {n : ℕ}, P (𝟙 (.mk n)))
(hc₁ : ∀ {n m : ℕ} (u : .mk (m + 1) ⟶ .mk n) (i : Fin (m + 2)), P u → (P (δ i ≫ u)))
(hc₂ : ∀ {n m : ℕ} (u : .mk m ⟶ .mk n) (i : Fin (m + 1)), P u → (P (σ i ≫ u ))) :
P = ⊤ := by
ext; constructor
· simp
· intro _
apply hom_induction' <;> assumption

/-- An induction principle for reasonning about objects in SimplexCategoryGenRel. This should be
used instead of identifying an object with `mk` of its len.-/
@[elab_as_elim, cases_eliminator, induction_eliminator]
protected def rec {P : SimplexCategoryGenRel → Sort*}
(H : ∀ n : ℕ, P (.mk n)) :
∀ x : SimplexCategoryGenRel, P x := by
intro x
exact H x.len

/-- A basic ext lemma for objects of SimplexCategoryGenRel --/
@[ext]
lemma ext {x y : SimplexCategoryGenRel} (h : x.len = y.len) : x = y := by
cases x
cases y
simp only [mk_len] at h
congr

end InductionPrinciples

section SimplicialIdentities

@[reassoc]
theorem δ_comp_δ {n} {i j : Fin (n + 2)} (H : i ≤ j) :
δ i ≫ δ j.succ = δ j ≫ δ i.castSucc := by
apply CategoryTheory.Quotient.sound
exact FreeSimplexQuiver.homRel.δ_comp_δ H

@[reassoc]
theorem δ_comp_σ_of_le {n} {i : Fin (n + 2)} {j : Fin (n + 1)} (H : i ≤ j.castSucc) :
δ i.castSucc ≫ σ j.succ = σ j ≫ δ i := by
apply CategoryTheory.Quotient.sound
exact FreeSimplexQuiver.homRel.δ_comp_σ_of_le H

@[reassoc]
theorem δ_comp_σ_self {n} {i : Fin (n + 1)} :
δ i.castSucc ≫ σ i = 𝟙 (mk n) := by
apply CategoryTheory.Quotient.sound
exact FreeSimplexQuiver.homRel.δ_comp_σ_self

@[reassoc]
theorem δ_comp_σ_succ {n} {i : Fin (n + 1)} : δ i.succ ≫ σ i = 𝟙 (mk n) := by
apply CategoryTheory.Quotient.sound
exact FreeSimplexQuiver.homRel.δ_comp_σ_succ

@[reassoc]
theorem δ_comp_σ_of_gt {n} {i : Fin (n + 2)} {j : Fin (n + 1)} (H : j.castSucc < i) :
δ i.succ ≫ σ j.castSucc = σ j ≫ δ i := by
apply CategoryTheory.Quotient.sound
exact FreeSimplexQuiver.homRel.δ_comp_σ_of_gt H

@[reassoc]
theorem σ_comp_σ {n} {i j : Fin (n + 1)} (H : i ≤ j) :
σ i.castSucc ≫ σ j = σ j.succ ≫ σ i := by
apply CategoryTheory.Quotient.sound
exact FreeSimplexQuiver.homRel.σ_comp_σ H

/-- A version of δ_comp_δ with indices in ℕ satisfying relevant inequalities. -/
lemma δ_comp_δ_nat {n} (i j : ℕ) (hi : i < n + 2) (hj : j < n + 2) (H : i ≤ j) :
δ ⟨i, hi⟩ ≫ δ ⟨j + 1, by omega⟩ = δ ⟨j, hj⟩ ≫ δ ⟨i, by omega⟩ :=
δ_comp_δ (n := n) (i := ⟨i, by omega⟩) (j := ⟨j, by omega⟩) (by simpa)

/-- A version of σ_comp_σ with indices in ℕ satisfying relevant inequalities. -/
lemma σ_comp_σ_nat {n} (i j : ℕ) (hi : i < n + 1) (hj : j < n + 1) (H : i ≤ j) :
σ ⟨i, by omega⟩ ≫ σ ⟨j, hj⟩ = σ ⟨j + 1, by omega⟩ ≫ σ ⟨i, hi⟩ :=
σ_comp_σ (n := n) (i := ⟨i, by omega⟩) (j := ⟨j, by omega⟩) (by simpa)

end SimplicialIdentities

/-- The canonical functor from `SimplexCategoryGenRel` to SimplexCategory, which exists as the
simplicial identities hold in `SimplexCategory`. -/
def toSimplexCategory : SimplexCategoryGenRel ⥤ SimplexCategory :=
CategoryTheory.Quotient.lift _
(Paths.lift
{ obj := .mk
map f := match f with
| FreeSimplexQuiver.Hom.δ i => SimplexCategory.δ i
| FreeSimplexQuiver.Hom.σ i => SimplexCategory.σ i })
(fun _ _ _ _ h ↦ by
cases h with
| δ_comp_δ H => exact SimplexCategory.δ_comp_δ H
| δ_comp_σ_of_le H => exact SimplexCategory.δ_comp_σ_of_le H
| δ_comp_σ_self => exact SimplexCategory.δ_comp_σ_self
| δ_comp_σ_succ => exact SimplexCategory.δ_comp_σ_succ
| δ_comp_σ_of_gt H => exact SimplexCategory.δ_comp_σ_of_gt H
| σ_comp_σ H => exact SimplexCategory.σ_comp_σ H)

@[simp]
lemma toSimplexCategory_obj_mk (n : ℕ) : toSimplexCategory.obj (mk n) = .mk n := rfl

@[simp]
lemma toSimplexCategory_map_δ {n : ℕ} (i : Fin (n + 2)) : toSimplexCategory.map (δ i) =
SimplexCategory.δ i := rfl

@[simp]
lemma toSimplexCategory_map_σ {n : ℕ} (i : Fin (n + 1)) : toSimplexCategory.map (σ i) =
SimplexCategory.σ i := rfl

@[simp]
lemma toSimplexCategory_len {x : SimplexCategoryGenRel} : (toSimplexCategory.obj x).len = x.len :=
rfl

end SimplexCategoryGenRel
Loading