feat(library/logic/axioms): use diaconescu to prove em

With the new "noncomputable" feature we can use Hilbert's choice without
being concerned it may accidentaly "leak" inside definitions we don't
want to use it.
This commit is contained in:
Leonardo de Moura 2015-07-29 09:53:31 -07:00
parent 0bda39c8ac
commit 6dbcf86fd4
3 changed files with 49 additions and 64 deletions

View file

@ -3,9 +3,53 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura Author: Leonardo de Moura
Excluded middle Prove excluded middle using: excluded middle follows from Hilbert's choice operator, function extensionality,
and Prop extensionality
Remark: This axiom can be derived from propext funext and hilbert.
See examples/diaconescu
-/ -/
axiom em (a : Prop) : a ¬a import logic.axioms.hilbert logic.eq
open eq.ops
section
parameter p : Prop
private definition U (x : Prop) : Prop := x = true p
private definition V (x : Prop) : Prop := x = false p
private noncomputable definition u := epsilon U
private noncomputable definition v := epsilon V
private lemma u_def : U u :=
epsilon_spec (exists.intro true (or.inl rfl))
private lemma v_def : V v :=
epsilon_spec (exists.intro false (or.inl rfl))
private lemma not_uv_or_p : ¬(u = v) p :=
or.elim u_def
(assume Hut : u = true,
or.elim v_def
(assume Hvf : v = false,
have Hne : ¬(u = v), from Hvf⁻¹ ▸ Hut⁻¹ ▸ true_ne_false,
or.inl Hne)
(assume Hp : p, or.inr Hp))
(assume Hp : p, or.inr Hp)
private lemma p_implies_uv : p → u = v :=
assume Hp : p,
have Hpred : U = V, from
funext (take x : Prop,
have Hl : (x = true p) → (x = false p), from
assume A, or.inr Hp,
have Hr : (x = false p) → (x = true p), from
assume A, or.inr Hp,
show (x = true p) = (x = false p), from
propext (iff.intro Hl Hr)),
have H' : epsilon U = epsilon V, from Hpred ▸ rfl,
show u = v, from H'
theorem em : p ¬p :=
have H : ¬(u = v) → ¬p, from mt p_implies_uv,
or.elim not_uv_or_p
(assume Hne : ¬(u = v), or.inr (H Hne))
(assume Hp : p, or.inl Hp)
end

View file

@ -1,58 +0,0 @@
/-
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.hilbert logic.eq
open eq.ops
/-
Diaconescus theorem: excluded middle follows from Hilbert's choice operator, function extensionality,
and Prop extensionality
-/
section
parameter p : Prop
private definition U (x : Prop) : Prop := x = true p
private definition V (x : Prop) : Prop := x = false p
private noncomputable definition u := epsilon U
private noncomputable definition v := epsilon V
private lemma u_def : U u :=
epsilon_spec (exists.intro true (or.inl rfl))
private lemma v_def : V v :=
epsilon_spec (exists.intro false (or.inl rfl))
private lemma not_uv_or_p : ¬(u = v) p :=
or.elim u_def
(assume Hut : u = true,
or.elim v_def
(assume Hvf : v = false,
have Hne : ¬(u = v), from Hvf⁻¹ ▸ Hut⁻¹ ▸ true_ne_false,
or.inl Hne)
(assume Hp : p, or.inr Hp))
(assume Hp : p, or.inr Hp)
private lemma p_implies_uv : p → u = v :=
assume Hp : p,
have Hpred : U = V, from
funext (take x : Prop,
have Hl : (x = true p) → (x = false p), from
assume A, or.inr Hp,
have Hr : (x = false p) → (x = true p), from
assume A, or.inr Hp,
show (x = true p) = (x = false p), from
propext (iff.intro Hl Hr)),
have H' : epsilon U = epsilon V, from Hpred ▸ rfl,
show u = v, from H'
theorem em : p ¬p :=
have H : ¬(u = v) → ¬p, from mt p_implies_uv,
or.elim not_uv_or_p
(assume Hne : ¬(u = v), or.inr (H Hne))
(assume Hp : p, or.inl Hp)
end

View file

@ -1,4 +1,3 @@
quot.sound : ∀ {A : Type} [s : setoid A] {a b : A}, setoid.r a b → quot.mk a = quot.mk b quot.sound : ∀ {A : Type} [s : setoid A] {a b : A}, setoid.r a b → quot.mk a = quot.mk b
em : ∀ (a : Prop), a ¬a
propext : ∀ {a b : Prop}, (a ↔ b) → a = b propext : ∀ {a b : Prop}, (a ↔ b) → a = b
strong_indefinite_description : Π {A : Type} (P : A → Prop), nonempty A → { (x : A)| Exists P → P x} strong_indefinite_description : Π {A : Type} (P : A → Prop), nonempty A → { (x : A)| Exists P → P x}