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.
|
2016-02-02 18:44:01 +00:00
|
|
|
|
Authors: Floris van Doorn, Jakob von Raumer
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
|
Ported from Coq HoTT
|
2014-12-12 04:14:53 +00:00
|
|
|
|
Theorems about products
|
|
|
|
|
-/
|
|
|
|
|
|
2016-03-03 15:48:27 +00:00
|
|
|
|
open eq equiv is_equiv is_trunc prod prod.ops unit
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-08-07 16:37:05 +00:00
|
|
|
|
variables {A A' B B' C D : Type} {P Q : A → Type}
|
2014-12-12 04:14:53 +00:00
|
|
|
|
{a a' a'' : A} {b b₁ b₂ b' b'' : B} {u v w : A × B}
|
|
|
|
|
|
|
|
|
|
namespace prod
|
|
|
|
|
|
2015-08-07 16:37:05 +00:00
|
|
|
|
/- Paths in a product space -/
|
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
protected definition eta [unfold 3] (u : A × B) : (pr₁ u, pr₂ u) = u :=
|
2017-01-07 13:53:15 +00:00
|
|
|
|
by cases u; reflexivity
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
definition pair_eq [unfold 7 8] (pa : a = a') (pb : b = b') : (a, b) = (a', b') :=
|
2017-01-07 13:53:15 +00:00
|
|
|
|
ap011 prod.mk pa pb
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2017-05-24 21:12:26 +00:00
|
|
|
|
definition prod_eq [unfold 3 4 5 6] (H₁ : u.1 = v.1) (H₂ : u.2 = v.2) : 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
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
definition eq_pr1 [unfold 5] (p : u = v) : u.1 = v.1 :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
ap pr1 p
|
|
|
|
|
|
2015-10-22 22:41:55 +00:00
|
|
|
|
definition eq_pr2 [unfold 5] (p : u = v) : u.2 = v.2 :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
ap pr2 p
|
|
|
|
|
|
|
|
|
|
namespace ops
|
|
|
|
|
postfix `..1`:(max+1) := eq_pr1
|
|
|
|
|
postfix `..2`:(max+1) := eq_pr2
|
|
|
|
|
end ops
|
|
|
|
|
open ops
|
|
|
|
|
|
2015-12-19 23:13:02 +00:00
|
|
|
|
protected definition ap_pr1 (p : u = v) : ap pr1 p = p..1 := idp
|
|
|
|
|
protected definition ap_pr2 (p : u = v) : ap pr2 p = p..2 := idp
|
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
definition pair_prod_eq (p : u.1 = v.1) (q : u.2 = v.2)
|
|
|
|
|
: ((prod_eq p q)..1, (prod_eq p q)..2) = (p, q) :=
|
2015-10-22 22:41:55 +00:00
|
|
|
|
by induction u; induction v; esimp at *; induction p; induction q; reflexivity
|
2015-08-06 20:37:52 +00:00
|
|
|
|
|
|
|
|
|
definition prod_eq_pr1 (p : u.1 = v.1) (q : u.2 = v.2) : (prod_eq p q)..1 = p :=
|
|
|
|
|
(pair_prod_eq p q)..1
|
|
|
|
|
|
|
|
|
|
definition prod_eq_pr2 (p : u.1 = v.1) (q : u.2 = v.2) : (prod_eq p q)..2 = q :=
|
|
|
|
|
(pair_prod_eq p q)..2
|
|
|
|
|
|
|
|
|
|
definition prod_eq_eta (p : u = v) : prod_eq (p..1) (p..2) = p :=
|
|
|
|
|
by induction p; induction u; reflexivity
|
|
|
|
|
|
2015-08-07 16:37:05 +00:00
|
|
|
|
-- the uncurried version of prod_eq. We will prove that this is an equivalence
|
2017-05-19 17:36:43 +00:00
|
|
|
|
definition prod_eq_unc [unfold 5] (H : u.1 = v.1 × u.2 = v.2) : u = v :=
|
|
|
|
|
by cases H with H₁ H₂; exact prod_eq H₁ H₂
|
2015-08-06 20:37:52 +00:00
|
|
|
|
|
|
|
|
|
definition pair_prod_eq_unc : Π(pq : u.1 = v.1 × u.2 = v.2),
|
|
|
|
|
((prod_eq_unc pq)..1, (prod_eq_unc pq)..2) = pq
|
|
|
|
|
| pair_prod_eq_unc (pq₁, pq₂) := pair_prod_eq pq₁ pq₂
|
|
|
|
|
|
|
|
|
|
definition prod_eq_unc_pr1 (pq : u.1 = v.1 × u.2 = v.2) : (prod_eq_unc pq)..1 = pq.1 :=
|
|
|
|
|
(pair_prod_eq_unc pq)..1
|
|
|
|
|
|
|
|
|
|
definition prod_eq_unc_pr2 (pq : u.1 = v.1 × u.2 = v.2) : (prod_eq_unc pq)..2 = pq.2 :=
|
|
|
|
|
(pair_prod_eq_unc pq)..2
|
|
|
|
|
|
|
|
|
|
definition prod_eq_unc_eta (p : u = v) : prod_eq_unc (p..1, p..2) = p :=
|
|
|
|
|
prod_eq_eta p
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition is_equiv_prod_eq [instance] [constructor] (u v : A × B)
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: is_equiv (prod_eq_unc : u.1 = v.1 × u.2 = v.2 → u = v) :=
|
|
|
|
|
adjointify prod_eq_unc
|
|
|
|
|
(λp, (p..1, p..2))
|
|
|
|
|
prod_eq_unc_eta
|
|
|
|
|
pair_prod_eq_unc
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_eq_equiv [constructor] (u v : A × B) : (u = v) ≃ (u.1 = v.1 × u.2 = v.2) :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
(equiv.mk prod_eq_unc _)⁻¹ᵉ
|
|
|
|
|
|
2017-05-19 17:36:43 +00:00
|
|
|
|
definition pair_eq_pair_equiv [constructor] (a a' : A) (b b' : B) :
|
|
|
|
|
((a, b) = (a', b')) ≃ (a = a' × b = b') :=
|
|
|
|
|
prod_eq_equiv (a, b) (a', b')
|
|
|
|
|
|
2017-01-07 13:53:15 +00:00
|
|
|
|
definition ap_prod_mk_left (p : a = a') : ap (λa, prod.mk a b) p = prod_eq p idp :=
|
|
|
|
|
ap_eq_ap011_left prod.mk p b
|
|
|
|
|
|
|
|
|
|
definition ap_prod_mk_right (p : b = b') : ap (λb, prod.mk a b) p = prod_eq idp p :=
|
|
|
|
|
ap_eq_ap011_right prod.mk a p
|
|
|
|
|
|
2017-06-02 16:13:20 +00:00
|
|
|
|
definition pair_eq_eta {A B : Type} {u v : A × B}
|
|
|
|
|
(p : u = v) : pair_eq (p..1) (p..2) = prod.eta u ⬝ p ⬝ (prod.eta v)⁻¹ :=
|
|
|
|
|
by induction p; induction u; reflexivity
|
|
|
|
|
|
|
|
|
|
definition prod_eq_eq {A B : Type} {u v : A × B}
|
|
|
|
|
{p₁ q₁ : u.1 = v.1} {p₂ q₂ : u.2 = v.2} (α₁ : p₁ = q₁) (α₂ : p₂ = q₂)
|
|
|
|
|
: prod_eq p₁ p₂ = prod_eq q₁ q₂ :=
|
|
|
|
|
by cases α₁; cases α₂; reflexivity
|
|
|
|
|
|
|
|
|
|
definition prod_eq_assemble {A B : Type} {u v : A × B}
|
|
|
|
|
{p q : u = v} (α₁ : p..1 = q..1) (α₂ : p..2 = q..2) : p = q :=
|
|
|
|
|
(prod_eq_eta p)⁻¹ ⬝ prod.prod_eq_eq α₁ α₂ ⬝ prod_eq_eta q
|
|
|
|
|
|
|
|
|
|
definition eq_pr1_concat {A B : Type} {u v w : A × B}
|
|
|
|
|
(p : u = v) (q : v = w)
|
|
|
|
|
: (p ⬝ q)..1 = p..1 ⬝ q..1 :=
|
|
|
|
|
by cases q; reflexivity
|
|
|
|
|
|
|
|
|
|
definition eq_pr2_concat {A B : Type} {u v w : A × B}
|
|
|
|
|
(p : u = v) (q : v = w)
|
|
|
|
|
: (p ⬝ q)..2 = p..2 ⬝ q..2 :=
|
|
|
|
|
by cases q; reflexivity
|
|
|
|
|
|
2015-12-19 23:13:02 +00:00
|
|
|
|
/- Groupoid structure -/
|
|
|
|
|
definition prod_eq_inv (p : a = a') (q : b = b') : (prod_eq p q)⁻¹ = prod_eq p⁻¹ q⁻¹ :=
|
|
|
|
|
by cases p; cases q; reflexivity
|
|
|
|
|
|
|
|
|
|
definition prod_eq_concat (p : a = a') (p' : a' = a'') (q : b = b') (q' : b' = b'')
|
|
|
|
|
: prod_eq p q ⬝ prod_eq p' q' = prod_eq (p ⬝ p') (q ⬝ q') :=
|
|
|
|
|
by cases p; cases q; cases p'; cases q'; reflexivity
|
|
|
|
|
|
2017-01-07 13:53:15 +00:00
|
|
|
|
definition prod_eq_concat_idp (p : a = a') (q : b = b')
|
|
|
|
|
: prod_eq p idp ⬝ prod_eq idp q = prod_eq p q :=
|
|
|
|
|
by cases p; cases q; reflexivity
|
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
/- Transport -/
|
|
|
|
|
|
2015-08-07 16:37:05 +00:00
|
|
|
|
definition prod_transport (p : a = a') (u : P a × Q a)
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: p ▸ u = (p ▸ u.1, p ▸ u.2) :=
|
|
|
|
|
by induction p; induction u; reflexivity
|
|
|
|
|
|
2015-12-19 23:13:02 +00:00
|
|
|
|
definition prod_eq_transport (p : a = a') (q : b = b') {R : A × B → Type} (r : R (a, b))
|
|
|
|
|
: (prod_eq p q) ▸ r = p ▸ q ▸ r :=
|
|
|
|
|
by induction p; induction q; reflexivity
|
|
|
|
|
|
2015-08-07 16:37:05 +00:00
|
|
|
|
/- Pathovers -/
|
|
|
|
|
|
|
|
|
|
definition etao (p : a = a') (bc : P a × Q a) : bc =[p] (p ▸ bc.1, p ▸ bc.2) :=
|
|
|
|
|
by induction p; induction bc; apply idpo
|
|
|
|
|
|
|
|
|
|
definition prod_pathover (p : a = a') (u : P a × Q a) (v : P a' × Q a')
|
|
|
|
|
(r : u.1 =[p] v.1) (s : u.2 =[p] v.2) : u =[p] v :=
|
|
|
|
|
begin
|
|
|
|
|
induction u, induction v, esimp at *, induction r,
|
|
|
|
|
induction s using idp_rec_on,
|
|
|
|
|
apply idpo
|
|
|
|
|
end
|
|
|
|
|
|
2017-06-02 16:13:20 +00:00
|
|
|
|
open prod.ops
|
|
|
|
|
definition prod_pathover_equiv {A : Type} {B C : A → Type} {a a' : A} (p : a = a')
|
|
|
|
|
(x : B a × C a) (x' : B a' × C a') : x =[p] x' ≃ x.1 =[p] x'.1 × x.2 =[p] x'.2 :=
|
|
|
|
|
begin
|
|
|
|
|
fapply equiv.MK,
|
|
|
|
|
{ intro q, induction q, constructor: constructor },
|
|
|
|
|
{ intro v, induction v with q r, exact prod_pathover _ _ _ q r },
|
|
|
|
|
{ intro v, induction v with q r, induction x with b c, induction x' with b' c',
|
|
|
|
|
esimp at *, induction q, refine idp_rec_on r _, reflexivity },
|
|
|
|
|
{ intro q, induction q, induction x with b c, reflexivity }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
|
2015-08-07 16:37:05 +00:00
|
|
|
|
/-
|
|
|
|
|
TODO:
|
|
|
|
|
* define the projections from the type u =[p] v
|
|
|
|
|
* show that the uncurried version of prod_pathover is an equivalence
|
|
|
|
|
-/
|
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
/- Functorial action -/
|
|
|
|
|
|
|
|
|
|
variables (f : A → A') (g : B → B')
|
|
|
|
|
definition prod_functor [unfold 7] (u : A × B) : A' × B' :=
|
|
|
|
|
(f u.1, g u.2)
|
|
|
|
|
|
|
|
|
|
definition ap_prod_functor (p : u.1 = v.1) (q : u.2 = v.2)
|
|
|
|
|
: ap (prod_functor f g) (prod_eq p q) = prod_eq (ap f p) (ap g q) :=
|
|
|
|
|
by induction u; induction v; esimp at *; induction p; induction q; reflexivity
|
|
|
|
|
|
2016-02-08 11:07:53 +00:00
|
|
|
|
/- Helpers for functions of two arguments -/
|
|
|
|
|
definition ap_diagonal {a a' : A} (p : a = a')
|
|
|
|
|
: ap (λx : A, (x,x)) p = prod_eq p p :=
|
|
|
|
|
by cases p; constructor
|
|
|
|
|
|
|
|
|
|
definition ap_binary (m : A → B → C) (p : a = a') (q : b = b')
|
|
|
|
|
: ap (λz : A × B, m z.1 z.2) (prod_eq p q)
|
2017-01-07 13:53:15 +00:00
|
|
|
|
= ap (m a) q ⬝ ap (λx : A, m x b') p :=
|
2016-02-08 11:07:53 +00:00
|
|
|
|
by cases p; cases q; constructor
|
|
|
|
|
|
2017-01-07 13:53:15 +00:00
|
|
|
|
definition ap_prod_elim {A B C : Type} {a a' : A} {b b' : B} (m : A → B → C)
|
|
|
|
|
(p : a = a') (q : b = b') : ap (prod.rec m) (prod_eq p q)
|
|
|
|
|
= ap (m a) q ⬝ ap (λx : A, m x b') p :=
|
|
|
|
|
by cases p; cases q; constructor
|
|
|
|
|
|
|
|
|
|
definition ap_prod_elim_idp {A B C : Type} {a a' : A} (m : A → B → C)
|
|
|
|
|
(p : a = a') (b : B) : ap (prod.rec m) (prod_eq p idp) = ap (λx : A, m x b) p :=
|
|
|
|
|
ap_prod_elim m p idp ⬝ !idp_con
|
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
/- Equivalences -/
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition is_equiv_prod_functor [instance] [constructor] [H : is_equiv f] [H : is_equiv g]
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: is_equiv (prod_functor f g) :=
|
|
|
|
|
begin
|
|
|
|
|
apply adjointify _ (prod_functor f⁻¹ g⁻¹),
|
|
|
|
|
intro u, induction u, rewrite [▸*,right_inv f,right_inv g],
|
|
|
|
|
intro u, induction u, rewrite [▸*,left_inv f,left_inv g],
|
|
|
|
|
end
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_equiv_prod_of_is_equiv [constructor] [H : is_equiv f] [H : is_equiv g]
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: A × B ≃ A' × B' :=
|
|
|
|
|
equiv.mk (prod_functor f g) _
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_equiv_prod [constructor] (f : A ≃ A') (g : B ≃ B') : A × B ≃ A' × B' :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
equiv.mk (prod_functor f g) _
|
|
|
|
|
|
2016-11-27 22:13:38 +00:00
|
|
|
|
-- rename
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_equiv_prod_left [constructor] (g : B ≃ B') : A × B ≃ A × B' :=
|
2016-04-11 17:11:59 +00:00
|
|
|
|
prod_equiv_prod equiv.rfl g
|
2015-08-06 20:37:52 +00:00
|
|
|
|
|
2016-11-27 22:13:38 +00:00
|
|
|
|
-- rename
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_equiv_prod_right [constructor] (f : A ≃ A') : A × B ≃ A' × B :=
|
2016-04-11 17:11:59 +00:00
|
|
|
|
prod_equiv_prod f equiv.rfl
|
2015-08-06 20:37:52 +00:00
|
|
|
|
|
2014-12-12 04:14:53 +00:00
|
|
|
|
/- Symmetry -/
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition is_equiv_flip [instance] [constructor] (A B : Type)
|
|
|
|
|
: is_equiv (flip : A × B → B × A) :=
|
2014-12-12 04:14:53 +00:00
|
|
|
|
adjointify flip
|
|
|
|
|
flip
|
|
|
|
|
(λu, destruct u (λb a, idp))
|
|
|
|
|
(λu, destruct u (λa b, idp))
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_comm_equiv [constructor] (A B : Type) : A × B ≃ B × A :=
|
2014-12-12 04:14:53 +00:00
|
|
|
|
equiv.mk flip _
|
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
/- Associativity -/
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_assoc_equiv [constructor] (A B C : Type) : A × (B × C) ≃ (A × B) × C :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply equiv.MK,
|
|
|
|
|
{ intro z, induction z with a z, induction z with b c, exact (a, b, c)},
|
|
|
|
|
{ intro z, induction z with z c, induction z with a b, exact (a, (b, c))},
|
|
|
|
|
{ intro z, induction z with z c, induction z with a b, reflexivity},
|
|
|
|
|
{ intro z, induction z with a z, induction z with b c, reflexivity},
|
|
|
|
|
end
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_contr_equiv [constructor] (A B : Type) [H : is_contr B] : A × B ≃ A :=
|
2015-05-26 13:56:41 +00:00
|
|
|
|
equiv.MK pr1
|
|
|
|
|
(λx, (x, !center))
|
|
|
|
|
(λx, idp)
|
|
|
|
|
(λx, by cases x with a b; exact pair_eq idp !center_eq)
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition prod_unit_equiv [constructor] (A : Type) : A × unit ≃ A :=
|
2015-05-26 13:56:41 +00:00
|
|
|
|
!prod_contr_equiv
|
|
|
|
|
|
2016-02-02 18:44:01 +00:00
|
|
|
|
definition prod_empty_equiv (A : Type) : A × empty ≃ empty :=
|
|
|
|
|
begin
|
|
|
|
|
fapply equiv.MK,
|
|
|
|
|
{ intro x, cases x with a e, cases e },
|
|
|
|
|
{ intro e, cases e },
|
|
|
|
|
{ intro e, cases e },
|
|
|
|
|
{ intro x, cases x with a e, cases e }
|
|
|
|
|
end
|
|
|
|
|
|
2015-08-06 20:37:52 +00:00
|
|
|
|
/- Universal mapping properties -/
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition is_equiv_prod_rec [instance] [constructor] (P : A × B → Type)
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: is_equiv (prod.rec : (Πa b, P (a, b)) → Πu, P u) :=
|
|
|
|
|
adjointify _
|
|
|
|
|
(λg a b, g (a, b))
|
|
|
|
|
(λg, eq_of_homotopy (λu, by induction u;reflexivity))
|
|
|
|
|
(λf, idp)
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition equiv_prod_rec [constructor] (P : A × B → Type) : (Πa b, P (a, b)) ≃ (Πu, P u) :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
equiv.mk prod.rec _
|
|
|
|
|
|
2016-06-23 20:49:54 +00:00
|
|
|
|
definition imp_imp_equiv_prod_imp [constructor] (A B C : Type) : (A → B → C) ≃ (A × B → C) :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
!equiv_prod_rec
|
|
|
|
|
|
|
|
|
|
definition prod_corec_unc [unfold 4] {P Q : A → Type} (u : (Πa, P a) × (Πa, Q a)) (a : A)
|
|
|
|
|
: P a × Q a :=
|
|
|
|
|
(u.1 a, u.2 a)
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition is_equiv_prod_corec [constructor] (P Q : A → Type)
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: is_equiv (prod_corec_unc : (Πa, P a) × (Πa, Q a) → Πa, P a × Q a) :=
|
|
|
|
|
adjointify _
|
|
|
|
|
(λg, (λa, (g a).1, λa, (g a).2))
|
|
|
|
|
(by intro g; apply eq_of_homotopy; intro a; esimp; induction (g a); reflexivity)
|
|
|
|
|
(by intro h; induction h with f g; reflexivity)
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition equiv_prod_corec [constructor] (P Q : A → Type)
|
|
|
|
|
: ((Πa, P a) × (Πa, Q a)) ≃ (Πa, P a × Q a) :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
equiv.mk _ !is_equiv_prod_corec
|
|
|
|
|
|
2015-08-07 14:44:57 +00:00
|
|
|
|
definition imp_prod_imp_equiv_imp_prod [constructor] (A B C : Type)
|
|
|
|
|
: (A → B) × (A → C) ≃ (A → (B × C)) :=
|
2015-08-06 20:37:52 +00:00
|
|
|
|
!equiv_prod_corec
|
|
|
|
|
|
2015-09-11 23:05:48 +00:00
|
|
|
|
theorem is_trunc_prod (A B : Type) (n : trunc_index) [HA : is_trunc n A] [HB : is_trunc n B]
|
2015-08-06 20:37:52 +00:00
|
|
|
|
: is_trunc n (A × B) :=
|
|
|
|
|
begin
|
|
|
|
|
revert A B HA HB, induction n with n IH, all_goals intro A B HA HB,
|
|
|
|
|
{ fapply is_contr.mk,
|
|
|
|
|
exact (!center, !center),
|
|
|
|
|
intro u, apply prod_eq, all_goals apply center_eq},
|
|
|
|
|
{ apply is_trunc_succ_intro, intro u v,
|
|
|
|
|
apply is_trunc_equiv_closed_rev, apply prod_eq_equiv,
|
|
|
|
|
exact IH _ _ _ _}
|
|
|
|
|
end
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
end prod
|
2015-08-06 20:37:52 +00:00
|
|
|
|
|
2015-09-11 22:53:08 +00:00
|
|
|
|
attribute prod.is_trunc_prod [instance] [priority 1510]
|
2015-12-10 19:37:11 +00:00
|
|
|
|
|
2016-04-22 19:12:25 +00:00
|
|
|
|
namespace prod
|
|
|
|
|
/- pointed products -/
|
|
|
|
|
open pointed
|
|
|
|
|
definition pointed_prod [instance] [constructor] (A B : Type) [H1 : pointed A] [H2 : pointed B]
|
|
|
|
|
: pointed (A × B) :=
|
|
|
|
|
pointed.mk (pt,pt)
|
|
|
|
|
|
|
|
|
|
definition pprod [constructor] (A B : Type*) : Type* :=
|
|
|
|
|
pointed.mk' (A × B)
|
|
|
|
|
|
|
|
|
|
infixr ` ×* `:35 := pprod
|
|
|
|
|
|
|
|
|
|
definition ppr1 [constructor] {A B : Type*} : A ×* B →* A :=
|
|
|
|
|
pmap.mk pr1 idp
|
|
|
|
|
|
|
|
|
|
definition ppr2 [constructor] {A B : Type*} : A ×* B →* B :=
|
|
|
|
|
pmap.mk pr2 idp
|
|
|
|
|
|
|
|
|
|
definition tprod [constructor] {n : trunc_index} (A B : n-Type) : n-Type :=
|
|
|
|
|
trunctype.mk (A × B) _
|
|
|
|
|
|
|
|
|
|
infixr `×t`:30 := tprod
|
|
|
|
|
|
|
|
|
|
definition ptprod [constructor] {n : ℕ₋₂} (A B : n-Type*) : n-Type* :=
|
|
|
|
|
ptrunctype.mk' n (A × B)
|
|
|
|
|
|
2017-06-02 16:13:20 +00:00
|
|
|
|
definition pprod_functor [constructor] {A B C D : Type*} (f : A →* C) (g : B →* D) : A ×* B →* C ×* D :=
|
|
|
|
|
pmap.mk (prod_functor f g) (prod_eq (respect_pt f) (respect_pt g))
|
2016-04-22 19:12:25 +00:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end prod
|