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-10-21 21:08:07 +00:00
|
|
|
import general_notation .prop
|
2014-08-22 23:36:47 +00:00
|
|
|
|
2014-10-05 18:11:48 +00:00
|
|
|
-- logic.eq
|
2014-08-22 23:36:47 +00:00
|
|
|
-- ====================
|
|
|
|
|
|
|
|
-- Equality.
|
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-10-21 21:08:07 +00:00
|
|
|
notation a = b := eq a b
|
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
|
2014-10-09 01:41:18 +00:00
|
|
|
theorem proof_irrel {a : Prop} (H₁ H₂ : a) : H₁ = H₂ :=
|
2014-10-05 17:19:50 +00:00
|
|
|
rfl
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-09-05 01:41:06 +00:00
|
|
|
namespace eq
|
2014-10-03 01:25:00 +00:00
|
|
|
variables {A : Type}
|
|
|
|
variables {a b c : A}
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem id_refl (H₁ : a = a) : H₁ = (eq.refl a) :=
|
2014-10-10 21:52:21 +00:00
|
|
|
rfl
|
2014-08-17 21:41:23 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem irrel (H₁ H₂ : a = b) : H₁ = H₂ :=
|
2014-10-09 01:41:18 +00:00
|
|
|
!proof_irrel
|
2014-08-15 03:12:54 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem subst {P : A → Prop} (H₁ : a = b) (H₂ : P a) : P b :=
|
|
|
|
rec H₂ H₁
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem trans (H₁ : a = b) (H₂ : b = c) : a = c :=
|
|
|
|
subst H₂ H₁
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-03 01:25:00 +00:00
|
|
|
theorem symm (H : a = b) : b = a :=
|
2014-09-10 02:15:11 +00:00
|
|
|
subst H (refl a)
|
2014-10-10 23:33:58 +00:00
|
|
|
|
|
|
|
namespace ops
|
2014-10-21 21:08:07 +00:00
|
|
|
notation H `⁻¹` := symm H
|
|
|
|
notation H1 ⬝ H2 := trans H1 H2
|
|
|
|
notation H1 ▸ H2 := subst H1 H2
|
2014-10-10 23:33:58 +00:00
|
|
|
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-10-05 17:19:50 +00:00
|
|
|
definition rec_on {A : Type} {a a' : A} {B : Πa' : A, a = a' → Type} (H₁ : a = a') (H₂ : B a (refl a)) : B a' H₁ :=
|
|
|
|
eq.rec (λH₁ : a = a, show B a H₁, from H₂) H₁ H₁
|
2014-09-26 23:36:04 +00:00
|
|
|
|
|
|
|
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-10-10 21:52:21 +00:00
|
|
|
rfl
|
2014-09-10 02:15:11 +00:00
|
|
|
|
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
|
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
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)⁻¹
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
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 :=
|
2014-09-26 23:36:04 +00:00
|
|
|
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-10-10 21:52:21 +00:00
|
|
|
rfl
|
2014-09-10 02:15:11 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem rec_on_compose {A : Type} {a b c : A} {P : A → Type} (H₁ : a = b) (H₂ : b = c)
|
2014-09-10 02:15:11 +00:00
|
|
|
(u : P a) :
|
2014-10-05 17:19:50 +00:00
|
|
|
rec_on H₂ (rec_on H₁ u) = rec_on (trans H₁ H₂) u :=
|
|
|
|
(show ∀ H₂ : b = c, rec_on H₂ (rec_on H₁ u) = rec_on (trans H₁ H₂) u,
|
|
|
|
from rec_on H₂ (take (H₂ : b = b), rec_on_id H₂ _))
|
|
|
|
H₂
|
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-10-05 17:19:50 +00:00
|
|
|
section
|
|
|
|
variables {A B C D E F : Type}
|
|
|
|
variables {a a' : A} {b b' : B} {c c' : C} {d d' : D} {e e' : E} {f f' : F}
|
|
|
|
|
|
|
|
theorem congr_fun {B : A → Type} {f g : Π x, B x} (H : f = g) (a : A) : f a = g a :=
|
|
|
|
H ▸ rfl
|
|
|
|
|
|
|
|
theorem congr_arg (f : A → B) (H : a = a') : f a = f a' :=
|
|
|
|
H ▸ rfl
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr {f g : A → B} (H₁ : f = g) (H₂ : a = a') : f a = g a' :=
|
|
|
|
H₁ ▸ H₂ ▸ rfl
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr_arg2 (f : A → B → C) (Ha : a = a') (Hb : b = b') : f a b = f a' b' :=
|
|
|
|
congr (congr_arg f Ha) Hb
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr_arg3 (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
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr_arg4 (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
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr_arg5 (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
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr2 (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
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr3 (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
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr4 (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
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem congr5 (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
|
|
|
|
end
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
section
|
2014-10-09 01:41:18 +00:00
|
|
|
variables {A : Type} {B : A → Type} {C : Πa, B a → Type} {D : Πa b, C a b → Type} {R : Type}
|
2014-10-10 21:52:21 +00:00
|
|
|
variables {a₁ a₂ : A}
|
|
|
|
{b₁ : B a₁} {b₂ : B a₂}
|
|
|
|
{c₁ : C a₁ b₁} {c₂ : C a₂ b₂}
|
2014-10-09 01:41:18 +00:00
|
|
|
{d₁ : D a₁ b₁ c₁} {d₂ : D a₂ b₂ c₂}
|
|
|
|
|
2014-10-10 21:52:21 +00:00
|
|
|
theorem congr_arg2_dep (f : Πa, B a → R) (H₁ : a₁ = a₂) (H₂ : eq.rec_on H₁ b₁ = b₂)
|
2014-10-09 01:41:18 +00:00
|
|
|
: f a₁ b₁ = f a₂ b₂ :=
|
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
|
|
|
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-10-05 17:19:50 +00:00
|
|
|
theorem congr_arg3_dep (f : Πa b, C a b → R) (H₁ : a₁ = a₂) (H₂ : eq.rec_on H₁ b₁ = b₂)
|
2014-10-09 01:41:18 +00:00
|
|
|
(H₃ : eq.rec_on (congr_arg2_dep C H₁ H₂) c₁ = c₂) : f a₁ b₁ c₁ = f a₂ b₂ c₂ :=
|
2014-10-05 17:19:50 +00:00
|
|
|
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₂),
|
2014-10-09 01:41:18 +00:00
|
|
|
have H₃' : eq.rec_on H₂ c₁ = c₂, from rec_on_irrel H₂ _ c₁ ⬝ H₃,
|
2014-10-05 17:19:50 +00:00
|
|
|
congr_arg2_dep (f a₁) H₂ H₃')
|
|
|
|
b₂ H₂ c₂ H₃
|
2014-10-09 01:41:18 +00:00
|
|
|
|
|
|
|
-- for the moment the following theorem is commented out, because it takes long to prove
|
|
|
|
-- theorem congr_arg4_dep (f : Πa b c, D a b c → R) (H₁ : a₁ = a₂) (H₂ : eq.rec_on H₁ b₁ = b₂)
|
2014-10-10 21:52:21 +00:00
|
|
|
-- (H₃ : eq.rec_on (congr_arg2_dep C H₁ H₂) c₁ = c₂)
|
2014-10-09 01:41:18 +00:00
|
|
|
-- (H₄ : eq.rec_on (congr_arg3_dep D H₁ H₂ H₃) d₁ = d₂) : f a₁ b₁ c₁ d₁ = f a₂ b₂ c₂ d₂ :=
|
|
|
|
-- eq.rec_on H₁
|
|
|
|
-- (λ b₂ H₂ c₂ H₃ d₂ (H₄ : _),
|
|
|
|
-- have H₃' [visible] : eq.rec_on H₂ c₁ = c₂, from rec_on_irrel H₂ _ c₁ ⬝ H₃,
|
|
|
|
-- have H₄' : rec_on (congr_arg2_dep (D a₁) H₂ H₃') d₁ = d₂, from rec_on_irrel _ _ d₁ ⬝ H₄,
|
|
|
|
-- congr_arg3_dep (f a₁) H₂ H₃' H₄')
|
|
|
|
-- b₂ H₂ c₂ H₃ d₂ H₄
|
2014-10-05 17:19:50 +00:00
|
|
|
end
|
|
|
|
|
|
|
|
section
|
|
|
|
variables {A B : Type} {C : A → B → Type} {R : Type}
|
|
|
|
variables {a₁ a₂ : A} {b₁ b₂ : B} {c₁ : C a₁ b₁} {c₂ : C a₂ b₂}
|
|
|
|
theorem congr_arg3_ndep_dep (f : Πa b, C a b → R) (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₃
|
|
|
|
end
|
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
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
section
|
|
|
|
variables {a b c : Prop}
|
|
|
|
|
|
|
|
theorem eqmp (H₁ : a = b) (H₂ : a) : b :=
|
|
|
|
H₁ ▸ H₂
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem eqmpr (H₁ : a = b) (H₂ : b) : a :=
|
|
|
|
H₁⁻¹ ▸ H₂
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem eq_true_elim (H : a = true) : a :=
|
|
|
|
H⁻¹ ▸ trivial
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem eq_false_elim (H : a = false) : ¬a :=
|
|
|
|
assume Ha : a, H ▸ Ha
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem imp_trans (H₁ : a → b) (H₂ : b → c) : a → c :=
|
|
|
|
assume Ha, H₂ (H₁ Ha)
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem imp_eq_trans (H₁ : a → b) (H₂ : b = c) : a → c :=
|
|
|
|
assume Ha, H₂ ▸ (H₁ Ha)
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem eq_imp_trans (H₁ : a = b) (H₂ : b → c) : a → c :=
|
|
|
|
assume Ha, H₂ (H₁ ▸ Ha)
|
|
|
|
end
|
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-10-21 21:08:07 +00:00
|
|
|
notation a ≠ b := ne a b
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-09-05 01:41:06 +00:00
|
|
|
namespace ne
|
2014-10-03 01:25:00 +00:00
|
|
|
variable {A : Type}
|
|
|
|
variables {a b : A}
|
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem intro : (a = b → false) → a ≠ b :=
|
|
|
|
assume H, H
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem elim : a ≠ b → a = b → false :=
|
|
|
|
assume H₁ H₂, H₁ H₂
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem irrefl : a ≠ a → false :=
|
|
|
|
assume H, H rfl
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem symm : a ≠ b → b ≠ a :=
|
|
|
|
assume (H : a ≠ b) (H₁ : b = a), H (H₁⁻¹)
|
2014-09-05 01:41:06 +00:00
|
|
|
end ne
|
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
section
|
|
|
|
variables {A : Type} {a b c : A}
|
|
|
|
|
|
|
|
theorem a_neq_a_elim : a ≠ a → false :=
|
|
|
|
assume H, H rfl
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem eq_ne_trans : a = b → b ≠ c → a ≠ c :=
|
|
|
|
assume H₁ H₂, H₁⁻¹ ▸ H₂
|
2014-08-01 00:48:51 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem ne_eq_trans : a ≠ b → b = c → a ≠ c :=
|
|
|
|
assume H₁ H₂, H₂ ▸ H₁
|
|
|
|
end
|
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
|
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
section
|
|
|
|
variables {p : Prop}
|
|
|
|
|
|
|
|
theorem p_ne_false : p → p ≠ false :=
|
|
|
|
assume (Hp : p) (Heq : p = false), Heq ▸ Hp
|
2014-08-04 05:58:12 +00:00
|
|
|
|
2014-10-05 17:19:50 +00:00
|
|
|
theorem p_ne_true : ¬p → p ≠ true :=
|
|
|
|
assume (Hnp : ¬p) (Heq : p = true), absurd trivial (Heq ▸ Hnp)
|
|
|
|
end
|
2014-09-05 01:41:06 +00:00
|
|
|
|
|
|
|
theorem true_ne_false : ¬true = false :=
|
|
|
|
assume H : true = false,
|
|
|
|
H ▸ trivial
|
2014-10-09 01:41:18 +00:00
|
|
|
|
|
|
|
inductive subsingleton [class] (A : Type) : Prop :=
|
|
|
|
intro : (∀ a b : A, a = b) -> subsingleton A
|
|
|
|
|
|
|
|
namespace subsingleton
|
|
|
|
definition elim {A : Type} (H : subsingleton A) : ∀(a b : A), a = b :=
|
|
|
|
rec (fun p, p) H
|
|
|
|
end subsingleton
|
|
|
|
|
2014-10-10 21:52:21 +00:00
|
|
|
protected definition prop.subsingleton [instance] (P : Prop) : subsingleton P :=
|
2014-10-09 01:41:18 +00:00
|
|
|
subsingleton.intro (λa b, !proof_irrel)
|