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
|
|
|
|
|
definition TypeU := (Type U)
|
2013-12-30 05:59:57 +00:00
|
|
|
|
|
2014-01-05 20:05:08 +00:00
|
|
|
|
variable Bool : Type
|
|
|
|
|
-- The following builtin declarations can be removed as soon as Lean supports inductive datatypes and match expressions
|
|
|
|
|
builtin true : Bool
|
|
|
|
|
builtin false : Bool
|
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
|
|
|
|
|
|
|
|
|
|
builtin eq {A : (Type U)} (a b : A) : Bool
|
|
|
|
|
infix 50 = : eq
|
|
|
|
|
|
2014-01-16 23:07:51 +00:00
|
|
|
|
definition neq {A : TypeU} (a b : A) := ¬ (a = b)
|
|
|
|
|
infix 50 ≠ : neq
|
2014-01-16 10:05:09 +00:00
|
|
|
|
|
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-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-01-22 16:52:31 +00:00
|
|
|
|
definition Exists (A : TypeU) (P : A → Bool) := ¬ (∀ x, ¬ (P x))
|
|
|
|
|
|
|
|
|
|
definition nonempty (A : TypeU) := ∃ x : A, true
|
|
|
|
|
|
2014-01-22 17:18:40 +00:00
|
|
|
|
-- If we have an element of type A, then A is nonempty
|
|
|
|
|
theorem nonempty_intro {A : TypeU} (a : A) : nonempty A
|
2014-01-22 16:52:31 +00:00
|
|
|
|
:= assume H : (∀ x, ¬ true), (H a)
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-01-09 02:29:02 +00:00
|
|
|
|
theorem em (a : Bool) : a ∨ ¬ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Hna : ¬ a, Hna
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
|
|
|
|
axiom case (P : Bool → Bool) (H1 : P true) (H2 : P false) (a : Bool) : P a
|
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
axiom refl {A : TypeU} (a : A) : a = a
|
2013-12-29 22:01:30 +00:00
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
axiom subst {A : TypeU} {a b : A} {P : A → Bool} (H1 : P a) (H2 : a = b) : P b
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 01:25:14 +00:00
|
|
|
|
-- Function extensionality
|
2014-01-17 02:06:25 +00:00
|
|
|
|
axiom funext {A : TypeU} {B : A → TypeU} {f g : ∀ x : A, B x} (H : ∀ x : A, f x = g x) : f = g
|
2014-01-08 08:38:39 +00:00
|
|
|
|
|
2014-01-09 01:25:14 +00:00
|
|
|
|
-- Forall extensionality
|
2014-01-16 23:07:51 +00:00
|
|
|
|
axiom allext {A : TypeU} {B C : A → Bool} (H : ∀ x : A, B x = C x) : (∀ x : A, B x) = (∀ x : A, C x)
|
2014-01-09 00:19:11 +00:00
|
|
|
|
|
2014-01-22 16:21:11 +00:00
|
|
|
|
-- Epsilon (Hilbert's operator)
|
2014-01-22 16:52:31 +00:00
|
|
|
|
variable eps {A : TypeU} (H : nonempty A) (P : A → Bool) : A
|
2014-01-22 16:21:11 +00:00
|
|
|
|
alias ε : eps
|
2014-01-22 17:18:40 +00:00
|
|
|
|
axiom eps_ax {A : TypeU} (H : nonempty A) {P : A → Bool} (a : A) : P a → P (ε H P)
|
2014-01-22 16:21:11 +00:00
|
|
|
|
|
2014-01-19 20:20:09 +00:00
|
|
|
|
-- Proof irrelevance
|
|
|
|
|
axiom proof_irrel {a : Bool} (H1 H2 : a) : H1 = H2
|
|
|
|
|
|
2014-01-09 00:19:11 +00:00
|
|
|
|
-- Alias for subst where we can provide P explicitly, but keep A,a,b implicit
|
2014-01-17 02:06:25 +00:00
|
|
|
|
theorem substp {A : TypeU} {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-01-31 03:11:58 +00:00
|
|
|
|
theorem eps_th {A : TypeU} {P : A → Bool} (a : A) : P a → P (ε (nonempty_intro a) P)
|
|
|
|
|
:= assume H : P a, @eps_ax A (nonempty_intro a) P a H
|
|
|
|
|
|
|
|
|
|
theorem eps_singleton {A : TypeU} (H : nonempty A) (a : A) : ε H (λ x, x = a) = a
|
|
|
|
|
:= let P := λ x, x = a,
|
|
|
|
|
Ha : P a := refl a
|
|
|
|
|
in eps_ax H a Ha
|
|
|
|
|
|
2014-01-31 03:38:51 +00:00
|
|
|
|
-- A function space (∀ x : A, B x) is nonempty if forall a : A, we have nonempty (B a)
|
|
|
|
|
theorem nonempty_fun {A : TypeU} {B : A → TypeU} (Hn : ∀ a, nonempty (B a)) : nonempty (∀ x, B x)
|
|
|
|
|
:= nonempty_intro (λ x, ε (Hn x) (λ y, true))
|
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
-- if-then-else expression
|
|
|
|
|
definition ite {A : TypeU} (c : Bool) (a b : A) : A
|
|
|
|
|
:= ε (nonempty_intro a) (λ r, (c → r = a) ∧ (¬ c → r = b))
|
|
|
|
|
notation 45 if _ then _ else _ : ite
|
|
|
|
|
|
|
|
|
|
-- We will mark 'not' as opaque later
|
2014-01-10 23:20:55 +00:00
|
|
|
|
theorem not_intro {a : Bool} (H : a → false) : ¬ a
|
|
|
|
|
:= H
|
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
theorem eta {A : TypeU} {B : A → TypeU} (f : ∀ x : A, B x) : (λ x : A, f x) = f
|
2014-01-09 01:25:14 +00:00
|
|
|
|
:= funext (λ x : A, refl (f x))
|
|
|
|
|
|
2014-01-19 19:24:20 +00:00
|
|
|
|
-- create default rewrite rule set
|
|
|
|
|
(* mk_rewrite_rule_set() *)
|
|
|
|
|
|
2014-01-06 03:10:21 +00:00
|
|
|
|
theorem trivial : true
|
|
|
|
|
:= refl true
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-06 03:10:21 +00:00
|
|
|
|
theorem absurd {a : Bool} (H1 : a) (H2 : ¬ a) : false
|
2014-01-08 08:38:39 +00:00
|
|
|
|
:= H2 H1
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
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-01-15 22:42:00 +00:00
|
|
|
|
theorem boolcomplete (a : Bool) : a = true ∨ a = false
|
|
|
|
|
:= case (λ x, x = true ∨ x = false) trivial trivial a
|
2014-01-09 02:29:02 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem false_elim (a : Bool) (H : false) : a
|
2014-01-09 02:29:02 +00:00
|
|
|
|
:= case (λ x, x) trivial H a
|
|
|
|
|
|
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-01-16 17:00:01 +00:00
|
|
|
|
theorem not_not_eq (a : Bool) : ¬ ¬ a ↔ a
|
|
|
|
|
:= case (λ x, ¬ ¬ x ↔ x) trivial trivial a
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
add_rewrite not_not_eq
|
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_not_elim {a : Bool} (H : ¬ ¬ a) : a
|
|
|
|
|
:= (not_not_eq a) ◂ H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-08 08:38:39 +00:00
|
|
|
|
theorem mt {a b : Bool} (H1 : a → b) (H2 : ¬ b) : ¬ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Ha : a, absurd (H1 Ha) H2
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-08 08:38:39 +00:00
|
|
|
|
theorem contrapos {a b : Bool} (H : a → b) : ¬ b → ¬ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume Hnb : ¬ b, mt H Hnb
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem absurd_elim {a : Bool} (b : Bool) (H1 : a) (H2 : ¬ a) : b
|
|
|
|
|
:= false_elim b (absurd H1 H2)
|
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-01-05 19:25:58 +00:00
|
|
|
|
(have ¬ ¬ a :
|
2014-01-12 02:36:37 +00:00
|
|
|
|
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-06 03:10:21 +00:00
|
|
|
|
theorem resolve1 {a b : Bool} (H1 : a ∨ b) (H2 : ¬ a) : b
|
2014-01-08 08:38:39 +00:00
|
|
|
|
:= H1 H2
|
2014-01-06 03:10:21 +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 00:52:43 +00:00
|
|
|
|
-- Recall that or is defined as ¬ a → b
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem or_introl {a : Bool} (H : a) (b : Bool) : a ∨ b
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume H1 : ¬ a, absurd_elim b H H1
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem or_intror {b : Bool} (a : Bool) (H : b) : a ∨ b
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume H1 : ¬ a, 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,
|
|
|
|
|
absurd (have c : H3 (have b : resolve1 H1 (have ¬ a : (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-17 02:06:25 +00:00
|
|
|
|
theorem symm {A : TypeU} {a b : A} (H : a = b) : b = a
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= subst (refl a) H
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-27 02:52:41 +00:00
|
|
|
|
theorem eqmpr {a b : Bool} (H1 : a = b) (H2 : b) : a
|
|
|
|
|
:= (symm H1) ◂ H2
|
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
theorem trans {A : TypeU} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= subst H1 H2
|
2014-01-03 06:47:45 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem ne_symm {A : TypeU} {a b : A} (H : a ≠ b) : b ≠ a
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume H1 : b = a, H (symm H1)
|
2014-01-03 06:47:45 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem eq_ne_trans {A : TypeU} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= subst H2 (symm H1)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem ne_eq_trans {A : TypeU} {a b c : A} (H1 : a ≠ b) (H2 : b = c) : a ≠ c
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= subst H1 H2
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem eqt_elim {a : Bool} (H : a = true) : a
|
2014-01-06 05:39:31 +00:00
|
|
|
|
:= (symm H) ◂ trivial
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-15 22:42:00 +00:00
|
|
|
|
theorem eqf_elim {a : Bool} (H : a = false) : ¬ a
|
2014-01-11 19:13:54 +00:00
|
|
|
|
:= not_intro (λ Ha : a, H ◂ Ha)
|
|
|
|
|
|
2014-01-18 10:47:11 +00:00
|
|
|
|
theorem congr1 {A B : TypeU} {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
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
theorem congr2 {A B : TypeU} {a b : A} (f : A → B) (H : a = b) : f a = f b
|
2014-01-16 00:34:23 +00:00
|
|
|
|
:= substp (fun x : A, f a = f x) (refl (f a)) H
|
|
|
|
|
|
2014-01-17 02:06:25 +00:00
|
|
|
|
theorem congr {A B : TypeU} {f g : A → B} {a b : A} (H1 : f = g) (H2 : a = b) : f a = g b
|
2014-01-16 23:07:51 +00:00
|
|
|
|
:= subst (congr2 f H2) (congr1 b H1)
|
2014-01-16 00:34:23 +00:00
|
|
|
|
|
2014-01-09 00:52:43 +00:00
|
|
|
|
-- Recall that exists is defined as ¬ ∀ x : A, ¬ P x
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem exists_elim {A : TypeU} {P : A → Bool} {B : Bool} (H1 : Exists A P) (H2 : ∀ (a : A) (H : P a), B) : B
|
2014-01-06 03:10:21 +00:00
|
|
|
|
:= refute (λ R : ¬ B,
|
2014-01-12 02:36:37 +00:00
|
|
|
|
absurd (take a : A, mt (assume H : P a, H2 a H) R)
|
2014-01-05 20:05:08 +00:00
|
|
|
|
H1)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem exists_intro {A : TypeU} {P : A → Bool} (a : A) (H : P a) : Exists A P
|
2014-01-12 02:36:37 +00:00
|
|
|
|
:= assume H1 : (∀ x : A, ¬ P x),
|
2014-01-08 08:38:39 +00:00
|
|
|
|
absurd H (H1 a)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-25 02:52:50 +00:00
|
|
|
|
theorem nonempty_elim {A : TypeU} (H1 : nonempty A) {B : Bool} (H2 : A → B) : B
|
|
|
|
|
:= obtain (w : A) (Hw : true), from H1,
|
|
|
|
|
H2 w
|
|
|
|
|
|
2014-01-22 16:52:31 +00:00
|
|
|
|
theorem nonempty_ex_intro {A : TypeU} {P : A → Bool} (H : ∃ x, P x) : nonempty A
|
2014-01-22 17:41:07 +00:00
|
|
|
|
:= obtain (w : A) (Hw : P w), from H,
|
|
|
|
|
exists_intro w trivial
|
2014-01-22 16:52:31 +00:00
|
|
|
|
|
|
|
|
|
theorem exists_to_eps {A : TypeU} {P : A → Bool} (H : ∃ x, P x) : P (ε (nonempty_ex_intro H) P)
|
2014-01-22 17:41:07 +00:00
|
|
|
|
:= obtain (w : A) (Hw : P w), from H,
|
feat(library/elaborator): compensate the lack of eta-reduction (and eta-expanded normal forms) in the kernel normalizer
Before this commit, the elaborator was solving constraints of the form
ctx |- (?m x) == (f x)
as
?m <- (fun x : A, f x) where A is the domain of f.
In our kernel, the terms f and (fun x, f x) are not definitionally equal.
So, the solution above is not the only one. Another possible solution is
?m <- f
Depending of the circumstances we want ?m <- (fun x : A, f x) OR ?m <- f.
For example, when Lean is elaborating the eta-theorem in kernel.lean, the first solution should be used:
?m <- (fun x : A, f x)
When we are elaborating the axiom_of_choice theorem, we need to use the second one:
?m <- f
Of course, we can always provide the parameters explicitly and bypass the elaborator.
However, this goes against the idea that the elaborator can do mechanical steps for us.
This commit addresses this issue by creating a case-split
?m <- (fun x : A, f x)
OR
?m <- f
Another solution is to implement eta-expanded normal forms in the Kernel.
With this change, we were able to cleanup the following "hacks" in kernel.lean:
@eps_ax A (nonempty_ex_intro H) P w Hw
@axiom_of_choice A B P H
where we had to explicitly provided the implicit arguments
This commit also improves the imitation step for Pi-terms that are actually arrows.
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-01-22 21:18:27 +00:00
|
|
|
|
eps_ax (nonempty_ex_intro H) w Hw
|
2014-01-22 16:21:11 +00:00
|
|
|
|
|
2014-01-22 19:04:27 +00:00
|
|
|
|
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)
|
2014-01-22 16:21:11 +00:00
|
|
|
|
:= exists_intro
|
2014-01-22 17:18:40 +00:00
|
|
|
|
(λ x, ε (nonempty_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)
|
2014-01-22 16:21:11 +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-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-27 03:50:12 +00:00
|
|
|
|
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)
|
|
|
|
|
|
2014-01-22 19:04:27 +00:00
|
|
|
|
theorem skolem_th {A : TypeU} {B : A → TypeU} {P : ∀ x : A, B x → Bool} :
|
2014-01-22 17:41:07 +00:00
|
|
|
|
(∀ x, ∃ y, P x y) ↔ ∃ f, (∀ x, P x (f x))
|
|
|
|
|
:= iff_intro
|
feat(library/elaborator): compensate the lack of eta-reduction (and eta-expanded normal forms) in the kernel normalizer
Before this commit, the elaborator was solving constraints of the form
ctx |- (?m x) == (f x)
as
?m <- (fun x : A, f x) where A is the domain of f.
In our kernel, the terms f and (fun x, f x) are not definitionally equal.
So, the solution above is not the only one. Another possible solution is
?m <- f
Depending of the circumstances we want ?m <- (fun x : A, f x) OR ?m <- f.
For example, when Lean is elaborating the eta-theorem in kernel.lean, the first solution should be used:
?m <- (fun x : A, f x)
When we are elaborating the axiom_of_choice theorem, we need to use the second one:
?m <- f
Of course, we can always provide the parameters explicitly and bypass the elaborator.
However, this goes against the idea that the elaborator can do mechanical steps for us.
This commit addresses this issue by creating a case-split
?m <- (fun x : A, f x)
OR
?m <- f
Another solution is to implement eta-expanded normal forms in the Kernel.
With this change, we were able to cleanup the following "hacks" in kernel.lean:
@eps_ax A (nonempty_ex_intro H) P w Hw
@axiom_of_choice A B P H
where we had to explicitly provided the implicit arguments
This commit also improves the imitation step for Pi-terms that are actually arrows.
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-01-22 21:18:27 +00:00
|
|
|
|
(λ H : (∀ x, ∃ y, P x y), axiom_of_choice H)
|
2014-01-22 17:41:07 +00:00
|
|
|
|
(λ H : (∃ f, (∀ x, P x (f x))),
|
2014-01-22 19:04:27 +00:00
|
|
|
|
take x, obtain (fw : ∀ x, B x) (Hw : ∀ x, P x (fw x)), from H,
|
2014-01-22 17:41:07 +00:00
|
|
|
|
exists_intro (fw x) (Hw x))
|
|
|
|
|
|
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-01-16 17:00:01 +00:00
|
|
|
|
theorem neq_elim {A : TypeU} {a b : A} (H : a ≠ b) : a = b ↔ false
|
2014-01-15 23:30:16 +00:00
|
|
|
|
:= eqf_intro H
|
|
|
|
|
|
2014-01-19 18:34:18 +00:00
|
|
|
|
theorem eq_id {A : TypeU} (a : A) : (a = a) ↔ true
|
|
|
|
|
:= 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-01-31 03:11:58 +00:00
|
|
|
|
add_rewrite eq_id iff_id
|
|
|
|
|
|
|
|
|
|
-- Remark: ordered rewriting + assoc + comm + left_comm sorts a term lexicographically
|
|
|
|
|
theorem left_comm {A : TypeU} {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)) :
|
|
|
|
|
∀ 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-01-09 16:33:52 +00:00
|
|
|
|
:= eqt_intro (case (λ x : Bool, true ∨ x) trivial 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-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
|
|
|
|
|
:= case (λ x, (x → true) ↔ true) trivial trivial a
|
2014-01-11 21:35:20 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem imp_truel (a : Bool) : (true → a) ↔ a
|
|
|
|
|
:= case (λ x, (true → x) ↔ x) trivial trivial a
|
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
|
|
|
|
|
:= case (λ x, (false → x) ↔ true) trivial trivial a
|
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
|
|
|
|
|
:= iff_intro
|
|
|
|
|
(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-20 07:26:34 +00:00
|
|
|
|
theorem not_true : ¬ true ↔ false
|
|
|
|
|
:= trivial
|
|
|
|
|
|
|
|
|
|
theorem not_false : ¬ false ↔ true
|
|
|
|
|
:= trivial
|
|
|
|
|
|
2014-01-23 20:54:25 +00:00
|
|
|
|
theorem not_neq {A : TypeU} (a b : A) : ¬ (a ≠ b) ↔ a = b
|
|
|
|
|
:= not_not_eq (a = b)
|
|
|
|
|
|
2014-01-31 03:11:58 +00:00
|
|
|
|
add_rewrite not_true not_false not_neq
|
|
|
|
|
|
2014-01-23 20:54:25 +00:00
|
|
|
|
theorem not_neq_elim {A : TypeU} {a b : A} (H : ¬ (a ≠ b)) : a = b
|
|
|
|
|
:= (not_neq a b) ◂ H
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_and (a b : Bool) : ¬ (a ∧ b) ↔ ¬ a ∨ ¬ b
|
|
|
|
|
:= case (λ x, ¬ (x ∧ b) ↔ ¬ x ∨ ¬ b)
|
|
|
|
|
(case (λ y, ¬ (true ∧ y) ↔ ¬ true ∨ ¬ y) trivial trivial b)
|
|
|
|
|
(case (λ y, ¬ (false ∧ y) ↔ ¬ false ∨ ¬ y) trivial trivial b)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
a
|
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_and_elim {a b : Bool} (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b
|
|
|
|
|
:= (not_and a b) ◂ H
|
2014-01-01 19:35:21 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_or (a b : Bool) : ¬ (a ∨ b) ↔ ¬ a ∧ ¬ b
|
|
|
|
|
:= case (λ x, ¬ (x ∨ b) ↔ ¬ x ∧ ¬ b)
|
|
|
|
|
(case (λ y, ¬ (true ∨ y) ↔ ¬ true ∧ ¬ y) trivial trivial b)
|
|
|
|
|
(case (λ y, ¬ (false ∨ y) ↔ ¬ false ∧ ¬ y) trivial trivial b)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
a
|
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_or_elim {a b : Bool} (H : ¬ (a ∨ b)) : ¬ a ∧ ¬ b
|
|
|
|
|
:= (not_or a b) ◂ H
|
2014-01-01 19:35:21 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_iff (a b : Bool) : ¬ (a ↔ b) ↔ (¬ a ↔ b)
|
|
|
|
|
:= case (λ x, ¬ (x ↔ b) ↔ ((¬ x) ↔ b))
|
|
|
|
|
(case (λ y, ¬ (true ↔ y) ↔ ((¬ true) ↔ y)) trivial trivial b)
|
|
|
|
|
(case (λ y, ¬ (false ↔ y) ↔ ((¬ false) ↔ y)) trivial trivial b)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
a
|
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_iff_elim {a b : Bool} (H : ¬ (a ↔ b)) : (¬ a) ↔ b
|
2014-01-09 16:33:52 +00:00
|
|
|
|
:= (not_iff a b) ◂ H
|
2014-01-01 19:35:21 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_implies (a b : Bool) : ¬ (a → b) ↔ a ∧ ¬ b
|
|
|
|
|
:= case (λ x, ¬ (x → b) ↔ x ∧ ¬ b)
|
|
|
|
|
(case (λ y, ¬ (true → y) ↔ true ∧ ¬ y) trivial trivial b)
|
|
|
|
|
(case (λ y, ¬ (false → y) ↔ false ∧ ¬ y) trivial trivial b)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
a
|
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_implies_elim {a b : Bool} (H : ¬ (a → b)) : a ∧ ¬ b
|
|
|
|
|
:= (not_implies a b) ◂ H
|
2014-01-01 19:35:21 +00:00
|
|
|
|
|
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-01-25 02:52:50 +00:00
|
|
|
|
theorem exists_rem {A : TypeU} (H : nonempty 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,
|
|
|
|
|
nonempty_elim H (λ w, exists_intro w Hr))
|
|
|
|
|
|
|
|
|
|
theorem forall_rem {A : TypeU} (H : nonempty A) (p : Bool) : (∀ x : A, p) ↔ p
|
|
|
|
|
:= iff_intro
|
|
|
|
|
(assume Hl : (∀ x : A, p),
|
|
|
|
|
nonempty_elim H (λ w, Hl w))
|
|
|
|
|
(assume Hr : p,
|
|
|
|
|
take x, Hr)
|
|
|
|
|
|
2014-01-16 17:00:01 +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)
|
2014-01-09 01:25:14 +00:00
|
|
|
|
:= congr2 (Exists A) (funext H)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-16 17:00:01 +00:00
|
|
|
|
theorem not_forall (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 : not_congr (allext (λ x : A, symm (not_not_eq (P x))))
|
|
|
|
|
... = ∃ x : A, ¬ P x : refl (∃ x : A, ¬ P x)
|
2013-12-29 10:44:49 +00:00
|
|
|
|
|
2014-01-09 16:33:52 +00:00
|
|
|
|
theorem not_forall_elim {A : (Type U)} {P : A → Bool} (H : ¬ (∀ x : A, P x)) : ∃ x : A, ¬ P x
|
|
|
|
|
:= (not_forall A P) ◂ H
|
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-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-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-01-14 02:21:22 +00:00
|
|
|
|
:= calc (∀ x, φ x ∨ p) = (∀ x, p ∨ φ x) : allext (λ x, or_comm (φ x) p)
|
2014-01-16 23:07:51 +00:00
|
|
|
|
... = (p ∨ ∀ x, φ x) : forall_or_distributer p φ
|
|
|
|
|
... = ((∀ x, φ x) ∨ p) : or_comm p (∀ x, φ x)
|
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_and_distributel {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, φ x ∧ p) ↔ (∃ x, φ x) ∧ p
|
2014-01-14 02:21:22 +00:00
|
|
|
|
:= 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-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-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-01-31 03:11:58 +00:00
|
|
|
|
theorem if_true {A : TypeU} (a b : A) : (if true then a else b) = a
|
|
|
|
|
:= calc (if true then a else b) = ε (nonempty_intro a) (λ r, (true → r = a) ∧ (¬ true → r = b)) : refl (if true then a else b)
|
|
|
|
|
... = ε (nonempty_intro a) (λ r, r = a) : by simp
|
|
|
|
|
... = a : eps_singleton (nonempty_intro a) a
|
|
|
|
|
|
|
|
|
|
theorem if_false {A : TypeU} (a b : A) : (if false then a else b) = b
|
|
|
|
|
:= calc (if false then a else b) = ε (nonempty_intro a) (λ r, (false → r = a) ∧ (¬ false → r = b)) : refl (if false then a else b)
|
|
|
|
|
... = ε (nonempty_intro a) (λ r, r = b) : by simp
|
|
|
|
|
... = b : eps_singleton (nonempty_intro a) b
|
|
|
|
|
|
|
|
|
|
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)
|
|
|
|
|
|
|
|
|
|
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
|
|
|
|
|
set_opaque ite true
|