lean2/library/logic/axioms/prop_decidable.lean
Leonardo de Moura dbaf81e16d refactor(library): remove unnecessary 'standard' subdirectory
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-08-23 18:08:09 -07:00

22 lines
1.1 KiB
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
----------------------------------------------------------------------------------------------------
import logic.axioms.classical logic.axioms.hilbert logic.classes.decidable
using decidable inhabited nonempty
-- Excluded middle + Hilbert implies every proposition is decidable
-- First, we show that (decidable a) is inhabited for any 'a' using the excluded middle
theorem decidable_inhabited [instance] (a : Prop) : inhabited (decidable a) :=
nonempty_imp_inhabited
(or_elim (em a)
(assume Ha, nonempty_intro (inl Ha))
(assume Hna, nonempty_intro (inr Hna)))
-- Note that decidable_inhabited is marked as an instance, and it is silently used
-- for synthesizing the implicit argument in the following 'epsilon'
theorem prop_decidable [instance] (a : Prop) : decidable a :=
epsilon (λd, true)