module 1Lab.HIT.Truncation where

Propositional truncationπŸ”—

Let AA be a type. The propositional truncation of AA is a type which represents the proposition β€œA is inhabited”. In MLTT, propositional truncations can not be constructed without postulates, even in the presence of impredicative prop. However, Cubical Agda provides a tool to define them: higher inductive types.

data βˆ₯_βˆ₯ {β„“} (A : Type β„“) : Type β„“ where
  inc    : A β†’ βˆ₯ A βˆ₯
  squash : is-prop βˆ₯ A βˆ₯

The two constructors that generate βˆ₯_βˆ₯ state precisely that the truncation is inhabited when A is (inc), and that it is a proposition (squash).

is-prop-βˆ₯-βˆ₯ : βˆ€ {β„“} {A : Type β„“} β†’ is-prop βˆ₯ A βˆ₯
is-prop-βˆ₯-βˆ₯ = squash

The eliminator for βˆ₯_βˆ₯ says that you can eliminate onto PP whenever it is a family of propositions, by providing a case for inc.

βˆ₯-βˆ₯-elim : βˆ€ {β„“ β„“'} {A : Type β„“}
             {P : βˆ₯ A βˆ₯ β†’ Type β„“'}
         β†’ ((x : _) β†’ is-prop (P x))
         β†’ ((x : A) β†’ P (inc x))
         β†’ (x : βˆ₯ A βˆ₯) β†’ P x
βˆ₯-βˆ₯-elim pprop incc (inc x) = incc x
βˆ₯-βˆ₯-elim pprop incc (squash x y i) =
  is-propβ†’pathp (Ξ» j β†’ pprop (squash x y j)) (βˆ₯-βˆ₯-elim pprop incc x)
                                             (βˆ₯-βˆ₯-elim pprop incc y)
                                             i

The propositional truncation can be called the free proposition on a type, because it satisfies the universal property that a left adjoint would have. Specifically, let B be a proposition. We have:

βˆ₯-βˆ₯-univ : βˆ€ {β„“} {A : Type β„“} {B : Type β„“}
         β†’ is-prop B β†’ (βˆ₯ A βˆ₯ β†’ B) ≃ (A β†’ B)
βˆ₯-βˆ₯-univ {A = A} {B = B} bprop = Isoβ†’Equiv (inc' , iso rec (Ξ» _ β†’ refl) beta) where
  inc' : (x : βˆ₯ A βˆ₯ β†’ B) β†’ A β†’ B
  inc' f x = f (inc x)

  rec : (f : A β†’ B) β†’ βˆ₯ A βˆ₯ β†’ B
  rec f (inc x) = f x
  rec f (squash x y i) = bprop (rec f x) (rec f y) i

  beta : _
  beta f = funext (βˆ₯-βˆ₯-elim (Ξ» _ β†’ is-propβ†’is-set bprop _ _) (Ξ» _ β†’ refl))

Furthermore, as required of a free construction, the propositional truncation extends to a functor:

βˆ₯-βˆ₯-map : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'}
        β†’ (A β†’ B) β†’ βˆ₯ A βˆ₯ β†’ βˆ₯ B βˆ₯
βˆ₯-βˆ₯-map f (inc x)        = inc (f x)
βˆ₯-βˆ₯-map f (squash x y i) = squash (βˆ₯-βˆ₯-map f x) (βˆ₯-βˆ₯-map f y) i

Using the propositional truncation, we can define the existential quantifier as a truncated Ξ£.

βˆƒ : βˆ€ {a b} (A : Type a) (B : A β†’ Type b) β†’ Type _
βˆƒ A B = βˆ₯ Ξ£ A B βˆ₯

syntax βˆƒ A (Ξ» x β†’ B) = βˆƒ[ x ∈ A ] B

Note that if PP is already a proposition, then truncating it does nothing:

is-propβ†’equivβˆ₯-βˆ₯ : βˆ€ {β„“} {P : Type β„“} β†’ is-prop P β†’ P ≃ βˆ₯ P βˆ₯
is-propβ†’equivβˆ₯-βˆ₯ pprop = prop-ext pprop squash inc (βˆ₯-βˆ₯-proj pprop)

In fact, an alternative definition of is-prop is given by β€œbeing equivalent to your own truncation”:

is-prop≃equivβˆ₯-βˆ₯ : βˆ€ {β„“} {P : Type β„“}
               β†’ is-prop P ≃ (P ≃ βˆ₯ P βˆ₯)
is-prop≃equivβˆ₯-βˆ₯ {P = P} =
  prop-ext is-prop-is-prop eqv-prop is-propβ†’equivβˆ₯-βˆ₯ inv
  where
    inv : (P ≃ βˆ₯ P βˆ₯) β†’ is-prop P
    inv eqv = equivβ†’is-hlevel 1 ((eqv e⁻¹) .fst) ((eqv e⁻¹) .snd) squash

    eqv-prop : is-prop (P ≃ βˆ₯ P βˆ₯)
    eqv-prop x y = Ξ£-path (Ξ» i p β†’ squash (x .fst p) (y .fst p) i)
                          (is-equiv-is-prop _ _ _)

Throughout the 1Lab, we use the words β€œmere” and β€œmerely” to modify a type to mean its propositional truncation. This terminology is adopted from the HoTT book. For example, a type XX is said merely equivalent to YY if the type βˆ₯X≑Yβˆ₯\| X \equiv Y \| is inhabited.

Maps into setsπŸ”—

The elimination principle for βˆ₯Aβˆ₯\| A \| says that we can only use the AA inside in a way that doesn’t matter: the motive of elimination must be a family of propositions, so our use of AA must not matter in a very strong sense. Often, it’s useful to relax this requirement slightly: Can we map out of βˆ₯Aβˆ₯\| A \| using a constant function?

The answer is yes, provided we’re mapping into a set! In that case, the image of ff is a proposition, so that we can map from βˆ₯Aβˆ₯β†’im⁑fβ†’B\| A \| \to \operatorname*{im}f \to B. In the next section, we’ll see a more general method for mapping into types that aren’t sets.

From the discussion in 1Lab.Counterexamples.Sigma, we know the definition of image, or more properly of (βˆ’1)(-1)-image:

private
  image : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'} β†’ (A β†’ B) β†’ Type _
  image {A = A} {B = B} f = Ξ£[ b ∈ B ] βˆƒ[ a ∈ A ] (f a ≑ b)

To see that the image indeed implements the concept of image, we define a way to factor any map through its image. By the definition of image, we have that the map f-image is always surjective, and since βˆƒ is a family of props, the first projection out of image is an embedding. Thus we factor a map ff as Aβ† im⁑fβ†ͺBA \twoheadrightarrow \operatorname*{im}f \hookrightarrow B.

f-image
  : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'}
  β†’ (f : A β†’ B) β†’ A β†’ image f
f-image f x = f x , inc (x , refl)

We now prove the theorem that will let us map out of a propositional truncation using a constant function into sets: if BB is a set, and f:Aβ†’Bf : A \to B is a constant function, then im⁑f\operatorname*{im}f is a proposition.

is-constant→image-is-prop
  : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'}
  β†’ is-set B
  β†’ (f : A β†’ B) β†’ (βˆ€ x y β†’ f x ≑ f y) β†’ is-prop (image f)

This is intuitively true (if the function is constant, then there is at most one thing in the image!), but formalising it turns out to be slightly tricky, and the requirement that BB be a set is perhaps unexpected.

A sketch of the proof is as follows. Suppose that we have some (a,x)(a, x) and (b,y)(b, y) in the image. We know, morally, that xx (respectively yy) give us some fβˆ—(a):Af^*(a) : A and p:f(fβˆ—a)=ap : f(f^*a) = a (resp q:f(fβˆ—(b))=bq : f(f^*(b)) = b) β€” which would establish that a≑ba \equiv b, as we need, since we have a=f(fβˆ—(a))=f(fβˆ—(b))=ba = f(f^*(a)) = f(f^*(b)) = b, where the middle equation is by constancy of ff β€” but pp and qq are hidden under propositional truncations, so we crucially need to use the fact that BB is a set so that a=ba = b is a proposition.

is-constant→image-is-prop bset f fconst (a , x) (b , y) =
  Ξ£-prop-path (Ξ» _ β†’ squash)
    (βˆ₯-βˆ₯-elimβ‚‚ (Ξ» _ _ β†’ bset _ _)
      (Ξ» { (f*a , p) (f*b , q) β†’ sym p Β·Β· fconst f*a f*b Β·Β· q }) x y)

Using the image factorisation, we can project from a propositional truncation onto a set using a constant map.

βˆ₯-βˆ₯-rec-set : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'}
            β†’ is-set B
            β†’ (f : A β†’ B)
            β†’ (βˆ€ x y β†’ f x ≑ f y)
            β†’ βˆ₯ A βˆ₯ β†’ B
βˆ₯-βˆ₯-rec-set {A = A} {B} bset f fconst x =
  βˆ₯-βˆ₯-elim {P = Ξ» _ β†’ image f}
    (λ _ → is-constant→image-is-prop bset f fconst) (f-image f) x .fst

Maps into groupoidsπŸ”—

We can push this idea further: as discussed in (Kraus 2015), in general, functions βˆ₯Aβˆ₯β†’B\| A \| \to B are equivalent to coherently constant functions Aβ†’BA \to B. This involves an infinite tower of conditions, each relating to the previous one, which isn’t something we can easily formulate in the language of type theory.

However, when BB is an nn-type, it is enough to ask for the first nn levels of the tower. In the case of sets, we’ve seen that the naΓ―ve notion of constancy is enough. We now deal with the case of groupoids, which requires an additional step: we ask for a function f:Aβ†’Bf : A \to B equipped with a witness of constancy constx,y:fx≑fy\mathrm{const}_{x,y} : f x \equiv f y and a coherence cohx,y,z:constx,yβˆ™consty,z≑constx,z\mathrm{coh}_{x,y,z} : \mathrm{const}_{x,y} \bullet \mathrm{const}_{y,z} \equiv \mathrm{const}_{x,z}.

This time, we cannot hope to show that the image of ff is a proposition: the image of a map βŠ€β†’S1\top \to S^1 is S1S^1. Instead, we use the following higher inductive type, which can be thought of as the β€œcodiscrete groupoid” on AA:

data βˆ₯_βˆ₯Β³ {β„“} (A : Type β„“) : Type β„“ where
  inc : A β†’ βˆ₯ A βˆ₯Β³
  iconst : βˆ€ a b β†’ inc a ≑ inc b
  icoh : βˆ€ a b c β†’ PathP (Ξ» i β†’ inc a ≑ iconst b c i) (iconst a b) (iconst a c)
  squash : is-groupoid βˆ₯ A βˆ₯Β³

The recursion principle for this type says exactly that any coherently constant function into a groupoid factors through βˆ₯Aβˆ₯3\| A \|^3!

βˆ₯-βˆ₯Β³-rec
  : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'}
  β†’ is-groupoid B
  β†’ (f : A β†’ B)
  β†’ (fconst : βˆ€ x y β†’ f x ≑ f y)
  β†’ (βˆ€ x y z β†’ fconst x y βˆ™ fconst y z ≑ fconst x z)
  β†’ βˆ₯ A βˆ₯Β³ β†’ B
βˆ₯-βˆ₯Β³-rec {A = A} {B} bgrpd f fconst fcoh = go where
  go : βˆ₯ A βˆ₯Β³ β†’ B
  go (inc x) = f x
  go (iconst a b i) = fconst a b i
  go (icoh a b c i j) = βˆ™β†’square (sym (fcoh a b c)) i j
  go (squash x y a b p q i j k) = bgrpd
    (go x) (go y)
    (Ξ» i β†’ go (a i)) (Ξ» i β†’ go (b i))
    (Ξ» i j β†’ go (p i j)) (Ξ» i j β†’ go (q i j))
    i j k

All that remains to show is that βˆ₯Aβˆ₯3\| A \|^3 is a proposition1, which mainly requires writing more elimination principles.

βˆ₯-βˆ₯Β³-is-prop : βˆ€ {β„“} {A : Type β„“} β†’ is-prop βˆ₯ A βˆ₯Β³
βˆ₯-βˆ₯Β³-is-prop = is-contr-if-inhabitedβ†’is-prop $
  βˆ₯-βˆ₯Β³-elim-prop (Ξ» _ β†’ hlevel 1)
    (Ξ» a β†’ contr (inc a) (βˆ₯-βˆ₯Β³-elim-set (Ξ» _ β†’ squash _ _) (iconst a) (icoh a)))

Hence we get the desired recursion principle for the usual propositional truncation.

βˆ₯-βˆ₯-rec-groupoid
  : βˆ€ {β„“ β„“'} {A : Type β„“} {B : Type β„“'}
  β†’ is-groupoid B
  β†’ (f : A β†’ B)
  β†’ (fconst : βˆ€ x y β†’ f x ≑ f y)
  β†’ (βˆ€ x y z β†’ fconst x y βˆ™ fconst y z ≑ fconst x z)
  β†’ βˆ₯ A βˆ₯ β†’ B
βˆ₯-βˆ₯-rec-groupoid bgrpd f fconst fcoh =
  βˆ₯-βˆ₯Β³-rec bgrpd f fconst fcoh ∘ βˆ₯-βˆ₯-rec βˆ₯-βˆ₯Β³-is-prop inc
As we hinted at above, this method generalises (externally) to nn-types; we sketch the details of the next level for the curious reader.

The next coherence involves a tetrahedron all of whose faces are coh\mathrm{coh}, or, since we’re doing cubical type theory, a β€œcubical tetrahedron”:

data βˆ₯_βˆ₯⁴ {β„“} (A : Type β„“) : Type β„“ where
  inc : A β†’ βˆ₯ A βˆ₯⁴
  iconst : βˆ€ a b β†’ inc a ≑ inc b
  icoh : βˆ€ a b c β†’ PathP (Ξ» i β†’ inc a ≑ iconst b c i) (iconst a b) (iconst a c)
  iassoc : βˆ€ a b c d β†’ PathP (Ξ» i β†’ PathP (Ξ» j β†’ inc a ≑ icoh b c d i j)
                                          (iconst a b) (icoh a c d i))
                             (icoh a b c) (icoh a b d)
  squash : is-hlevel βˆ₯ A βˆ₯⁴ 4

βˆ₯-βˆ₯⁴-rec
  : βˆ€ {β„“} {A : Type β„“} {β„“'} {B : Type β„“'}
  β†’ is-hlevel B 4
  β†’ (f : A β†’ B)
  β†’ (fconst : βˆ€ a b β†’ f a ≑ f b)
  β†’ (fcoh : βˆ€ a b c β†’ PathP (Ξ» i β†’ f a ≑ fconst b c i) (fconst a b) (fconst a c))
  β†’ (βˆ€ a b c d β†’ PathP (Ξ» i β†’ PathP (Ξ» j β†’ f a ≑ fcoh b c d i j)
                                    (fconst a b) (fcoh a c d i))
                       (fcoh a b c) (fcoh a b d))
  β†’ βˆ₯ A βˆ₯⁴ β†’ B
unquoteDef βˆ₯-βˆ₯⁴-rec = make-rec-n 4 βˆ₯-βˆ₯⁴-rec (quote βˆ₯_βˆ₯⁴)

  1. in fact, it’s even a propositional truncation of AA, in that it satisfies the same universal property as βˆ₯Aβˆ₯\| A \|β†©οΈŽ


References