Martin Escardo

In univalent logic, as opposed to Curry-Howard logic, a proposition is
a subsingleton or a type such that any two of its elements are
identified.

https://www.newton.ac.uk/files/seminar/20170711100011001-1442677.pdf
https://unimath.github.io/bham2017/uf.pdf

\begin{code}

{-# OPTIONS --safe --without-K #-}

module UF.Subsingletons where

open import MLTT.Plus-Properties
open import MLTT.Spartan
open import MLTT.Unit-Properties
open import UF.Base

is-prop : 𝓀 Μ‡ β†’ 𝓀 Μ‡
is-prop X = (x y : X) β†’ x = y

is-prop-valued-family : {X : 𝓀 Μ‡ } β†’ (X β†’ π“₯ Μ‡ ) β†’ 𝓀 βŠ” π“₯ Μ‡
is-prop-valued-family A = βˆ€ x β†’ is-prop (A x)

\end{code}

And of course we could adopt a terminology borrowed from topos logic:

\begin{code}

is-truth-value is-subsingleton : 𝓀 Μ‡ β†’ 𝓀 Μ‡
is-truth-value  = is-prop
is-subsingleton = is-prop

Ξ£-is-prop : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ }
          β†’ is-prop X
          β†’ ((x : X) β†’ is-prop (A x))
          β†’ is-prop (Ξ£ A)
Ξ£-is-prop {𝓀} {π“₯} {X} {A} i j (x , a) (y , b) =
 to-Σ-= (i x y , j y (transport A (i x y) a) b)

\end{code}

Next we define singleton (or contractible types). The terminology
"contractible" is due to Voevodsky. I currently prefer the terminology
"singleton type", because it makes more sense when we consider
univalent type theory as interesting on its own right independently of
its homotopical (originally motivating) models. Also it emphasizes
that we don't require homotopy theory as a prerequisite to understand
univalent type theory.

\begin{code}

is-central : (X : 𝓀 Μ‡ ) β†’ X β†’ 𝓀 Μ‡
is-central X c = (x : X) β†’ c = x

is-singleton : 𝓀 Μ‡ β†’ 𝓀 Μ‡
is-singleton X = Ξ£ c κž‰ X , is-central X c

center : {X : 𝓀 Μ‡ } β†’ is-singleton X β†’ X
center = pr₁

centrality : {X : 𝓀 Μ‡ } (i : is-singleton X) β†’ is-central X (center i)
centrality = prβ‚‚

\end{code}

For compatibility with the homotopical terminology:

\begin{code}

is-center-of-contraction-of : (X : 𝓀 Μ‡ ) β†’ X β†’ 𝓀 Μ‡
is-center-of-contraction-of = is-central

is-contr : 𝓀 Μ‡ β†’ 𝓀 Μ‡
is-contr = is-singleton

πŸ™-is-singleton : is-singleton (πŸ™ {𝓀})
πŸ™-is-singleton = ⋆ , (Ξ» (x : πŸ™) β†’ (πŸ™-all-⋆ x)⁻¹)

singletons-are-props : {X : 𝓀 Μ‡ } β†’ is-singleton X β†’ is-prop X
singletons-are-props (c , Ο†) x y = x =⟨ (Ο† x) ⁻¹ ⟩
                                   c =⟨ Ο† y ⟩
                                   y ∎

prop-criterion' : {X : 𝓀 Μ‡ }
                β†’ (X β†’ is-singleton X)
                β†’ is-prop X
prop-criterion' Ο† x = singletons-are-props (Ο† x) x

prop-criterion : {X : 𝓀 Μ‡ } β†’ (X β†’ is-prop X) β†’ is-prop X
prop-criterion Ο† x = Ο† x x

pointed-props-are-singletons : {X : 𝓀 Μ‡ }
                             β†’ X
                             β†’ is-prop X
                             β†’ is-singleton X
pointed-props-are-singletons x h = x , h x

\end{code}

The two prototypical propositions:

\begin{code}

𝟘-is-prop : is-prop (𝟘 {𝓀})
𝟘-is-prop {𝓀} x y = unique-from-𝟘 {𝓀} {𝓀} x

πŸ™-is-prop : is-prop (πŸ™ {𝓀})
πŸ™-is-prop {𝓀} ⋆ ⋆ = refl {𝓀}

singleton-type : {X : 𝓀 Μ‡ } (x : X) β†’ 𝓀 Μ‡
singleton-type x = Ξ£ y κž‰ type-of x , x = y

singleton-center : {X : 𝓀 Μ‡ } (x : X) β†’ singleton-type x
singleton-center x = (x , refl)

singleton-types-are-singletons'' : {X : 𝓀 Μ‡ } {x x' : X} (r : x = x')
                                 β†’ singleton-center x = (x' , r)
singleton-types-are-singletons'' {𝓀} {X} = J A (Ξ» x β†’ refl)
 where
  A : (x x' : X) β†’ x = x' β†’ 𝓀 Μ‡
  A x x' r = singleton-center x =[ Ξ£ x' κž‰ X , x = x' ] (x' , r)

singleton-types-are-singletons : {X : 𝓀 Μ‡ } (xβ‚€ : X)
                               β†’ is-singleton (singleton-type xβ‚€)
singleton-types-are-singletons xβ‚€ =
 singleton-center xβ‚€ ,
 (Ξ» t β†’ singleton-types-are-singletons'' (prβ‚‚ t))

singleton-types-are-singletons'
 : {X : 𝓀 Μ‡ } {x : X}
 β†’ is-central (singleton-type x) (singleton-center x)
singleton-types-are-singletons' {𝓀} {X} (y , refl) = refl

singleton-types-are-props : {X : 𝓀 Μ‡ } (x : X) β†’ is-prop (singleton-type x)
singleton-types-are-props x =
 singletons-are-props (singleton-types-are-singletons x)

singleton-type' : {X : 𝓀 Μ‡ } β†’ X β†’ 𝓀 Μ‡
singleton-type' x = Ξ£ y κž‰ type-of x , y = x

singleton'-center : {X : 𝓀 Μ‡ } (x : X) β†’ singleton-type' x
singleton'-center x = (x , refl)

Γ—-prop-criterion-necessity : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ }
                           β†’ is-prop (X Γ— Y)
                           β†’ (Y β†’ is-prop X) Γ— (X β†’ is-prop Y)
Γ—-prop-criterion-necessity i = (Ξ» y x x' β†’ ap pr₁ (i (x , y) (x' , y))) ,
                               (Ξ» x y y' β†’ ap prβ‚‚ (i (x , y) (x  , y')))

Γ—-prop-criterion : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ }
                 β†’ (Y β†’ is-prop X) Γ— (X β†’ is-prop Y)
                 β†’ is-prop (X Γ— Y)
Γ—-prop-criterion (i , j) (x , y) (x' , y') = to-Ξ£-= (i y x x' , j x _ _)

Γ—-𝟘-is-prop : {X : 𝓀 Μ‡ } β†’ is-prop (X Γ— 𝟘 {π“₯})
Γ—-𝟘-is-prop (x , z) _ = 𝟘-elim z

𝟘-Γ—-is-prop : {X : 𝓀 Μ‡ } β†’ is-prop (𝟘 {π“₯} Γ— X)
𝟘-Γ—-is-prop (z , x) _ = 𝟘-elim z

Γ—-is-prop : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ }
          β†’ is-prop X
          β†’ is-prop Y
          β†’ is-prop (X Γ— Y)
Γ—-is-prop i j = Γ—-prop-criterion ((Ξ» _ β†’ i) , (Ξ» _ β†’ j))

to-subtype-= : {X : 𝓦 Μ‡ } {A : X β†’ π“₯ Μ‡ }
               {x y : X} {a : A x} {b : A y}
             β†’ ((x : X) β†’ is-prop (A x))
             β†’ x = y
             β†’ (x , a) = (y , b)
to-subtype-= {𝓀} {π“₯} {X} {A} {x} {y} {a} {b} s p =
 to-Σ-= (p , s y (transport A p a) b)

subtypes-of-props-are-props' : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ } (m : X β†’ Y)
                             β†’ left-cancellable m
                             β†’ is-prop Y
                             β†’ is-prop X
subtypes-of-props-are-props' m lc i x x' = lc (i (m x) (m x'))

pr₁-lc : {X : 𝓀 Μ‡ } {Y : X β†’ π“₯ Μ‡ }
       β†’ ({x : X} β†’ is-prop (Y x))
       β†’ left-cancellable (pr₁ {𝓀} {π“₯} {X} {Y})
pr₁-lc f p = to-Ξ£-= (p , (f _ _))

subsets-of-props-are-props : (X : 𝓀 Μ‡ ) (Y : X β†’ π“₯ Μ‡ )
                         β†’ is-prop X
                         β†’ ({x : X} β†’ is-prop (Y x))
                         β†’ is-prop (Ξ£ x κž‰ X , Y x)
subsets-of-props-are-props X Y h p =
 subtypes-of-props-are-props' pr₁ (pr₁-lc p) h

inl-lc-is-section : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ }
                    {x x' : X}
                    (p : inl {𝓀} {π“₯} {X} {Y} x = inl x')
                  β†’ p = ap inl (inl-lc p)
inl-lc-is-section refl = refl

inr-lc-is-section : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ } {y y' : Y}
                    (p : inr {𝓀} {π“₯} {X} {Y} y = inr y')
                  β†’ p = ap inr (inr-lc p)
inr-lc-is-section refl = refl

\end{code}

The following says that, in particular, for any proposition P, we have
that P + Β¬ P is a proposition, or that the decidability of a
proposition is a proposition:

\begin{code}

sum-of-contradictory-props : {P : 𝓀 Μ‡ } {Q : π“₯ Μ‡ }
                           β†’ is-prop P
                           β†’ is-prop Q
                           β†’ (P β†’ Q β†’ 𝟘 {𝓦})
                           β†’ is-prop (P + Q)
sum-of-contradictory-props {𝓀} {π“₯} {𝓦} {P} {Q} i j f = Ξ³
 where
  Ξ³ : (x y : P + Q) β†’ x = y
  Ξ³ (inl p) (inl p') = ap inl (i p p')
  Ξ³ (inl p) (inr q)  = 𝟘-elim {𝓀 βŠ” π“₯} {𝓦} (f p q)
  γ (inr q) (inl p)  = 𝟘-elim (f p q)
  Ξ³ (inr q) (inr q') = ap inr (j q q')

sum-of-contradictory-props' : {P : 𝓀 Μ‡ } {Q : π“₯ Μ‡ }
                            β†’ (is-prop P Γ— is-prop Q Γ— (P β†’ Q β†’ 𝟘 {𝓦}))
                            β†’ is-prop (P + Q)
sum-of-contradictory-props' (i , j , f) = sum-of-contradictory-props i j f

sum-of-contradictory-props'-converse : {P : 𝓀 Μ‡ } {Q : π“₯ Μ‡ }
                                     β†’ is-prop (P + Q)
                                     β†’ (is-prop P Γ— is-prop Q Γ— (P β†’ Q β†’ 𝟘 {𝓦}))
sum-of-contradictory-props'-converse k =
 (Ξ» p p' β†’ inl-lc (k (inl p) (inl p'))) ,
 (Ξ» q q' β†’ inr-lc (k (inr q) (inr q'))) ,
 (Ξ» p q β†’ 𝟘-elim (+disjoint (k (inl p) (inr q))))

\end{code}

Formulation of propositional extensionality:

\begin{code}

propext : βˆ€ 𝓀 β†’ 𝓀 ⁺ Μ‡
propext 𝓀 = {P Q : 𝓀 Μ‡ } β†’ is-prop P β†’ is-prop Q β†’ (P β†’ Q) β†’ (Q β†’ P) β†’ P = Q

PropExt : 𝓀ω
PropExt = βˆ€ 𝓀 β†’ propext 𝓀

Prop-Ext : 𝓀ω
Prop-Ext = βˆ€ {𝓀} β†’ propext 𝓀

\end{code}

Without assuming excluded middle, we have that there are no truth
values other than 𝟘 and πŸ™:

\begin{code}

no-props-other-than-𝟘-or-πŸ™ : propext 𝓀
                           β†’ Β¬ (Ξ£ P κž‰ 𝓀 Μ‡ , is-prop P Γ— (P β‰  𝟘) Γ— (P β‰  πŸ™))
no-props-other-than-𝟘-or-πŸ™ pe (P , i , f , g) = 𝟘-elim (Ο† u)
 where
  u : Β¬ P
  u p = g l
   where
    l : P = πŸ™
    l = pe i πŸ™-is-prop unique-to-πŸ™ (Ξ» _ β†’ p)

  Ο† : ¬¬ P
  Ο† u = f l
   where
    l : P = 𝟘
    l = pe i 𝟘-is-prop (Ξ» p β†’ 𝟘-elim (u p)) 𝟘-elim

\end{code}

Notice how we used 𝟘-elim above to coerce a hypothetical value in 𝟘
{𝓀₀}, arising from negation, to a value in 𝟘 {𝓀}. Otherwise "u" would
have sufficed in place of "Ξ» p β†’ 𝟘-elim (u p)". The same technique is
used in the following construction.

\begin{code}

𝟘-is-not-πŸ™ : 𝟘 {𝓀} β‰  πŸ™ {𝓀}
𝟘-is-not-πŸ™ p = 𝟘-elim (Idtofun (p ⁻¹) ⋆)

universe-has-two-distinct-points : has-two-distinct-points (𝓀 Μ‡ )
universe-has-two-distinct-points = ((𝟘 , πŸ™) , 𝟘-is-not-πŸ™)

\end{code}

Unique existence.

\begin{code}

βˆƒ! : {X : 𝓀 Μ‡ } (A : X β†’ π“₯ Μ‡ ) β†’ 𝓀 βŠ” π“₯ Μ‡
βˆƒ! A = is-singleton (Ξ£ A)

existsUnique : (X : 𝓀 Μ‡ ) (A : X β†’ π“₯ Μ‡ ) β†’ 𝓀 βŠ” π“₯ Μ‡
existsUnique X A = βˆƒ! A

syntax existsUnique X (Ξ» x β†’ b) = βˆƒ! x κž‰ X , b

witness-uniqueness : {X : 𝓀 Μ‡ } (A : X β†’ π“₯ Μ‡ )
                   β†’ (βˆƒ! x κž‰ X , A x)
                   β†’ (x y : X) β†’ A x β†’ A y β†’ x = y
witness-uniqueness A e x y a b = ap pr₁ (singletons-are-props e (x , a) (y , b))

infixr -1 existsUnique

βˆƒ!-intro : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ } (x : X) (a : A x)
         β†’ ((Οƒ : Ξ£ A) β†’ (x , a) = Οƒ)
         β†’ βˆƒ! A
βˆƒ!-intro x a o = (x , a) , o

βˆƒ!-witness : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ } β†’ βˆƒ! A β†’ X
βˆƒ!-witness ((x , a) , o) = x

βˆƒ!-is-witness : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ }
                (u : βˆƒ! A)
              β†’ A (βˆƒ!-witness u)
βˆƒ!-is-witness ((x , a) , o) = a

description : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ } β†’ βˆƒ! A β†’ Ξ£ A
description (Οƒ , o) = Οƒ

βˆƒ!-uniqueness' : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ }
                 (u : βˆƒ! A)
               β†’ (Οƒ : Ξ£ A)
               β†’ description u = Οƒ
βˆƒ!-uniqueness' ((x , a) , o) = o

βˆƒ!-uniqueness : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ }
                (u : βˆƒ! A)
              β†’ (x : X)
                (a : A x)
              β†’ description u = (x , a)
βˆƒ!-uniqueness u x a = βˆƒ!-uniqueness' u (x , a)

βˆƒ!-uniqueness'' : {X : 𝓀 Μ‡ } {A : X β†’ π“₯ Μ‡ }
                  (u : βˆƒ! A)
                β†’ (Οƒ Ο‰ : Ξ£ A)
                β†’ Οƒ = Ο‰
βˆƒ!-uniqueness'' u Οƒ Ο‰ = βˆƒ!-uniqueness' u Οƒ ⁻¹ βˆ™ βˆƒ!-uniqueness' u Ο‰

\end{code}

Added 5 March 2020 by Tom de Jong.

\begin{code}

+-is-prop : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ }
          β†’ is-prop X
          β†’ is-prop Y
          β†’ (X β†’ Β¬ Y)
          β†’ is-prop (X + Y)
+-is-prop = sum-of-contradictory-props

+-is-prop' : {X : 𝓀 Μ‡ } {Y : π“₯ Μ‡ }
           β†’ is-prop X
           β†’ is-prop Y
           β†’ (Y β†’ Β¬ X)
           β†’ is-prop (X + Y)
+-is-prop' {𝓀} {π“₯} {X} {Y} i j f = +-is-prop i j (Ξ» y x β†’ f x y)

\end{code}

Added 16th June 2020 by Martin Escardo. (Should have added this ages
ago to avoid boiler-plate code.)

\begin{code}

×₃-is-prop : {π“₯β‚€ π“₯₁ π“₯β‚‚ : Universe}
             {Xβ‚€ : π“₯β‚€ Μ‡ }
             {X₁ : π“₯₁ Μ‡ }
             {Xβ‚‚ : π“₯β‚‚ Μ‡ }
           β†’ is-prop Xβ‚€
           β†’ is-prop X₁
           β†’ is-prop Xβ‚‚
           β†’ is-prop (Xβ‚€ Γ— X₁ Γ— Xβ‚‚)
×₃-is-prop iβ‚€ i₁ iβ‚‚ =
 Γ—-is-prop iβ‚€ (Γ—-is-prop i₁ iβ‚‚)

Γ—β‚„-is-prop : {π“₯β‚€ π“₯₁ π“₯β‚‚ π“₯₃ : Universe}
             {Xβ‚€ : π“₯β‚€ Μ‡ }
             {X₁ : π“₯₁ Μ‡ }
             {Xβ‚‚ : π“₯β‚‚ Μ‡ }
             {X₃ : π“₯₃ Μ‡ }
           β†’ is-prop Xβ‚€
           β†’ is-prop X₁
           β†’ is-prop Xβ‚‚
           β†’ is-prop X₃
           β†’ is-prop (Xβ‚€ Γ— X₁ Γ— Xβ‚‚ Γ— X₃)
Γ—β‚„-is-prop iβ‚€ i₁ iβ‚‚ i₃ =
 Γ—-is-prop iβ‚€ (×₃-is-prop i₁ iβ‚‚ i₃)

Γ—β‚…-is-prop : {π“₯β‚€ π“₯₁ π“₯β‚‚ π“₯₃ π“₯β‚„ : Universe}
             {Xβ‚€ : π“₯β‚€ Μ‡ }
             {X₁ : π“₯₁ Μ‡ }
             {Xβ‚‚ : π“₯β‚‚ Μ‡ }
             {X₃ : π“₯₃ Μ‡ }
             {Xβ‚„ : π“₯β‚„ Μ‡ }
           β†’ is-prop Xβ‚€
           β†’ is-prop X₁
           β†’ is-prop Xβ‚‚
           β†’ is-prop X₃
           β†’ is-prop Xβ‚„
           β†’ is-prop (Xβ‚€ Γ— X₁ Γ— Xβ‚‚ Γ— X₃ Γ— Xβ‚„)
Γ—β‚…-is-prop iβ‚€ i₁ iβ‚‚ i₃ iβ‚„ =
 Γ—-is-prop iβ‚€ (Γ—β‚„-is-prop i₁ iβ‚‚ i₃ iβ‚„)

×₆-is-prop : {π“₯β‚€ π“₯₁ π“₯β‚‚ π“₯₃ π“₯β‚„ π“₯β‚… : Universe}
             {Xβ‚€ : π“₯β‚€ Μ‡ }
             {X₁ : π“₯₁ Μ‡ }
             {Xβ‚‚ : π“₯β‚‚ Μ‡ }
             {X₃ : π“₯₃ Μ‡ }
             {Xβ‚„ : π“₯β‚„ Μ‡ }
             {Xβ‚… : π“₯β‚… Μ‡ }
           β†’ is-prop Xβ‚€
           β†’ is-prop X₁
           β†’ is-prop Xβ‚‚
           β†’ is-prop X₃
           β†’ is-prop Xβ‚„
           β†’ is-prop Xβ‚…
           β†’ is-prop (Xβ‚€ Γ— X₁ Γ— Xβ‚‚ Γ— X₃ Γ— Xβ‚„ Γ— Xβ‚…)
×₆-is-prop iβ‚€ i₁ iβ‚‚ i₃ iβ‚„ iβ‚… =
 Γ—-is-prop iβ‚€ (Γ—β‚…-is-prop i₁ iβ‚‚ i₃ iβ‚„ iβ‚…)

×₇-is-prop : {π“₯β‚€ π“₯₁ π“₯β‚‚ π“₯₃ π“₯β‚„ π“₯β‚… π“₯₆ : Universe}
             {Xβ‚€ : π“₯β‚€ Μ‡ }
             {X₁ : π“₯₁ Μ‡ }
             {Xβ‚‚ : π“₯β‚‚ Μ‡ }
             {X₃ : π“₯₃ Μ‡ }
             {Xβ‚„ : π“₯β‚„ Μ‡ }
             {Xβ‚… : π“₯β‚… Μ‡ }
             {X₆ : π“₯₆ Μ‡ }
           β†’ is-prop Xβ‚€
           β†’ is-prop X₁
           β†’ is-prop Xβ‚‚
           β†’ is-prop X₃
           β†’ is-prop Xβ‚„
           β†’ is-prop Xβ‚…
           β†’ is-prop X₆
           β†’ is-prop (Xβ‚€ Γ— X₁ Γ— Xβ‚‚ Γ— X₃ Γ— Xβ‚„ Γ— Xβ‚… Γ— X₆)
×₇-is-prop iβ‚€ i₁ iβ‚‚ i₃ iβ‚„ iβ‚… i₆ =
 Γ—-is-prop iβ‚€ (×₆-is-prop i₁ iβ‚‚ i₃ iβ‚„ iβ‚… i₆)

Γ—β‚ˆ-is-prop : {π“₯β‚€ π“₯₁ π“₯β‚‚ π“₯₃ π“₯β‚„ π“₯β‚… π“₯₆ π“₯₇ : Universe}
             {Xβ‚€ : π“₯β‚€ Μ‡ }
             {X₁ : π“₯₁ Μ‡ }
             {Xβ‚‚ : π“₯β‚‚ Μ‡ }
             {X₃ : π“₯₃ Μ‡ }
             {Xβ‚„ : π“₯β‚„ Μ‡ }
             {Xβ‚… : π“₯β‚… Μ‡ }
             {X₆ : π“₯₆ Μ‡ }
             {X₇ : π“₯₇ Μ‡ }
           β†’ is-prop Xβ‚€
           β†’ is-prop X₁
           β†’ is-prop Xβ‚‚
           β†’ is-prop X₃
           β†’ is-prop Xβ‚„
           β†’ is-prop Xβ‚…
           β†’ is-prop X₆
           β†’ is-prop X₇ β†’ is-prop (Xβ‚€ Γ— X₁ Γ— Xβ‚‚ Γ— X₃ Γ— Xβ‚„ Γ— Xβ‚… Γ— X₆ Γ— X₇)
Γ—β‚ˆ-is-prop iβ‚€ i₁ iβ‚‚ i₃ iβ‚„ iβ‚… i₆ i₇ =
 Γ—-is-prop iβ‚€ (×₇-is-prop i₁ iβ‚‚ i₃ iβ‚„ iβ‚… i₆ i₇)

\end{code}