2014-07-19 20:09:47 +01: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
|
2014-08-29 23:45:57 -04:00
|
|
|
|
|
|
|
-- logic.axioms.prop_decidable
|
|
|
|
-- ===========================
|
2014-07-31 18:40:09 -07:00
|
|
|
|
2014-11-30 20:34:12 -08:00
|
|
|
import logic.axioms.classical logic.axioms.hilbert
|
2014-09-03 16:00:38 -07:00
|
|
|
open decidable inhabited nonempty
|
2014-07-19 20:09:47 +01:00
|
|
|
|
|
|
|
-- Excluded middle + Hilbert implies every proposition is decidable
|
|
|
|
|
2014-07-21 02:42:11 +01:00
|
|
|
-- First, we show that (decidable a) is inhabited for any 'a' using the excluded middle
|
2014-08-20 15:49:44 -07:00
|
|
|
theorem decidable_inhabited [instance] (a : Prop) : inhabited (decidable a) :=
|
2014-08-14 20:12:54 -07:00
|
|
|
nonempty_imp_inhabited
|
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
|
|
|
|
2014-08-20 15:49:44 -07:00
|
|
|
-- Note that decidable_inhabited is marked as an instance, and it is silently used
|
2014-07-21 02:42:11 +01:00
|
|
|
-- for synthesizing the implicit argument in the following 'epsilon'
|
2014-07-28 19:58:57 -07:00
|
|
|
theorem prop_decidable [instance] (a : Prop) : decidable a :=
|
2014-08-14 20:12:54 -07:00
|
|
|
epsilon (λd, true)
|