Martin Escardo 2011.
\begin{code}
{-# OPTIONS --safe --without-K #-}
module NotionsOfDecidability.Decidable where
open import MLTT.Spartan
open import MLTT.Two-Properties
open import UF.Equiv
open import UF.PropTrunc
open import UF.Subsingletons
open import UF.SubtypeClassifier
open import UF.Logic
¬¬-elim : {A : π€ Μ } β is-decidable A β ¬¬ A β A
¬¬-elim (inl a) f = a
¬¬-elim (inr g) f = π-elim(f g)
map-decidable : {A : π€ Μ } {B : π₯ Μ }
              β (A β B)
              β (B β A)
              β is-decidable A
              β is-decidable B
map-decidable f g (inl x) = inl (f x)
map-decidable f g (inr h) = inr (Ξ» y β h (g y))
map-decidable-β : {A : π€ Μ } {B : π₯ Μ }
                β (A β B)
                β (is-decidable A β is-decidable B)
map-decidable-β (f , g) = map-decidable f g ,
                          map-decidable g f
decidability-is-closed-under-β : {A : π€ Μ } {B : π₯ Μ }
                               β (A β B)
                               β is-decidable A
                               β is-decidable B
decidability-is-closed-under-β (f , e) = map-decidable f (inverse f e)
map-decidable' : {A : π€ Μ } {B : π₯ Μ }
               β (A β Β¬ B)
               β (Β¬ A β B)
               β is-decidable A
               β is-decidable B
map-decidable' f g (inl x) = inr (f x)
map-decidable' f g (inr h) = inl (g h)
empty-is-decidable : {X : π€ Μ } β is-empty X β is-decidable X
empty-is-decidable = inr
π-is-decidable : is-decidable (π {π€})
π-is-decidable = empty-is-decidable π-elim
pointed-is-decidable : {X : π€ Μ } β X β is-decidable X
pointed-is-decidable = inl
π-is-decidable : is-decidable (π {π€})
π-is-decidable = pointed-is-decidable β
equivs-are-decidable : {X : π€ Μ } {Y : π₯ Μ } (π : X β Y)
                     β each-fiber-of β π β is-decidable
equivs-are-decidable π y = inl (β π ββ»ΒΉ y , inverses-are-sections' π y)
id-is-decidable : {X : π€ Μ } β each-fiber-of (id {π€} {X}) is-decidable
id-is-decidable x = inl (x , refl)
decidable-closed-under-Ξ£ : {X : π€ Μ } {Y : X β π₯ Μ }
                         β is-prop X
                         β is-decidable X
                         β ((x : X) β is-decidable (Y x))
                         β is-decidable (Ξ£ Y)
decidable-closed-under-Ξ£ {π€} {π₯} {X} {Y} isp d e = g d
 where
  g : is-decidable X β is-decidable (Ξ£ Y)
  g (inl x) = h (e x)
   where
    Ο : Ξ£ Y β Y x
    Ο (x' , y) = transport Y (isp x' x) y
    h : is-decidable(Y x) β is-decidable (Ξ£ Y)
    h (inl y) = inl (x , y)
    h (inr v) = inr (contrapositive Ο v)
  g (inr u) = inr (contrapositive prβ u)
Γ-preserves-decidability : {A : π€ Μ } {B : π₯ Μ }
                         β is-decidable A
                         β is-decidable B
                         β is-decidable (A Γ B)
Γ-preserves-decidability (inl a) (inl b) = inl (a , b)
Γ-preserves-decidability (inl a) (inr v) = inr (Ξ» c β v (prβ c))
Γ-preserves-decidability (inr u) _       = inr (Ξ» c β u (prβ c))
+-preserves-decidability : {A : π€ Μ } {B : π₯ Μ }
                         β is-decidable A
                         β is-decidable B
                         β is-decidable (A + B)
+-preserves-decidability (inl a) _       = inl (inl a)
+-preserves-decidability (inr u) (inl b) = inl (inr b)
+-preserves-decidability (inr u) (inr v) = inr (cases u v)
\end{code}
The following was added by Ayberk Tosun on 2024-05-28.
\begin{code}
module _ (pt : propositional-truncations-exist) where
 open Disjunction pt
 open PropositionalTruncation pt using (β£_β£; β₯β₯-rec)
 β¨-preserves-decidability : (P : Ξ© π€) (Q : Ξ© π₯)
                          β is-decidable (P holds)
                          β is-decidable (Q holds)
                          β is-decidable ((P β¨ Q) holds)
 β¨-preserves-decidability P Q Ο Ο =
  cases caseβ caseβ (+-preserves-decidability Ο Ο)
   where
    caseβ : P holds + Q holds β is-decidable ((P β¨ Q) holds)
    caseβ (inl p) = inl β£ inl p β£
    caseβ (inr q) = inl β£ inr q β£
    caseβ : Β¬ (P holds + Q holds) β is-decidable ((P β¨ Q) holds)
    caseβ = inr β β₯β₯-rec π-is-prop
\end{code}
End of addition.
\begin{code}
β-preserves-decidability : {A : π€ Μ } {B : π₯ Μ }
                         β is-decidable A
                         β is-decidable B
                         β is-decidable (A β B)
β-preserves-decidability d       (inl b) = inl (Ξ» _ β b)
β-preserves-decidability (inl a) (inr v) = inr (Ξ» f β v (f a))
β-preserves-decidability (inr u) (inr v) = inl (Ξ» a β π-elim (u a))
β-preserves-decidability' : {A : π€ Μ } {B : π₯ Μ }
                          β (Β¬ B β  is-decidable A)
                          β is-decidable B
                          β is-decidable (A β B)
β-preserves-decidability' Ο (inl b) = inl (Ξ» _ β b)
β-preserves-decidability' {π€} {π₯} {A} {B} Ο (inr v) = Ξ³ (Ο v)
 where
  Ξ³ : is-decidable A β is-decidable (A β B)
  Ξ³ (inl a) = inr (Ξ» f β v (f a))
  Ξ³ (inr u) = inl (Ξ» a β π-elim (u a))
β-preserves-decidability'' : {A : π€ Μ } {B : π₯ Μ }
                           β is-decidable A
                           β (A β is-decidable B)
                           β is-decidable (A β B)
β-preserves-decidability'' {π€} {π₯} {A} {B} (inl a) Ο = Ξ³ (Ο a)
 where
  Ξ³ : is-decidable B β is-decidable (A β B)
  Ξ³ (inl b) = inl (Ξ» _ β b)
  Ξ³ (inr v) = inr (Ξ» f β v (f a))
β-preserves-decidability'' (inr u) Ο = inl (Ξ» a β π-elim (u a))
Β¬-preserves-decidability : {A : π€ Μ }
                         β is-decidable A
                         β is-decidable(Β¬ A)
Β¬-preserves-decidability d = β-preserves-decidability d π-is-decidable
which-of : {A : π€ Μ } {B : π₯ Μ }
         β A + B
         β Ξ£ b κ π , (b οΌ β β A)
                   Γ (b οΌ β β B)
which-of (inl a) = β ,
                   (Ξ» (r : β οΌ β) β a) ,
                   (Ξ» (p : β οΌ β) β π-elim (zero-is-not-one p))
which-of (inr b) = β ,
                   (Ξ» (p : β οΌ β) β π-elim (zero-is-not-one (p β»ΒΉ))) ,
                   (Ξ» (r : β οΌ β) β b)
\end{code}
The following is a special case we are interested in:
\begin{code}
boolean-value : {A : π€ Μ }
              β is-decidable A
              β Ξ£ b κ π , (b οΌ β β   A)
                        Γ (b οΌ β β Β¬ A)
boolean-value = which-of
module _ {X : π€ Μ } {Aβ : X β π₯ Μ } {Aβ : X β π¦ Μ }
         (h : (x : X) β Aβ x + Aβ x)
       where
 indicator : Ξ£ p κ (X β π) , ((x : X) β (p x οΌ β β Aβ x)
                                      Γ (p x οΌ β β Aβ x))
 indicator = (Ξ» x β prβ(lemmaβ x)) , (Ξ» x β prβ(lemmaβ x))
  where
   lemmaβ : (x : X) β (Aβ x + Aβ x) β Ξ£ b κ π , (b οΌ β β Aβ x)
                                              Γ (b οΌ β β Aβ x)
   lemmaβ x = which-of
   lemmaβ : (x : X) β Ξ£ b κ π , (b οΌ β β Aβ x) Γ (b οΌ β β Aβ x)
   lemmaβ x = lemmaβ x (h x)
 indicator-map : X β π
 indicator-map = prβ indicator
 indicator-property : (x : X) β (indicator-map x οΌ β β Aβ x)
                              Γ (indicator-map x οΌ β β Aβ x)
 indicator-property = prβ indicator
 indicator-propertyβ : (x : X) β indicator-map x οΌ β β Aβ x
 indicator-propertyβ x = prβ (indicator-property x)
 indicator-propertyβ : (x : X) β indicator-map x οΌ β β Aβ x
 indicator-propertyβ x = prβ (indicator-property x)
module _ {X : π€ Μ } (A : X β π₯ Μ )
         (Ξ΄ : (x : X) β A x + Β¬ A x)
       where
 private
  f : (x : X) β is-decidable (A x) β π
  f x (inl a) = β
  f x (inr Ξ½) = β
  fβ : (x : X) (d : is-decidable (A x)) β f x d οΌ β β A x
  fβ x (inl a) e = a
  fβ x (inr Ξ½) e = π-elim (one-is-not-zero e)
  fβ : (x : X) (d : is-decidable (A x)) β f x d οΌ β β Β¬ A x
  fβ x (inl a) e = π-elim (zero-is-not-one e)
  fβ x (inr Ξ½) e = Ξ½
  fβ-back : (x : X) (d : is-decidable (A x)) β A x β f x d οΌ β
  fβ-back x (inl a) a' = refl
  fβ-back x (inr Ξ½) a' = π-elim (Ξ½ a')
  fβ-back : (x : X) (d : is-decidable (A x)) β Β¬ A x β f x d οΌ β
  fβ-back x (inl a) Ξ½' = π-elim (Ξ½' a)
  fβ-back x (inr Ξ½) Ξ½' = refl
  Ο : X β π
  Ο x = f x (Ξ΄ x)
 characteristic-map : X β π
 characteristic-map = Ο
 characteristic-map-propertyβ : (x : X) β Ο x οΌ β β A x
 characteristic-map-propertyβ x = fβ x (Ξ΄ x)
 characteristic-map-propertyβ : (x : X) β Ο x οΌ β β Β¬ A x
 characteristic-map-propertyβ x = fβ x (Ξ΄ x)
 characteristic-map-propertyβ-back : (x : X) β A x β Ο x οΌ β
 characteristic-map-propertyβ-back x = fβ-back x (Ξ΄ x)
 characteristic-map-propertyβ-back : (x : X) β Β¬ A x β Ο x οΌ β
 characteristic-map-propertyβ-back x = fβ-back x (Ξ΄ x)
\end{code}
Added by Tom de Jong, November 2021.
\begin{code}
decidable-β : {X : π€ Μ } {Y : π₯ Μ }
            β X β Y
            β is-decidable X
            β is-decidable Y
decidable-β {π€} {π₯} {X} {Y} (f , g) (inl  x) = inl (f x)
decidable-β {π€} {π₯} {X} {Y} (f , g) (inr nx) = inr (nx β g)
decidable-cong : {X : π€ Μ } {Y : π₯ Μ }
               β X β Y
               β is-decidable X
               β is-decidable Y
decidable-cong e = decidable-β (β e β , β e ββ»ΒΉ)
\end{code}
Added by Tom de Jong in January 2022.
\begin{code}
all-types-are-¬¬-decidable : (X : π€ Μ ) β ¬¬ (is-decidable X)
all-types-are-¬¬-decidable X h = claimβ claimβ
 where
  claimβ : Β¬ X
  claimβ x = h (inl x)
  claimβ : ¬¬ X
  claimβ nx = h (inr nx)
¬¬-stable-if-decidable : (X : π€ Μ ) β is-decidable X β ¬¬-stable X
¬¬-stable-if-decidable X = ¬¬-elim
\end{code}
Added 21th August 2024 by Alice Laroche.
\begin{code}
module _ (pt : propositional-truncations-exist) where
 open PropositionalTruncation pt
 decidable-inhabited-types-are-pointed : {X : π€ Μ } β β₯ X β₯ β is-decidable X β X
 decidable-inhabited-types-are-pointed β£xβ£ (inl x)  = x
 decidable-inhabited-types-are-pointed β£xβ£ (inr Β¬x) =
  π-elim (β₯β₯-rec π-is-prop Β¬x β£xβ£)
\end{code}
End of addition.
Added by Martin Escardo 17th September 2024. The propositional
truncation of a decidable type can be constructed with no assumptions
and it has split support.
\begin{code}
β₯_β₯β¨_β© : (X : π€ Μ ) β is-decidable X β π€β Μ
β₯ X β₯β¨ inl x β© = π
β₯ X β₯β¨ inr Ξ½ β© = π
β₯β₯β¨_β©-is-prop : {X : π€ Μ } (Ξ΄ : is-decidable X) β is-prop β₯ X β₯β¨ Ξ΄ β©
β₯β₯β¨ inl x β©-is-prop = π-is-prop
β₯β₯β¨ inr Ξ½ β©-is-prop = π-is-prop
β₯β₯β¨_β©-is-decidable : {X : π€ Μ } (Ξ΄ : is-decidable X) β is-decidable β₯ X β₯β¨ Ξ΄ β©
β₯β₯β¨ inl x β©-is-decidable = π-is-decidable
β₯β₯β¨ inr Ξ½ β©-is-decidable = π-is-decidable
β£_β£β¨_β© : {X : π€ Μ } β X β (Ξ΄ : is-decidable X) β β₯ X β₯β¨ Ξ΄ β©
β£ x β£β¨ inl _ β© = β
β£ x β£β¨ inr Ξ½ β© = Ξ½ x
\end{code}
Notice that the induction principle doesn't require the family A to be
prop-valued.
\begin{code}
β₯β₯β¨_β©-induction : {X : π€ Μ } (Ξ΄ : is-decidable X)
                 (A : β₯ X β₯β¨ Ξ΄ β© β π₯ Μ )
               β ((x : X) β A β£ x β£β¨ Ξ΄ β©)
               β (s : β₯ X β₯β¨ Ξ΄ β©) β A s
β₯β₯β¨ inl x β©-induction A f β = f x
β₯β₯β¨ inr Ξ½ β©-induction A f s = π-elim s
\end{code}
But the induction equation does.
\begin{code}
β₯β₯β¨_β©-induction-equation : {X : π€ Μ }
                          (Ξ΄ : is-decidable X)
                          (A : β₯ X β₯β¨ Ξ΄ β© β π₯ Μ )
                        β ((s : β₯ X β₯β¨ Ξ΄ β©) β is-prop (A s))
                        β (f : (x : X) β A β£ x β£β¨ Ξ΄ β©)
                          (x : X)
                        β β₯β₯β¨ Ξ΄ β©-induction A f β£ x β£β¨ Ξ΄ β© οΌ f x
β₯β₯β¨ inl x β©-induction-equation A A-is-prop f x' = A-is-prop β (f x) (f x')
β₯β₯β¨ inr Ξ½ β©-induction-equation A A-is-prop f x  = π-elim (Ξ½ x)
β₯β₯β¨_β©-rec : {X : π€ Μ } (Ξ΄ : is-decidable X) {A : π₯ Μ }
          β (X β A) β β₯ X β₯β¨ Ξ΄ β© β A
β₯β₯β¨ Ξ΄ β©-rec {A} = β₯β₯β¨ Ξ΄ β©-induction (Ξ» _ β A)
β£β£β¨_β©-exit : {X : π€ Μ } (Ξ΄ : is-decidable X) β β₯ X β₯β¨ Ξ΄ β© β X
β£β£β¨ Ξ΄ β©-exit = β₯β₯β¨ Ξ΄ β©-rec id
β£β£β¨_β©-exit-is-section : {X : π€ Μ } (Ξ΄ : is-decidable X) (s : β₯ X β₯β¨ Ξ΄ β©)
                     β β£ β£β£β¨ Ξ΄ β©-exit s β£β¨ Ξ΄ β© οΌ s
β£β£β¨ inl x β©-exit-is-section β = refl
β£β£β¨ inr Ξ½ β©-exit-is-section s = π-elim s
infix 0 β₯_β₯β¨_β©
infix 0 β£_β£β¨_β©
module propositional-truncation-of-decidable-type
        (pt : propositional-truncations-exist)
       where
 open propositional-truncations-exist pt public
 module _ {X : π€ Μ } (Ξ΄ : is-decidable X) where
  β₯β₯β¨_β©-to-β₯β₯ : β₯ X β₯β¨ Ξ΄ β© β β₯ X β₯
  β₯β₯β¨_β©-to-β₯β₯ = β₯β₯β¨ Ξ΄ β©-rec β£_β£
  β₯β₯-to-β₯β₯β¨_β© : β₯ X β₯ β β₯ X β₯β¨ Ξ΄ β©
  β₯β₯-to-β₯β₯β¨_β© = β₯β₯-rec (β₯β₯β¨ Ξ΄ β©-is-prop) β£_β£β¨ Ξ΄ β©
  decidable-types-have-split-support : β₯ X β₯ β X
  decidable-types-have-split-support s = β£β£β¨ Ξ΄ β©-exit (β₯β₯-to-β₯β₯β¨_β© s)
\end{code}
Added by Fredrik Bakke 22nd August 2025. Negations of decidable types are
decidable.
\begin{code}
decidable-types-are-closed-under-negations : {X : π€ Μ }
                                           β is-decidable X
                                           β is-decidable (Β¬ X)
decidable-types-are-closed-under-negations (inl x) = inr (Ξ» nx β nx x)
decidable-types-are-closed-under-negations (inr nx) = inl nx
\end{code}