The category Discrete PUnit
#
We define star : C ⥤ Discrete PUnit
sending everything to PUnit.star
,
show that any two functors to Discrete PUnit
are naturally isomorphic,
and construct the equivalence (Discrete PUnit ⥤ C) ≌ C
.
The constant functor sending everything to PUnit.star
.
Equations
- CategoryTheory.Functor.star C = (CategoryTheory.Functor.const C).obj { as := PUnit.unit }
Instances For
@[simp]
def
CategoryTheory.Functor.punitExt
{C : Type u}
[Category.{v, u} C]
(F G : Functor C (Discrete PUnit.{w + 1}))
:
Any two functors to Discrete PUnit
are isomorphic.
Equations
- F.punitExt G = CategoryTheory.NatIso.ofComponents (fun (X : C) => CategoryTheory.eqToIso ⋯) ⋯
Instances For
theorem
CategoryTheory.Functor.punit_ext'
{C : Type u}
[Category.{v, u} C]
(F G : Functor C (Discrete PUnit.{w + 1}))
:
Any two functors to Discrete PUnit
are equal.
You probably want to use punitExt
instead of this.
@[reducible, inline]
The functor from Discrete PUnit
sending everything to the given object.
Equations
Instances For
Functors from Discrete PUnit
are equivalent to the category itself.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
CategoryTheory.Functor.equiv_functor_obj
{C : Type u}
[Category.{v, u} C]
(F : Functor (Discrete PUnit.{w + 1}) C)
:
@[simp]
@[simp]
theorem
CategoryTheory.Functor.equiv_functor_map
{C : Type u}
[Category.{v, u} C]
{X✝ Y✝ : Functor (Discrete PUnit.{w + 1}) C}
(θ : X✝ ⟶ Y✝)
:
@[simp]
@[simp]
theorem
CategoryTheory.Functor.equiv_unitIso
{C : Type u}
[Category.{v, u} C]
:
equiv.unitIso = NatIso.ofComponents
(fun (x : Functor (Discrete PUnit.{w + 1}) C) =>
Discrete.natIso fun (x_1 : Discrete PUnit.{w + 1}) =>
Iso.refl (((Functor.id (Functor (Discrete PUnit.{w + 1}) C)).obj x).obj x_1))
⋯
A category being equivalent to PUnit
is equivalent to it having a unique morphism between
any two objects. (In fact, such a category is also a groupoid;
see CategoryTheory.Groupoid.ofHomUnique
)