2014-08-01 00:48:51 +00:00
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
-- Released under Apache 2.0 license as described in the file LICENSE.
2014-09-26 23:36:04 +00:00
-- Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn
2014-08-22 23:36:47 +00:00
-- logic.connectives.eq
-- ====================
-- Equality.
2014-08-01 00:48:51 +00:00
2014-08-28 00:46:07 +00:00
import .prop
2014-08-01 00:48:51 +00:00
-- eq
-- --
inductive eq {A : Type} (a : A) : A → Prop :=
2014-08-22 22:46:10 +00:00
refl : eq a a
2014-08-01 00:48:51 +00:00
2014-08-22 23:36:47 +00:00
infix `=` := eq
2014-09-17 21:39:05 +00:00
definition rfl {A : Type} {a : A} := eq.refl a
2014-08-17 21:41:23 +00:00
2014-09-26 23:36:04 +00:00
-- proof irrelevance is built in
theorem proof_irrel {a : Prop} {H1 H2 : a} : H1 = H2 := rfl
2014-09-05 01:41:06 +00:00
namespace eq
2014-09-10 02:15:11 +00:00
theorem id_refl {A : Type} {a : A} (H1 : a = a) : H1 = (eq.refl a) :=
2014-09-26 23:36:04 +00:00
proof_irrel
2014-08-17 21:41:23 +00:00
2014-09-10 02:15:11 +00:00
theorem irrel {A : Type} {a b : A} (H1 H2 : a = b) : H1 = H2 :=
2014-09-26 23:36:04 +00:00
proof_irrel
2014-08-15 03:12:54 +00:00
2014-09-10 02:15:11 +00:00
theorem subst {A : Type} {a b : A} {P : A → Prop} (H1 : a = b) (H2 : P a) : P b :=
rec H2 H1
2014-08-01 00:48:51 +00:00
2014-09-10 02:15:11 +00:00
theorem trans {A : Type} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c :=
subst H2 H1
2014-08-01 00:48:51 +00:00
2014-09-10 02:15:11 +00:00
theorem symm {A : Type} {a b : A} (H : a = b) : b = a :=
subst H (refl a)
2014-10-02 00:51:17 +00:00
namespace ops
postfix `⁻¹` := symm
infixr `⬝` := trans
infixr `▸` := subst
end ops
2014-09-05 01:41:06 +00:00
end eq
calc_subst eq.subst
calc_refl eq.refl
calc_trans eq.trans
2014-08-15 03:12:54 +00:00
2014-10-02 00:51:17 +00:00
open eq.ops
2014-08-27 00:30:27 +00:00
2014-09-04 23:36:06 +00:00
namespace eq
2014-09-10 02:15:11 +00:00
-- eq_rec with arguments swapped, for transporting an element of a dependent type
2014-09-26 23:36:04 +00:00
-- definition rec_on {A : Type} {a1 a2 : A} {B : A → Type} (H1 : a1 = a2) (H2 : B a1) : B a2 :=
-- eq.rec H2 H1
definition rec_on {A : Type} {a a' : A} {B : Πa' : A, a = a' → Type} (H1 : a = a') (H2 : B a (refl a)) : B a' H1 :=
eq.rec (λH1 : a = a, show B a H1, from H2) H1 H1
theorem rec_on_id {A : Type} {a : A} {B : Πa' : A, a = a' → Type} (H : a = a) (b : B a H) : rec_on H b = b :=
2014-09-10 02:15:11 +00:00
refl (rec_on rfl b)
2014-09-26 23:36:04 +00:00
theorem rec_on_constant {A : Type} {a a' : A} {B : Type} (H : a = a') (b : B) : rec_on H b = b :=
rec_on H (λ(H' : a = a), rec_on_id H' b) H
theorem rec_on_constant2 {A : Type} {a₁ a₂ a₃ a₄ : A} {B : Type} (H₁ : a₁ = a₂) (H₂ : a₃ = a₄) (b : B) : rec_on H₁ b = rec_on H₂ b :=
rec_on_constant H₁ b ⬝ rec_on_constant H₂ b ⁻¹
theorem rec_on_irrel {A B : Type} {a a' : A} {f : A → B} {D : B → Type} (H : a = a') (H' : f a = f a') (b : D (f a)) : rec_on H b = rec_on H' b :=
rec_on H (λ(H : a = a) (H' : f a = f a), rec_on_id H b ⬝ rec_on_id H' b⁻¹) H H'
2014-09-10 02:15:11 +00:00
theorem rec_id {A : Type} {a : A} {B : A → Type} (H : a = a) (b : B a) : rec b H = b :=
2014-09-26 23:36:04 +00:00
id_refl H⁻¹ ▸ refl (eq.rec b (refl a))
2014-09-10 02:15:11 +00:00
theorem rec_on_compose {A : Type} {a b c : A} {P : A → Type} (H1 : a = b) (H2 : b = c)
(u : P a) :
rec_on H2 (rec_on H1 u) = rec_on (trans H1 H2) u :=
(show ∀(H2 : b = c), rec_on H2 (rec_on H1 u) = rec_on (trans H1 H2) u,
from rec_on H2 (take (H2 : b = b), rec_on_id H2 _))
H2
2014-09-04 23:36:06 +00:00
end eq
2014-08-01 00:48:51 +00:00
2014-09-26 23:36:04 +00:00
open eq
2014-08-20 02:32:44 +00:00
theorem congr_fun {A : Type} {B : A → Type} {f g : Π x, B x} (H : f = g) (a : A) : f a = g a :=
2014-09-10 02:15:11 +00:00
H ▸ rfl
2014-08-01 00:48:51 +00:00
2014-08-20 02:32:44 +00:00
theorem congr_arg {A : Type} {B : Type} {a b : A} (f : A → B) (H : a = b) : f a = f b :=
2014-09-10 02:15:11 +00:00
H ▸ rfl
2014-08-01 00:48:51 +00:00
2014-09-26 23:36:04 +00:00
theorem congr {A : Type} {B : Type} {f g : A → B} {a b : A} (H1 : f = g) (H2 : a = b) :
f a = g b :=
H1 ▸ H2 ▸ rfl
theorem congr_arg2 {A B C : Type} {a a' : A} {b b' : B} (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
congr (congr_arg f Ha) Hb
theorem congr_arg3 {A B C D : Type} {a a' : A} {b b' : B} {c c' : C} (f : A → B → C → D) (Ha : a = a') (Hb : b = b') (Hc : c = c') : f a b c = f a' b' c' :=
congr (congr_arg2 f Ha Hb) Hc
theorem congr_arg4 {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D} (f : A → B → C → D → E) (Ha : a = a') (Hb : b = b') (Hc : c = c') (Hd : d = d') : f a b c d = f a' b' c' d' :=
congr (congr_arg3 f Ha Hb Hc) Hd
theorem congr_arg5 {A B C D E F : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D} {e e' : E} (f : A → B → C → D → E → F) (Ha : a = a') (Hb : b = b') (Hc : c = c') (Hd : d = d') (He : e = e') : f a b c d e = f a' b' c' d' e' :=
congr (congr_arg4 f Ha Hb Hc Hd) He
theorem congr2 {A B C : Type} {a a' : A} {b b' : B} (f f' : A → B → C) (Hf : f = f') (Ha : a = a') (Hb : b = b') : f a b = f' a' b' :=
Hf ▸ congr_arg2 f Ha Hb
theorem congr3 {A B C D : Type} {a a' : A} {b b' : B} {c c' : C} (f f' : A → B → C → D) (Hf : f = f') (Ha : a = a') (Hb : b = b') (Hc : c = c') : f a b c = f' a' b' c' :=
Hf ▸ congr_arg3 f Ha Hb Hc
theorem congr4 {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D} (f f' : A → B → C → D → E) (Hf : f = f') (Ha : a = a') (Hb : b = b') (Hc : c = c') (Hd : d = d') : f a b c d = f' a' b' c' d' :=
Hf ▸ congr_arg4 f Ha Hb Hc Hd
theorem congr5 {A B C D E F : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D} {e e' : E} (f f' : A → B → C → D → E → F) (Hf : f = f') (Ha : a = a') (Hb : b = b') (Hc : c = c') (Hd : d = d') (He : e = e') : f a b c d e = f' a' b' c' d' e' :=
Hf ▸ congr_arg5 f Ha Hb Hc Hd He
theorem congr_arg2_dep {A : Type} {B : A → Type} {C : Type} {a₁ a₂ : A}
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
{b₁ : B a₁} {b₂ : B a₂} (f : Πa, B a → C) (H₁ : a₁ = a₂) (H₂ : eq.rec_on H₁ b₁ = b₂) :
f a₁ b₁ = f a₂ b₂ :=
eq.rec_on H₁
(λ (b₂ : B a₁) (H₁ : a₁ = a₁) (H₂ : eq.rec_on H₁ b₁ = b₂),
calc
f a₁ b₁ = f a₁ (eq.rec_on H₁ b₁) : {(eq.rec_on_id H₁ b₁)⁻¹}
... = f a₁ b₂ : {H₂})
b₂ H₁ H₂
2014-09-26 23:36:04 +00:00
theorem congr_arg3_dep {A : Type} {B : A → Type} {C : Πa, B a → Type} {D : Type} {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} {c₁ : C a₁ b₁} {c₂ : C a₂ b₂} (f : Πa b, C a b → D)
(H₁ : a₁ = a₂) (H₂ : eq.rec_on H₁ b₁ = b₂) (H₃ : eq.rec_on (congr_arg2_dep C H₁ H₂) c₁ = c₂) :
f a₁ b₁ c₁ = f a₂ b₂ c₂ :=
eq.rec_on H₁
(λ (b₂ : B a₁) (H₂ : b₁ = b₂) (c₂ : C a₁ b₂)
(H₃ : (rec_on (congr_arg2_dep C (refl a₁) H₂) c₁) = c₂),
have H₃' : eq.rec_on H₂ c₁ = c₂,
from (rec_on_irrel H₂ (congr_arg2_dep C (refl a₁) H₂) c₁⁻¹) ▸ H₃,
congr_arg2_dep (f a₁) H₂ H₃')
b₂ H₂ c₂ H₃
theorem congr_arg3_ndep_dep {A B : Type} {C : A → B → Type} {D : Type} {a₁ a₂ : A} {b₁ b₂ : B} {c₁ : C a₁ b₁} {c₂ : C a₂ b₂} (f : Πa b, C a b → D)
(H₁ : a₁ = a₂) (H₂ : b₁ = b₂) (H₃ : eq.rec_on (congr_arg2 C H₁ H₂) c₁ = c₂) :
f a₁ b₁ c₁ = f a₂ b₂ c₂ :=
congr_arg3_dep f H₁ (rec_on_constant H₁ b₁ ⬝ H₂) H₃
2014-08-01 00:48:51 +00:00
theorem equal_f {A : Type} {B : A → Type} {f g : Π x, B x} (H : f = g) : ∀x, f x = g x :=
2014-08-20 02:32:44 +00:00
take x, congr_fun H x
2014-08-01 00:48:51 +00:00
theorem eqmp {a b : Prop} (H1 : a = b) (H2 : a) : b :=
H1 ▸ H2
theorem eqmpr {a b : Prop} (H1 : a = b) (H2 : b) : a :=
2014-08-22 23:36:47 +00:00
H1⁻¹ ▸ H2
2014-08-01 00:48:51 +00:00
2014-08-30 03:45:57 +00:00
theorem eq_true_elim {a : Prop} (H : a = true) : a :=
2014-08-22 23:36:47 +00:00
H⁻¹ ▸ trivial
2014-08-01 00:48:51 +00:00
2014-08-30 03:45:57 +00:00
theorem eq_false_elim {a : Prop} (H : a = false) : ¬a :=
2014-08-22 23:36:47 +00:00
assume Ha : a, H ▸ Ha
2014-08-01 00:48:51 +00:00
theorem imp_trans {a b c : Prop} (H1 : a → b) (H2 : b → c) : a → c :=
assume Ha, H2 (H1 Ha)
theorem imp_eq_trans {a b c : Prop} (H1 : a → b) (H2 : b = c) : a → c :=
2014-08-22 23:36:47 +00:00
assume Ha, H2 ▸ (H1 Ha)
2014-08-01 00:48:51 +00:00
theorem eq_imp_trans {a b c : Prop} (H1 : a = b) (H2 : b → c) : a → c :=
2014-08-22 23:36:47 +00:00
assume Ha, H2 (H1 ▸ Ha)
2014-08-01 00:48:51 +00:00
-- ne
-- --
2014-09-17 21:39:05 +00:00
definition ne {A : Type} (a b : A) := ¬(a = b)
2014-08-22 23:36:47 +00:00
infix `≠` := ne
2014-08-01 00:48:51 +00:00
2014-09-05 01:41:06 +00:00
namespace ne
2014-09-10 02:15:11 +00:00
theorem intro {A : Type} {a b : A} (H : a = b → false) : a ≠ b :=
H
2014-08-01 00:48:51 +00:00
2014-09-10 02:15:11 +00:00
theorem elim {A : Type} {a b : A} (H1 : a ≠ b) (H2 : a = b) : false :=
H1 H2
2014-08-01 00:48:51 +00:00
2014-09-10 02:15:11 +00:00
theorem irrefl {A : Type} {a : A} (H : a ≠ a) : false :=
H rfl
2014-08-01 00:48:51 +00:00
2014-09-10 02:15:11 +00:00
theorem symm {A : Type} {a b : A} (H : a ≠ b) : b ≠ a :=
assume H1 : b = a, H (H1⁻¹)
2014-09-05 01:41:06 +00:00
end ne
theorem a_neq_a_elim {A : Type} {a : A} (H : a ≠ a) : false :=
H rfl
2014-08-01 00:48:51 +00:00
2014-08-04 05:58:12 +00:00
theorem eq_ne_trans {A : Type} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c :=
H1⁻¹ ▸ H2
2014-08-01 00:48:51 +00:00
2014-08-04 05:58:12 +00:00
theorem ne_eq_trans {A : Type} {a b c : A} (H1 : a ≠ b) (H2 : b = c) : a ≠ c :=
H2 ▸ H1
2014-08-01 00:48:51 +00:00
calc_trans eq_ne_trans
2014-08-04 05:58:12 +00:00
calc_trans ne_eq_trans
theorem p_ne_false {p : Prop} (Hp : p) : p ≠ false :=
assume Heq : p = false, Heq ▸ Hp
theorem p_ne_true {p : Prop} (Hnp : ¬p) : p ≠ true :=
2014-08-28 01:34:09 +00:00
assume Heq : p = true, absurd trivial (Heq ▸ Hnp)
2014-09-05 01:41:06 +00:00
theorem true_ne_false : ¬true = false :=
assume H : true = false,
H ▸ trivial