2014-07-19 19:09:47 +00: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-07-21 01:42:11 +00:00
|
|
|
import classical hilbert decidable
|
|
|
|
using decidable
|
2014-07-19 19:09:47 +00:00
|
|
|
|
|
|
|
-- Excluded middle + Hilbert implies every proposition is decidable
|
|
|
|
|
2014-07-21 01:42:11 +00:00
|
|
|
-- First, we show that (decidable a) is inhabited for any 'a' using the excluded middle
|
2014-07-22 16:43:18 +00:00
|
|
|
theorem inhabited_decidable [instance] (a : Prop) : inhabited (decidable a)
|
2014-07-19 19:09:47 +00:00
|
|
|
:= or_elim (em a)
|
2014-07-21 01:42:11 +00:00
|
|
|
(assume Ha, inhabited_intro (inl Ha))
|
|
|
|
(assume Hna, inhabited_intro (inr Hna))
|
2014-07-19 19:09:47 +00:00
|
|
|
|
2014-07-21 01:42:11 +00:00
|
|
|
-- Note that inhabited_decidable is marked as an instance, and it is silently used
|
|
|
|
-- for synthesizing the implicit argument in the following 'epsilon'
|
2014-07-22 16:53:05 +00:00
|
|
|
theorem prop_decidable [instance] (a : Prop) : decidable a
|
2014-07-21 01:42:11 +00:00
|
|
|
:= epsilon (λ d, true)
|