2014-12-12 04:14:53 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Floris van Doorn. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Author: Floris van Doorn
|
|
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
|
Ported from Coq HoTT
|
2014-12-12 04:14:53 +00:00
|
|
|
|
Theorems about products
|
|
|
|
|
-/
|
|
|
|
|
|
2015-05-26 13:56:41 +00:00
|
|
|
|
open eq equiv is_equiv is_trunc prod prod.ops unit
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
variables {A A' B B' C D : Type}
|
|
|
|
|
{a a' a'' : A} {b b₁ b₂ b' b'' : B} {u v w : A × B}
|
|
|
|
|
|
|
|
|
|
namespace prod
|
|
|
|
|
|
2015-05-26 13:56:41 +00:00
|
|
|
|
protected definition eta (u : A × B) : (pr₁ u, pr₂ u) = u :=
|
2015-02-24 21:27:57 +00:00
|
|
|
|
by cases u; apply idp
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-05-26 13:56:41 +00:00
|
|
|
|
definition pair_eq (pa : a = a') (pb : b = b') : (a, b) = (a', b') :=
|
2015-02-24 21:27:57 +00:00
|
|
|
|
by cases pa; cases pb; apply idp
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-24 21:27:57 +00:00
|
|
|
|
definition prod_eq (H₁ : pr₁ u = pr₁ v) (H₂ : pr₂ u = pr₂ v) : u = v :=
|
2015-05-26 13:56:41 +00:00
|
|
|
|
by cases u; cases v; exact pair_eq H₁ H₂
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
/- Symmetry -/
|
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
definition is_equiv_flip [instance] (A B : Type) : is_equiv (@flip A B) :=
|
2014-12-12 04:14:53 +00:00
|
|
|
|
adjointify flip
|
|
|
|
|
flip
|
|
|
|
|
(λu, destruct u (λb a, idp))
|
|
|
|
|
(λu, destruct u (λa b, idp))
|
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
definition prod_comm_equiv (A B : Type) : A × B ≃ B × A :=
|
2014-12-12 04:14:53 +00:00
|
|
|
|
equiv.mk flip _
|
|
|
|
|
|
2015-05-26 13:56:41 +00:00
|
|
|
|
definition prod_contr_equiv (A B : Type) [H : is_contr B] : A × B ≃ A :=
|
|
|
|
|
equiv.MK pr1
|
|
|
|
|
(λx, (x, !center))
|
|
|
|
|
(λx, idp)
|
|
|
|
|
(λx, by cases x with a b; exact pair_eq idp !center_eq)
|
|
|
|
|
|
|
|
|
|
definition prod_unit_equiv (A : Type) : A × unit ≃ A :=
|
|
|
|
|
!prod_contr_equiv
|
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
-- is_trunc_prod is defined in sigma
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
end prod
|