open import 1Lab.HLevel.Retracts open import 1Lab.Type.Sigma open import 1Lab.HLevel open import 1Lab.Equiv open import 1Lab.Path open import 1Lab.Type 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 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 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 = ∥ Σ B ∥ syntax ∃ A (λ x → B) = ∃[ x ∈ A ] B
Note that if 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 (∥-∥-elim (λ x → pprop) λ x → x)
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 _ _ _)
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! However, the witness of constancy we use must be very coherent indeed. In particular, we need enough coherence on top of a family of paths to ensure that the image of is a proposition; Then we can map from .
From the discussion in 1Lab.Counterexamples.Sigma, we know the definition of image, or more properly of -image:
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 as .
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 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 crucially, the
is-constant→image-is-prop bset f f-const (a , x) (b , y) = Σ-prop-path (λ _ → squash) (∥-∥-elim₂ (λ _ _ → bset _ _) (λ { (f*a , p) (f*b , q) → sym p ·· f-const 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 ℓ'} → (f : A → B) → (∀ x y → f x ≡ f y) → is-set B → ∥ A ∥ → B ∥-∥-rec-set {A = A} {B} f f-const bset x = ∥-∥-elim {P = λ _ → image f} (λ _ → is-constant→image-is-prop bset f f-const) (f-image f) x .fst