2014-07-05 05:22:26 +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-08-01 01:40:09 +00:00
|
|
|
|
|
2014-08-30 03:45:57 +00:00
|
|
|
|
-- logic.axioms.classical
|
|
|
|
|
-- ======================
|
|
|
|
|
|
2014-10-05 17:50:13 +00:00
|
|
|
|
import logic.quantifiers logic.cast algebra.relation
|
2014-08-01 01:40:09 +00:00
|
|
|
|
|
2014-10-02 00:51:17 +00:00
|
|
|
|
open eq.ops
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
axiom prop_complete (a : Prop) : a = true ∨ a = false
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-08-30 03:45:57 +00:00
|
|
|
|
theorem cases (P : Prop → Prop) (H1 : P true) (H2 : P false) (a : Prop) : P a :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim (prop_complete a)
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(assume Ht : a = true, Ht⁻¹ ▸ H1)
|
|
|
|
|
(assume Hf : a = false, Hf⁻¹ ▸ H2)
|
|
|
|
|
|
2014-08-30 03:45:57 +00:00
|
|
|
|
theorem cases_on (a : Prop) {P : Prop → Prop} (H1 : P true) (H2 : P false) : P a :=
|
|
|
|
|
cases P H1 H2 a
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-08-30 03:45:57 +00:00
|
|
|
|
-- this supercedes the em in decidable
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem em (a : Prop) : a ∨ ¬a :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim (prop_complete a)
|
|
|
|
|
(assume Ht : a = true, or.inl (eq_true_elim Ht))
|
|
|
|
|
(assume Hf : a = false, or.inr (eq_false_elim Hf))
|
2014-07-29 02:58:57 +00:00
|
|
|
|
|
|
|
|
|
theorem prop_complete_swapped (a : Prop) : a = false ∨ a = true :=
|
2014-08-30 03:45:57 +00:00
|
|
|
|
cases (λ x, x = false ∨ x = true)
|
2014-09-05 04:25:21 +00:00
|
|
|
|
(or.inr rfl)
|
|
|
|
|
(or.inl rfl)
|
2014-07-29 02:58:57 +00:00
|
|
|
|
a
|
|
|
|
|
|
|
|
|
|
theorem propext {a b : Prop} (Hab : a → b) (Hba : b → a) : a = b :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim (prop_complete a)
|
|
|
|
|
(assume Hat, or.elim (prop_complete b)
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(assume Hbt, Hat ⬝ Hbt⁻¹)
|
2014-09-02 02:44:04 +00:00
|
|
|
|
(assume Hbf, false_elim (Hbf ▸ (Hab (eq_true_elim Hat)))))
|
2014-09-05 04:25:21 +00:00
|
|
|
|
(assume Haf, or.elim (prop_complete b)
|
2014-09-02 02:44:04 +00:00
|
|
|
|
(assume Hbt, false_elim (Haf ▸ (Hba (eq_true_elim Hbt))))
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(assume Hbf, Haf ⬝ Hbf⁻¹))
|
|
|
|
|
|
|
|
|
|
theorem iff_to_eq {a b : Prop} (H : a ↔ b) : a = b :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.elim (assume H1 H2, propext H1 H2) H
|
2014-07-29 02:58:57 +00:00
|
|
|
|
|
|
|
|
|
theorem iff_eq_eq {a b : Prop} : (a ↔ b) = (a = b) :=
|
|
|
|
|
propext
|
|
|
|
|
(assume H, iff_to_eq H)
|
|
|
|
|
(assume H, eq_to_iff H)
|
|
|
|
|
|
2014-09-03 23:00:38 +00:00
|
|
|
|
open relation
|
2014-08-30 03:54:28 +00:00
|
|
|
|
theorem iff_congruence [instance] (P : Prop → Prop) : congruence iff iff P :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
congruence.mk
|
2014-08-20 02:32:44 +00:00
|
|
|
|
(take (a b : Prop),
|
|
|
|
|
assume H : a ↔ b,
|
2014-09-05 01:41:06 +00:00
|
|
|
|
show P a ↔ P b, from eq_to_iff (iff_to_eq H ▸ eq.refl (P a)))
|