2014-12-22 14:54:01 -05:00
|
|
|
/-
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
Author: Leonardo de Moura
|
2015-04-08 19:28:13 -04:00
|
|
|
|
|
|
|
Excluded middle + Hilbert implies every proposition is decidable.
|
2014-12-22 14:54:01 -05:00
|
|
|
-/
|
2015-06-14 17:33:22 -07:00
|
|
|
import logic.axioms.prop_complete logic.axioms.hilbert data.sum
|
|
|
|
open decidable inhabited nonempty sum
|
2014-07-19 20:09:47 +01:00
|
|
|
|
2015-06-26 19:14:36 -07:00
|
|
|
theorem decidable_inhabited [instance] [priority 0] (a : Prop) : inhabited (decidable a) :=
|
2014-12-22 14:54:01 -05:00
|
|
|
inhabited_of_nonempty
|
2014-09-04 21:25:21 -07:00
|
|
|
(or.elim (em a)
|
2014-09-04 16:36:06 -07:00
|
|
|
(assume Ha, nonempty.intro (inl Ha))
|
|
|
|
(assume Hna, nonempty.intro (inr Hna)))
|
2014-07-19 20:09:47 +01:00
|
|
|
|
2015-06-26 19:14:36 -07:00
|
|
|
theorem prop_decidable [instance] [priority 0] (a : Prop) : decidable a :=
|
2015-04-08 19:28:13 -04:00
|
|
|
arbitrary (decidable a)
|
2015-06-14 17:33:22 -07:00
|
|
|
|
|
|
|
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
|