2014-02-04 22:15:33 +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-01-05 20:05:08 +00:00
|
|
|
|
import macros
|
2014-01-31 03:11:58 +00:00
|
|
|
|
import tactic
|
2013-12-29 21:03:32 +00:00
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
universe U ≥ 1
|
2014-02-02 02:27:14 +00:00
|
|
|
|
definition TypeU := (Type U)
|
|
|
|
|
|
|
|
|
|
-- create default rewrite rule set
|
|
|
|
|
(* mk_rewrite_rule_set() *)
|
2013-12-30 05:59:57 +00:00
|
|
|
|
|
2014-01-05 20:05:08 +00:00
|
|
|
|
variable Bool : Type
|
2014-02-02 02:27:14 +00:00
|
|
|
|
-- Heterogeneous equality
|
refactor(kernel): add heterogeneous equality back to expr
The main motivation is that we will be able to move equalities between universes.
For example, suppose we have
A : (Type i)
B : (Type i)
H : @eq (Type j) A B
where j > i
We didn't find any trick for deducing (@eq (Type i) A B) from H.
Before this commit, heterogeneous equality as a constant with type
heq : {A B : (Type U)} : A -> B -> Bool
So, from H, we would only be able to deduce
(@heq (Type j) (Type j) A B)
Not being able to move the equality back to a smaller universe is
problematic in several cases. I list some instances in the end of the commit message.
With this commit, Heterogeneous equality is a special kind of expression.
It is not a constant anymore. From H, we can deduce
H1 : A == B
That is, we are essentially "erasing" the universes when we move to heterogeneous equality.
Now, since A and B have (Type i), we can deduce (@eq (Type i) A B) from H1. The proof term is
(to_eq (Type i) A B (to_heq (Type j) A B H)) : (@eq (Type i) A B)
So, it remains to explain why we need this feature.
For example, suppose we want to state the Pi extensionality axiom.
axiom hpiext {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)} :
A = A' → (∀ x x', x == x' → B x == B' x') → (∀ x, B x) == (∀ x, B' x)
This axiom produces an "inflated" equality at (Type U) when we treat heterogeneous
equality as a constant. The conclusion
(∀ x, B x) == (∀ x, B' x)
is syntax sugar for
(@heq (Type U) (Type U) (∀ x : A, B x) (∀ x : A', B' x))
Even if A, A', B, B' live in a much smaller universe.
As I described above, it doesn't seem to be a way to move this equality back to a smaller universe.
So, if we wanted to keep the heterogeneous equality as a constant, it seems we would
have to support axiom schemas. That is, hpiext would be parametrized by the universes where
A, A', B and B'. Another possibility would be to have universe polymorphism like Agda.
None of the solutions seem attractive.
So, we decided to have heterogeneous equality as a special kind of expression.
And use the trick above to move equalities back to the right universe.
BTW, the parser is not creating the new heterogeneous equalities yet.
Moreover, kernel.lean still contains a constant name heq2 that is the heterogeneous
equality as a constant.
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-02-07 18:07:08 +00:00
|
|
|
|
variable heq2 {A B : (Type U)} (a : A) (b : B) : Bool
|
|
|
|
|
infixl 50 == : heq2
|
2014-02-02 02:27:14 +00:00
|
|
|
|
|
|
|
|
|
-- Reflexivity for heterogeneous equality
|
|
|
|
|
axiom hrefl {A : (Type U)} (a : A) : a == a
|
|
|
|
|
|
|
|
|
|
-- Homogeneous equality
|
|
|
|
|
definition eq {A : (Type U)} (a b : A) := a == b
|
|
|
|
|
infix 50 = : eq
|
|
|
|
|
|
|
|
|
|
theorem refl {A : (Type U)} (a : A) : a = a
|
|
|
|
|
:= hrefl a
|
|
|
|
|
|
|
|
|
|
theorem heq_eq {A : (Type U)} (a b : A) : (a == b) = (a = b)
|
|
|
|
|
:= refl (a == b)
|
|
|
|
|
|
|
|
|
|
definition true : Bool
|
|
|
|
|
:= (λ x : Bool, x) = (λ x : Bool, x)
|
|
|
|
|
|
|
|
|
|
theorem trivial : true
|
|
|
|
|
:= refl (λ x : Bool, x)
|
|
|
|
|
|
|
|
|
|
set_opaque true true
|
|
|
|
|
|
|
|
|
|
definition false : Bool
|
|
|
|
|
:= ∀ x : Bool, x
|
|
|
|
|
|
|
|
|
|
set_opaque false true
|
|
|
|
|
alias ⊤ : true
|
|
|
|
|
alias ⊥ : false
|
2013-12-30 05:59:57 +00:00
|
|
|
|
|
2014-01-09 00:52:43 +00:00
|
|
|
|
definition not (a : Bool) := a → false
|
2014-01-05 20:05:08 +00:00
|
|
|
|
notation 40 ¬ _ : not
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-01-09 00:52:43 +00:00
|
|
|
|
definition or (a b : Bool) := ¬ a → b
|
2014-01-05 20:05:08 +00:00
|
|
|
|
infixr 30 || : or
|
|
|
|
|
infixr 30 \/ : or
|
|
|
|
|
infixr 30 ∨ : or
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-01-09 00:52:43 +00:00
|
|
|
|
definition and (a b : Bool) := ¬ (a → ¬ b)
|
2014-01-05 20:05:08 +00:00
|
|
|
|
infixr 35 && : and
|
|
|
|
|
infixr 35 /\ : and
|
|
|
|
|
infixr 35 ∧ : and
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
definition implies (a b : Bool) := a → b
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
definition neq {A : (Type U)} (a b : A) := ¬ (a = b)
|
2014-01-16 23:07:51 +00:00
|
|
|
|
infix 50 ≠ : neq
|
2014-01-16 10:05:09 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem a_neq_a_elim {A : (Type U)} {a : A} (H : a ≠ a) : false
|
|
|
|
|
:= H (refl a)
|
|
|
|
|
|
2014-01-16 23:07:51 +00:00
|
|
|
|
definition iff (a b : Bool) := a = b
|
2014-01-16 10:05:09 +00:00
|
|
|
|
infixr 25 <-> : iff
|
2014-01-16 17:00:01 +00:00
|
|
|
|
infixr 25 ↔ : iff
|
2014-01-16 10:05:09 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem em (a : Bool) : a ∨ ¬ a
|
|
|
|
|
:= assume Hna : ¬ a, Hna
|
|
|
|
|
|
|
|
|
|
theorem not_intro {a : Bool} (H : a → false) : ¬ a
|
|
|
|
|
:= H
|
|
|
|
|
|
|
|
|
|
theorem absurd {a : Bool} (H1 : a) (H2 : ¬ a) : false
|
|
|
|
|
:= H2 H1
|
|
|
|
|
|
2014-01-09 00:52:43 +00:00
|
|
|
|
-- The Lean parser has special treatment for the constant exists.
|
|
|
|
|
-- It allows us to write
|
|
|
|
|
-- exists x y : A, P x y and ∃ x y : A, P x y
|
|
|
|
|
-- as syntax sugar for
|
|
|
|
|
-- exists A (fun x : A, exists A (fun y : A, P x y))
|
|
|
|
|
-- That is, it treats the exists as an extra binder such as fun and forall.
|
|
|
|
|
-- It also provides an alias (Exists) that should be used when we
|
|
|
|
|
-- want to treat exists as a constant.
|
2014-02-02 02:27:14 +00:00
|
|
|
|
definition Exists (A : (Type U)) (P : A → Bool)
|
|
|
|
|
:= ¬ (∀ x, ¬ (P x))
|
2014-01-22 16:52:31 +00:00
|
|
|
|
|
2014-02-06 01:33:06 +00:00
|
|
|
|
definition exists_unique {A : (Type U)} (p : A → Bool)
|
|
|
|
|
:= ∃ x, p x ∧ ∀ y, y ≠ x → ¬ p y
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
axiom case (P : Bool → Bool) (H1 : P true) (H2 : P false) (a : Bool) : P a
|
2014-01-22 16:52:31 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem false_elim (a : Bool) (H : false) : a
|
|
|
|
|
:= case (λ x, x) trivial H a
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem mt {a b : Bool} (H1 : a → b) (H2 : ¬ b) : ¬ a
|
|
|
|
|
:= assume Ha : a, absurd (H1 Ha) H2
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem contrapos {a b : Bool} (H : a → b) : ¬ b → ¬ a
|
|
|
|
|
:= assume Hnb : ¬ b, mt H Hnb
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem absurd_elim {a : Bool} (b : Bool) (H1 : a) (H2 : ¬ a) : b
|
|
|
|
|
:= false_elim b (absurd H1 H2)
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
-- Recall that or is defined as ¬ a → b
|
|
|
|
|
theorem or_introl {a : Bool} (H : a) (b : Bool) : a ∨ b
|
|
|
|
|
:= assume H1 : ¬ a, absurd_elim b H H1
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem or_intror {b : Bool} (a : Bool) (H : b) : a ∨ b
|
|
|
|
|
:= assume H1 : ¬ a, H
|
2014-01-08 08:38:39 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem boolcomplete (a : Bool) : a = true ∨ a = false
|
|
|
|
|
:= case (λ x, x = true ∨ x = false)
|
|
|
|
|
(or_introl (refl true) (true = false))
|
|
|
|
|
(or_intror (false = true) (refl false))
|
|
|
|
|
a
|
2014-01-09 00:19:11 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem boolcomplete_swapped (a : Bool) : a = false ∨ a = true
|
|
|
|
|
:= case (λ x, x = false ∨ x = true)
|
|
|
|
|
(or_intror (true = false) (refl true))
|
|
|
|
|
(or_introl (refl false) (false = true))
|
|
|
|
|
a
|
2014-01-22 16:21:11 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem resolve1 {a b : Bool} (H1 : a ∨ b) (H2 : ¬ a) : b
|
|
|
|
|
:= H1 H2
|
|
|
|
|
|
|
|
|
|
axiom subst {A : (Type U)} {a b : A} {P : A → Bool} (H1 : P a) (H2 : a = b) : P b
|
2014-01-19 20:20:09 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
-- Alias for subst where we provide P explicitly, but keep A,a,b implicit
|
|
|
|
|
theorem substp {A : (Type U)} {a b : A} (P : A → Bool) (H1 : P a) (H2 : a = b) : P b
|
2014-01-09 00:19:11 +00:00
|
|
|
|
:= subst H1 H2
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem symm {A : (Type U)} {a b : A} (H : a = b) : b = a
|
|
|
|
|
:= subst (refl a) H
|
2014-01-31 03:11:58 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem trans {A : (Type U)} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c
|
|
|
|
|
:= subst H1 H2
|
2014-01-31 03:11:58 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem congr1 {A B : (Type U)} {f g : A → B} (a : A) (H : f = g) : f a = g a
|
|
|
|
|
:= substp (fun h : A → B, f a = h a) (refl (f a)) H
|
2014-01-31 03:38:51 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem congr2 {A B : (Type U)} {a b : A} (f : A → B) (H : a = b) : f a = f b
|
|
|
|
|
:= substp (fun x : A, f a = f x) (refl (f a)) H
|
2014-01-31 03:11:58 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem congr {A B : (Type U)} {f g : A → B} {a b : A} (H1 : f = g) (H2 : a = b) : f a = g b
|
|
|
|
|
:= subst (congr2 f H2) (congr1 b H1)
|
2014-01-10 23:20:55 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem true_ne_false : ¬ true = false
|
|
|
|
|
:= assume H : true = false,
|
|
|
|
|
subst trivial H
|
2014-01-09 01:25:14 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem absurd_not_true (H : ¬ true) : false
|
|
|
|
|
:= absurd trivial H
|
2014-01-19 19:24:20 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem not_false_trivial : ¬ false
|
|
|
|
|
:= assume H : false, H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
-- "equality modus pones"
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem eqmp {a b : Bool} (H1 : a = b) (H2 : a) : b
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= subst H2 H1
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-08 08:38:39 +00:00
|
|
|
|
infixl 100 <| : eqmp
|
|
|
|
|
infixl 100 ◂ : eqmp
|
2014-01-03 18:33:57 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem eqmpr {a b : Bool} (H1 : a = b) (H2 : b) : a
|
|
|
|
|
:= (symm H1) ◂ H2
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem imp_trans {a b c : Bool} (H1 : a → b) (H2 : b → c) : a → c
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Ha, H2 (H1 Ha)
|
2014-01-02 18:53:14 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem imp_eq_trans {a b c : Bool} (H1 : a → b) (H2 : b = c) : a → c
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Ha, H2 ◂ (H1 Ha)
|
2014-01-02 18:53:14 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem eq_imp_trans {a b c : Bool} (H1 : a = b) (H2 : b → c) : a → c
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Ha, H2 (H1 ◂ Ha)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem to_eq {A : TypeU} {a b : A} (H : a == b) : a = b
|
|
|
|
|
:= (heq_eq a b) ◂ H
|
|
|
|
|
|
|
|
|
|
theorem to_heq {A : TypeU} {a b : A} (H : a = b) : a == b
|
|
|
|
|
:= (symm (heq_eq a b)) ◂ H
|
|
|
|
|
|
|
|
|
|
theorem iff_eliml {a b : Bool} (H : a ↔ b) : a → b
|
|
|
|
|
:= (λ Ha : a, eqmp H Ha)
|
|
|
|
|
|
|
|
|
|
theorem iff_elimr {a b : Bool} (H : a ↔ b) : b → a
|
|
|
|
|
:= (λ Hb : b, eqmpr H Hb)
|
|
|
|
|
|
|
|
|
|
theorem ne_symm {A : TypeU} {a b : A} (H : a ≠ b) : b ≠ a
|
|
|
|
|
:= assume H1 : b = a, H (symm H1)
|
|
|
|
|
|
|
|
|
|
theorem eq_ne_trans {A : TypeU} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c
|
|
|
|
|
:= subst H2 (symm H1)
|
|
|
|
|
|
|
|
|
|
theorem ne_eq_trans {A : TypeU} {a b c : A} (H1 : a ≠ b) (H2 : b = c) : a ≠ c
|
|
|
|
|
:= subst H1 H2
|
|
|
|
|
|
|
|
|
|
theorem eqt_elim {a : Bool} (H : a = true) : a
|
|
|
|
|
:= (symm H) ◂ trivial
|
|
|
|
|
|
|
|
|
|
theorem eqf_elim {a : Bool} (H : a = false) : ¬ a
|
|
|
|
|
:= not_intro (λ Ha : a, H ◂ Ha)
|
|
|
|
|
|
|
|
|
|
theorem heqt_elim {a : Bool} (H : a == true) : a
|
|
|
|
|
:= eqt_elim (to_eq H)
|
|
|
|
|
|
|
|
|
|
theorem not_true : (¬ true) = false
|
|
|
|
|
:= let aux : ¬ (¬ true) = true
|
|
|
|
|
:= assume H : (¬ true) = true,
|
|
|
|
|
absurd_not_true (subst trivial (symm H))
|
|
|
|
|
in resolve1 (boolcomplete (¬ true)) aux
|
|
|
|
|
|
|
|
|
|
theorem not_false : (¬ false) = true
|
|
|
|
|
:= let aux : ¬ (¬ false) = false
|
|
|
|
|
:= assume H : (¬ false) = false,
|
|
|
|
|
subst not_false_trivial H
|
|
|
|
|
in resolve1 (boolcomplete_swapped (¬ false)) aux
|
|
|
|
|
|
|
|
|
|
add_rewrite not_true not_false
|
|
|
|
|
|
|
|
|
|
theorem not_not_eq (a : Bool) : (¬ ¬ a) = a
|
|
|
|
|
:= case (λ x, (¬ ¬ x) = x)
|
|
|
|
|
(calc (¬ ¬ true) = (¬ false) : { not_true }
|
|
|
|
|
... = true : not_false)
|
|
|
|
|
(calc (¬ ¬ false) = (¬ true) : { not_false }
|
|
|
|
|
... = false : not_true)
|
|
|
|
|
a
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
add_rewrite not_not_eq
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem not_neq {A : TypeU} (a b : A) : ¬ (a ≠ b) ↔ a = b
|
|
|
|
|
:= not_not_eq (a = b)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
add_rewrite not_neq
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem not_neq_elim {A : TypeU} {a b : A} (H : ¬ (a ≠ b)) : a = b
|
|
|
|
|
:= (not_neq a b) ◂ H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem not_not_elim {a : Bool} (H : ¬ ¬ a) : a
|
|
|
|
|
:= (not_not_eq a) ◂ H
|
2014-01-06 03:10:21 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_imp_eliml {a b : Bool} (Hnab : ¬ (a → b)) : a
|
|
|
|
|
:= not_not_elim
|
2014-02-06 15:50:22 +00:00
|
|
|
|
(show ¬ ¬ a,
|
|
|
|
|
from assume Hna : ¬ a, absurd (assume Ha : a, absurd_elim b Ha Hna)
|
|
|
|
|
Hnab)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_imp_elimr {a b : Bool} (H : ¬ (a → b)) : ¬ b
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Hb : b, absurd (assume Ha : a, Hb)
|
|
|
|
|
H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 00:52:43 +00:00
|
|
|
|
-- Recall that and is defined as ¬ (a → ¬ b)
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem and_intro {a b : Bool} (H1 : a) (H2 : b) : a ∧ b
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume H : a → ¬ b, absurd H2 (H H1)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem and_eliml {a b : Bool} (H : a ∧ b) : a
|
|
|
|
|
:= not_imp_eliml H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem and_elimr {a b : Bool} (H : a ∧ b) : b
|
|
|
|
|
:= not_not_elim (not_imp_elimr H)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem or_elim {a b c : Bool} (H1 : a ∨ b) (H2 : a → c) (H3 : b → c) : c
|
|
|
|
|
:= not_not_elim
|
2014-01-12 02:36:37 +00:00
|
|
|
|
(assume H : ¬ c,
|
2014-02-06 15:50:22 +00:00
|
|
|
|
absurd (H3 (resolve1 H1 (mt (assume Ha : a, H2 Ha) H)))
|
2013-12-29 10:44:49 +00:00
|
|
|
|
H)
|
|
|
|
|
|
2014-01-06 03:10:21 +00:00
|
|
|
|
theorem refute {a : Bool} (H : ¬ a → false) : a
|
2014-01-09 16:33:52 +00:00
|
|
|
|
:= or_elim (em a) (λ H1 : a, H1) (λ H1 : ¬ a, false_elim a (H H1))
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem boolext {a b : Bool} (Hab : a → b) (Hba : b → a) : a = b
|
2014-01-09 16:33:52 +00:00
|
|
|
|
:= or_elim (boolcomplete a)
|
2014-01-15 22:42:00 +00:00
|
|
|
|
(λ Hat : a = true, or_elim (boolcomplete b)
|
|
|
|
|
(λ Hbt : b = true, trans Hat (symm Hbt))
|
|
|
|
|
(λ Hbf : b = false, false_elim (a = b) (subst (Hab (eqt_elim Hat)) Hbf)))
|
|
|
|
|
(λ Haf : a = false, or_elim (boolcomplete b)
|
|
|
|
|
(λ Hbt : b = true, false_elim (a = b) (subst (Hba (eqt_elim Hbt)) Haf))
|
|
|
|
|
(λ Hbf : b = false, trans Haf (symm Hbf)))
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
-- Another name for boolext
|
2014-01-17 18:14:57 +00:00
|
|
|
|
theorem iff_intro {a b : Bool} (Hab : a → b) (Hba : b → a) : a ↔ b
|
2014-01-09 17:00:05 +00:00
|
|
|
|
:= boolext Hab Hba
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem eqt_intro {a : Bool} (H : a) : a = true
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H1 : a, trivial)
|
|
|
|
|
(assume H2 : true, H)
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem eqf_intro {a : Bool} (H : ¬ a) : a = false
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H1 : a, absurd H1 H)
|
|
|
|
|
(assume H2 : false, false_elim a H2)
|
2014-01-11 19:13:54 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem a_neq_a {A : (Type U)} (a : A) : (a ≠ a) ↔ false
|
|
|
|
|
:= boolext (assume H, a_neq_a_elim H)
|
|
|
|
|
(assume H, false_elim (a ≠ a) H)
|
2014-01-15 23:30:16 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem eq_id {A : (Type U)} (a : A) : (a = a) ↔ true
|
2014-01-19 18:34:18 +00:00
|
|
|
|
:= eqt_intro (refl a)
|
|
|
|
|
|
2014-01-27 03:50:12 +00:00
|
|
|
|
theorem iff_id (a : Bool) : (a ↔ a) ↔ true
|
|
|
|
|
:= eqt_intro (refl a)
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem neq_elim {A : (Type U)} {a b : A} (H : a ≠ b) : a = b ↔ false
|
|
|
|
|
:= eqf_intro H
|
|
|
|
|
|
|
|
|
|
theorem neq_to_not_eq {A : (Type U)} {a b : A} : a ≠ b ↔ ¬ a = b
|
|
|
|
|
:= refl (a ≠ b)
|
|
|
|
|
|
|
|
|
|
add_rewrite eq_id iff_id neq_to_not_eq
|
2014-01-31 03:11:58 +00:00
|
|
|
|
|
|
|
|
|
-- Remark: ordered rewriting + assoc + comm + left_comm sorts a term lexicographically
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem left_comm {A : (Type U)} {R : A -> A -> A} (comm : ∀ x y, R x y = R y x) (assoc : ∀ x y z, R (R x y) z = R x (R y z)) :
|
2014-01-31 03:11:58 +00:00
|
|
|
|
∀ x y z, R x (R y z) = R y (R x z)
|
|
|
|
|
:= take x y z, calc R x (R y z) = R (R x y) z : symm (assoc x y z)
|
|
|
|
|
... = R (R y x) z : { comm x y }
|
|
|
|
|
... = R y (R x z) : assoc y x z
|
|
|
|
|
|
2014-01-16 23:07:51 +00:00
|
|
|
|
theorem or_comm (a b : Bool) : (a ∨ b) = (b ∨ a)
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, or_elim H (λ H1, or_intror b H1) (λ H2, or_introl H2 a))
|
|
|
|
|
(assume H, or_elim H (λ H1, or_intror a H1) (λ H2, or_introl H2 b))
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_assoc (a b c : Bool) : (a ∨ b) ∨ c ↔ a ∨ (b ∨ c)
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H : (a ∨ b) ∨ c,
|
2014-01-09 17:00:05 +00:00
|
|
|
|
or_elim H (λ H1 : a ∨ b, or_elim H1 (λ Ha : a, or_introl Ha (b ∨ c))
|
|
|
|
|
(λ Hb : b, or_intror a (or_introl Hb c)))
|
|
|
|
|
(λ Hc : c, or_intror a (or_intror b Hc)))
|
2014-01-12 02:36:37 +00:00
|
|
|
|
(assume H : a ∨ (b ∨ c),
|
2014-01-09 16:33:52 +00:00
|
|
|
|
or_elim H (λ Ha : a, (or_introl (or_introl Ha b) c))
|
2014-01-09 17:00:05 +00:00
|
|
|
|
(λ H1 : b ∨ c, or_elim H1 (λ Hb : b, or_introl (or_intror a Hb) c)
|
|
|
|
|
(λ Hc : c, or_intror (a ∨ b) Hc)))
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_id (a : Bool) : a ∨ a ↔ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, or_elim H (λ H1, H1) (λ H2, H2))
|
|
|
|
|
(assume H, or_introl H a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_falsel (a : Bool) : a ∨ false ↔ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, or_elim H (λ H1, H1) (λ H2, false_elim a H2))
|
|
|
|
|
(assume H, or_introl H false)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_falser (a : Bool) : false ∨ a ↔ a
|
2014-01-19 05:11:12 +00:00
|
|
|
|
:= trans (or_comm false a) (or_falsel a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_truel (a : Bool) : true ∨ a ↔ true
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= boolext (assume H : true ∨ a, trivial)
|
|
|
|
|
(assume H : true, or_introl trivial a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_truer (a : Bool) : a ∨ true ↔ true
|
2014-01-19 05:11:12 +00:00
|
|
|
|
:= trans (or_comm a true) (or_truel a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_tauto (a : Bool) : a ∨ ¬ a ↔ true
|
2014-01-09 16:33:52 +00:00
|
|
|
|
:= eqt_intro (em a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
theorem or_left_comm (a b c : Bool) : a ∨ (b ∨ c) ↔ b ∨ (a ∨ c)
|
|
|
|
|
:= left_comm or_comm or_assoc a b c
|
|
|
|
|
|
|
|
|
|
add_rewrite or_comm or_assoc or_id or_falsel or_falser or_truel or_truer or_tauto or_left_comm
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem resolve2 {a b : Bool} (H1 : a ∨ b) (H2 : ¬ b) : a
|
|
|
|
|
:= resolve1 ((or_comm a b) ◂ H1) H2
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_comm (a b : Bool) : a ∧ b ↔ b ∧ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, and_intro (and_elimr H) (and_eliml H))
|
|
|
|
|
(assume H, and_intro (and_elimr H) (and_eliml H))
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_id (a : Bool) : a ∧ a ↔ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, and_eliml H)
|
|
|
|
|
(assume H, and_intro H H)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_assoc (a b c : Bool) : (a ∧ b) ∧ c ↔ a ∧ (b ∧ c)
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, and_intro (and_eliml (and_eliml H)) (and_intro (and_elimr (and_eliml H)) (and_elimr H)))
|
|
|
|
|
(assume H, and_intro (and_intro (and_eliml H) (and_eliml (and_elimr H))) (and_elimr (and_elimr H)))
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_truer (a : Bool) : a ∧ true ↔ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H : a ∧ true, and_eliml H)
|
|
|
|
|
(assume H : a, and_intro H trivial)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_truel (a : Bool) : true ∧ a ↔ a
|
2014-01-09 16:33:52 +00:00
|
|
|
|
:= trans (and_comm true a) (and_truer a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_falsel (a : Bool) : a ∧ false ↔ false
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, and_elimr H)
|
|
|
|
|
(assume H, false_elim (a ∧ false) H)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_falser (a : Bool) : false ∧ a ↔ false
|
2014-01-19 05:11:12 +00:00
|
|
|
|
:= trans (and_comm false a) (and_falsel a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_absurd (a : Bool) : a ∧ ¬ a ↔ false
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H, absurd (and_eliml H) (and_elimr H))
|
|
|
|
|
(assume H, false_elim (a ∧ ¬ a) H)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
theorem and_left_comm (a b c : Bool) : a ∧ (b ∧ c) ↔ b ∧ (a ∧ c)
|
|
|
|
|
:= left_comm and_comm and_assoc a b c
|
|
|
|
|
|
|
|
|
|
add_rewrite and_comm and_assoc and_id and_falsel and_falser and_truel and_truer and_absurd and_left_comm
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem imp_truer (a : Bool) : (a → true) ↔ true
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= boolext (assume H, trivial)
|
|
|
|
|
(assume H Ha, trivial)
|
2014-01-11 21:35:20 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem imp_truel (a : Bool) : (true → a) ↔ a
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= boolext (assume H : true → a, H trivial)
|
|
|
|
|
(assume Ha H, Ha)
|
2014-01-11 21:35:20 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem imp_falser (a : Bool) : (a → false) ↔ ¬ a
|
2014-01-11 21:35:20 +00:00
|
|
|
|
:= refl _
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem imp_falsel (a : Bool) : (false → a) ↔ true
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= boolext (assume H, trivial)
|
|
|
|
|
(assume H Hf, false_elim a Hf)
|
2014-01-11 21:35:20 +00:00
|
|
|
|
|
2014-01-28 23:56:26 +00:00
|
|
|
|
theorem imp_id (a : Bool) : (a → a) ↔ true
|
|
|
|
|
:= eqt_intro (λ H : a, H)
|
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
add_rewrite imp_truer imp_truel imp_falser imp_falsel imp_id
|
|
|
|
|
|
2014-01-25 02:52:50 +00:00
|
|
|
|
theorem imp_or (a b : Bool) : (a → b) ↔ ¬ a ∨ b
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= boolext
|
2014-01-25 02:52:50 +00:00
|
|
|
|
(assume H : a → b,
|
|
|
|
|
(or_elim (em a)
|
|
|
|
|
(λ Ha : a, or_intror (¬ a) (H Ha))
|
|
|
|
|
(λ Hna : ¬ a, or_introl Hna b)))
|
|
|
|
|
(assume H : ¬ a ∨ b,
|
|
|
|
|
assume Ha : a,
|
|
|
|
|
resolve1 H ((symm (not_not_eq a)) ◂ Ha))
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_congr {a b : Bool} (H : a ↔ b) : ¬ a ↔ ¬ b
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= congr2 not H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
-- Recall that exists is defined as ¬ ∀ x : A, ¬ P x
|
|
|
|
|
theorem exists_elim {A : TypeU} {P : A → Bool} {B : Bool} (H1 : Exists A P) (H2 : ∀ (a : A) (H : P a), B) : B
|
|
|
|
|
:= refute (λ R : ¬ B,
|
|
|
|
|
absurd (take a : A, mt (assume H : P a, H2 a H) R)
|
|
|
|
|
H1)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem exists_intro {A : TypeU} {P : A → Bool} (a : A) (H : P a) : Exists A P
|
|
|
|
|
:= assume H1 : (∀ x : A, ¬ P x),
|
|
|
|
|
absurd H (H1 a)
|
2014-01-01 19:35:21 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_exists (A : (Type U)) (P : A → Bool) : ¬ (∃ x : A, P x) ↔ (∀ x : A, ¬ P x)
|
2014-01-16 10:05:09 +00:00
|
|
|
|
:= calc (¬ ∃ x : A, P x) = ¬ ¬ ∀ x : A, ¬ P x : refl (¬ ∃ x : A, P x)
|
|
|
|
|
... = ∀ x : A, ¬ P x : not_not_eq (∀ x : A, ¬ P x)
|
2013-12-30 02:30:41 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_exists_elim {A : (Type U)} {P : A → Bool} (H : ¬ ∃ x : A, P x) : ∀ x : A, ¬ P x
|
|
|
|
|
:= (not_exists A P) ◂ H
|
2014-01-01 19:35:21 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem exists_unfold1 {A : TypeU} {P : A → Bool} (a : A) (H : ∃ x : A, P x) : P a ∨ (∃ x : A, x ≠ a ∧ P x)
|
|
|
|
|
:= exists_elim H
|
2013-12-30 02:30:41 +00:00
|
|
|
|
(λ (w : A) (H1 : P w),
|
2014-01-09 16:33:52 +00:00
|
|
|
|
or_elim (em (w = a))
|
|
|
|
|
(λ Heq : w = a, or_introl (subst H1 Heq) (∃ x : A, x ≠ a ∧ P x))
|
|
|
|
|
(λ Hne : w ≠ a, or_intror (P a) (exists_intro w (and_intro Hne H1))))
|
2013-12-30 02:30:41 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem exists_unfold2 {A : TypeU} {P : A → Bool} (a : A) (H : P a ∨ (∃ x : A, x ≠ a ∧ P x)) : ∃ x : A, P x
|
|
|
|
|
:= or_elim H
|
|
|
|
|
(λ H1 : P a, exists_intro a H1)
|
2013-12-30 04:33:31 +00:00
|
|
|
|
(λ H2 : (∃ x : A, x ≠ a ∧ P x),
|
2014-01-09 16:33:52 +00:00
|
|
|
|
exists_elim H2
|
2013-12-30 02:30:41 +00:00
|
|
|
|
(λ (w : A) (Hw : w ≠ a ∧ P w),
|
2014-01-09 16:33:52 +00:00
|
|
|
|
exists_intro w (and_elimr Hw)))
|
2013-12-30 02:30:41 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem exists_unfold {A : TypeU} (P : A → Bool) (a : A) : (∃ x : A, P x) ↔ (P a ∨ (∃ x : A, x ≠ a ∧ P x))
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= boolext (assume H : (∃ x : A, P x), exists_unfold1 a H)
|
|
|
|
|
(assume H : (P a ∨ (∃ x : A, x ≠ a ∧ P x)), exists_unfold2 a H)
|
2014-01-01 19:00:32 +00:00
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
definition inhabited (A : (Type U))
|
|
|
|
|
:= ∃ x : A, true
|
|
|
|
|
|
|
|
|
|
-- If we have an element of type A, then A is inhabited
|
|
|
|
|
theorem inhabited_intro {A : TypeU} (a : A) : inhabited A
|
|
|
|
|
:= assume H : (∀ x, ¬ true), absurd_not_true (H a)
|
|
|
|
|
|
|
|
|
|
theorem inhabited_elim {A : TypeU} (H1 : inhabited A) {B : Bool} (H2 : A → B) : B
|
|
|
|
|
:= obtain (w : A) (Hw : true), from H1,
|
|
|
|
|
H2 w
|
|
|
|
|
|
|
|
|
|
theorem inhabited_ex_intro {A : TypeU} {P : A → Bool} (H : ∃ x, P x) : inhabited A
|
|
|
|
|
:= obtain (w : A) (Hw : P w), from H,
|
|
|
|
|
exists_intro w trivial
|
|
|
|
|
|
|
|
|
|
-- If a function space is non-empty, then for every 'a' in the domain, the range (B a) is not empty
|
|
|
|
|
theorem inhabited_range {A : TypeU} {B : A → TypeU} (H : inhabited (∀ x, B x)) (a : A) : inhabited (B a)
|
|
|
|
|
:= refute (λ N : ¬ inhabited (B a),
|
|
|
|
|
let s1 : ¬ ∃ x : B a, true := N,
|
|
|
|
|
s2 : ∀ x : B a, false := take x : B a, absurd_not_true (not_exists_elim s1 x),
|
|
|
|
|
s3 : ∃ y : (∀ x, B x), true := H
|
|
|
|
|
in obtain (w : (∀ x, B x)) (Hw : true), from s3,
|
|
|
|
|
let s4 : B a := w a
|
|
|
|
|
in s2 s4)
|
|
|
|
|
|
|
|
|
|
theorem exists_rem {A : TypeU} (H : inhabited A) (p : Bool) : (∃ x : A, p) ↔ p
|
|
|
|
|
:= iff_intro
|
|
|
|
|
(assume Hl : (∃ x : A, p),
|
|
|
|
|
obtain (w : A) (Hw : p), from Hl,
|
|
|
|
|
Hw)
|
|
|
|
|
(assume Hr : p,
|
|
|
|
|
inhabited_elim H (λ w, exists_intro w Hr))
|
|
|
|
|
|
|
|
|
|
theorem forall_rem {A : TypeU} (H : inhabited A) (p : Bool) : (∀ x : A, p) ↔ p
|
|
|
|
|
:= iff_intro
|
|
|
|
|
(assume Hl : (∀ x : A, p),
|
|
|
|
|
inhabited_elim H (λ w, Hl w))
|
|
|
|
|
(assume Hr : p,
|
|
|
|
|
take x, Hr)
|
|
|
|
|
|
2014-01-11 21:48:28 +00:00
|
|
|
|
-- Congruence theorems for contextual simplification
|
|
|
|
|
|
|
|
|
|
-- Simplify a → b, by first simplifying a to c using the fact that ¬ b is true, and then
|
|
|
|
|
-- b to d using the fact that c is true
|
|
|
|
|
theorem imp_congrr {a b c d : Bool} (H_ac : ∀ (H_nb : ¬ b), a = c) (H_bd : ∀ (H_c : c), b = d) : (a → b) = (c → d)
|
|
|
|
|
:= or_elim (em b)
|
|
|
|
|
(λ H_b : b,
|
|
|
|
|
or_elim (em c)
|
|
|
|
|
(λ H_c : c,
|
2014-01-31 03:11:58 +00:00
|
|
|
|
calc (a → b) = (a → true) : { eqt_intro H_b }
|
|
|
|
|
... = true : imp_truer a
|
|
|
|
|
... = (c → true) : symm (imp_truer c)
|
|
|
|
|
... = (c → b) : { symm (eqt_intro H_b) }
|
|
|
|
|
... = (c → d) : { H_bd H_c })
|
2014-01-11 21:48:28 +00:00
|
|
|
|
(λ H_nc : ¬ c,
|
|
|
|
|
calc (a → b) = (a → true) : { eqt_intro H_b }
|
|
|
|
|
... = true : imp_truer a
|
|
|
|
|
... = (false → d) : symm (imp_falsel d)
|
|
|
|
|
... = (c → d) : { symm (eqf_intro H_nc) }))
|
|
|
|
|
(λ H_nb : ¬ b,
|
|
|
|
|
or_elim (em c)
|
|
|
|
|
(λ H_c : c,
|
|
|
|
|
calc (a → b) = (c → b) : { H_ac H_nb }
|
|
|
|
|
... = (c → d) : { H_bd H_c })
|
|
|
|
|
(λ H_nc : ¬ c,
|
|
|
|
|
calc (a → b) = (c → b) : { H_ac H_nb }
|
|
|
|
|
... = (false → b) : { eqf_intro H_nc }
|
|
|
|
|
... = true : imp_falsel b
|
|
|
|
|
... = (false → d) : symm (imp_falsel d)
|
|
|
|
|
... = (c → d) : { symm (eqf_intro H_nc) }))
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- Simplify a → b, by first simplifying b to d using the fact that a is true, and then
|
|
|
|
|
-- b to d using the fact that ¬ d is true.
|
|
|
|
|
-- This kind of congruence seems to be useful in very rare cases.
|
2014-01-21 22:03:51 +00:00
|
|
|
|
theorem imp_congrl {a b c d : Bool} (H_bd : ∀ (H_a : a), b = d) (H_ac : ∀ (H_nd : ¬ d), a = c) : (a → b) = (c → d)
|
2014-01-11 21:48:28 +00:00
|
|
|
|
:= or_elim (em a)
|
|
|
|
|
(λ H_a : a,
|
|
|
|
|
or_elim (em d)
|
|
|
|
|
(λ H_d : d,
|
|
|
|
|
calc (a → b) = (a → d) : { H_bd H_a }
|
|
|
|
|
... = (a → true) : { eqt_intro H_d }
|
|
|
|
|
... = true : imp_truer a
|
|
|
|
|
... = (c → true) : symm (imp_truer c)
|
|
|
|
|
... = (c → d) : { symm (eqt_intro H_d) })
|
|
|
|
|
(λ H_nd : ¬ d,
|
|
|
|
|
calc (a → b) = (c → b) : { H_ac H_nd }
|
|
|
|
|
... = (c → d) : { H_bd H_a }))
|
|
|
|
|
(λ H_na : ¬ a,
|
|
|
|
|
or_elim (em d)
|
|
|
|
|
(λ H_d : d,
|
|
|
|
|
calc (a → b) = (false → b) : { eqf_intro H_na }
|
|
|
|
|
... = true : imp_falsel b
|
|
|
|
|
... = (c → true) : symm (imp_truer c)
|
|
|
|
|
... = (c → d) : { symm (eqt_intro H_d) })
|
|
|
|
|
(λ H_nd : ¬ d,
|
|
|
|
|
calc (a → b) = (false → b) : { eqf_intro H_na }
|
|
|
|
|
... = true : imp_falsel b
|
|
|
|
|
... = (false → d) : symm (imp_falsel d)
|
|
|
|
|
... = (a → d) : { symm (eqf_intro H_na) }
|
|
|
|
|
... = (c → d) : { H_ac H_nd }))
|
|
|
|
|
|
|
|
|
|
-- (Common case) simplify a to c, and then b to d using the fact that c is true
|
|
|
|
|
theorem imp_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_c : c), b = d) : (a → b) = (c → d)
|
|
|
|
|
:= imp_congrr (λ H, H_ac) H_bd
|
|
|
|
|
|
|
|
|
|
-- In the following theorems we are using the fact that a ∨ b is defined as ¬ a → b
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_congrr {a b c d : Bool} (H_ac : ∀ (H_nb : ¬ b), a = c) (H_bd : ∀ (H_nc : ¬ c), b = d) : a ∨ b ↔ c ∨ d
|
2014-01-11 21:48:28 +00:00
|
|
|
|
:= imp_congrr (λ H_nb : ¬ b, congr2 not (H_ac H_nb)) H_bd
|
2014-01-21 22:03:51 +00:00
|
|
|
|
theorem or_congrl {a b c d : Bool} (H_bd : ∀ (H_na : ¬ a), b = d) (H_ac : ∀ (H_nd : ¬ d), a = c) : a ∨ b ↔ c ∨ d
|
|
|
|
|
:= imp_congrl H_bd (λ H_nd : ¬ d, congr2 not (H_ac H_nd))
|
2014-01-11 21:48:28 +00:00
|
|
|
|
-- (Common case) simplify a to c, and then b to d using the fact that ¬ c is true
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem or_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_nc : ¬ c), b = d) : a ∨ b ↔ c ∨ d
|
2014-01-11 21:48:28 +00:00
|
|
|
|
:= or_congrr (λ H, H_ac) H_bd
|
|
|
|
|
|
|
|
|
|
-- In the following theorems we are using the fact hat a ∧ b is defined as ¬ (a → ¬ b)
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_congrr {a b c d : Bool} (H_ac : ∀ (H_b : b), a = c) (H_bd : ∀ (H_c : c), b = d) : a ∧ b ↔ c ∧ d
|
2014-01-11 21:48:28 +00:00
|
|
|
|
:= congr2 not (imp_congrr (λ (H_nnb : ¬ ¬ b), H_ac (not_not_elim H_nnb)) (λ H_c : c, congr2 not (H_bd H_c)))
|
2014-01-21 22:03:51 +00:00
|
|
|
|
theorem and_congrl {a b c d : Bool} (H_bd : ∀ (H_a : a), b = d) (H_ac : ∀ (H_d : d), a = c) : a ∧ b ↔ c ∧ d
|
|
|
|
|
:= congr2 not (imp_congrl (λ H_a : a, congr2 not (H_bd H_a)) (λ (H_nnd : ¬ ¬ d), H_ac (not_not_elim H_nnd)))
|
2014-01-11 21:48:28 +00:00
|
|
|
|
-- (Common case) simplify a to c, and then b to d using the fact that c is true
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem and_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_c : c), b = d) : a ∧ b ↔ c ∧ d
|
2014-01-11 21:48:28 +00:00
|
|
|
|
:= and_congrr (λ H, H_ac) H_bd
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem not_and (a b : Bool) : ¬ (a ∧ b) ↔ ¬ a ∨ ¬ b
|
|
|
|
|
:= boolext (assume H, or_elim (em a)
|
|
|
|
|
(assume Ha, or_elim (em b)
|
|
|
|
|
(assume Hb, absurd_elim (¬ a ∨ ¬ b) (and_intro Ha Hb) H)
|
|
|
|
|
(assume Hnb, or_intror (¬ a) Hnb))
|
|
|
|
|
(assume Hna, or_introl Hna (¬ b)))
|
|
|
|
|
(assume (H : ¬ a ∨ ¬ b) (N : a ∧ b),
|
|
|
|
|
or_elim H
|
|
|
|
|
(assume Hna, absurd (and_eliml N) Hna)
|
|
|
|
|
(assume Hnb, absurd (and_elimr N) Hnb))
|
|
|
|
|
|
|
|
|
|
theorem not_and_elim {a b : Bool} (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b
|
|
|
|
|
:= (not_and a b) ◂ H
|
|
|
|
|
|
|
|
|
|
theorem not_or (a b : Bool) : ¬ (a ∨ b) ↔ ¬ a ∧ ¬ b
|
|
|
|
|
:= boolext (assume H, or_elim (em a)
|
|
|
|
|
(assume Ha, absurd_elim (¬ a ∧ ¬ b) (or_introl Ha b) H)
|
|
|
|
|
(assume Hna, or_elim (em b)
|
|
|
|
|
(assume Hb, absurd_elim (¬ a ∧ ¬ b) (or_intror a Hb) H)
|
|
|
|
|
(assume Hnb, and_intro Hna Hnb)))
|
|
|
|
|
(assume (H : ¬ a ∧ ¬ b) (N : a ∨ b),
|
|
|
|
|
or_elim N
|
|
|
|
|
(assume Ha, absurd Ha (and_eliml H))
|
|
|
|
|
(assume Hb, absurd Hb (and_elimr H)))
|
|
|
|
|
|
|
|
|
|
theorem not_or_elim {a b : Bool} (H : ¬ (a ∨ b)) : ¬ a ∧ ¬ b
|
|
|
|
|
:= (not_or a b) ◂ H
|
|
|
|
|
|
|
|
|
|
theorem not_implies (a b : Bool) : ¬ (a → b) ↔ a ∧ ¬ b
|
|
|
|
|
:= calc (¬ (a → b)) = ¬ (¬ a ∨ b) : { imp_or a b }
|
|
|
|
|
... = ¬ ¬ a ∧ ¬ b : not_or (¬ a) b
|
|
|
|
|
... = a ∧ ¬ b : by simp
|
|
|
|
|
|
|
|
|
|
theorem not_implies_elim {a b : Bool} (H : ¬ (a → b)) : a ∧ ¬ b
|
|
|
|
|
:= (not_implies a b) ◂ H
|
|
|
|
|
|
|
|
|
|
theorem a_eq_not_a (a : Bool) : (a = ¬ a) ↔ false
|
|
|
|
|
:= boolext (λ H, or_elim (em a)
|
|
|
|
|
(λ Ha, absurd Ha (subst Ha H))
|
|
|
|
|
(λ Hna, absurd (subst Hna (symm H)) Hna))
|
|
|
|
|
(λ H, false_elim (a = ¬ a) H)
|
|
|
|
|
|
|
|
|
|
theorem a_iff_not_a (a : Bool) : (a ↔ ¬ a) ↔ false
|
|
|
|
|
:= a_eq_not_a a
|
|
|
|
|
|
|
|
|
|
theorem true_eq_false : (true = false) ↔ false
|
|
|
|
|
:= subst (a_eq_not_a true) not_true
|
|
|
|
|
|
|
|
|
|
theorem true_iff_false : (true ↔ false) ↔ false
|
|
|
|
|
:= true_eq_false
|
|
|
|
|
|
|
|
|
|
theorem false_eq_true : (false = true) ↔ false
|
|
|
|
|
:= subst (a_eq_not_a false) not_false
|
|
|
|
|
|
|
|
|
|
theorem false_iff_true : (false ↔ true) ↔ false
|
|
|
|
|
:= false_eq_true
|
|
|
|
|
|
|
|
|
|
theorem a_iff_true (a : Bool) : (a ↔ true) ↔ a
|
|
|
|
|
:= boolext (λ H, eqt_elim H)
|
|
|
|
|
(λ H, eqt_intro H)
|
|
|
|
|
|
|
|
|
|
theorem a_iff_false (a : Bool) : (a ↔ false) ↔ ¬ a
|
|
|
|
|
:= boolext (λ H, eqf_elim H)
|
|
|
|
|
(λ H, eqf_intro H)
|
|
|
|
|
|
|
|
|
|
add_rewrite a_eq_not_a a_iff_not_a true_eq_false true_iff_false false_eq_true false_iff_true a_iff_true a_iff_false
|
|
|
|
|
|
|
|
|
|
theorem not_iff (a b : Bool) : ¬ (a ↔ b) ↔ (¬ a ↔ b)
|
|
|
|
|
:= or_elim (em b)
|
|
|
|
|
(λ Hb, calc (¬ (a ↔ b)) = (¬ (a ↔ true)) : { eqt_intro Hb }
|
|
|
|
|
... = ¬ a : { a_iff_true a }
|
|
|
|
|
... = ¬ a ↔ true : { symm (a_iff_true (¬ a)) }
|
|
|
|
|
... = ¬ a ↔ b : { symm (eqt_intro Hb) })
|
|
|
|
|
(λ Hnb, calc (¬ (a ↔ b)) = (¬ (a ↔ false)) : { eqf_intro Hnb }
|
|
|
|
|
... = ¬ ¬ a : { a_iff_false a }
|
|
|
|
|
... = ¬ a ↔ false : { symm (a_iff_false (¬ a)) }
|
|
|
|
|
... = ¬ a ↔ b : { symm (eqf_intro Hnb) })
|
|
|
|
|
|
|
|
|
|
theorem not_iff_elim {a b : Bool} (H : ¬ (a ↔ b)) : (¬ a) ↔ b
|
|
|
|
|
:= (not_iff a b) ◂ H
|
|
|
|
|
|
2014-01-16 23:07:51 +00:00
|
|
|
|
theorem forall_or_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∀ x, p ∨ φ x) = (p ∨ ∀ x, φ x)
|
2014-01-14 02:21:22 +00:00
|
|
|
|
:= boolext
|
|
|
|
|
(assume H : (∀ x, p ∨ φ x),
|
|
|
|
|
or_elim (em p)
|
|
|
|
|
(λ Hp : p, or_introl Hp (∀ x, φ x))
|
|
|
|
|
(λ Hnp : ¬ p, or_intror p (take x,
|
|
|
|
|
resolve1 (H x) Hnp)))
|
|
|
|
|
(assume H : (p ∨ ∀ x, φ x),
|
|
|
|
|
take x,
|
|
|
|
|
or_elim H
|
|
|
|
|
(λ H1 : p, or_introl H1 (φ x))
|
|
|
|
|
(λ H2 : (∀ x, φ x), or_intror p (H2 x)))
|
|
|
|
|
|
2014-01-16 23:07:51 +00:00
|
|
|
|
theorem forall_or_distributel {A : Type} (p : Bool) (φ : A → Bool) : (∀ x, φ x ∨ p) = ((∀ x, φ x) ∨ p)
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= boolext
|
|
|
|
|
(assume H : (∀ x, φ x ∨ p),
|
|
|
|
|
or_elim (em p)
|
|
|
|
|
(λ Hp : p, or_intror (∀ x, φ x) Hp)
|
|
|
|
|
(λ Hnp : ¬ p, or_introl (take x, resolve2 (H x) Hnp) p))
|
|
|
|
|
(assume H : (∀ x, φ x) ∨ p,
|
|
|
|
|
take x,
|
|
|
|
|
or_elim H
|
|
|
|
|
(λ H1 : (∀ x, φ x), or_introl (H1 x) p)
|
|
|
|
|
(λ H2 : p, or_intror (φ x) H2))
|
2014-01-14 02:21:22 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem forall_and_distribute {A : TypeU} (φ ψ : A → Bool) : (∀ x, φ x ∧ ψ x) ↔ (∀ x, φ x) ∧ (∀ x, ψ x)
|
2014-01-14 02:21:22 +00:00
|
|
|
|
:= boolext
|
|
|
|
|
(assume H : (∀ x, φ x ∧ ψ x),
|
|
|
|
|
and_intro (take x, and_eliml (H x)) (take x, and_elimr (H x)))
|
|
|
|
|
(assume H : (∀ x, φ x) ∧ (∀ x, ψ x),
|
|
|
|
|
take x, and_intro (and_eliml H x) (and_elimr H x))
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem exists_and_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, p ∧ φ x) ↔ p ∧ ∃ x, φ x
|
2014-01-14 02:21:22 +00:00
|
|
|
|
:= boolext
|
|
|
|
|
(assume H : (∃ x, p ∧ φ x),
|
|
|
|
|
obtain (w : A) (Hw : p ∧ φ w), from H,
|
|
|
|
|
and_intro (and_eliml Hw) (exists_intro w (and_elimr Hw)))
|
|
|
|
|
(assume H : (p ∧ ∃ x, φ x),
|
|
|
|
|
obtain (w : A) (Hw : φ w), from (and_elimr H),
|
|
|
|
|
exists_intro w (and_intro (and_eliml H) Hw))
|
|
|
|
|
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem exists_or_distribute {A : TypeU} (φ ψ : A → Bool) : (∃ x, φ x ∨ ψ x) ↔ (∃ x, φ x) ∨ (∃ x, ψ x)
|
2014-01-14 02:21:22 +00:00
|
|
|
|
:= boolext
|
|
|
|
|
(assume H : (∃ x, φ x ∨ ψ x),
|
|
|
|
|
obtain (w : A) (Hw : φ w ∨ ψ w), from H,
|
|
|
|
|
or_elim Hw
|
|
|
|
|
(λ Hw1 : φ w, or_introl (exists_intro w Hw1) (∃ x, ψ x))
|
|
|
|
|
(λ Hw2 : ψ w, or_intror (∃ x, φ x) (exists_intro w Hw2)))
|
|
|
|
|
(assume H : (∃ x, φ x) ∨ (∃ x, ψ x),
|
|
|
|
|
or_elim H
|
|
|
|
|
(λ H1 : (∃ x, φ x),
|
|
|
|
|
obtain (w : A) (Hw : φ w), from H1,
|
|
|
|
|
exists_intro w (or_introl Hw (ψ w)))
|
|
|
|
|
(λ H2 : (∃ x, ψ x),
|
|
|
|
|
obtain (w : A) (Hw : ψ w), from H2,
|
|
|
|
|
exists_intro w (or_intror (φ w) Hw)))
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem eq_exists_intro {A : (Type U)} {P Q : A → Bool} (H : ∀ x : A, P x ↔ Q x) : (∃ x : A, P x) ↔ (∃ x : A, Q x)
|
|
|
|
|
:= boolext
|
|
|
|
|
(assume Hex, obtain w Pw, from Hex, exists_intro w ((H w) ◂ Pw))
|
|
|
|
|
(assume Hex, obtain w Qw, from Hex, exists_intro w ((symm (H w)) ◂ Qw))
|
|
|
|
|
|
|
|
|
|
theorem not_forall (A : (Type U)) (P : A → Bool) : ¬ (∀ x : A, P x) ↔ (∃ x : A, ¬ P x)
|
|
|
|
|
:= boolext
|
|
|
|
|
(assume H, refute (λ N : ¬ (∃ x, ¬ P x),
|
|
|
|
|
absurd (take x, not_not_elim (not_exists_elim N x)) H))
|
|
|
|
|
(assume (H : ∃ x, ¬ P x) (N : ∀ x, P x),
|
|
|
|
|
obtain w Hw, from H,
|
|
|
|
|
absurd (N w) Hw)
|
|
|
|
|
|
|
|
|
|
theorem not_forall_elim {A : (Type U)} {P : A → Bool} (H : ¬ (∀ x : A, P x)) : ∃ x : A, ¬ P x
|
|
|
|
|
:= (not_forall A P) ◂ H
|
|
|
|
|
|
|
|
|
|
theorem exists_and_distributel {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, φ x ∧ p) ↔ (∃ x, φ x) ∧ p
|
|
|
|
|
:= calc (∃ x, φ x ∧ p) = (∃ x, p ∧ φ x) : eq_exists_intro (λ x, and_comm (φ x) p)
|
|
|
|
|
... = (p ∧ (∃ x, φ x)) : exists_and_distributer p φ
|
|
|
|
|
... = ((∃ x, φ x) ∧ p) : and_comm p (∃ x, φ x)
|
2014-01-25 02:52:50 +00:00
|
|
|
|
|
|
|
|
|
theorem exists_imp_distribute {A : TypeU} (φ ψ : A → Bool) : (∃ x, φ x → ψ x) ↔ ((∀ x, φ x) → (∃ x, ψ x))
|
|
|
|
|
:= calc (∃ x, φ x → ψ x) = (∃ x, ¬ φ x ∨ ψ x) : eq_exists_intro (λ x, imp_or (φ x) (ψ x))
|
|
|
|
|
... = (∃ x, ¬ φ x) ∨ (∃ x, ψ x) : exists_or_distribute _ _
|
|
|
|
|
... = ¬ (∀ x, φ x) ∨ (∃ x, ψ x) : { symm (not_forall A φ) }
|
|
|
|
|
... = (∀ x, φ x) → (∃ x, ψ x) : symm (imp_or _ _)
|
|
|
|
|
|
2014-02-02 02:27:14 +00:00
|
|
|
|
theorem forall_uninhabited {A : (Type U)} {B : A → Bool} (H : ¬ inhabited A) : ∀ x, B x
|
|
|
|
|
:= refute (λ N : ¬ (∀ x, B x),
|
|
|
|
|
obtain w Hw, from not_forall_elim N,
|
|
|
|
|
absurd (inhabited_intro w) H)
|
|
|
|
|
|
|
|
|
|
theorem allext {A : (Type U)} {B C : A → Bool} (H : ∀ x : A, B x = C x) : (∀ x : A, B x) = (∀ x : A, C x)
|
|
|
|
|
:= boolext
|
|
|
|
|
(assume Hl, take x, (H x) ◂ (Hl x))
|
|
|
|
|
(assume Hr, take x, (symm (H x)) ◂ (Hr x))
|
|
|
|
|
|
|
|
|
|
-- Up to this point, we proved all theorems using just reflexivity, substitution and case (proof by cases)
|
|
|
|
|
|
|
|
|
|
-- Function extensionality
|
|
|
|
|
axiom funext {A : (Type U)} {B : A → (Type U)} {f g : ∀ x : A, B x} (H : ∀ x : A, f x = g x) : f = g
|
|
|
|
|
|
|
|
|
|
-- Eta is a consequence of function extensionality
|
|
|
|
|
theorem eta {A : TypeU} {B : A → TypeU} (f : ∀ x : A, B x) : (λ x : A, f x) = f
|
|
|
|
|
:= funext (λ x : A, refl (f x))
|
|
|
|
|
|
|
|
|
|
-- Epsilon (Hilbert's operator)
|
|
|
|
|
variable eps {A : TypeU} (H : inhabited A) (P : A → Bool) : A
|
|
|
|
|
alias ε : eps
|
|
|
|
|
axiom eps_ax {A : TypeU} (H : inhabited A) {P : A → Bool} (a : A) : P a → P (ε H P)
|
|
|
|
|
|
|
|
|
|
theorem eps_th {A : TypeU} {P : A → Bool} (a : A) : P a → P (ε (inhabited_intro a) P)
|
|
|
|
|
:= assume H : P a, @eps_ax A (inhabited_intro a) P a H
|
|
|
|
|
|
|
|
|
|
theorem eps_singleton {A : TypeU} (H : inhabited A) (a : A) : ε H (λ x, x = a) = a
|
|
|
|
|
:= let P := λ x, x = a,
|
|
|
|
|
Ha : P a := refl a
|
|
|
|
|
in eps_ax H a Ha
|
|
|
|
|
|
|
|
|
|
-- A function space (∀ x : A, B x) is inhabited if forall a : A, we have inhabited (B a)
|
|
|
|
|
theorem inhabited_fun {A : TypeU} {B : A → TypeU} (Hn : ∀ a, inhabited (B a)) : inhabited (∀ x, B x)
|
|
|
|
|
:= inhabited_intro (λ x, ε (Hn x) (λ y, true))
|
|
|
|
|
|
|
|
|
|
theorem exists_to_eps {A : TypeU} {P : A → Bool} (H : ∃ x, P x) : P (ε (inhabited_ex_intro H) P)
|
|
|
|
|
:= obtain (w : A) (Hw : P w), from H,
|
|
|
|
|
eps_ax (inhabited_ex_intro H) w Hw
|
|
|
|
|
|
|
|
|
|
theorem axiom_of_choice {A : TypeU} {B : A → TypeU} {R : ∀ x : A, B x → Bool} (H : ∀ x, ∃ y, R x y) : ∃ f, ∀ x, R x (f x)
|
|
|
|
|
:= exists_intro
|
|
|
|
|
(λ x, ε (inhabited_ex_intro (H x)) (λ y, R x y)) -- witness for f
|
|
|
|
|
(λ x, exists_to_eps (H x)) -- proof that witness satisfies ∀ x, R x (f x)
|
|
|
|
|
|
|
|
|
|
theorem skolem_th {A : TypeU} {B : A → TypeU} {P : ∀ x : A, B x → Bool} :
|
|
|
|
|
(∀ x, ∃ y, P x y) ↔ ∃ f, (∀ x, P x (f x))
|
|
|
|
|
:= iff_intro
|
|
|
|
|
(λ H : (∀ x, ∃ y, P x y), axiom_of_choice H)
|
|
|
|
|
(λ H : (∃ f, (∀ x, P x (f x))),
|
|
|
|
|
take x, obtain (fw : ∀ x, B x) (Hw : ∀ x, P x (fw x)), from H,
|
|
|
|
|
exists_intro (fw x) (Hw x))
|
|
|
|
|
|
|
|
|
|
-- if-then-else expression, we define it using Hilbert's operator
|
|
|
|
|
definition ite {A : TypeU} (c : Bool) (a b : A) : A
|
|
|
|
|
:= ε (inhabited_intro a) (λ r, (c → r = a) ∧ (¬ c → r = b))
|
|
|
|
|
notation 45 if _ then _ else _ : ite
|
2014-01-31 06:13:34 +00:00
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
theorem if_true {A : TypeU} (a b : A) : (if true then a else b) = a
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= calc (if true then a else b) = ε (inhabited_intro a) (λ r, (true → r = a) ∧ (¬ true → r = b)) : refl (if true then a else b)
|
|
|
|
|
... = ε (inhabited_intro a) (λ r, r = a) : by simp
|
|
|
|
|
... = a : eps_singleton (inhabited_intro a) a
|
2014-01-31 03:11:58 +00:00
|
|
|
|
|
|
|
|
|
theorem if_false {A : TypeU} (a b : A) : (if false then a else b) = b
|
2014-02-02 02:27:14 +00:00
|
|
|
|
:= calc (if false then a else b) = ε (inhabited_intro a) (λ r, (false → r = a) ∧ (¬ false → r = b)) : refl (if false then a else b)
|
|
|
|
|
... = ε (inhabited_intro a) (λ r, r = b) : by simp
|
|
|
|
|
... = b : eps_singleton (inhabited_intro a) b
|
2014-01-31 03:11:58 +00:00
|
|
|
|
|
|
|
|
|
theorem if_a_a {A : TypeU} (c : Bool) (a: A) : (if c then a else a) = a
|
|
|
|
|
:= or_elim (em c)
|
|
|
|
|
(λ H : c, calc (if c then a else a) = (if true then a else a) : { eqt_intro H }
|
|
|
|
|
... = a : if_true a a)
|
|
|
|
|
(λ H : ¬ c, calc (if c then a else a) = (if false then a else a) : { eqf_intro H }
|
|
|
|
|
... = a : if_false a a)
|
|
|
|
|
|
2014-01-31 06:38:35 +00:00
|
|
|
|
add_rewrite if_true if_false if_a_a
|
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
theorem if_congr {A : TypeU} {b c : Bool} {x y u v : A} (H_bc : b = c)
|
|
|
|
|
(H_xu : ∀ (H_c : c), x = u) (H_yv : ∀ (H_nc : ¬ c), y = v) :
|
|
|
|
|
(if b then x else y) = if c then u else v
|
|
|
|
|
:= or_elim (em c)
|
|
|
|
|
(λ H_c : c, calc
|
|
|
|
|
(if b then x else y) = if c then x else y : { H_bc }
|
|
|
|
|
... = if true then x else y : { eqt_intro H_c }
|
|
|
|
|
... = x : if_true _ _
|
|
|
|
|
... = u : H_xu H_c
|
|
|
|
|
... = if true then u else v : symm (if_true _ _)
|
|
|
|
|
... = if c then u else v : { symm (eqt_intro H_c) })
|
|
|
|
|
(λ H_nc : ¬ c, calc
|
|
|
|
|
(if b then x else y) = if c then x else y : { H_bc }
|
|
|
|
|
... = if false then x else y : { eqf_intro H_nc }
|
|
|
|
|
... = y : if_false _ _
|
|
|
|
|
... = v : H_yv H_nc
|
|
|
|
|
... = if false then u else v : symm (if_false _ _)
|
|
|
|
|
... = if c then u else v : { symm (eqf_intro H_nc) })
|
|
|
|
|
|
|
|
|
|
theorem if_imp_then {a b c : Bool} (H : if a then b else c) : a → b
|
|
|
|
|
:= assume Ha : a, eqt_elim (calc b = if true then b else c : symm (if_true b c)
|
|
|
|
|
... = if a then b else c : { symm (eqt_intro Ha) }
|
|
|
|
|
... = true : eqt_intro H)
|
|
|
|
|
|
|
|
|
|
theorem if_imp_else {a b c : Bool} (H : if a then b else c) : ¬ a → c
|
|
|
|
|
:= assume Hna : ¬ a, eqt_elim (calc c = if false then b else c : symm (if_false b c)
|
|
|
|
|
... = if a then b else c : { symm (eqf_intro Hna) }
|
|
|
|
|
... = true : eqt_intro H)
|
|
|
|
|
|
|
|
|
|
theorem app_if_distribute {A B : TypeU} (c : Bool) (f : A → B) (a b : A) : f (if c then a else b) = if c then f a else f b
|
|
|
|
|
:= or_elim (em c)
|
|
|
|
|
(λ Hc : c , calc
|
|
|
|
|
f (if c then a else b) = f (if true then a else b) : { eqt_intro Hc }
|
|
|
|
|
... = f a : { if_true a b }
|
|
|
|
|
... = if true then f a else f b : symm (if_true (f a) (f b))
|
|
|
|
|
... = if c then f a else f b : { symm (eqt_intro Hc) })
|
|
|
|
|
(λ Hnc : ¬ c, calc
|
|
|
|
|
f (if c then a else b) = f (if false then a else b) : { eqf_intro Hnc }
|
|
|
|
|
... = f b : { if_false a b }
|
|
|
|
|
... = if false then f a else f b : symm (if_false (f a) (f b))
|
|
|
|
|
... = if c then f a else f b : { symm (eqf_intro Hnc) })
|
|
|
|
|
|
|
|
|
|
theorem eq_if_distributer {A : TypeU} (c : Bool) (a b v : A) : (v = (if c then a else b)) = if c then v = a else v = b
|
|
|
|
|
:= app_if_distribute c (eq v) a b
|
|
|
|
|
|
|
|
|
|
theorem eq_if_distributel {A : TypeU} (c : Bool) (a b v : A) : ((if c then a else b) = v) = if c then a = v else b = v
|
|
|
|
|
:= app_if_distribute c (λ x, x = v) a b
|
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
set_opaque exists true
|
|
|
|
|
set_opaque not true
|
|
|
|
|
set_opaque or true
|
|
|
|
|
set_opaque and true
|
2014-01-31 03:11:58 +00:00
|
|
|
|
set_opaque implies true
|
2014-02-02 02:27:14 +00:00
|
|
|
|
set_opaque ite true
|
|
|
|
|
set_opaque eq true
|
|
|
|
|
|
|
|
|
|
definition injective {A B : (Type U)} (f : A → B) := ∀ x1 x2, f x1 = f x2 → x1 = x2
|
|
|
|
|
definition non_surjective {A B : (Type U)} (f : A → B) := ∃ y, ∀ x, ¬ f x = y
|
|
|
|
|
|
|
|
|
|
-- The set of individuals, we need to assert the existence of one infinite set
|
|
|
|
|
variable ind : Type
|
|
|
|
|
-- ind is infinite, i.e., there is a function f s.t. f is injective, and not surjective
|
|
|
|
|
axiom infinity : ∃ f : ind → ind, injective f ∧ non_surjective f
|
|
|
|
|
|
2014-02-04 09:46:50 +00:00
|
|
|
|
-- Pair extensionality
|
|
|
|
|
axiom pairext {A : (Type U)} {B : A → (Type U)} {a b : sig x, B x}
|
|
|
|
|
(H1 : proj1 a = proj1 b) (H2 : proj2 a == proj2 b)
|
|
|
|
|
: a = b
|
|
|
|
|
|
2014-02-04 22:15:33 +00:00
|
|
|
|
-- Heterogeneous equality axioms and theorems
|
|
|
|
|
|
|
|
|
|
-- In the following definitions the type of A and B cannot be (Type U)
|
|
|
|
|
-- because A = B would be @eq (Type U+1) A B, and
|
|
|
|
|
-- the type of eq is (∀T : (Type U), T → T → bool).
|
|
|
|
|
-- So, we define M a universe smaller than U.
|
|
|
|
|
universe M ≥ 1
|
|
|
|
|
definition TypeM := (Type M)
|
|
|
|
|
|
|
|
|
|
-- We can "type-cast" an A expression into a B expression, if we can prove that A = B
|
|
|
|
|
variable cast {A B : (Type M)} : A = B → A → B
|
|
|
|
|
axiom cast_heq {A B : (Type M)} (H : A = B) (a : A) : cast H a == a
|
|
|
|
|
|
|
|
|
|
-- Heterogeneous equality satisfies the usual properties: symmetry, transitivity, congruence and function extensionality
|
|
|
|
|
axiom hsymm {A B : (Type U)} {a : A} {b : B} : a == b → b == a
|
|
|
|
|
axiom htrans {A B C : (Type U)} {a : A} {b : B} {c : C} : a == b → b == c → a == c
|
|
|
|
|
axiom hcongr {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)} {f : ∀ x, B x} {f' : ∀ x, B' x} {a : A} {a' : A'} :
|
|
|
|
|
f == f' → a == a' → f a == f' a'
|
|
|
|
|
axiom hfunext {A A' : (Type M)} {B : A → (Type U)} {B' : A' → (Type U)} {f : ∀ x, B x} {f' : ∀ x, B' x} :
|
|
|
|
|
A = A' → (∀ x x', x == x' → f x == f' x') → f == f'
|
|
|
|
|
|
|
|
|
|
-- Heterogeneous version of the allext theorem
|
|
|
|
|
theorem hallext {A A' : (Type M)} {B : A → Bool} {B' : A' → Bool}
|
|
|
|
|
(Ha : A = A') (Hb : ∀ x x', x == x' → B x = B' x') : (∀ x, B x) = (∀ x, B' x)
|
|
|
|
|
:= boolext
|
|
|
|
|
(assume (H : ∀ x : A, B x),
|
|
|
|
|
take x' : A', (Hb (cast (symm Ha) x') x' (cast_heq (symm Ha) x')) ◂ (H (cast (symm Ha) x')))
|
|
|
|
|
(assume (H : ∀ x' : A', B' x'),
|
|
|
|
|
take x : A, (symm (Hb x (cast Ha x) (hsymm (cast_heq Ha x)))) ◂ (H (cast Ha x)))
|
|
|
|
|
|
|
|
|
|
-- Simpler version of hfunext axiom, we use it to build proofs
|
|
|
|
|
theorem hsfunext {A : (Type M)} {B B' : A → (Type U)} {f : ∀ x, B x} {f' : ∀ x, B' x} :
|
|
|
|
|
(∀ x, f x == f' x) → f == f'
|
|
|
|
|
:= λ Hb,
|
|
|
|
|
hfunext (refl A) (λ (x x' : A) (Hx : x == x'),
|
|
|
|
|
let s1 : f x == f' x := Hb x,
|
|
|
|
|
s2 : f' x == f' x' := hcongr (hrefl f') Hx
|
|
|
|
|
in htrans s1 s2)
|
|
|
|
|
|
|
|
|
|
-- Some theorems that are useful for applying simplifications.
|
|
|
|
|
theorem cast_eq {A : (Type M)} (H : A = A) (a : A) : cast H a = a
|
|
|
|
|
:= to_eq (cast_heq H a)
|
|
|
|
|
|
|
|
|
|
theorem cast_trans {A B C : (Type M)} (Hab : A = B) (Hbc : B = C) (a : A) : cast Hbc (cast Hab a) = cast (trans Hab Hbc) a
|
|
|
|
|
:= let s1 : cast Hbc (cast Hab a) == cast Hab a := cast_heq Hbc (cast Hab a),
|
|
|
|
|
s2 : cast Hab a == a := cast_heq Hab a,
|
|
|
|
|
s3 : cast (trans Hab Hbc) a == a := cast_heq (trans Hab Hbc) a,
|
|
|
|
|
s4 : cast Hbc (cast Hab a) == cast (trans Hab Hbc) a := htrans (htrans s1 s2) (hsymm s3)
|
|
|
|
|
in to_eq s4
|
|
|
|
|
|
|
|
|
|
theorem cast_pull {A : (Type M)} {B B' : A → (Type M)}
|
|
|
|
|
(f : ∀ x, B x) (a : A) (Hb : (∀ x, B x) = (∀ x, B' x)) (Hba : (B a) = (B' a)) :
|
|
|
|
|
cast Hb f a = cast Hba (f a)
|
|
|
|
|
:= let s1 : cast Hb f a == f a := hcongr (cast_heq Hb f) (hrefl a),
|
|
|
|
|
s2 : cast Hba (f a) == f a := cast_heq Hba (f a)
|
|
|
|
|
in to_eq (htrans s1 (hsymm s2))
|
|
|
|
|
|
|
|
|
|
-- Proof irrelevance is true in the set theoretic model we have for Lean.
|
|
|
|
|
axiom proof_irrel {a : Bool} (H1 H2 : a) : H1 = H2
|
|
|
|
|
|
|
|
|
|
-- A more general version of proof_irrel that can be be derived using proof_irrel, heq axioms and boolext/iff_intro
|
|
|
|
|
theorem hproof_irrel {a b : Bool} (H1 : a) (H2 : b) : H1 == H2
|
|
|
|
|
:= let H1b : b := cast (by simp) H1,
|
|
|
|
|
H1_eq_H1b : H1 == H1b := hsymm (cast_heq (by simp) H1),
|
|
|
|
|
H1b_eq_H2 : H1b == H2 := to_heq (proof_irrel H1b H2)
|
|
|
|
|
in htrans H1_eq_H1b H1b_eq_H2
|