lean2/library/logic/axioms/prop_decidable.lean
Leonardo de Moura ff5022c2f4 feat(library/data/axioms): add 'type_decidable'
In Lean standard mode, Hilbert choice implies that all types are
decidable.
2015-06-14 17:33:22 -07:00

24 lines
847 B
Text

/-
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura
Excluded middle + Hilbert implies every proposition is decidable.
-/
import logic.axioms.prop_complete logic.axioms.hilbert data.sum
open decidable inhabited nonempty sum
theorem decidable_inhabited [instance] (a : Prop) : inhabited (decidable a) :=
inhabited_of_nonempty
(or.elim (em a)
(assume Ha, nonempty.intro (inl Ha))
(assume Hna, nonempty.intro (inr Hna)))
theorem prop_decidable [instance] (a : Prop) : decidable a :=
arbitrary (decidable a)
theorem type_decidable (A : Type) : A + (A → false) :=
match prop_decidable (nonempty A) with
| inl Hp := sum.inl (inhabited.value (inhabited_of_nonempty Hp))
| inr Hn := sum.inr (λ a, absurd (nonempty.intro a) Hn)
end