module Cat.Instances.Discrete whereDiscrete categoriesπ
Given a groupoid , we can build the category with space of objects and a single morphism whenever .
Disc : (A : Type β) β is-groupoid A β Precategory β β
Disc A A-grpd .Ob = A
Disc A A-grpd .Hom = _β‘_
Disc A A-grpd .Hom-set = A-grpd
Disc A A-grpd .id = refl
Disc A A-grpd ._β_ p q = q β p
Disc A A-grpd .idr _ = β-idl _
Disc A A-grpd .idl _ = β-idr _
Disc A A-grpd .assoc _ _ _ = sym (β-assoc _ _ _)
Disc' : Set β β Precategory β β
Disc' A = Disc β£ A β£ h where abstract
  h : is-groupoid β£ A β£
  h = is-hlevel-suc 2 (A .is-tr)Clearly this is a univalent category:
Disc-is-category : β {A : Type β} {A-grpd} β is-category (Disc A A-grpd)
Disc-is-category .to-path is = is .to
Disc-is-category .to-path-over is = β
-pathp _ _ _ Ξ» i j β is .to (i β§ j)We can lift any function between the underlying types to a functor between discrete categories. This is because every function automatically respects equality in a functorial way.
lift-disc
  : β {A : Set β} {B : Set β'}
  β (β£ A β£ β β£ B β£)
  β Functor (Disc' A) (Disc' B)
lift-disc f .Fβ = f
lift-disc f .Fβ = ap f
lift-disc f .F-id = refl
lift-disc f .F-β p q = ap-β f q pDiagrams in Disc(X)π
If
is a discrete type, then it is in particular in
Set, and we
can make diagrams of shape
in some category
,
using the decidable equality on
.
We note that the decidable equality is redundant information:
The construction Disc' above
extends into a left
adjoint to the Ob functor.
Disc-diagram
  : β {X : Set β} β¦ _ : Discrete β£ X β£ β¦
  β (β£ X β£ β Ob C)
  β Functor (Disc' X) C
Disc-diagram {C = C} {X = X} β¦ d β¦ f = F where
  module C = Precategory C
  P : β£ X β£ β β£ X β£ β Type _
  P x y = C.Hom (f x) (f y)
  go : β {x y : β£ X β£} β x β‘ y β Dec (x β‘α΅’ y) β P x y
  go {x} {.x} p (yes reflα΅’) = C.id
  go {x} {y}  p (no Β¬p)     = absurd (Β¬p (Idβpath.from p))The object part of the functor is the provided , and the decidable equality is used to prove that respects equality. This is why itβs redundant: automatically respects equality, because itβs a function! However, by using the decision procedure, we get better computational behaviour: Very often, will be , and substitution along is easy to deal with.
  F : Functor _ _
  F .Fβ = f
  F .Fβ {x} {y} p = go p (x β‘α΅’? y)Proving that our our is functorial involves a bunch of tedious computations with equalities and a whole waterfall of absurd cases:
  F .F-id {x} = refl
  F .F-β  {x} {y} {z} f g =
    J (Ξ» y g β β {z} (f : y β‘ z) β go (g β f) (x β‘α΅’? z) β‘ go f (y β‘α΅’? z) C.β go g (x β‘α΅’? y))
      (Ξ» f β J (Ξ» z f β go (refl β f) (x β‘α΅’? z) β‘ go f (x β‘α΅’? z) C.β C.id) (sym (C.idr _)) f)
      g f