module 1Lab.Path.IdentitySystem where

Identity systems🔗

An identity system is a way of characterising the path spaces of a particular type, without necessarily having to construct a full encode-decode equivalence. Essentially, the data of an identity system is precisely the data required to implement path induction, a.k.a. the J eliminator. Any type with the data of an identity system satisfies its own J, and conversely, if the type satisfies J, it is an identity system.

We unravel the definition of being an identity system into the following data, using a translation that takes advantage of cubical type theory’s native support for paths-over-paths:

record
  is-identity-system { ℓ'} {A : Type }
    (R : A  A  Type ℓ')
    (refl :  a  R a a)
    : Type (  ℓ')
  where
  no-eta-equality
  field
    to-path      :  {a b}  R a b  a  b
    to-path-over
      :  {a b} (p : R a b)
       PathP  i  R a (to-path p i)) (refl a) p

  is-contr-ΣR :  {a}  is-contr (Σ A (R a))
  is-contr-ΣR .centre    = _ , refl _
  is-contr-ΣR .paths x i = to-path (x .snd) i , to-path-over (x .snd) i

open is-identity-system public

As mentioned before, the data of an identity system gives us exactly what is required to prove J for the relation This is essentially the decomposition of J into contractibility of singletons, but with singletons replaced by

IdsJ
  :  { ℓ' ℓ''} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a} {a : A}
   is-identity-system R r
   (P :  b  R a b  Type ℓ'')
   P a (r a)
    {b} s  P b s
IdsJ ids P pr s =
  transport  i  P (ids .to-path s i) (ids .to-path-over s i)) pr
IdsJ-refl
  :  { ℓ' ℓ''} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a} {a : A}
   (ids : is-identity-system R r)
   (P :  b  R a b  Type ℓ'')
   (x : P a (r a))
   IdsJ ids P x (r a)  x
IdsJ-refl {R = R} {r = r} {a = a} ids P x =
  transport  i  P (ids .to-path (r a) i) (ids .to-path-over (r a) i)) x ≡⟨⟩
  subst P'  i  ids .to-path (r a) i , ids .to-path-over (r a) i) x      ≡⟨ ap  e  subst P' e x) lemma 
  subst P' refl x                                                          ≡⟨ transport-refl x 
  x 
  where
    P' : Σ _ (R a)  Type _
    P' (b , r) = P b r

    lemma : Σ-pathp (ids .to-path (r a)) (ids .to-path-over (r a))  refl
    lemma = is-contr→is-set (is-contr-ΣR ids) _ _ _ _

to-path-refl-coh
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a}
   (ids : is-identity-system R r)
    a
   (Σ-pathp (ids .to-path (r a)) (ids .to-path-over (r a)))  refl
to-path-refl-coh {r = r} ids a =
  is-contr→is-set (is-contr-ΣR ids) _ _
    (Σ-pathp (ids .to-path (r a)) (ids .to-path-over (r a)))
    refl

to-path-refl
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a} {a : A}
   (ids : is-identity-system R r)
   ids .to-path (r a)  refl
to-path-refl {r = r} {a = a} ids = ap (ap fst) $ to-path-refl-coh ids a

to-path-over-refl
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a} {a : A}
   (ids : is-identity-system R r)
   PathP  i  PathP  j  R a (to-path-refl {a = a} ids i j)) (r a) (r a))
      (ids .to-path-over (r a))
      refl
to-path-over-refl {a = a} ids = ap (ap snd) $ to-path-refl-coh ids a

Note that for any the type of identity system data on is a proposition. This is because it is exactly equivalent to the type being contractible for every which is a proposition by standard results.

contr→identity-system
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a}
   (∀ {a}  is-contr (Σ _ (R a)))
   is-identity-system R r
contr→identity-system {R = R} {r} c = ids where
  paths' :  {a} (p : Σ _ (R a))  (a , r a)  p
  paths' _ = is-contr→is-prop c _ _

  ids : is-identity-system R r
  ids .to-path p = ap fst (paths' (_ , p))
  ids .to-path-over p = ap snd (paths' (_ , p))

If we have a relation together with reflexivity witness then any equivalence equips with the structure of an identity system, by contractibility of singletons. Of course if we do not particularly care about the specific reflexivity witness, we can simply define as

equiv-path→identity-system
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a}
   (∀ {a b}  R a b  (a  b))
   is-identity-system R r
equiv-path→identity-system eqv = contr→identity-system $
  Equiv→is-hlevel 0 ((total  _  eqv .fst) , equiv→total (eqv .snd)))
    (contr _ Singleton-is-contr)

Conversely, any identity system implies an equivalence

identity-system-gives-path
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a}
   is-identity-system R r
    {a b}  R a b  (a  b)
identity-system-gives-path {R = R} {r = r} ids =
  Iso→Equiv (ids .to-path , iso from ri li) where
    from :  {a b}  a  b  R a b
    from {a = a} p = transport  i  R a (p i)) (r a)

    ri :  {a b}  is-right-inverse (from {a} {b}) (ids .to-path)
    ri = J  y p  ids .to-path (from p)  p)
           ( ap (ids .to-path) (transport-refl _)
            to-path-refl ids)

    li :  {a b}  is-left-inverse (from {a} {b}) (ids .to-path)
    li = IdsJ ids  y p  from (ids .to-path p)  p)
          ( ap from (to-path-refl ids)
           transport-refl _ )

In subtypes🔗

Let be an embedding. If is an identity system on then it can be pulled back along to an identity system on

module
  _ { ℓ' ℓ''} {A : Type } {B : Type ℓ'}
    {R : B  B  Type ℓ''} {r :  b  R b b}
    (ids : is-identity-system R r)
    (f : A  B)
  where

  pullback-identity-system
    : is-identity-system  x y  R (f .fst x) (f .fst y))  _  r _)
  pullback-identity-system .to-path {a} {b} x = ap fst $
    f .snd (f .fst b) (a , ids .to-path x) (b , refl)
  pullback-identity-system .to-path-over {a} {b} p i =
    comp
       j  R (f .fst a) (f .snd (f .fst b) (a , ids .to-path p) (b , refl) i .snd (~ j)))
      ( i) λ where
      k (k = i0)  ids .to-path-over p (~ k)
      k (i = i0)  ids .to-path-over p (~ k  i)
      k (i = i1)  p

This is actually part of an equivalence: if the equality identity system on (thus any identity system) can be pulled back along then is an embedding.

identity-system→embedding
  :  { ℓ'} {A : Type } {B : Type ℓ'}
   (f : A  B)
   is-identity-system  x y  f x  f y)  _  refl)
   is-embedding f
identity-system→embedding f ids = cancellable→embedding
  (identity-system-gives-path ids)

Univalence🔗

Note that univalence is precisely the statement that equivalences are an identity system on the universe:

univalence-identity-system
  :  {}  is-identity-system {A = Type } _≃_ λ _  id , id-equiv
univalence-identity-system .to-path = ua
univalence-identity-system .to-path-over p =
  Σ-prop-pathp  _  is-equiv-is-prop) $ funextP $ λ a  path→ua-pathp p refl
Path-identity-system
  :  {} {A : Type }  is-identity-system (Path A)  _  refl)
Path-identity-system .to-path p = p
Path-identity-system .to-path-over p i j = p (i  j)

is-identity-system-is-prop
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a}
   is-prop (is-identity-system R r)
is-identity-system-is-prop {A = A} {R} {r} =
  retract→is-hlevel 1 from to cancel λ x y i a  is-contr-is-prop (x a) (y a) i
  where
    to : is-identity-system R r   x  is-contr (Σ A (R x))
    to ids x = is-contr-ΣR ids

    from : (∀ x  is-contr (Σ A (R x)))  is-identity-system R r
    from x = contr→identity-system (x _)

    cancel'
      :  (x : is-identity-system R r) {a b} (s : R a b)
       PathP  i  (a , r a)  (b , s))
        (is-contr-ΣR (from (to x)) .paths (b , s))
        (is-contr-ΣR x .paths (b , s))
    cancel' x s = is-prop→squarep  _ _  is-contr→is-prop (is-contr-ΣR x)) _ _ _ _

    cancel : is-left-inverse from to
    cancel x i .to-path s = ap fst (cancel' x s i)
    cancel x i .to-path-over s = ap snd (cancel' x s i)

instance
  H-Level-is-identity-system
    :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  a  R a a} {n}
     H-Level (is-identity-system R r) (suc n)
  H-Level-is-identity-system = prop-instance is-identity-system-is-prop

identity-system→hlevel
  :  { ℓ'} {A : Type } n {R : A  A  Type ℓ'} {r :  x  R x x}
   is-identity-system R r
   (∀ x y  is-hlevel (R x y) n)
   is-hlevel A (suc n)
identity-system→hlevel zero ids hl x y = ids .to-path (hl _ _ .centre)
identity-system→hlevel (suc n) ids hl x y =
  Equiv→is-hlevel (suc n) (identity-system-gives-path ids e⁻¹) (hl x y)

Sets and Hedberg’s theorem🔗

We now apply the general theory of identity systems to something a lot more mundane: recognising sets. An immediate consequence of having an identity system on a type is that, if is pointwise an then is an Now, if is a reflexive family of propositions, then all we need for to be an identity system is that by the previous observation, this implies is a set.

set-identity-system
  :  { ℓ'} {A : Type } {R : A  A  Type ℓ'} {r :  x  R x x}
   (∀ x y  is-prop (R x y))
   (∀ {x y}  R x y  x  y)
   is-identity-system R r
set-identity-system rprop rpath .to-path = rpath
set-identity-system rprop rpath .to-path-over p =
  is-prop→pathp  i  rprop _ _) _ p

If is a type with ¬¬-stable equality, then by the theorem above, the pointwise double negation of its identity types is an identity system: and so, if a type has decidable (thus ¬¬-stable) equality, it is a set. This is known as Hedberg’s theorem.

¬¬-stable-identity-system
  :  {} {A : Type }
   (∀ {x y}  ¬ ¬ Path A x y  x  y)
   is-identity-system  x y  ¬ ¬ Path A x y) λ a k  k refl
¬¬-stable-identity-system = set-identity-system λ x y f g 
  funext λ h  absurd (g h)

opaque
  Discrete→is-set :  {} {A : Type }  Discrete A  is-set A
  Discrete→is-set {A = A} dec =
    identity-system→hlevel 1 (¬¬-stable-identity-system stable) λ x y f g 
      funext λ h  absurd (g h)
    where
      stable : {x y : A}  ¬ ¬ x  y  x  y
      stable {x = x} {y = y} ¬¬p with dec {x} {y}
      ... | yes p = p
      ... | no ¬p = absurd (¬¬p ¬p)