module 1Lab.HIT.Truncation where
Propositional truncation🔗
Let be a type. The propositional truncation of 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
: A → ∥ A ∥
inc : is-prop ∥ A ∥ squash
The two constructors that generate ∥_∥
state
precisely that the truncation is inhabited when A
is (inc
), and that it is a proposition (squash
).
: ∀ {ℓ} {A : Type ℓ} → is-prop ∥ A ∥
is-prop-∥-∥ = squash is-prop-∥-∥
instance
: ∀ {n} {ℓ} {A : Type ℓ} → H-Level ∥ A ∥ (suc n)
H-Level-truncation = prop-instance squash H-Level-truncation
The eliminator for ∥_∥
says that you
can eliminate onto
whenever it is a family of propositions, by providing a case for inc
.
: ∀ {ℓ ℓ'} {A : Type ℓ}
∥-∥-elim {P : ∥ A ∥ → Type ℓ'}
→ ((x : _) → is-prop (P x))
→ ((x : A) → P (inc x))
→ (x : ∥ A ∥) → P x
(inc x) = incc x
∥-∥-elim pprop incc (squash x y i) =
∥-∥-elim pprop incc (λ j → pprop (squash x y j)) (∥-∥-elim pprop incc x)
is-prop→pathp (∥-∥-elim pprop incc y)
i
: ∀ {ℓ ℓ₁ ℓ₂} {A : Type ℓ} {B : Type ℓ₁}
∥-∥-elim₂ {P : ∥ A ∥ → ∥ B ∥ → Type ℓ₂}
→ (∀ x y → is-prop (P x y))
→ (∀ x y → P (inc x) (inc y))
→ ∀ x y → P x y
{A = A} {B} {P} pprop work x y = go x y where
∥-∥-elim₂ : ∀ x y → P x y
go (inc x) (inc x₁) = work x x₁
go (inc x) (squash y y₁ i) =
go (λ i → pprop (inc x) (squash y y₁ i))
is-prop→pathp (go (inc x) y) (go (inc x) y₁) i
(squash x x₁ i) z =
go (λ i → pprop (squash x x₁ i) z)
is-prop→pathp (go x z) (go x₁ z) i
: ∀ {ℓ ℓ'} {A : Type ℓ} {P : Type ℓ'}
∥-∥-rec → is-prop P
→ (A → P)
→ (x : ∥ A ∥) → P
= ∥-∥-elim (λ _ → pprop)
∥-∥-rec pprop
: ∀ {ℓ} {A : Type ℓ} → is-prop A → ∥ A ∥ → A
∥-∥-out = ∥-∥-rec ap λ x → x
∥-∥-out ap
: ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ''} {P : Type ℓ'}
∥-∥-rec₂ → is-prop P
→ (A → B → P)
→ (x : ∥ A ∥) (y : ∥ B ∥) → P
= ∥-∥-elim₂ (λ _ _ → pprop)
∥-∥-rec₂ pprop
: ∀ {ℓ} {A : Type ℓ} ⦃ _ : H-Level A 1 ⦄ → ∥ A ∥ → A
∥-∥-out! = ∥-∥-out (hlevel 1)
∥-∥-out!
instance
Inductive-∥∥: ∀ {ℓ ℓ' ℓm} {A : Type ℓ} {P : ∥ A ∥ → Type ℓ'} ⦃ i : Inductive (∀ x → P (inc x)) ℓm ⦄
→ ⦃ _ : ∀ {x} → H-Level (P x) 1 ⦄
→ Inductive (∀ x → P x) ℓm
= record
Inductive-∥∥ ⦃ i ⦄ { methods = i .Inductive.methods
; from = λ f → ∥-∥-elim (λ x → hlevel 1) (i .Inductive.from f)
}
: ∀ {ℓ} {A : Type ℓ} → ⦃ Dec A ⦄ → Dec ∥ A ∥
Dec-∥∥ = yes (inc a)
Dec-∥∥ ⦃ yes a ⦄ = no (rec! ¬a) Dec-∥∥ ⦃ no ¬a ⦄
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:
: ∀ {ℓ} {A : Type ℓ} {B : Type ℓ}
∥-∥-univ → is-prop B → (∥ A ∥ → B) ≃ (A → B)
{A = A} {B = B} bprop = Iso→Equiv (inc' , iso rec (λ _ → refl) beta) where
∥-∥-univ : (x : ∥ A ∥ → B) → A → B
inc' = f (inc x)
inc' f x
: (f : A → B) → ∥ A ∥ → B
rec (inc x) = f x
rec f (squash x y i) = bprop (rec f x) (rec f y) i
rec f
: _
beta = funext (∥-∥-elim (λ _ → is-prop→is-set bprop _ _) (λ _ → refl)) beta f
Furthermore, as required of a free construction, the propositional truncation extends to a functor:
: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'}
∥-∥-map → (A → B) → ∥ A ∥ → ∥ B ∥
(inc x) = inc (f x)
∥-∥-map f (squash x y i) = squash (∥-∥-map f x) (∥-∥-map f y) i ∥-∥-map f
: ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''}
∥-∥-map₂ → (A → B → C) → ∥ A ∥ → ∥ B ∥ → ∥ C ∥
(inc x) (inc y) = inc (f x y)
∥-∥-map₂ f (squash x y i) z = squash (∥-∥-map₂ f x z) (∥-∥-map₂ f y z) i
∥-∥-map₂ f (squash y z i) = squash (∥-∥-map₂ f x y) (∥-∥-map₂ f x z) i ∥-∥-map₂ f x
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
: ∀ {a b} (A : Type a) (B : A → Type b) → Type _
∃-syntax = ∃
∃-syntax
syntax ∃-syntax A (λ x → B) = ∃[ x ∈ A ] B
infix 5 ∃-syntax
Note that if is already a proposition, then truncating it does nothing:
: ∀ {ℓ} {P : Type ℓ} → is-prop P → P ≃ ∥ P ∥
is-prop→equiv∥-∥ = prop-ext pprop squash inc (∥-∥-out pprop) is-prop→equiv∥-∥ pprop
In fact, an alternative definition of is-prop
is given by “being equivalent to
your own truncation”:
: ∀ {ℓ} {P : Type ℓ}
is-prop≃equiv∥-∥ → is-prop P ≃ (P ≃ ∥ P ∥)
{P = P} =
is-prop≃equiv∥-∥
prop-ext is-prop-is-prop eqv-prop is-prop→equiv∥-∥ invwhere
: (P ≃ ∥ P ∥) → is-prop P
inv = equiv→is-hlevel 1 ((eqv e⁻¹) .fst) ((eqv e⁻¹) .snd) squash
inv eqv
: is-prop (P ≃ ∥ P ∥)
eqv-prop = Σ-path (λ i p → squash (x .fst p) (y .fst p) i)
eqv-prop x y (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 is said merely equivalent to if the type is inhabited.
Maps into sets🔗
The elimination principle for says that we can only use the inside in a way that doesn’t matter: the motive of elimination must be a family of propositions, so our use of must not matter in a very strong sense. Often, it’s useful to relax this requirement slightly: Can we map out of using a constant function?
The answer is yes, provided we’re mapping into a set! In that case, the image of is a proposition, so that we can map from 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
: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _
image {A = A} {B = B} f = Σ[ b ∈ B ] ∃[ a ∈ A ] (f a ≡ b) image
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
image-inc
is always surjective, and
since ∃
is a family of props, the first projection out of
image
is an embedding. Thus we
factor a map
as
image-inc: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'}
→ (f : A → B) → A → image f
= f x , inc (x , refl) image-inc f x
We now prove the theorem that will let us map out of a propositional truncation using a constant function into sets: if is a set, and is a constant function, then 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 be a set is perhaps unexpected.
A sketch of the proof is as follows. Suppose that we have some and in the image. We know, morally, that (respectively give us some and (resp — which would establish that as we need, since we have where the middle equation is by constancy of — but and are hidden under propositional truncations, so we crucially need to use the fact that is a set so that is a proposition.
(a , x) (b , y) =
is-constant→image-is-prop bset f fconst (λ _ → squash)
Σ-prop-path (∥-∥-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.
: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'}
∥-∥-rec-set → is-set B
→ (f : A → B)
→ (∀ x y → f x ≡ f y)
→ ∥ A ∥ → B
{A = A} {B} bset f fconst x =
∥-∥-rec-set {P = λ _ → image f}
∥-∥-elim (λ _ → is-constant→image-is-prop bset f fconst) (image-inc f) x .fst
∥-∥-rec-set!: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} ⦃ _ : H-Level B 2 ⦄
→ (f : A → B) (c : ∀ x y → f x ≡ f y)
→ ∥ A ∥ → B
{A = A} f c x = go x .fst where
∥-∥-rec-set! : ∥ A ∥ → image f
go (inc x) = f x , inc (x , refl)
go (squash x y i) = is-constant→image-is-prop (hlevel 2) f c (go x) (go y) i go
Maps into groupoids🔗
We can push this idea further: as discussed in (Kraus 2015), in general, functions are equivalent to coherently constant functions 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 is an it is enough to ask for the first 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 equipped with a witness of constancy and a coherence
This time, we cannot hope to show that the image of is a proposition: the image of a map is Instead, we use the following higher inductive type, which can be thought of as the “codiscrete groupoid” on
data ∥_∥³ {ℓ} (A : Type ℓ) : Type ℓ where
: A → ∥ A ∥³
inc : ∀ a b → inc a ≡ inc b
iconst : ∀ a b c → PathP (λ i → inc a ≡ iconst b c i) (iconst a b) (iconst a c)
icoh : is-groupoid ∥ A ∥³ squash
The recursion principle for this type says exactly that any coherently constant function into a groupoid factors through
∥-∥³-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
{A = A} {B} bgrpd f fconst fcoh = go where
∥-∥³-rec : ∥ 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 (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 is a proposition1, which mainly requires writing more elimination principles.
∥-∥³-elim-set: ∀ {ℓ} {A : Type ℓ} {ℓ'} {P : ∥ A ∥³ → Type ℓ'}
→ (∀ a → is-set (P a))
→ (f : (a : A) → P (inc a))
→ (∀ a b → PathP (λ i → P (iconst a b i)) (f a) (f b))
→ ∀ a → P a
unquoteDef ∥-∥³-elim-set = make-elim-n 2 ∥-∥³-elim-set (quote ∥_∥³)
∥-∥³-elim-prop: ∀ {ℓ} {A : Type ℓ} {ℓ'} {P : ∥ A ∥³ → Type ℓ'}
→ (∀ a → is-prop (P a))
→ (f : (a : A) → P (inc a))
→ ∀ a → P a
unquoteDef ∥-∥³-elim-prop = make-elim-n 1 ∥-∥³-elim-prop (quote ∥_∥³)
: ∀ {ℓ} {A : Type ℓ} → is-prop ∥ A ∥³
∥-∥³-is-prop = is-contr-if-inhabited→is-prop $
∥-∥³-is-prop (λ _ → hlevel 1)
∥-∥³-elim-prop (λ 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 we sketch the details of the next level for the curious reader.
The next coherence involves a tetrahedron all of whose faces are or, since we’re doing cubical type theory, a “cubical tetrahedron”:
data ∥_∥⁴ {ℓ} (A : Type ℓ) : Type ℓ where
: A → ∥ A ∥⁴
inc : ∀ a b → inc a ≡ inc b
iconst : ∀ a b c → PathP (λ i → inc a ≡ iconst b c i) (iconst a b) (iconst a c)
icoh : ∀ a b c d → PathP (λ i → PathP (λ j → inc a ≡ icoh b c d i j)
iassoc (iconst a b) (icoh a c d i))
(icoh a b c) (icoh a b d)
: is-hlevel ∥ A ∥⁴ 4
squash
∥-∥⁴-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 ∥_∥⁴)
open import Meta.Idiom
open import Meta.Bind
instance
: Map (eff ∥_∥)
Map-∥∥ .Map.map = ∥-∥-map
Map-∥∥
: Idiom (eff ∥_∥)
Idiom-∥∥ .Idiom.pure = inc
Idiom-∥∥ .Idiom._<*>_ {A = A} {B = B} = go where
Idiom-∥∥ : ∥ (A → B) ∥ → ∥ A ∥ → ∥ B ∥
go (inc f) (inc x) = inc (f x)
go (inc f) (squash x y i) = squash (go (inc f) x) (go (inc f) y) i
go (squash f g i) y = squash (go f y) (go g y) i
go
: Bind (eff ∥_∥)
Bind-∥∥ .Bind._>>=_ {A = A} {B = B} = go where
Bind-∥∥ : ∥ A ∥ → (A → ∥ B ∥) → ∥ B ∥
go (inc x) f = f x
go (squash x y i) f = squash (go x f) (go y f) i go
is-embedding→image-inc-is-equiv: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B}
→ is-embedding f
→ is-equiv (image-inc f)
{f = f} f-emb =
is-embedding→image-inc-is-equiv
is-iso→is-equiv $(λ im → fst $ ∥-∥-out (f-emb _) (im .snd))
iso (λ im → Σ-prop-path! (snd $ ∥-∥-out (f-emb _) (im .snd)))
(λ _ → refl)
is-embedding→image-equiv: ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B}
→ is-embedding f
→ A ≃ image f
{f = f} f-emb =
is-embedding→image-equiv image-inc f , is-embedding→image-inc-is-equiv f-emb
in fact, it’s even a propositional truncation of in that it satisfies the same universal property as ↩︎
References
- Kraus, Nicolai. 2015. “The General Universal Property of the Propositional Truncation.” https://doi.org/10.4230/LIPICS.TYPES.2014.111.