feat(library/hott): multiple changes in the HoTT library

Jakob accidentally undid some of my changes in commit aad4592, reverted that;
made style changes in multiple files;
in types/sigma: finished porting Coq-HoTT, and finished unfinished proof;
in axioms/funext: rename path_forall, make arguments implicit and make instance visible
This commit is contained in:
Floris van Doorn 2014-12-04 23:35:38 -05:00
parent 5278f70dea
commit fec45abda5
10 changed files with 893 additions and 723 deletions

View file

@ -6,7 +6,7 @@
import .natural_transformation hott.path import .natural_transformation hott.path
import data.unit data.sigma data.prod data.empty data.bool hott.types.prod import data.unit data.sigma data.prod data.empty data.bool hott.types.prod hott.types.sigma
open path prod eq eq.ops open path prod eq eq.ops
@ -94,7 +94,7 @@ namespace precategory
definition prod_precategory {obC obD : Type} (C : precategory obC) (D : precategory obD) definition prod_precategory {obC obD : Type} (C : precategory obC) (D : precategory obD)
: precategory (obC × obD) := : precategory (obC × obD) :=
mk (λ a b, hom (pr1 a) (pr1 b) × hom (pr2 a) (pr2 b)) mk (λ a b, hom (pr1 a) (pr1 b) × hom (pr2 a) (pr2 b))
(λ a b, trunc_prod nat.zero (!homH) (!homH)) (λ a b, !trunc_prod)
(λ a b c g f, (pr1 g ∘ pr1 f , pr2 g ∘ pr2 f) ) (λ a b c g f, (pr1 g ∘ pr1 f , pr2 g ∘ pr2 f) )
(λ a, (id, id)) (λ a, (id, id))
(λ a b c d h g f, pair_path !assoc !assoc ) (λ a b c d h g f, pair_path !assoc !assoc )

View file

@ -16,20 +16,19 @@ inductive funext [class] : Type :=
namespace funext namespace funext
context
universe variables l k universe variables l k
parameters [F : funext.{l k}] {A : Type.{l}} {P : A → Type.{k}} (f g : Π x, P x) variables [F : funext.{l k}] {A : Type.{l}} {P : A → Type.{k}}
protected definition ap [instance] : is_equiv (@apD10 A P f g) := include F
protected definition ap [instance] (f g : Π x, P x) : is_equiv (@apD10 A P f g) :=
rec_on F (λ(H : Π A P f g, _), !H) rec_on F (λ(H : Π A P f g, _), !H)
definition path_forall : f g → f ≈ g := definition path_pi {f g : Π x, P x} : f g → f ≈ g :=
@is_equiv.inv _ _ (@apD10 A P f g) ap is_equiv.inv (@apD10 A P f g)
end omit F
definition path_pi2 [F : funext] {A B : Type} {P : A → B → Type}
definition path_forall2 [F : funext] {A B : Type} {P : A → B → Type}
(f g : Πx y, P x y) : (Πx y, f x y ≈ g x y) → f ≈ g := (f g : Πx y, P x y) : (Πx y, f x y ≈ g x y) → f ≈ g :=
λ E, path_forall f g (λx, path_forall (f x) (g x) (E x)) λ E, path_pi (λx, path_pi (E x))
end funext end funext

View file

@ -30,7 +30,7 @@ namespace is_equiv
variables {A B C : Type} (f : A → B) (g : B → C) {f' : A → B} variables {A B C : Type} (f : A → B) (g : B → C) {f' : A → B}
-- The identity function is an equivalence. -- The identity function is an equivalence.
protected definition id_is_equiv : (@is_equiv A A id) := is_equiv.mk id (λa, idp) (λa, idp) (λa, idp) definition id_is_equiv : (@is_equiv A A id) := is_equiv.mk id (λa, idp) (λa, idp) (λa, idp)
-- The composition of two equivalences is, again, an equivalence. -- The composition of two equivalences is, again, an equivalence.
protected definition compose [Hf : is_equiv f] [Hg : is_equiv g] : (is_equiv (g ∘ f)) := protected definition compose [Hf : is_equiv f] [Hg : is_equiv g] : (is_equiv (g ∘ f)) :=

View file

@ -18,15 +18,15 @@ namespace is_equiv
definition precomp_closed [instance] {A B : Type} (f : A → B) [F : funext] [Hf : is_equiv f] (C : Type) definition precomp_closed [instance] {A B : Type} (f : A → B) [F : funext] [Hf : is_equiv f] (C : Type)
: is_equiv (precomp f C) := : is_equiv (precomp f C) :=
adjointify (precomp f C) (λh, h ∘ f⁻¹) adjointify (precomp f C) (λh, h ∘ f⁻¹)
(λh, path_forall _ _ (λx, ap h (sect f x))) (λh, path_pi (λx, ap h (sect f x)))
(λg, path_forall _ _ (λy, ap g (retr f y))) (λg, path_pi (λy, ap g (retr f y)))
--Postcomposing with an equivalence is an equivalence --Postcomposing with an equivalence is an equivalence
definition postcomp_closed [instance] {A B : Type} (f : A → B) [F : funext] [Hf : is_equiv f] (C : Type) definition postcomp_closed [instance] {A B : Type} (f : A → B) [F : funext] [Hf : is_equiv f] (C : Type)
: is_equiv (postcomp f C) := : is_equiv (postcomp f C) :=
adjointify (postcomp f C) (λl, f⁻¹ ∘ l) adjointify (postcomp f C) (λl, f⁻¹ ∘ l)
(λh, path_forall _ _ (λx, retr f (h x))) (λh, path_pi (λx, retr f (h x)))
(λg, path_forall _ _ (λy, sect f (g y))) (λg, path_pi (λy, sect f (g y)))
--Conversely, if pre- or post-composing with a function is always an equivalence, --Conversely, if pre- or post-composing with a function is always an equivalence,
--then that function is also an equivalence. It's convenient to know --then that function is also an equivalence. It's convenient to know

View file

@ -19,162 +19,158 @@ inductive path.{l} {A : Type.{l}} (a : A) : A → Type.{l} :=
idpath : path a a idpath : path a a
namespace path namespace path
variables {A B C : Type} {P : A → Type} {x y z t : A}
notation a ≈ b := path a b notation a ≈ b := path a b
notation x ≈ y `:>`:50 A:49 := @path A x y notation x ≈ y `:>`:50 A:49 := @path A x y
definition idp {A : Type} {a : A} := idpath a definition idp {a : A} := idpath a
-- unbased path induction -- unbased path induction
definition rec' [reducible] {A : Type} {P : Π (a b : A), (a ≈ b) -> Type} definition rec' [reducible] {P : Π (a b : A), (a ≈ b) -> Type}
(H : Π (a : A), P a a idp) {a b : A} (p : a ≈ b) : P a b p := (H : Π (a : A), P a a idp) {a b : A} (p : a ≈ b) : P a b p :=
path.rec (H a) p path.rec (H a) p
definition rec_on' [reducible] {A : Type} {P : Π (a b : A), (a ≈ b) -> Type} {a b : A} (p : a ≈ b) definition rec_on' [reducible] {P : Π (a b : A), (a ≈ b) -> Type} {a b : A} (p : a ≈ b)
(H : Π (a : A), P a a idp) : P a b p := (H : Π (a : A), P a a idp) : P a b p :=
path.rec (H a) p path.rec (H a) p
-- Concatenation and inverse -- Concatenation and inverse
-- ------------------------- -- -------------------------
definition concat {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : x ≈ z := definition concat (p : x ≈ y) (q : y ≈ z) : x ≈ z :=
path.rec (λu, u) q p path.rec (λu, u) q p
definition inverse {A : Type} {x y : A} (p : x ≈ y) : y ≈ x := definition inverse (p : x ≈ y) : y ≈ x :=
path.rec (idpath x) p path.rec (idpath x) p
notation p₁ ⬝ p₂ := concat p₁ p₂ notation p₁ ⬝ p₂ := concat p₁ p₂
notation p ⁻¹ := inverse p notation p ⁻¹ := inverse p
-- In Coq, these are not needed, because concat and inv are kept transparent
-- definition inv_1 {A : Type} (x : A) : (idpath x)⁻¹ ≈ idpath x := idp
-- definition concat_11 {A : Type} (x : A) : idpath x ⬝ idpath x ≈ idpath x := idp
-- The 1-dimensional groupoid structure -- The 1-dimensional groupoid structure
-- ------------------------------------ -- ------------------------------------
-- The identity path is a right unit. -- The identity path is a right unit.
definition concat_p1 {A : Type} {x y : A} (p : x ≈ y) : p ⬝ idp ≈ p := definition concat_p1 (p : x ≈ y) : p ⬝ idp ≈ p :=
rec_on p idp rec_on p idp
-- The identity path is a right unit. -- The identity path is a right unit.
definition concat_1p {A : Type} {x y : A} (p : x ≈ y) : idp ⬝ p ≈ p := definition concat_1p (p : x ≈ y) : idp ⬝ p ≈ p :=
rec_on p idp rec_on p idp
-- Concatenation is associative. -- Concatenation is associative.
definition concat_p_pp {A : Type} {x y z t : A} (p : x ≈ y) (q : y ≈ z) (r : z ≈ t) : definition concat_p_pp (p : x ≈ y) (q : y ≈ z) (r : z ≈ t) :
p ⬝ (q ⬝ r) ≈ (p ⬝ q) ⬝ r := p ⬝ (q ⬝ r) ≈ (p ⬝ q) ⬝ r :=
rec_on r (rec_on q idp) rec_on r (rec_on q idp)
definition concat_pp_p {A : Type} {x y z t : A} (p : x ≈ y) (q : y ≈ z) (r : z ≈ t) : definition concat_pp_p (p : x ≈ y) (q : y ≈ z) (r : z ≈ t) :
(p ⬝ q) ⬝ r ≈ p ⬝ (q ⬝ r) := (p ⬝ q) ⬝ r ≈ p ⬝ (q ⬝ r) :=
rec_on r (rec_on q idp) rec_on r (rec_on q idp)
-- The left inverse law. -- The left inverse law.
definition concat_pV {A : Type} {x y : A} (p : x ≈ y) : p ⬝ p⁻¹ ≈ idp := definition concat_pV (p : x ≈ y) : p ⬝ p⁻¹ ≈ idp :=
rec_on p idp rec_on p idp
-- The right inverse law. -- The right inverse law.
definition concat_Vp {A : Type} {x y : A} (p : x ≈ y) : p⁻¹ ⬝ p ≈ idp := definition concat_Vp (p : x ≈ y) : p⁻¹ ⬝ p ≈ idp :=
rec_on p idp rec_on p idp
-- Several auxiliary theorems about canceling inverses across associativity. These are somewhat -- Several auxiliary theorems about canceling inverses across associativity. These are somewhat
-- redundant, following from earlier theorems. -- redundant, following from earlier theorems.
definition concat_V_pp {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : p⁻¹ ⬝ (p ⬝ q) ≈ q := definition concat_V_pp (p : x ≈ y) (q : y ≈ z) : p⁻¹ ⬝ (p ⬝ q) ≈ q :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition concat_p_Vp {A : Type} {x y z : A} (p : x ≈ y) (q : x ≈ z) : p ⬝ (p⁻¹ ⬝ q) ≈ q := definition concat_p_Vp (p : x ≈ y) (q : x ≈ z) : p ⬝ (p⁻¹ ⬝ q) ≈ q :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition concat_pp_V {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : (p ⬝ q) ⬝ q⁻¹ ≈ p := definition concat_pp_V (p : x ≈ y) (q : y ≈ z) : (p ⬝ q) ⬝ q⁻¹ ≈ p :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition concat_pV_p {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) : (p ⬝ q⁻¹) ⬝ q ≈ p := definition concat_pV_p (p : x ≈ z) (q : y ≈ z) : (p ⬝ q⁻¹) ⬝ q ≈ p :=
rec_on q (take p, rec_on p idp) p rec_on q (take p, rec_on p idp) p
-- Inverse distributes over concatenation -- Inverse distributes over concatenation
definition inv_pp {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : (p ⬝ q)⁻¹ ≈ q⁻¹ ⬝ p⁻¹ := definition inv_pp (p : x ≈ y) (q : y ≈ z) : (p ⬝ q)⁻¹ ≈ q⁻¹ ⬝ p⁻¹ :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition inv_Vp {A : Type} {x y z : A} (p : y ≈ x) (q : y ≈ z) : (p⁻¹ ⬝ q)⁻¹ ≈ q⁻¹ ⬝ p := definition inv_Vp (p : y ≈ x) (q : y ≈ z) : (p⁻¹ ⬝ q)⁻¹ ≈ q⁻¹ ⬝ p :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
-- universe metavariables -- universe metavariables
definition inv_pV {A : Type} {x y z : A} (p : x ≈ y) (q : z ≈ y) : (p ⬝ q⁻¹)⁻¹ ≈ q ⬝ p⁻¹ := definition inv_pV (p : x ≈ y) (q : z ≈ y) : (p ⬝ q⁻¹)⁻¹ ≈ q ⬝ p⁻¹ :=
rec_on p (take q, rec_on q idp) q rec_on p (take q, rec_on q idp) q
definition inv_VV {A : Type} {x y z : A} (p : y ≈ x) (q : z ≈ y) : (p⁻¹ ⬝ q⁻¹)⁻¹ ≈ q ⬝ p := definition inv_VV (p : y ≈ x) (q : z ≈ y) : (p⁻¹ ⬝ q⁻¹)⁻¹ ≈ q ⬝ p :=
rec_on p (rec_on q idp) rec_on p (rec_on q idp)
-- Inverse is an involution. -- Inverse is an involution.
definition inv_V {A : Type} {x y : A} (p : x ≈ y) : p⁻¹⁻¹ ≈ p := definition inv_V (p : x ≈ y) : p⁻¹⁻¹ ≈ p :=
rec_on p idp rec_on p idp
-- Theorems for moving things around in equations -- Theorems for moving things around in equations
-- ---------------------------------------------- -- ----------------------------------------------
definition moveR_Mp {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) : definition moveR_Mp (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) :
p ≈ (r⁻¹ ⬝ q) → (r ⬝ p) ≈ q := p ≈ (r⁻¹ ⬝ q) → (r ⬝ p) ≈ q :=
rec_on r (take p h, concat_1p _ ⬝ h ⬝ concat_1p _) p rec_on r (take p h, concat_1p _ ⬝ h ⬝ concat_1p _) p
definition moveR_pM {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) : definition moveR_pM (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) :
r ≈ q ⬝ p⁻¹ → r ⬝ p ≈ q := r ≈ q ⬝ p⁻¹ → r ⬝ p ≈ q :=
rec_on p (take q h, (concat_p1 _ ⬝ h ⬝ concat_p1 _)) q rec_on p (take q h, (concat_p1 _ ⬝ h ⬝ concat_p1 _)) q
definition moveR_Vp {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) (r : x ≈ y) : definition moveR_Vp (p : x ≈ z) (q : y ≈ z) (r : x ≈ y) :
p ≈ r ⬝ q → r⁻¹ ⬝ p ≈ q := p ≈ r ⬝ q → r⁻¹ ⬝ p ≈ q :=
rec_on r (take q h, concat_1p _ ⬝ h ⬝ concat_1p _) q rec_on r (take q h, concat_1p _ ⬝ h ⬝ concat_1p _) q
definition moveR_pV {A : Type} {x y z : A} (p : z ≈ x) (q : y ≈ z) (r : y ≈ x) : definition moveR_pV (p : z ≈ x) (q : y ≈ z) (r : y ≈ x) :
r ≈ q ⬝ p → r ⬝ p⁻¹ ≈ q := r ≈ q ⬝ p → r ⬝ p⁻¹ ≈ q :=
rec_on p (take r h, concat_p1 _ ⬝ h ⬝ concat_p1 _) r rec_on p (take r h, concat_p1 _ ⬝ h ⬝ concat_p1 _) r
definition moveL_Mp {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) : definition moveL_Mp (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) :
r⁻¹ ⬝ q ≈ p → q ≈ r ⬝ p := r⁻¹ ⬝ q ≈ p → q ≈ r ⬝ p :=
rec_on r (take p h, (concat_1p _)⁻¹ ⬝ h ⬝ (concat_1p _)⁻¹) p rec_on r (take p h, (concat_1p _)⁻¹ ⬝ h ⬝ (concat_1p _)⁻¹) p
definition moveL_pM {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) : definition moveL_pM (p : x ≈ z) (q : y ≈ z) (r : y ≈ x) :
q ⬝ p⁻¹ ≈ r → q ≈ r ⬝ p := q ⬝ p⁻¹ ≈ r → q ≈ r ⬝ p :=
rec_on p (take q h, (concat_p1 _)⁻¹ ⬝ h ⬝ (concat_p1 _)⁻¹) q rec_on p (take q h, (concat_p1 _)⁻¹ ⬝ h ⬝ (concat_p1 _)⁻¹) q
definition moveL_Vp {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) (r : x ≈ y) : definition moveL_Vp (p : x ≈ z) (q : y ≈ z) (r : x ≈ y) :
r ⬝ q ≈ p → q ≈ r⁻¹ ⬝ p := r ⬝ q ≈ p → q ≈ r⁻¹ ⬝ p :=
rec_on r (take q h, (concat_1p _)⁻¹ ⬝ h ⬝ (concat_1p _)⁻¹) q rec_on r (take q h, (concat_1p _)⁻¹ ⬝ h ⬝ (concat_1p _)⁻¹) q
definition moveL_pV {A : Type} {x y z : A} (p : z ≈ x) (q : y ≈ z) (r : y ≈ x) : definition moveL_pV (p : z ≈ x) (q : y ≈ z) (r : y ≈ x) :
q ⬝ p ≈ r → q ≈ r ⬝ p⁻¹ := q ⬝ p ≈ r → q ≈ r ⬝ p⁻¹ :=
rec_on p (take r h, (concat_p1 _)⁻¹ ⬝ h ⬝ (concat_p1 _)⁻¹) r rec_on p (take r h, (concat_p1 _)⁻¹ ⬝ h ⬝ (concat_p1 _)⁻¹) r
definition moveL_1M {A : Type} {x y : A} (p q : x ≈ y) : definition moveL_1M (p q : x ≈ y) :
p ⬝ q⁻¹ ≈ idp → p ≈ q := p ⬝ q⁻¹ ≈ idp → p ≈ q :=
rec_on q (take p h, (concat_p1 _)⁻¹ ⬝ h) p rec_on q (take p h, (concat_p1 _)⁻¹ ⬝ h) p
definition moveL_M1 {A : Type} {x y : A} (p q : x ≈ y) : definition moveL_M1 (p q : x ≈ y) :
q⁻¹ ⬝ p ≈ idp → p ≈ q := q⁻¹ ⬝ p ≈ idp → p ≈ q :=
rec_on q (take p h, (concat_1p _)⁻¹ ⬝ h) p rec_on q (take p h, (concat_1p _)⁻¹ ⬝ h) p
definition moveL_1V {A : Type} {x y : A} (p : x ≈ y) (q : y ≈ x) : definition moveL_1V (p : x ≈ y) (q : y ≈ x) :
p ⬝ q ≈ idp → p ≈ q⁻¹ := p ⬝ q ≈ idp → p ≈ q⁻¹ :=
rec_on q (take p h, (concat_p1 _)⁻¹ ⬝ h) p rec_on q (take p h, (concat_p1 _)⁻¹ ⬝ h) p
definition moveL_V1 {A : Type} {x y : A} (p : x ≈ y) (q : y ≈ x) : definition moveL_V1 (p : x ≈ y) (q : y ≈ x) :
q ⬝ p ≈ idp → p ≈ q⁻¹ := q ⬝ p ≈ idp → p ≈ q⁻¹ :=
rec_on q (take p h, (concat_1p _)⁻¹ ⬝ h) p rec_on q (take p h, (concat_1p _)⁻¹ ⬝ h) p
definition moveR_M1 {A : Type} {x y : A} (p q : x ≈ y) : definition moveR_M1 (p q : x ≈ y) :
idp ≈ p⁻¹ ⬝ q → p ≈ q := idp ≈ p⁻¹ ⬝ q → p ≈ q :=
rec_on p (take q h, h ⬝ (concat_1p _)) q rec_on p (take q h, h ⬝ (concat_1p _)) q
definition moveR_1M {A : Type} {x y : A} (p q : x ≈ y) : definition moveR_1M (p q : x ≈ y) :
idp ≈ q ⬝ p⁻¹ → p ≈ q := idp ≈ q ⬝ p⁻¹ → p ≈ q :=
rec_on p (take q h, h ⬝ (concat_p1 _)) q rec_on p (take q h, h ⬝ (concat_p1 _)) q
definition moveR_1V {A : Type} {x y : A} (p : x ≈ y) (q : y ≈ x) : definition moveR_1V (p : x ≈ y) (q : y ≈ x) :
idp ≈ q ⬝ p → p⁻¹ ≈ q := idp ≈ q ⬝ p → p⁻¹ ≈ q :=
rec_on p (take q h, h ⬝ (concat_p1 _)) q rec_on p (take q h, h ⬝ (concat_p1 _)) q
definition moveR_V1 {A : Type} {x y : A} (p : x ≈ y) (q : y ≈ x) : definition moveR_V1 (p : x ≈ y) (q : y ≈ x) :
idp ≈ p ⬝ q → p⁻¹ ≈ q := idp ≈ p ⬝ q → p⁻¹ ≈ q :=
rec_on p (take q h, h ⬝ (concat_1p _)) q rec_on p (take q h, h ⬝ (concat_1p _)) q
@ -182,7 +178,7 @@ rec_on p (take q h, h ⬝ (concat_1p _)) q
-- Transport -- Transport
-- --------- -- ---------
definition transport [reducible] {A : Type} (P : A → Type) {x y : A} (p : x ≈ y) (u : P x) : P y := definition transport [reducible] (P : A → Type) {x y : A} (p : x ≈ y) (u : P x) : P y :=
path.rec_on p u path.rec_on p u
-- This idiom makes the operation right associative. -- This idiom makes the operation right associative.
@ -193,20 +189,20 @@ path.rec_on p idp
definition ap01 := ap definition ap01 := ap
definition homotopy [reducible] {A : Type} {P : A → Type} (f g : Πx, P x) : Type := definition homotopy [reducible] (f g : Πx, P x) : Type :=
Πx : A, f x ≈ g x Πx : A, f x ≈ g x
notation f g := homotopy f g notation f g := homotopy f g
definition apD10 {A} {B : A → Type} {f g : Πx, B x} (H : f ≈ g) : f g := definition apD10 {f g : Πx, P x} (H : f ≈ g) : f g :=
λx, path.rec_on H idp λx, path.rec_on H idp
definition ap10 {A B} {f g : A → B} (H : f ≈ g) : f g := apD10 H definition ap10 {f g : A → B} (H : f ≈ g) : f g := apD10 H
definition ap11 {A B} {f g : A → B} (H : f ≈ g) {x y : A} (p : x ≈ y) : f x ≈ g y := definition ap11 {f g : A → B} (H : f ≈ g) {x y : A} (p : x ≈ y) : f x ≈ g y :=
rec_on H (rec_on p idp) rec_on H (rec_on p idp)
definition apD {A:Type} {B : A → Type} (f : Πa:A, B a) {x y : A} (p : x ≈ y) : p ▹ (f x) ≈ f y := definition apD (f : Πa:A, P a) {x y : A} (p : x ≈ y) : p ▹ (f x) ≈ f y :=
rec_on p idp rec_on p idp
@ -221,19 +217,19 @@ calc_symm inverse
-- More theorems for moving things around in equations -- More theorems for moving things around in equations
-- --------------------------------------------------- -- ---------------------------------------------------
definition moveR_transport_p {A : Type} (P : A → Type) {x y : A} (p : x ≈ y) (u : P x) (v : P y) : definition moveR_transport_p (P : A → Type) {x y : A} (p : x ≈ y) (u : P x) (v : P y) :
u ≈ p⁻¹ ▹ v → p ▹ u ≈ v := u ≈ p⁻¹ ▹ v → p ▹ u ≈ v :=
rec_on p (take v, id) v rec_on p (take v, id) v
definition moveR_transport_V {A : Type} (P : A → Type) {x y : A} (p : y ≈ x) (u : P x) (v : P y) : definition moveR_transport_V (P : A → Type) {x y : A} (p : y ≈ x) (u : P x) (v : P y) :
u ≈ p ▹ v → p⁻¹ ▹ u ≈ v := u ≈ p ▹ v → p⁻¹ ▹ u ≈ v :=
rec_on p (take u, id) u rec_on p (take u, id) u
definition moveL_transport_V {A : Type} (P : A → Type) {x y : A} (p : x ≈ y) (u : P x) (v : P y) : definition moveL_transport_V (P : A → Type) {x y : A} (p : x ≈ y) (u : P x) (v : P y) :
p ▹ u ≈ v → u ≈ p⁻¹ ▹ v := p ▹ u ≈ v → u ≈ p⁻¹ ▹ v :=
rec_on p (take v, id) v rec_on p (take v, id) v
definition moveL_transport_p {A : Type} (P : A → Type) {x y : A} (p : y ≈ x) (u : P x) (v : P y) : definition moveL_transport_p (P : A → Type) {x y : A} (p : y ≈ x) (u : P x) (v : P y) :
p⁻¹ ▹ u ≈ v → u ≈ p ▹ v := p⁻¹ ▹ u ≈ v → u ≈ p ▹ v :=
rec_on p (take u, id) u rec_on p (take u, id) u
@ -244,25 +240,25 @@ rec_on p (take u, id) u
-- functorial. -- functorial.
-- Functions take identity paths to identity paths -- Functions take identity paths to identity paths
definition ap_1 {A B : Type} (x : A) (f : A → B) : (ap f idp) ≈ idp :> (f x ≈ f x) := idp definition ap_1 (x : A) (f : A → B) : (ap f idp) ≈ idp :> (f x ≈ f x) := idp
definition apD_1 {A B} (x : A) (f : Π x : A, B x) : apD f idp ≈ idp :> (f x ≈ f x) := idp definition apD_1 (x : A) (f : Π x : A, P x) : apD f idp ≈ idp :> (f x ≈ f x) := idp
-- Functions commute with concatenation. -- Functions commute with concatenation.
definition ap_pp {A B : Type} (f : A → B) {x y z : A} (p : x ≈ y) (q : y ≈ z) : definition ap_pp (f : A → B) {x y z : A} (p : x ≈ y) (q : y ≈ z) :
ap f (p ⬝ q) ≈ (ap f p) ⬝ (ap f q) := ap f (p ⬝ q) ≈ (ap f p) ⬝ (ap f q) :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition ap_p_pp {A B : Type} (f : A → B) {w x y z : A} (r : f w ≈ f x) (p : x ≈ y) (q : y ≈ z) : definition ap_p_pp (f : A → B) {w x y z : A} (r : f w ≈ f x) (p : x ≈ y) (q : y ≈ z) :
r ⬝ (ap f (p ⬝ q)) ≈ (r ⬝ ap f p) ⬝ (ap f q) := r ⬝ (ap f (p ⬝ q)) ≈ (r ⬝ ap f p) ⬝ (ap f q) :=
rec_on q (take p, rec_on p (concat_p_pp r idp idp)) p rec_on q (take p, rec_on p (concat_p_pp r idp idp)) p
definition ap_pp_p {A B : Type} (f : A → B) {w x y z : A} (p : x ≈ y) (q : y ≈ z) (r : f z ≈ f w) : definition ap_pp_p (f : A → B) {w x y z : A} (p : x ≈ y) (q : y ≈ z) (r : f z ≈ f w) :
(ap f (p ⬝ q)) ⬝ r ≈ (ap f p) ⬝ (ap f q ⬝ r) := (ap f (p ⬝ q)) ⬝ r ≈ (ap f p) ⬝ (ap f q ⬝ r) :=
rec_on q (rec_on p (take r, concat_pp_p _ _ _)) r rec_on q (rec_on p (take r, concat_pp_p _ _ _)) r
-- Functions commute with path inverses. -- Functions commute with path inverses.
definition inverse_ap {A B : Type} (f : A → B) {x y : A} (p : x ≈ y) : (ap f p)⁻¹ ≈ ap f (p⁻¹) := definition inverse_ap (f : A → B) {x y : A} (p : x ≈ y) : (ap f p)⁻¹ ≈ ap f (p⁻¹) :=
rec_on p idp rec_on p idp
definition ap_V {A B : Type} (f : A → B) {x y : A} (p : x ≈ y) : ap f (p⁻¹) ≈ (ap f p)⁻¹ := definition ap_V {A B : Type} (f : A → B) {x y : A} (p : x ≈ y) : ap f (p⁻¹) ≈ (ap f p)⁻¹ :=
@ -270,51 +266,51 @@ rec_on p idp
-- [ap] itself is functorial in the first argument. -- [ap] itself is functorial in the first argument.
definition ap_idmap {A : Type} {x y : A} (p : x ≈ y) : ap id p ≈ p := definition ap_idmap (p : x ≈ y) : ap id p ≈ p :=
rec_on p idp rec_on p idp
definition ap_compose {A B C : Type} (f : A → B) (g : B → C) {x y : A} (p : x ≈ y) : definition ap_compose (f : A → B) (g : B → C) {x y : A} (p : x ≈ y) :
ap (g ∘ f) p ≈ ap g (ap f p) := ap (g ∘ f) p ≈ ap g (ap f p) :=
rec_on p idp rec_on p idp
-- Sometimes we don't have the actual function [compose]. -- Sometimes we don't have the actual function [compose].
definition ap_compose' {A B C : Type} (f : A → B) (g : B → C) {x y : A} (p : x ≈ y) : definition ap_compose' (f : A → B) (g : B → C) {x y : A} (p : x ≈ y) :
ap (λa, g (f a)) p ≈ ap g (ap f p) := ap (λa, g (f a)) p ≈ ap g (ap f p) :=
rec_on p idp rec_on p idp
-- The action of constant maps. -- The action of constant maps.
definition ap_const {A B : Type} {x y : A} (p : x ≈ y) (z : B) : definition ap_const (p : x ≈ y) (z : B) :
ap (λu, z) p ≈ idp := ap (λu, z) p ≈ idp :=
rec_on p idp rec_on p idp
-- Naturality of [ap]. -- Naturality of [ap].
definition concat_Ap {A B : Type} {f g : A → B} (p : Π x, f x ≈ g x) {x y : A} (q : x ≈ y) : definition concat_Ap {f g : A → B} (p : Π x, f x ≈ g x) {x y : A} (q : x ≈ y) :
(ap f q) ⬝ (p y) ≈ (p x) ⬝ (ap g q) := (ap f q) ⬝ (p y) ≈ (p x) ⬝ (ap g q) :=
rec_on q (concat_1p _ ⬝ (concat_p1 _)⁻¹) rec_on q (concat_1p _ ⬝ (concat_p1 _)⁻¹)
-- Naturality of [ap] at identity. -- Naturality of [ap] at identity.
definition concat_A1p {A : Type} {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y) : definition concat_A1p {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y) :
(ap f q) ⬝ (p y) ≈ (p x) ⬝ q := (ap f q) ⬝ (p y) ≈ (p x) ⬝ q :=
rec_on q (concat_1p _ ⬝ (concat_p1 _)⁻¹) rec_on q (concat_1p _ ⬝ (concat_p1 _)⁻¹)
definition concat_pA1 {A : Type} {f : A → A} (p : Πx, x ≈ f x) {x y : A} (q : x ≈ y) : definition concat_pA1 {f : A → A} (p : Πx, x ≈ f x) {x y : A} (q : x ≈ y) :
(p x) ⬝ (ap f q) ≈ q ⬝ (p y) := (p x) ⬝ (ap f q) ≈ q ⬝ (p y) :=
rec_on q (concat_p1 _ ⬝ (concat_1p _)⁻¹) rec_on q (concat_p1 _ ⬝ (concat_1p _)⁻¹)
-- Naturality with other paths hanging around. -- Naturality with other paths hanging around.
definition concat_pA_pp {A B : Type} {f g : A → B} (p : Πx, f x ≈ g x) {x y : A} (q : x ≈ y) definition concat_pA_pp {f g : A → B} (p : Πx, f x ≈ g x) {x y : A} (q : x ≈ y)
{w z : B} (r : w ≈ f x) (s : g y ≈ z) : {w z : B} (r : w ≈ f x) (s : g y ≈ z) :
(r ⬝ ap f q) ⬝ (p y ⬝ s) ≈ (r ⬝ p x) ⬝ (ap g q ⬝ s) := (r ⬝ ap f q) ⬝ (p y ⬝ s) ≈ (r ⬝ p x) ⬝ (ap g q ⬝ s) :=
rec_on s (rec_on q idp) rec_on s (rec_on q idp)
definition concat_pA_p {A B : Type} {f g : A → B} (p : Πx, f x ≈ g x) {x y : A} (q : x ≈ y) definition concat_pA_p {f g : A → B} (p : Πx, f x ≈ g x) {x y : A} (q : x ≈ y)
{w : B} (r : w ≈ f x) : {w : B} (r : w ≈ f x) :
(r ⬝ ap f q) ⬝ p y ≈ (r ⬝ p x) ⬝ ap g q := (r ⬝ ap f q) ⬝ p y ≈ (r ⬝ p x) ⬝ ap g q :=
rec_on q idp rec_on q idp
-- TODO: try this using the simplifier, and compare proofs -- TODO: try this using the simplifier, and compare proofs
definition concat_A_pp {A B : Type} {f g : A → B} (p : Πx, f x ≈ g x) {x y : A} (q : x ≈ y) definition concat_A_pp {f g : A → B} (p : Πx, f x ≈ g x) {x y : A} (q : x ≈ y)
{z : B} (s : g y ≈ z) : {z : B} (s : g y ≈ z) :
(ap f q) ⬝ (p y ⬝ s) ≈ (p x) ⬝ (ap g q ⬝ s) := (ap f q) ⬝ (p y ⬝ s) ≈ (p x) ⬝ (ap g q ⬝ s) :=
rec_on s (rec_on q rec_on s (rec_on q
@ -325,32 +321,32 @@ rec_on s (rec_on q
-- This also works: -- This also works:
-- rec_on s (rec_on q (concat_1p _ ▹ idp)) -- rec_on s (rec_on q (concat_1p _ ▹ idp))
definition concat_pA1_pp {A : Type} {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y) definition concat_pA1_pp {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y)
{w z : A} (r : w ≈ f x) (s : y ≈ z) : {w z : A} (r : w ≈ f x) (s : y ≈ z) :
(r ⬝ ap f q) ⬝ (p y ⬝ s) ≈ (r ⬝ p x) ⬝ (q ⬝ s) := (r ⬝ ap f q) ⬝ (p y ⬝ s) ≈ (r ⬝ p x) ⬝ (q ⬝ s) :=
rec_on s (rec_on q idp) rec_on s (rec_on q idp)
definition concat_pp_A1p {A : Type} {g : A → A} (p : Πx, x ≈ g x) {x y : A} (q : x ≈ y) definition concat_pp_A1p {g : A → A} (p : Πx, x ≈ g x) {x y : A} (q : x ≈ y)
{w z : A} (r : w ≈ x) (s : g y ≈ z) : {w z : A} (r : w ≈ x) (s : g y ≈ z) :
(r ⬝ p x) ⬝ (ap g q ⬝ s) ≈ (r ⬝ q) ⬝ (p y ⬝ s) := (r ⬝ p x) ⬝ (ap g q ⬝ s) ≈ (r ⬝ q) ⬝ (p y ⬝ s) :=
rec_on s (rec_on q idp) rec_on s (rec_on q idp)
definition concat_pA1_p {A : Type} {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y) definition concat_pA1_p {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y)
{w : A} (r : w ≈ f x) : {w : A} (r : w ≈ f x) :
(r ⬝ ap f q) ⬝ p y ≈ (r ⬝ p x) ⬝ q := (r ⬝ ap f q) ⬝ p y ≈ (r ⬝ p x) ⬝ q :=
rec_on q idp rec_on q idp
definition concat_A1_pp {A : Type} {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y) definition concat_A1_pp {f : A → A} (p : Πx, f x ≈ x) {x y : A} (q : x ≈ y)
{z : A} (s : y ≈ z) : {z : A} (s : y ≈ z) :
(ap f q) ⬝ (p y ⬝ s) ≈ (p x) ⬝ (q ⬝ s) := (ap f q) ⬝ (p y ⬝ s) ≈ (p x) ⬝ (q ⬝ s) :=
rec_on s (rec_on q (concat_1p _ ▹ idp)) rec_on s (rec_on q (concat_1p _ ▹ idp))
definition concat_pp_A1 {A : Type} {g : A → A} (p : Πx, x ≈ g x) {x y : A} (q : x ≈ y) definition concat_pp_A1 {g : A → A} (p : Πx, x ≈ g x) {x y : A} (q : x ≈ y)
{w : A} (r : w ≈ x) : {w : A} (r : w ≈ x) :
(r ⬝ p x) ⬝ ap g q ≈ (r ⬝ q) ⬝ p y := (r ⬝ p x) ⬝ ap g q ≈ (r ⬝ q) ⬝ p y :=
rec_on q idp rec_on q idp
definition concat_p_A1p {A : Type} {g : A → A} (p : Πx, x ≈ g x) {x y : A} (q : x ≈ y) definition concat_p_A1p {g : A → A} (p : Πx, x ≈ g x) {x y : A} (q : x ≈ y)
{z : A} (s : g y ≈ z) : {z : A} (s : g y ≈ z) :
p x ⬝ (ap g q ⬝ s) ≈ q ⬝ (p y ⬝ s) := p x ⬝ (ap g q ⬝ s) ≈ q ⬝ (p y ⬝ s) :=
begin begin
@ -364,25 +360,26 @@ end
-- Application of paths between functions preserves the groupoid structure -- Application of paths between functions preserves the groupoid structure
definition apD10_1 {A} {B : A → Type} (f : Πx, B x) (x : A) : apD10 (idpath f) x ≈ idp := idp definition apD10_1 (f : Πx, P x) (x : A) : apD10 (idpath f) x ≈ idp := idp
definition apD10_pp {A} {B : A → Type} {f f' f'' : Πx, B x} (h : f ≈ f') (h' : f' ≈ f'') (x : A) : definition apD10_pp {f f' f'' : Πx, P x} (h : f ≈ f') (h' : f' ≈ f'') (x : A) :
apD10 (h ⬝ h') x ≈ apD10 h x ⬝ apD10 h' x := apD10 (h ⬝ h') x ≈ apD10 h x ⬝ apD10 h' x :=
rec_on h (take h', rec_on h' idp) h' rec_on h (take h', rec_on h' idp) h'
definition apD10_V {A : Type} {B : A → Type} {f g : Πx : A, B x} (h : f ≈ g) (x : A) : definition apD10_V {f g : Πx : A, P x} (h : f ≈ g) (x : A) :
apD10 (h⁻¹) x ≈ (apD10 h x)⁻¹ := apD10 (h⁻¹) x ≈ (apD10 h x)⁻¹ :=
rec_on h idp rec_on h idp
definition ap10_1 {A B} {f : A → B} (x : A) : ap10 (idpath f) x ≈ idp := idp definition ap10_1 {f : A → B} (x : A) : ap10 (idpath f) x ≈ idp := idp
definition ap10_pp {A B} {f f' f'' : A → B} (h : f ≈ f') (h' : f' ≈ f'') (x : A) : definition ap10_pp {f f' f'' : A → B} (h : f ≈ f') (h' : f' ≈ f'') (x : A) :
ap10 (h ⬝ h') x ≈ ap10 h x ⬝ ap10 h' x := apD10_pp h h' x ap10 (h ⬝ h') x ≈ ap10 h x ⬝ ap10 h' x := apD10_pp h h' x
definition ap10_V {A B} {f g : A→B} (h : f ≈ g) (x:A) : ap10 (h⁻¹) x ≈ (ap10 h x)⁻¹ := apD10_V h x definition ap10_V {f g : A → B} (h : f ≈ g) (x : A) : ap10 (h⁻¹) x ≈ (ap10 h x)⁻¹ :=
apD10_V h x
-- [ap10] also behaves nicely on paths produced by [ap] -- [ap10] also behaves nicely on paths produced by [ap]
definition ap_ap10 {A B C} (f g : A → B) (h : B → C) (p : f ≈ g) (a : A) : definition ap_ap10 (f g : A → B) (h : B → C) (p : f ≈ g) (a : A) :
ap h (ap10 p a) ≈ ap10 (ap (λ f', h ∘ f') p) a:= ap h (ap10 p a) ≈ ap10 (ap (λ f', h ∘ f') p) a:=
rec_on p idp rec_on p idp
@ -390,22 +387,22 @@ rec_on p idp
-- Transport and the groupoid structure of paths -- Transport and the groupoid structure of paths
-- --------------------------------------------- -- ---------------------------------------------
definition transport_1 {A : Type} (P : A → Type) {x : A} (u : P x) : definition transport_1 (P : A → Type) {x : A} (u : P x) :
idp ▹ u ≈ u := idp idp ▹ u ≈ u := idp
definition transport_pp {A : Type} (P : A → Type) {x y z : A} (p : x ≈ y) (q : y ≈ z) (u : P x) : definition transport_pp (P : A → Type) {x y z : A} (p : x ≈ y) (q : y ≈ z) (u : P x) :
p ⬝ q ▹ u ≈ q ▹ p ▹ u := p ⬝ q ▹ u ≈ q ▹ p ▹ u :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition transport_pV {A : Type} (P : A → Type) {x y : A} (p : x ≈ y) (z : P y) : definition transport_pV (P : A → Type) {x y : A} (p : x ≈ y) (z : P y) :
p ▹ p⁻¹ ▹ z ≈ z := p ▹ p⁻¹ ▹ z ≈ z :=
(transport_pp P (p⁻¹) p z)⁻¹ ⬝ ap (λr, transport P r z) (concat_Vp p) (transport_pp P (p⁻¹) p z)⁻¹ ⬝ ap (λr, transport P r z) (concat_Vp p)
definition transport_Vp {A : Type} (P : A → Type) {x y : A} (p : x ≈ y) (z : P x) : definition transport_Vp (P : A → Type) {x y : A} (p : x ≈ y) (z : P x) :
p⁻¹ ▹ p ▹ z ≈ z := p⁻¹ ▹ p ▹ z ≈ z :=
(transport_pp P p (p⁻¹) z)⁻¹ ⬝ ap (λr, transport P r z) (concat_pV p) (transport_pp P p (p⁻¹) z)⁻¹ ⬝ ap (λr, transport P r z) (concat_pV p)
definition transport_p_pp {A : Type} (P : A → Type) definition transport_p_pp (P : A → Type)
{x y z w : A} (p : x ≈ y) (q : y ≈ z) (r : z ≈ w) (u : P x) : {x y z w : A} (p : x ≈ y) (q : y ≈ z) (r : z ≈ w) (u : P x) :
ap (λe, e ▹ u) (concat_p_pp p q r) ⬝ (transport_pp P (p ⬝ q) r u) ⬝ ap (λe, e ▹ u) (concat_p_pp p q r) ⬝ (transport_pp P (p ⬝ q) r u) ⬝
ap (transport P r) (transport_pp P p q u) ap (transport P r) (transport_pp P p q u)
@ -414,53 +411,53 @@ definition transport_p_pp {A : Type} (P : A → Type)
rec_on r (rec_on q (rec_on p idp)) rec_on r (rec_on q (rec_on p idp))
-- Here is another coherence lemma for transport. -- Here is another coherence lemma for transport.
definition transport_pVp {A} (P : A → Type) {x y : A} (p : x ≈ y) (z : P x) : definition transport_pVp (P : A → Type) {x y : A} (p : x ≈ y) (z : P x) :
transport_pV P p (transport P p z) ≈ ap (transport P p) (transport_Vp P p z) := transport_pV P p (transport P p z) ≈ ap (transport P p) (transport_Vp P p z) :=
rec_on p idp rec_on p idp
-- Dependent transport in a doubly dependent type. -- Dependent transport in a doubly dependent type.
-- should B, C and y all be explicit here? -- should P, Q and y all be explicit here?
definition transportD {A : Type} (B : A → Type) (C : Π a : A, B a → Type) definition transportD (P : A → Type) (Q : Π a : A, P a → Type)
{x1 x2 : A} (p : x1 ≈ x2) (y : B x1) (z : C x1 y) : C x2 (p ▹ y) := {a a' : A} (p : a ≈ a') (b : P a) (z : Q a b) : Q a' (p ▹ b) :=
rec_on p z rec_on p z
-- In Coq the variables B, C and y are explicit, but in Lean we can probably have them implicit using the following notation -- In Coq the variables B, C and y are explicit, but in Lean we can probably have them implicit using the following notation
notation p `▹D`:65 x:64 := transportD _ _ p _ x notation p `▹D`:65 x:64 := transportD _ _ p _ x
-- Transporting along higher-dimensional paths -- Transporting along higher-dimensional paths
definition transport2 {A : Type} (P : A → Type) {x y : A} {p q : x ≈ y} (r : p ≈ q) (z : P x) : definition transport2 (P : A → Type) {x y : A} {p q : x ≈ y} (r : p ≈ q) (z : P x) :
p ▹ z ≈ q ▹ z := p ▹ z ≈ q ▹ z :=
ap (λp', p' ▹ z) r ap (λp', p' ▹ z) r
notation p `▹2`:65 x:64 := transport2 _ p _ x notation p `▹2`:65 x:64 := transport2 _ p _ x
-- An alternative definition. -- An alternative definition.
definition transport2_is_ap10 {A : Type} (Q : A → Type) {x y : A} {p q : x ≈ y} (r : p ≈ q) definition transport2_is_ap10 (Q : A → Type) {x y : A} {p q : x ≈ y} (r : p ≈ q)
(z : Q x) : (z : Q x) :
transport2 Q r z ≈ ap10 (ap (transport Q) r) z := transport2 Q r z ≈ ap10 (ap (transport Q) r) z :=
rec_on r idp rec_on r idp
definition transport2_p2p {A : Type} (P : A → Type) {x y : A} {p1 p2 p3 : x ≈ y} definition transport2_p2p (P : A → Type) {x y : A} {p1 p2 p3 : x ≈ y}
(r1 : p1 ≈ p2) (r2 : p2 ≈ p3) (z : P x) : (r1 : p1 ≈ p2) (r2 : p2 ≈ p3) (z : P x) :
transport2 P (r1 ⬝ r2) z ≈ transport2 P r1 z ⬝ transport2 P r2 z := transport2 P (r1 ⬝ r2) z ≈ transport2 P r1 z ⬝ transport2 P r2 z :=
rec_on r1 (rec_on r2 idp) rec_on r1 (rec_on r2 idp)
definition transport2_V {A : Type} (Q : A → Type) {x y : A} {p q : x ≈ y} (r : p ≈ q) (z : Q x) : definition transport2_V (Q : A → Type) {x y : A} {p q : x ≈ y} (r : p ≈ q) (z : Q x) :
transport2 Q (r⁻¹) z ≈ ((transport2 Q r z)⁻¹) := transport2 Q (r⁻¹) z ≈ ((transport2 Q r z)⁻¹) :=
rec_on r idp rec_on r idp
definition transportD2 {A : Type} (B C : A → Type) (D : Π(a:A), B a → C a → Type) definition transportD2 (B C : A → Type) (D : Π(a:A), B a → C a → Type)
{x1 x2 : A} (p : x1 ≈ x2) (y : B x1) (z : C x1) (w : D x1 y z) : D x2 (p ▹ y) (p ▹ z) := {x1 x2 : A} (p : x1 ≈ x2) (y : B x1) (z : C x1) (w : D x1 y z) : D x2 (p ▹ y) (p ▹ z) :=
rec_on p w rec_on p w
notation p `▹D2`:65 x:64 := transportD2 _ _ _ p _ _ x notation p `▹D2`:65 x:64 := transportD2 _ _ _ p _ _ x
definition concat_AT {A : Type} (P : A → Type) {x y : A} {p q : x ≈ y} {z w : P x} (r : p ≈ q) definition concat_AT (P : A → Type) {x y : A} {p q : x ≈ y} {z w : P x} (r : p ≈ q)
(s : z ≈ w) : (s : z ≈ w) :
ap (transport P p) s ⬝ transport2 P r w ≈ transport2 P r z ⬝ ap (transport P q) s := ap (transport P p) s ⬝ transport2 P r w ≈ transport2 P r z ⬝ ap (transport P q) s :=
rec_on r (concat_p1 _ ⬝ (concat_1p _)⁻¹) rec_on r (concat_p1 _ ⬝ (concat_1p _)⁻¹)
-- TODO (from Coq library): What should this be called? -- TODO (from Coq library): What should this be called?
definition ap_transport {A} {P Q : A → Type} {x y : A} (p : x ≈ y) (f : Πx, P x → Q x) (z : P x) : definition ap_transport {P Q : A → Type} {x y : A} (p : x ≈ y) (f : Πx, P x → Q x) (z : P x) :
f y (p ▹ z) ≈ (p ▹ (f x z)) := f y (p ▹ z) ≈ (p ▹ (f x z)) :=
rec_on p idp rec_on p idp
@ -478,34 +475,33 @@ subdirectory. Here we consider only the most basic cases.
-/ -/
-- Transporting in a constant fibration. -- Transporting in a constant fibration.
definition transport_const {A B : Type} {x1 x2 : A} (p : x1 ≈ x2) (y : B) : definition transport_const (p : x ≈ y) (z : B) : transport (λx, B) p z ≈ z :=
transport (λx, B) p y ≈ y :=
rec_on p idp rec_on p idp
definition transport2_const {A B : Type} {x1 x2 : A} {p q : x1 ≈ x2} (r : p ≈ q) (y : B) : definition transport2_const {p q : x ≈ y} (r : p ≈ q) (z : B) :
transport_const p y ≈ transport2 (λu, B) r y ⬝ transport_const q y := transport_const p z ≈ transport2 (λu, B) r z ⬝ transport_const q z :=
rec_on r (concat_1p _)⁻¹ rec_on r (concat_1p _)⁻¹
-- Transporting in a pulled back fibration. -- Transporting in a pulled back fibration.
-- TODO: P can probably be implicit -- TODO: P can probably be implicit
definition transport_compose {A B} {x y : A} (P : B → Type) (f : A → B) (p : x ≈ y) (z : P (f x)) : definition transport_compose (P : B → Type) (f : A → B) (p : x ≈ y) (z : P (f x)) :
transport (λx, P (f x)) p z ≈ transport P (ap f p) z := transport (P ∘ f) p z ≈ transport P (ap f p) z :=
rec_on p idp rec_on p idp
definition transport_precompose {A B C} (f : A → B) (g g' : B → C) (p : g ≈ g') : definition transport_precompose (f : A → B) (g g' : B → C) (p : g ≈ g') :
transport (λh : B → C, g ∘ f ≈ h ∘ f) p idp ≈ ap (λh, h ∘ f) p := transport (λh : B → C, g ∘ f ≈ h ∘ f) p idp ≈ ap (λh, h ∘ f) p :=
rec_on p idp rec_on p idp
definition apD10_ap_precompose {A B C} (f : A → B) (g g' : B → C) (p : g ≈ g') (a : A) : definition apD10_ap_precompose (f : A → B) (g g' : B → C) (p : g ≈ g') (a : A) :
apD10 (ap (λh : B → C, h ∘ f) p) a ≈ apD10 p (f a) := apD10 (ap (λh : B → C, h ∘ f) p) a ≈ apD10 p (f a) :=
rec_on p idp rec_on p idp
definition apD10_ap_postcompose {A B C} (f : B → C) (g g' : A → B) (p : g ≈ g') (a : A) : definition apD10_ap_postcompose (f : B → C) (g g' : A → B) (p : g ≈ g') (a : A) :
apD10 (ap (λh : A → B, f ∘ h) p) a ≈ ap f (apD10 p a) := apD10 (ap (λh : A → B, f ∘ h) p) a ≈ ap f (apD10 p a) :=
rec_on p idp rec_on p idp
-- A special case of [transport_compose] which seems to come up a lot. -- A special case of [transport_compose] which seems to come up a lot.
definition transport_idmap_ap A (P : A → Type) x y (p : x ≈ y) (u : P x) : definition transport_idmap_ap (P : A → Type) x y (p : x ≈ y) (u : P x) :
transport P p u ≈ transport (λz, z) (ap P p) u := transport P p u ≈ transport (λz, z) (ap P p) u :=
rec_on p idp rec_on p idp
@ -514,7 +510,7 @@ rec_on p idp
-- ------------------------------ -- ------------------------------
-- In a constant fibration, [apD] reduces to [ap], modulo [transport_const]. -- In a constant fibration, [apD] reduces to [ap], modulo [transport_const].
definition apD_const {A B} {x y : A} (f : A → B) (p: x ≈ y) : definition apD_const (f : A → B) (p: x ≈ y) :
apD f p ≈ transport_const p (f x) ⬝ ap f p := apD f p ≈ transport_const p (f x) ⬝ ap f p :=
rec_on p idp rec_on p idp
@ -523,75 +519,75 @@ rec_on p idp
-- ------------------------------------ -- ------------------------------------
-- Horizontal composition of 2-dimensional paths. -- Horizontal composition of 2-dimensional paths.
definition concat2 {A} {x y z : A} {p p' : x ≈ y} {q q' : y ≈ z} (h : p ≈ p') (h' : q ≈ q') : definition concat2 {p p' : x ≈ y} {q q' : y ≈ z} (h : p ≈ p') (h' : q ≈ q') :
p ⬝ q ≈ p' ⬝ q' := p ⬝ q ≈ p' ⬝ q' :=
rec_on h (rec_on h' idp) rec_on h (rec_on h' idp)
infixl `◾`:75 := concat2 infixl `◾`:75 := concat2
-- 2-dimensional path inversion -- 2-dimensional path inversion
definition inverse2 {A : Type} {x y : A} {p q : x ≈ y} (h : p ≈ q) : p⁻¹ ≈ q⁻¹ := definition inverse2 {p q : x ≈ y} (h : p ≈ q) : p⁻¹ ≈ q⁻¹ :=
rec_on h idp rec_on h idp
-- Whiskering -- Whiskering
-- ---------- -- ----------
definition whiskerL {A : Type} {x y z : A} (p : x ≈ y) {q r : y ≈ z} (h : q ≈ r) : p ⬝ q ≈ p ⬝ r := definition whiskerL (p : x ≈ y) {q r : y ≈ z} (h : q ≈ r) : p ⬝ q ≈ p ⬝ r :=
idp ◾ h idp ◾ h
definition whiskerR {A : Type} {x y z : A} {p q : x ≈ y} (h : p ≈ q) (r : y ≈ z) : p ⬝ r ≈ q ⬝ r := definition whiskerR {p q : x ≈ y} (h : p ≈ q) (r : y ≈ z) : p ⬝ r ≈ q ⬝ r :=
h ◾ idp h ◾ idp
-- Unwhiskering, a.k.a. cancelling -- Unwhiskering, a.k.a. cancelling
definition cancelL {A} {x y z : A} (p : x ≈ y) (q r : y ≈ z) : (p ⬝ q ≈ p ⬝ r) → (q ≈ r) := definition cancelL {x y z : A} (p : x ≈ y) (q r : y ≈ z) : (p ⬝ q ≈ p ⬝ r) → (q ≈ r) :=
rec_on p (take r, rec_on r (take q a, (concat_1p q)⁻¹ ⬝ a)) r q rec_on p (take r, rec_on r (take q a, (concat_1p q)⁻¹ ⬝ a)) r q
definition cancelR {A} {x y z : A} (p q : x ≈ y) (r : y ≈ z) : (p ⬝ r ≈ q ⬝ r) → (p ≈ q) := definition cancelR {x y z : A} (p q : x ≈ y) (r : y ≈ z) : (p ⬝ r ≈ q ⬝ r) → (p ≈ q) :=
rec_on r (rec_on p (take q a, a ⬝ concat_p1 q)) q rec_on r (rec_on p (take q a, a ⬝ concat_p1 q)) q
-- Whiskering and identity paths. -- Whiskering and identity paths.
definition whiskerR_p1 {A : Type} {x y : A} {p q : x ≈ y} (h : p ≈ q) : definition whiskerR_p1 {p q : x ≈ y} (h : p ≈ q) :
(concat_p1 p)⁻¹ ⬝ whiskerR h idp ⬝ concat_p1 q ≈ h := (concat_p1 p)⁻¹ ⬝ whiskerR h idp ⬝ concat_p1 q ≈ h :=
rec_on h (rec_on p idp) rec_on h (rec_on p idp)
definition whiskerR_1p {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : definition whiskerR_1p (p : x ≈ y) (q : y ≈ z) :
whiskerR idp q ≈ idp :> (p ⬝ q ≈ p ⬝ q) := whiskerR idp q ≈ idp :> (p ⬝ q ≈ p ⬝ q) :=
rec_on q idp rec_on q idp
definition whiskerL_p1 {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : definition whiskerL_p1 (p : x ≈ y) (q : y ≈ z) :
whiskerL p idp ≈ idp :> (p ⬝ q ≈ p ⬝ q) := whiskerL p idp ≈ idp :> (p ⬝ q ≈ p ⬝ q) :=
rec_on q idp rec_on q idp
definition whiskerL_1p {A : Type} {x y : A} {p q : x ≈ y} (h : p ≈ q) : definition whiskerL_1p {p q : x ≈ y} (h : p ≈ q) :
(concat_1p p) ⁻¹ ⬝ whiskerL idp h ⬝ concat_1p q ≈ h := (concat_1p p) ⁻¹ ⬝ whiskerL idp h ⬝ concat_1p q ≈ h :=
rec_on h (rec_on p idp) rec_on h (rec_on p idp)
definition concat2_p1 {A : Type} {x y : A} {p q : x ≈ y} (h : p ≈ q) : definition concat2_p1 {p q : x ≈ y} (h : p ≈ q) :
h ◾ idp ≈ whiskerR h idp :> (p ⬝ idp ≈ q ⬝ idp) := h ◾ idp ≈ whiskerR h idp :> (p ⬝ idp ≈ q ⬝ idp) :=
rec_on h idp rec_on h idp
definition concat2_1p {A : Type} {x y : A} {p q : x ≈ y} (h : p ≈ q) : definition concat2_1p {p q : x ≈ y} (h : p ≈ q) :
idp ◾ h ≈ whiskerL idp h :> (idp ⬝ p ≈ idp ⬝ q) := idp ◾ h ≈ whiskerL idp h :> (idp ⬝ p ≈ idp ⬝ q) :=
rec_on h idp rec_on h idp
-- TODO: note, 4 inductions -- TODO: note, 4 inductions
-- The interchange law for concatenation. -- The interchange law for concatenation.
definition concat_concat2 {A : Type} {x y z : A} {p p' p'' : x ≈ y} {q q' q'' : y ≈ z} definition concat_concat2 {p p' p'' : x ≈ y} {q q' q'' : y ≈ z}
(a : p ≈ p') (b : p' ≈ p'') (c : q ≈ q') (d : q' ≈ q'') : (a : p ≈ p') (b : p' ≈ p'') (c : q ≈ q') (d : q' ≈ q'') :
(a ◾ c) ⬝ (b ◾ d) ≈ (a ⬝ b) ◾ (c ⬝ d) := (a ◾ c) ⬝ (b ◾ d) ≈ (a ⬝ b) ◾ (c ⬝ d) :=
rec_on d (rec_on c (rec_on b (rec_on a idp))) rec_on d (rec_on c (rec_on b (rec_on a idp)))
definition concat_whisker {A} {x y z : A} (p p' : x ≈ y) (q q' : y ≈ z) (a : p ≈ p') (b : q ≈ q') : definition concat_whisker {x y z : A} (p p' : x ≈ y) (q q' : y ≈ z) (a : p ≈ p') (b : q ≈ q') :
(whiskerR a q) ⬝ (whiskerL p' b) ≈ (whiskerL p b) ⬝ (whiskerR a q') := (whiskerR a q) ⬝ (whiskerL p' b) ≈ (whiskerL p b) ⬝ (whiskerR a q') :=
rec_on b (rec_on a (concat_1p _)⁻¹) rec_on b (rec_on a (concat_1p _)⁻¹)
-- Structure corresponding to the coherence equations of a bicategory. -- Structure corresponding to the coherence equations of a bicategory.
-- The "pentagonator": the 3-cell witnessing the associativity pentagon. -- The "pentagonator": the 3-cell witnessing the associativity pentagon.
definition pentagon {A : Type} {v w x y z : A} (p : v ≈ w) (q : w ≈ x) (r : x ≈ y) (s : y ≈ z) : definition pentagon {v w x y z : A} (p : v ≈ w) (q : w ≈ x) (r : x ≈ y) (s : y ≈ z) :
whiskerL p (concat_p_pp q r s) whiskerL p (concat_p_pp q r s)
⬝ concat_p_pp p (q ⬝ r) s ⬝ concat_p_pp p (q ⬝ r) s
⬝ whiskerR (concat_p_pp p q r) s ⬝ whiskerR (concat_p_pp p q r) s
@ -599,11 +595,11 @@ definition pentagon {A : Type} {v w x y z : A} (p : v ≈ w) (q : w ≈ x) (r :
rec_on s (rec_on r (rec_on q (rec_on p idp))) rec_on s (rec_on r (rec_on q (rec_on p idp)))
-- The 3-cell witnessing the left unit triangle. -- The 3-cell witnessing the left unit triangle.
definition triangulator {A : Type} {x y z : A} (p : x ≈ y) (q : y ≈ z) : definition triangulator (p : x ≈ y) (q : y ≈ z) :
concat_p_pp p idp q ⬝ whiskerR (concat_p1 p) q ≈ whiskerL p (concat_1p q) := concat_p_pp p idp q ⬝ whiskerR (concat_p1 p) q ≈ whiskerL p (concat_1p q) :=
rec_on q (rec_on p idp) rec_on q (rec_on p idp)
definition eckmann_hilton {A : Type} {x:A} (p q : idp ≈ idp :> (x ≈ x)) : p ⬝ q ≈ q ⬝ p := definition eckmann_hilton {x:A} (p q : idp ≈ idp :> (x ≈ x)) : p ⬝ q ≈ q ⬝ p :=
(!whiskerR_p1 ◾ !whiskerL_1p)⁻¹ (!whiskerR_p1 ◾ !whiskerL_1p)⁻¹
⬝ (!concat_p1 ◾ !concat_p1) ⬝ (!concat_p1 ◾ !concat_p1)
⬝ (!concat_1p ◾ !concat_1p) ⬝ (!concat_1p ◾ !concat_1p)
@ -613,36 +609,35 @@ definition eckmann_hilton {A : Type} {x:A} (p q : idp ≈ idp :> (x ≈ x)) : p
⬝ (!whiskerL_1p ◾ !whiskerR_p1) ⬝ (!whiskerL_1p ◾ !whiskerR_p1)
-- The action of functions on 2-dimensional paths -- The action of functions on 2-dimensional paths
definition ap02 {A B : Type} (f:A → B) {x y : A} {p q : x ≈ y} (r : p ≈ q) : ap f p ≈ ap f q := definition ap02 (f:A → B) {x y : A} {p q : x ≈ y} (r : p ≈ q) : ap f p ≈ ap f q :=
rec_on r idp rec_on r idp
definition ap02_pp {A B} (f : A → B) {x y : A} {p p' p'' : x ≈ y} (r : p ≈ p') (r' : p' ≈ p'') : definition ap02_pp (f : A → B) {x y : A} {p p' p'' : x ≈ y} (r : p ≈ p') (r' : p' ≈ p'') :
ap02 f (r ⬝ r') ≈ ap02 f r ⬝ ap02 f r' := ap02 f (r ⬝ r') ≈ ap02 f r ⬝ ap02 f r' :=
rec_on r (rec_on r' idp) rec_on r (rec_on r' idp)
definition ap02_p2p {A B} (f : A → B) {x y z : A} {p p' : x ≈ y} {q q' :y ≈ z} (r : p ≈ p') definition ap02_p2p (f : A → B) {x y z : A} {p p' : x ≈ y} {q q' :y ≈ z} (r : p ≈ p')
(s : q ≈ q') : (s : q ≈ q') :
ap02 f (r ◾ s) ≈ ap_pp f p q ap02 f (r ◾ s) ≈ ap_pp f p q
⬝ (ap02 f r ◾ ap02 f s) ⬝ (ap02 f r ◾ ap02 f s)
⬝ (ap_pp f p' q')⁻¹ := ⬝ (ap_pp f p' q')⁻¹ :=
rec_on r (rec_on s (rec_on q (rec_on p idp))) rec_on r (rec_on s (rec_on q (rec_on p idp)))
-- rec_on r (rec_on s (rec_on p (rec_on q idp))) -- rec_on r (rec_on s (rec_on p (rec_on q idp)))
definition apD02 {A : Type} {B : A → Type} {x y : A} {p q : x ≈ y} (f : Π x, B x) (r : p ≈ q) : definition apD02 {p q : x ≈ y} (f : Π x, P x) (r : p ≈ q) :
apD f p ≈ transport2 B r (f x) ⬝ apD f q := apD f p ≈ transport2 P r (f x) ⬝ apD f q :=
rec_on r (concat_1p _)⁻¹ rec_on r (concat_1p _)⁻¹
-- And now for a lemma whose statement is much longer than its proof. -- And now for a lemma whose statement is much longer than its proof.
definition apD02_pp {A} (B : A → Type) (f : Π x:A, B x) {x y : A} definition apD02_pp (P : A → Type) (f : Π x:A, P x) {x y : A}
{p1 p2 p3 : x ≈ y} (r1 : p1 ≈ p2) (r2 : p2 ≈ p3) : {p1 p2 p3 : x ≈ y} (r1 : p1 ≈ p2) (r2 : p2 ≈ p3) :
apD02 f (r1 ⬝ r2) ≈ apD02 f r1 apD02 f (r1 ⬝ r2) ≈ apD02 f r1
⬝ whiskerL (transport2 B r1 (f x)) (apD02 f r2) ⬝ whiskerL (transport2 P r1 (f x)) (apD02 f r2)
⬝ concat_p_pp _ _ _ ⬝ concat_p_pp _ _ _
⬝ (whiskerR ((transport2_p2p B r1 r2 (f x))⁻¹) (apD f p3)) := ⬝ (whiskerR ((transport2_p2p P r1 r2 (f x))⁻¹) (apD f p3)) :=
rec_on r2 (rec_on r1 (rec_on p1 idp)) rec_on r2 (rec_on r1 (rec_on p1 idp))
end path
section namespace path
variables {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D} variables {A B C D E : Type} {a a' : A} {b b' : B} {c c' : C} {d d' : D}
theorem congr_arg2 (f : A → B → C) (Ha : a ≈ a') (Hb : b ≈ b') : f a b ≈ f a' b' := theorem congr_arg2 (f : A → B → C) (Ha : a ≈ a') (Hb : b ≈ b') : f a b ≈ f a' b' :=
@ -656,9 +651,9 @@ theorem congr_arg4 (f : A → B → C → D → E) (Ha : a ≈ a') (Hb : b ≈ b
: f a b c d ≈ f a' b' c' d' := : f a b c d ≈ f a' b' c' d' :=
rec_on Ha (congr_arg3 (f a) Hb Hc Hd) rec_on Ha (congr_arg3 (f a) Hb Hc Hd)
end end path
section namespace path
variables {A : Type} {B : A → Type} {C : Πa, B a → Type} {D : Πa b, C a b → Type} variables {A : Type} {B : A → Type} {C : Πa, B a → Type} {D : Πa b, C a b → Type}
{E : Πa b c, D a b c → Type} {F : Type} {E : Πa b c, D a b c → Type} {F : Type}
variables {a a' : A} variables {a a' : A}
@ -670,8 +665,6 @@ theorem dcongr_arg2 (f : Πa, B a → F) (Ha : a ≈ a') (Hb : Ha ▹ b ≈ b')
: f a b ≈ f a' b' := : f a b ≈ f a' b' :=
rec_on Hb (rec_on Ha idp) rec_on Hb (rec_on Ha idp)
end
/- From the Coq version: /- From the Coq version:
-- ** Tactics, hints, and aliases -- ** Tactics, hints, and aliases
@ -724,5 +717,4 @@ Ltac hott_simpl :=
autorewrite with paths in * |- * ; auto with path_hints. autorewrite with paths in * |- * ; auto with path_hints.
-/ -/
end path end path

View file

@ -78,12 +78,12 @@ namespace truncation
variables {A B : Type} variables {A B : Type}
-- maybe rename to is_trunc_succ.mk -- TODO: rename to is_trunc_succ
definition is_trunc_succ (A : Type) (n : trunc_index) [H : ∀x y : A, is_trunc n (x ≈ y)] definition is_trunc_succ (A : Type) (n : trunc_index) [H : ∀x y : A, is_trunc n (x ≈ y)]
: is_trunc n.+1 A := : is_trunc n.+1 A :=
is_trunc.mk (λ x y, !is_trunc.to_internal) is_trunc.mk (λ x y, !is_trunc.to_internal)
-- maybe rename to is_trunc_succ.elim -- TODO: rename to is_trunc_path
definition succ_is_trunc (n : trunc_index) [H : is_trunc (n.+1) A] (x y : A) : is_trunc n (x ≈ y) := definition succ_is_trunc (n : trunc_index) [H : is_trunc (n.+1) A] (x y : A) : is_trunc n (x ≈ y) :=
is_trunc.mk (!is_trunc.to_internal x y) is_trunc.mk (!is_trunc.to_internal x y)

View file

@ -10,30 +10,30 @@ import .sigma .pi
open path sigma sigma.ops equiv is_equiv open path sigma sigma.ops equiv is_equiv
inductive W {A : Type} (B : A → Type) := inductive Wtype {A : Type} (B : A → Type) :=
sup : Π(a : A), (B a → W B) → W B sup : Π(a : A), (B a → Wtype B) → Wtype B
namespace W namespace Wtype
notation `WW` binders `,` r:(scoped B, W B) := r notation `W` binders `,` r:(scoped B, Wtype B) := r
universe variables l k universe variables u v
variables {A A' : Type.{l}} {B B' : A → Type.{k}} {C : Π(a : A), B a → Type} variables {A A' : Type.{u}} {B B' : A → Type.{v}} {C : Π(a : A), B a → Type}
{a a' : A} {f : B a → W B} {f' : B a' → W B} {w w' : WW(a : A), B a} {a a' : A} {f : B a → W a, B a} {f' : B a' → W a, B a} {w w' : W(a : A), B a}
definition pr1 (w : WW(a : A), B a) : A := definition pr1 (w : W(a : A), B a) : A :=
W.rec_on w (λa f IH, a) Wtype.rec_on w (λa f IH, a)
definition pr2 (w : WW(a : A), B a) : B (pr1 w) → WW(a : A), B a := definition pr2 (w : W(a : A), B a) : B (pr1 w) → W(a : A), B a :=
W.rec_on w (λa f IH, f) Wtype.rec_on w (λa f IH, f)
namespace ops namespace ops
postfix `.1`:(max+1) := W.pr1 postfix `.1`:(max+1) := Wtype.pr1
postfix `.2`:(max+1) := W.pr2 postfix `.2`:(max+1) := Wtype.pr2
notation `⟨` a `,` f `⟩`:0 := W.sup a f --input ⟨ ⟩ as \< \> notation `⟨` a `,` f `⟩`:0 := Wtype.sup a f --input ⟨ ⟩ as \< \>
end ops end ops
open ops open ops
protected definition eta (w : WW a, B a) : ⟨w.1 , w.2⟩ ≈ w := protected definition eta (w : W a, B a) : ⟨w.1 , w.2⟩ ≈ w :=
cases_on w (λa f, idp) cases_on w (λa f, idp)
definition path_W_sup (p : a ≈ a') (q : p ▹ f ≈ f') : ⟨a, f⟩ ≈ ⟨a', f'⟩ := definition path_W_sup (p : a ≈ a') (q : p ▹ f ≈ f') : ⟨a, f⟩ ≈ ⟨a', f'⟩ :=
@ -112,18 +112,18 @@ namespace W
: transport (λx, B' x.1) (@path_W_uncurried A B w w' pq) ≈ transport B' pq.1 := : transport (λx, B' x.1) (@path_W_uncurried A B w w' pq) ≈ transport B' pq.1 :=
destruct pq transport_pr1_path_W destruct pq transport_pr1_path_W
definition isequiv_path_W /-[instance]-/ (w w' : W B) definition isequiv_path_W /-[instance]-/ (w w' : W a, B a)
: is_equiv (@path_W_uncurried A B w w') := : is_equiv (@path_W_uncurried A B w w') :=
adjointify path_W_uncurried adjointify path_W_uncurried
(λp, dpair (p..1) (p..2)) (λp, dpair (p..1) (p..2))
eta_path_W_uncurried eta_path_W_uncurried
sup_path_W_uncurried sup_path_W_uncurried
definition equiv_path_W (w w' : W B) : (Σ(p : w.1 ≈ w'.1), p ▹ w.2 ≈ w'.2) ≃ (w ≈ w') := definition equiv_path_W (w w' : W a, B a) : (Σ(p : w.1 ≈ w'.1), p ▹ w.2 ≈ w'.2) ≃ (w ≈ w') :=
equiv.mk path_W_uncurried !isequiv_path_W equiv.mk path_W_uncurried !isequiv_path_W
definition double_induction_on {P : W B → W B → Type} (w w' : W B) definition double_induction_on {P : (W a, B a) → (W a, B a) → Type} (w w' : W a, B a)
(H : ∀ (a a' : A) (f : B a → W B) (f' : B a' → W B), (H : ∀ (a a' : A) (f : B a → W a, B a) (f' : B a' → W a, B a),
(∀ (b : B a) (b' : B a'), P (f b) (f' b')) → P (sup a f) (sup a' f')) : P w w' := (∀ (b : B a) (b' : B a'), P (f b) (f' b')) → P (sup a f) (sup a' f')) : P w w' :=
begin begin
revert w', revert w',
@ -135,8 +135,8 @@ namespace W
/- truncatedness -/ /- truncatedness -/
open truncation open truncation
definition trunc_W [FUN : funext.{k (max 1 l k)}] (n : trunc_index) [HA : is_trunc (n.+1) A] definition trunc_W [FUN : funext.{v (max 1 u v)}] (n : trunc_index) [HA : is_trunc (n.+1) A]
: is_trunc (n.+1) (WW a, B a) := : is_trunc (n.+1) (W a, B a) :=
begin begin
fapply is_trunc_succ, intros (w, w'), fapply is_trunc_succ, intros (w, w'),
apply (double_induction_on w w'), intros (a, a', f, f', IH), apply (double_induction_on w w'), intros (a, a', f, f', IH),
@ -146,8 +146,8 @@ namespace W
fapply (succ_is_trunc n), fapply (succ_is_trunc n),
intro p, revert IH, generalize f', --change to revert after simpl intro p, revert IH, generalize f', --change to revert after simpl
apply (path.rec_on p), intros (f', IH), apply (path.rec_on p), intros (f', IH),
apply (@pi.trunc_path_pi FUN (B a) (λx, W B) n f f'), intro b, apply pi.trunc_path_pi, intro b,
apply IH apply IH
end end
end W end Wtype

View file

@ -7,52 +7,116 @@ Ported from Coq HoTT
Theorems about pi-types (dependent function spaces) Theorems about pi-types (dependent function spaces)
-/ -/
import ..trunc ..axioms.funext import ..trunc ..axioms.funext .sigma
open path equiv is_equiv funext open path equiv is_equiv funext
namespace pi namespace pi
universe variables l k universe variables l k
variables [H : funext.{l k}] {A : Type.{l}} {B : A → Type.{k}} {C : Πa, B a → Type} variables {A A' : Type.{l}} {B : A → Type.{k}} {C : Πa, B a → Type}
{D : Πa b, C a b → Type} {D : Πa b, C a b → Type}
{a a' a'' : A} {b b₁ b₂ : B a} {b' : B a'} {b'' : B a''} {f g h : Πa, B a} {a a' a'' : A} {b b₁ b₂ : B a} {b' : B a'} {b'' : B a''} {f g : Πa, B a}
include H
/- Paths -/ /- Paths -/
/- Paths [p : f ≈ g] in a function type [Πx:X, P x] are equivalent to functions taking values in path types, [H : Πx:X, f x ≈ g x], or concisely, [H : f == g]. /- Paths [p : f ≈ g] in a function type [Πx:X, P x] are equivalent to functions taking values in path types, [H : Πx:X, f x ≈ g x], or concisely, [H : f g].
This equivalence, however, is just the combination of [apD10] and function extensionality [funext], and as such, [path_forall], et seq. are given in axioms.funext and path: -/ This equivalence, however, is just the combination of [apD10] and function extensionality [funext], and as such, [path_forall], et seq. are given in axioms.funext and path: -/
/- Now we show how these things compute. -/ /- Now we show how these things compute. -/
definition equiv_apD10 : (f ≈ g) ≃ (f g) := definition apD10_path_pi [H : funext] (h : f g) : apD10 (path_pi h) h :=
apD10 (retr apD10 h)
definition path_pi_eta [H : funext] (p : f ≈ g) : path_pi (apD10 p) ≈ p :=
sect apD10 p
definition path_pi_idp [H : funext] : path_pi (λx : A, idpath (f x)) ≈ idpath f :=
!path_pi_eta
/- The identification of the path space of a dependent function space, up to equivalence, is of course just funext. -/
definition path_equiv_homotopy [H : funext] (f g : Πx, B x) : (f ≈ g) ≃ (f g) :=
equiv.mk _ !funext.ap equiv.mk _ !funext.ap
definition is_equiv_path_pi [instance] [H : funext] (f g : Πx, B x)
: is_equiv (@path_pi _ _ _ f g) :=
inv_closed apD10
definition homotopy_equiv_path [H : funext] (f g : Πx, B x) : (f g) ≃ (f ≈ g) :=
equiv.mk _ !is_equiv_path_pi
/- Transport -/
protected definition transport (p : a ≈ a') (f : Π(b : B a), C a b)
: (transport (λa, Π(b : B a), C a b) p f)
(λb, transport (C a') !transport_pV (transportD _ _ p _ (f (p⁻¹ ▹ b)))) :=
path.rec_on p (λx, idp)
/- A special case of [transport_pi] where the type [B] does not depend on [A],
and so it is just a fixed type [B]. -/
definition transport_constant {C : A → A' → Type} (p : a ≈ a') (f : Π(b : A'), C a b)
: (path.transport (λa, Π(b : A'), C a b) p f) (λb, path.transport (λa, C a b) p (f b)) :=
path.rec_on p (λx, idp)
/- Maps on paths -/
/- The action of maps given by lambda. -/
definition ap_lambdaD [H : funext] {C : A' → Type} (p : a ≈ a') (f : Πa b, C b) :
ap (λa b, f a b) p ≈ path_pi (λb, ap (λa, f a b) p) :=
begin
apply (path.rec_on p),
apply inverse,
apply path_pi_idp
end
/- Dependent paths -/
/- with more implicit arguments the conclusion of the following theorem is
(Π(b : B a), transportD B C p b (f b) ≈ g (path.transport B p b)) ≃
(path.transport (λa, Π(b : B a), C a b) p f ≈ g) -/
definition dpath_pi [H : funext] (p : a ≈ a') (f : Π(b : B a), C a b) (g : Π(b' : B a'), C a' b')
: (Π(b : B a), p ▹D (f b) ≈ g (p ▹ b)) ≃ (p ▹ f ≈ g) :=
path.rec_on p (λg, !homotopy_equiv_path) g
section open sigma.ops
/- more implicit arguments:
(Π(b : B a), path.transport C (sigma.path p idp) (f b) ≈ g (p ▹ b)) ≃
(Π(b : B a), transportD B (λ(a : A) (b : B a), C ⟨a, b⟩) p b (f b) ≈ g (path.transport B p b)) -/
definition dpath_pi_sigma {C : (Σa, B a) → Type} (p : a ≈ a')
(f : Π(b : B a), C ⟨a, b⟩) (g : Π(b' : B a'), C ⟨a', b'⟩) :
(Π(b : B a), (sigma.path p idp) ▹ (f b) ≈ g (p ▹ b)) ≃ (Π(b : B a), p ▹D (f b) ≈ g (p ▹ b)) :=
path.rec_on p (λg, !equiv.refl) g
end
/- truncation -/
open truncation open truncation
definition trunc_pi [instance] (B : A → Type.{k}) (n : trunc_index) definition trunc_pi [instance] [H : funext.{l k}] (B : A → Type.{k}) (n : trunc_index)
[H : Πa, is_trunc n (B a)] : is_trunc n (Πa, B a) := [H : a, is_trunc n (B a)] : is_trunc n (Πa, B a) :=
begin begin
reverts (B, H), reverts (B, H),
apply (truncation.trunc_index.rec_on n), apply (trunc_index.rec_on n),
intros (B, H), intros (B, H),
fapply is_contr.mk, fapply is_contr.mk,
intro a, apply center, apply H, --remove "apply H" when term synthesis works correctly intro a, apply center, apply H, --remove "apply H" when term synthesis works correctly
intro f, apply path_forall, intro f, apply path_pi,
intro x, apply (contr (f x)), intro x, apply (contr (f x)),
intros (n, IH, B, H), intros (n, IH, B, H),
fapply is_trunc_succ, intros (f, g), fapply is_trunc_succ, intros (f, g),
fapply trunc_equiv', fapply trunc_equiv',
apply equiv.symm, apply equiv_apD10, apply equiv.symm, apply path_equiv_homotopy,
apply IH, apply IH,
intro a, intro a,
show is_trunc n (f a ≈ g a), from show is_trunc n (f a ≈ g a), from
succ_is_trunc n (f a) (g a) succ_is_trunc n (f a) (g a)
end end
definition trunc_path_pi [instance] (n : trunc_index) (f g : Πa, B a) definition trunc_path_pi [instance] [H : funext.{l k}] (n : trunc_index) (f g : Πa, B a)
[H : ∀a, is_trunc n (f a ≈ g a)] : is_trunc n (f ≈ g) := [H : ∀a, is_trunc n (f a ≈ g a)] : is_trunc n (f ≈ g) :=
begin begin
apply trunc_equiv', apply equiv.symm, apply trunc_equiv', apply equiv.symm,
apply equiv_apD10, apply path_equiv_homotopy,
apply trunc_pi, exact H, apply trunc_pi, exact H,
end end

View file

@ -15,7 +15,8 @@ variables {A A' B B' C D : Type}
namespace prod namespace prod
definition eta_prod (u : A × B) : (pr₁ u , pr₂ u) ≈ u := -- prod.eta is already used for the eta rule for strict equality
protected definition peta (u : A × B) : (pr₁ u , pr₂ u) ≈ u :=
destruct u (λu1 u2, idp) destruct u (λu1 u2, idp)
definition pair_path (pa : a ≈ a') (pb : b ≈ b') : (a , b) ≈ (a' , b') := definition pair_path (pa : a ≈ a') (pb : b ≈ b') : (a , b) ≈ (a' , b') :=
@ -25,25 +26,22 @@ namespace prod
begin begin
apply (prod.rec_on u), intros (a₁, b₁), apply (prod.rec_on u), intros (a₁, b₁),
apply (prod.rec_on v), intros (a₂, b₂, H₁, H₂), apply (prod.rec_on v), intros (a₂, b₂, H₁, H₂),
apply (transport _ (eta_prod (a₁, b₁))), apply (transport _ (peta (a₁, b₁))),
apply (transport _ (eta_prod (a₂, b₂))), apply (transport _ (peta (a₂, b₂))),
apply (pair_path H₁ H₂), apply (pair_path H₁ H₂),
end end
/- Symmetry -/ /- Symmetry -/
definition isequiv_prod_symm [instance] (A B : Type) : is_equiv (@flip A B) := definition isequiv_flip [instance] (A B : Type) : is_equiv (@flip A B) :=
adjointify flip adjointify flip
flip flip
(λu, destruct u (λb a, idp)) (λu, destruct u (λb a, idp))
(λu, destruct u (λa b, idp)) (λu, destruct u (λa b, idp))
definition equiv_prod_symm (A B : Type) : A × B ≃ B × A := definition symm_equiv (A B : Type) : A × B ≃ B × A :=
equiv.mk flip _ equiv.mk flip _
-- Pairs preserve truncatedness -- trunc_prod is defined in sigma
definition trunc_prod [instance] {A B : Type} (n : trunc_index) :
(is_trunc n A) → (is_trunc n B) → is_trunc n (A × B)
:= sorry --Assignment for Floris
end prod end prod

View file

@ -7,193 +7,194 @@ Ported from Coq HoTT
Theorems about sigma-types (dependent sums) Theorems about sigma-types (dependent sums)
-/ -/
import ..trunc .prod import ..trunc .prod ..axioms.funext
open path sigma sigma.ops equiv is_equiv open path sigma sigma.ops equiv is_equiv
namespace sigma namespace sigma
-- remove the ₁'s (globally) infixr [local] ∘ := function.compose --remove
variables {A A' : Type} {B : A → Type} {B' : A' → Type} {C : Πa, B a → Type} variables {A A' : Type} {B : A → Type} {B' : A' → Type} {C : Πa, B a → Type}
{D : Πa b, C a b → Type} {D : Πa b, C a b → Type}
{a a' a'' : A} {b b₁ b₂ : B a} {b' : B a'} {b'' : B a''} {u v w : Σa, B a} {a a' a'' : A} {b b₁ b₂ : B a} {b' : B a'} {b'' : B a''} {u v w : Σa, B a}
definition eta_sigma (u : Σa, B a) : ⟨u.1 , u.2⟩ ≈ u := -- sigma.eta is already used for the eta rule for strict equality
protected definition peta (u : Σa, B a) : ⟨u.1 , u.2⟩ ≈ u :=
destruct u (λu1 u2, idp) destruct u (λu1 u2, idp)
definition eta2_sigma (u : Σa b, C a b) : ⟨u.1, u.2.1, u.2.2⟩ ≈ u := definition eta2 (u : Σa b, C a b) : ⟨u.1, u.2.1, u.2.2⟩ ≈ u :=
destruct u (λu1 u2, destruct u2 (λu21 u22, idp)) destruct u (λu1 u2, destruct u2 (λu21 u22, idp))
definition eta3_sigma (u : Σa b c, D a b c) : ⟨u.1, u.2.1, u.2.2.1, u.2.2.2⟩ ≈ u := definition eta3 (u : Σa b c, D a b c) : ⟨u.1, u.2.1, u.2.2.1, u.2.2.2⟩ ≈ u :=
destruct u (λu1 u2, destruct u2 (λu21 u22, destruct u22 (λu221 u222, idp))) destruct u (λu1 u2, destruct u2 (λu21 u22, destruct u22 (λu221 u222, idp)))
definition path_sigma_dpair (p : a ≈ a') (q : p ▹ b ≈ b') : dpair a b ≈ dpair a' b' := definition dpair_eq_dpair (p : a ≈ a') (q : p ▹ b ≈ b') : dpair a b ≈ dpair a' b' :=
path.rec_on p (λb b' q, path.rec_on q idp) b b' q path.rec_on p (λb b' q, path.rec_on q idp) b b' q
/- In Coq they often have to give u and v explicitly -/ /- In Coq they often have to give u and v explicitly -/
definition path_sigma (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) : u ≈ v := protected definition path (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) : u ≈ v :=
destruct u destruct u
(λu1 u2, destruct v (λ v1 v2, path_sigma_dpair)) (λu1 u2, destruct v (λ v1 v2, dpair_eq_dpair))
p q p q
/- Projections of paths from a total space -/ /- Projections of paths from a total space -/
definition pr1_path (p : u ≈ v) : u.1 ≈ v.1 := definition path_pr1 (p : u ≈ v) : u.1 ≈ v.1 :=
ap dpr1 p ap dpr1 p
postfix `..1`:(max+1) := pr1_path postfix `..1`:(max+1) := path_pr1
definition pr2_path (p : u ≈ v) : p..1 ▹ u.2 ≈ v.2 := definition path_pr2 (p : u ≈ v) : p..1 ▹ u.2 ≈ v.2 :=
path.rec_on p idp path.rec_on p idp
--Coq uses the following proof, which only computes if u,v are dpairs AND p is idp --Coq uses the following proof, which only computes if u,v are dpairs AND p is idp
--(transport_compose B dpr1 p u.2)⁻¹ ⬝ apD dpr2 p --(transport_compose B dpr1 p u.2)⁻¹ ⬝ apD dpr2 p
postfix `..2`:(max+1) := pr2_path postfix `..2`:(max+1) := path_pr2
definition dpair_path_sigma (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) definition dpair_sigma_path (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2)
: dpair (path_sigma p q)..1 (path_sigma p q)..2 ≈ ⟨p, q⟩ := : dpair (sigma.path p q)..1 (sigma.path p q)..2 ≈ ⟨p, q⟩ :=
begin begin
generalize q, generalize p, reverts (p, q),
apply (destruct u), intros (u1, u2), apply (destruct u), intros (u1, u2),
apply (destruct v), intros (v1, v2, p), generalize v2, apply (destruct v), intros (v1, v2, p), generalize v2, --change to revert
apply (path.rec_on p), intros (v2, q), apply (path.rec_on p), intros (v2, q),
apply (path.rec_on q), apply idp apply (path.rec_on q), apply idp
end end
definition pr1_path_sigma (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) : (path_sigma p q)..1 ≈ p := definition sigma_path_pr1 (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) : (sigma.path p q)..1 ≈ p :=
(!dpair_path_sigma)..1 (!dpair_sigma_path)..1
definition pr2_path_sigma (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) definition sigma_path_pr2 (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2)
: pr1_path_sigma p q ▹ (path_sigma p q)..2 ≈ q := : sigma_path_pr1 p q ▹ (sigma.path p q)..2 ≈ q :=
(!dpair_path_sigma)..2 (!dpair_sigma_path)..2
definition eta_path_sigma (p : u ≈ v) : path_sigma (p..1) (p..2) ≈ p := definition sigma_path_eta (p : u ≈ v) : sigma.path (p..1) (p..2) ≈ p :=
begin begin
apply (path.rec_on p), apply (path.rec_on p),
apply (destruct u), intros (u1, u2), apply (destruct u), intros (u1, u2),
apply idp apply idp
end end
definition transport_pr1_path_sigma {B' : A → Type} (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2) definition transport_dpr1_sigma_path {B' : A → Type} (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2)
: transport (λx, B' x.1) (path_sigma p q) ≈ transport B' p := : transport (λx, B' x.1) (sigma.path p q) ≈ transport B' p :=
begin begin
generalize q, generalize p, reverts (p, q),
apply (destruct u), intros (u1, u2), apply (destruct u), intros (u1, u2),
apply (destruct v), intros (v1, v2, p), generalize v2, apply (destruct v), intros (v1, v2, p), generalize v2,
apply (path.rec_on p), intros (v2, q), apply (path.rec_on p), intros (v2, q),
apply (path.rec_on q), apply idp apply (path.rec_on q), apply idp
end end
/- the uncurried version of path_sigma. We will prove that this is an equivalence -/ /- the uncurried version of sigma_path. We will prove that this is an equivalence -/
definition path_sigma_uncurried (pq : Σ(p : dpr1 u ≈ dpr1 v), p ▹ (dpr2 u) ≈ dpr2 v) : u ≈ v := definition sigma_path_uncurried (pq : Σ(p : dpr1 u ≈ dpr1 v), p ▹ (dpr2 u) ≈ dpr2 v) : u ≈ v :=
destruct pq path_sigma destruct pq sigma.path
definition dpair_path_sigma_uncurried (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2) definition dpair_sigma_path_uncurried (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2)
: dpair (path_sigma_uncurried pq)..1 (path_sigma_uncurried pq)..2 ≈ pq := : dpair (sigma_path_uncurried pq)..1 (sigma_path_uncurried pq)..2 ≈ pq :=
destruct pq dpair_path_sigma destruct pq dpair_sigma_path
definition pr1_path_sigma_uncurried (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2) definition sigma_path_pr1_uncurried (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2)
: (path_sigma_uncurried pq)..1 ≈ pq.1 := : (sigma_path_uncurried pq)..1 ≈ pq.1 :=
(!dpair_path_sigma_uncurried)..1 (!dpair_sigma_path_uncurried)..1
definition pr2_path_sigma_uncurried (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2) definition sigma_path_pr2_uncurried (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2)
: (pr1_path_sigma_uncurried pq) ▹ (path_sigma_uncurried pq)..2 ≈ pq.2 := : (sigma_path_pr1_uncurried pq) ▹ (sigma_path_uncurried pq)..2 ≈ pq.2 :=
(!dpair_path_sigma_uncurried)..2 (!dpair_sigma_path_uncurried)..2
definition eta_path_sigma_uncurried (p : u ≈ v) : path_sigma_uncurried (dpair p..1 p..2) ≈ p := definition sigma_path_eta_uncurried (p : u ≈ v) : sigma_path_uncurried (dpair p..1 p..2) ≈ p :=
!eta_path_sigma !sigma_path_eta
definition transport_pr1_path_sigma_uncurried {B' : A → Type} (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2) definition transport_sigma_path_dpr1_uncurried {B' : A → Type}
: transport (λx, B' x.1) (@path_sigma_uncurried A B u v pq) ≈ transport B' pq.1 := (pq : Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2)
destruct pq transport_pr1_path_sigma : transport (λx, B' x.1) (@sigma_path_uncurried A B u v pq) ≈ transport B' pq.1 :=
destruct pq transport_dpr1_sigma_path
definition isequiv_path_sigma /-[instance]-/ (u v : Σa, B a) definition is_equiv_sigma_path [instance] (u v : Σa, B a)
: is_equiv (@path_sigma_uncurried A B u v) := : is_equiv (@sigma_path_uncurried A B u v) :=
adjointify path_sigma_uncurried adjointify sigma_path_uncurried
(λp, ⟨p..1, p..2⟩) (λp, ⟨p..1, p..2⟩)
eta_path_sigma_uncurried sigma_path_eta_uncurried
dpair_path_sigma_uncurried dpair_sigma_path_uncurried
definition equiv_path_sigma (u v : Σa, B a) : (Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2) ≃ (u ≈ v) := definition equiv_sigma_path (u v : Σa, B a) : (Σ(p : u.1 ≈ v.1), p ▹ u.2 ≈ v.2) ≃ (u ≈ v) :=
equiv.mk path_sigma_uncurried !isequiv_path_sigma equiv.mk sigma_path_uncurried !is_equiv_sigma_path
definition path_sigma_dpair_pp_pp (p1 : a ≈ a' ) (q1 : p1 ▹ b ≈ b' ) definition dpair_eq_dpair_pp_pp (p1 : a ≈ a' ) (q1 : p1 ▹ b ≈ b' )
(p2 : a' ≈ a'') (q2 : p2 ▹ b' ≈ b'') : (p2 : a' ≈ a'') (q2 : p2 ▹ b' ≈ b'') :
path_sigma_dpair (p1 ⬝ p2) (transport_pp B p1 p2 b ⬝ ap (transport B p2) q1 ⬝ q2) dpair_eq_dpair (p1 ⬝ p2) (transport_pp B p1 p2 b ⬝ ap (transport B p2) q1 ⬝ q2)
path_sigma_dpair p1 q1 ⬝ path_sigma_dpair p2 q2 := dpair_eq_dpair p1 q1 ⬝ dpair_eq_dpair p2 q2 :=
begin begin
generalize q2, generalize q1, generalize b'', generalize p2, generalize b', reverts (b', p2, b'', q1, q2),
apply (path.rec_on p1), intros (b', p2), apply (path.rec_on p1), intros (b', p2),
apply (path.rec_on p2), intros (b'', q1), apply (path.rec_on p2), intros (b'', q1),
apply (path.rec_on q1), intro q2, apply (path.rec_on q1), intro q2,
apply (path.rec_on q2), apply idp apply (path.rec_on q2), apply idp
end end
definition path_sigma_pp_pp (p1 : u.1 ≈ v.1) (q1 : p1 ▹ u.2 ≈ v.2) definition sigma_path_pp_pp (p1 : u.1 ≈ v.1) (q1 : p1 ▹ u.2 ≈ v.2)
(p2 : v.1 ≈ w.1) (q2 : p2 ▹ v.2 ≈ w.2) : (p2 : v.1 ≈ w.1) (q2 : p2 ▹ v.2 ≈ w.2) :
path_sigma (p1 ⬝ p2) (transport_pp B p1 p2 u.2 ⬝ ap (transport B p2) q1 ⬝ q2) sigma.path (p1 ⬝ p2) (transport_pp B p1 p2 u.2 ⬝ ap (transport B p2) q1 ⬝ q2)
path_sigma p1 q1 ⬝ path_sigma p2 q2 := sigma.path p1 q1 ⬝ sigma.path p2 q2 :=
begin begin
generalize q2, generalize p2, generalize q1, generalize p1, reverts (p1, q1, p2, q2),
apply (destruct u), intros (u1, u2), apply (destruct u), intros (u1, u2),
apply (destruct v), intros (v1, v2), apply (destruct v), intros (v1, v2),
apply (destruct w), intros, apply (destruct w), intros,
apply path_sigma_dpair_pp_pp apply dpair_eq_dpair_pp_pp
end end
definition path_sigma_dpair_p1_1p (p : a ≈ a') (q : p ▹ b ≈ b') : definition dpair_eq_dpair_p1_1p (p : a ≈ a') (q : p ▹ b ≈ b') :
path_sigma_dpair p q ≈ path_sigma_dpair p idp ⬝ path_sigma_dpair idp q := dpair_eq_dpair p q ≈ dpair_eq_dpair p idp ⬝ dpair_eq_dpair idp q :=
begin begin
generalize q, generalize b', reverts (b', q),
apply (path.rec_on p), intros (b', q), apply (path.rec_on p), intros (b', q),
apply (path.rec_on q), apply idp apply (path.rec_on q), apply idp
end end
/- pr1_path commutes with the groupoid structure. -/ /- path_pr1 commutes with the groupoid structure. -/
definition pr1_path_1 (u : Σa, B a) : (idpath u)..1 ≈ idpath (u.1) := idp definition path_pr1_idp (u : Σa, B a) : (idpath u)..1 ≈ idpath (u.1) := idp
definition pr1_path_pp (p : u ≈ v) (q : v ≈ w) : (p ⬝ q) ..1 ≈ (p..1) ⬝ (q..1) := !ap_pp definition path_pr1_pp (p : u ≈ v) (q : v ≈ w) : (p ⬝ q) ..1 ≈ (p..1) ⬝ (q..1) := !ap_pp
definition pr1_path_V (p : u ≈ v) : p⁻¹ ..1 ≈ (p..1)⁻¹ := !ap_V definition path_pr1_V (p : u ≈ v) : p⁻¹ ..1 ≈ (p..1)⁻¹ := !ap_V
/- Applying dpair to one argument is the same as path_sigma_dpair with reflexivity in the first place. -/ /- Applying dpair to one argument is the same as dpair_eq_dpair with reflexivity in the first place. -/
definition ap_dpair (q : b₁ ≈ b₂) : ap (dpair a) q ≈ path_sigma_dpair idp q := definition ap_dpair (q : b₁ ≈ b₂) : ap (dpair a) q ≈ dpair_eq_dpair idp q :=
path.rec_on q idp path.rec_on q idp
/- Dependent transport is the same as transport along a path_sigma. -/ /- Dependent transport is the same as transport along a sigma_path. -/
definition transportD_is_transport (p : a ≈ a') (c : C a b) : definition transportD_eq_transport (p : a ≈ a') (c : C a b) :
p ▹D c ≈ transport (λu, C (u.1) (u.2)) (path_sigma_dpair p idp) c := p ▹D c ≈ transport (λu, C (u.1) (u.2)) (dpair_eq_dpair p idp) c :=
path.rec_on p idp path.rec_on p idp
definition path_path_sigma_path_sigma {p1 q1 : a ≈ a'} {p2 : p1 ▹ b ≈ b'} {q2 : q1 ▹ b ≈ b'} definition sigma_path_eq_sigma_path {p1 q1 : a ≈ a'} {p2 : p1 ▹ b ≈ b'} {q2 : q1 ▹ b ≈ b'}
(r : p1 ≈ q1) (s : r ▹ p2 ≈ q2) : path_sigma p1 p2 ≈ path_sigma q1 q2 := (r : p1 ≈ q1) (s : r ▹ p2 ≈ q2) : sigma.path p1 p2 ≈ sigma.path q1 q2 :=
path.rec_on r path.rec_on r
proof (λq2 s, path.rec_on s idp) qed proof (λq2 s, path.rec_on s idp) qed
q2 q2
s s
-- begin -- begin
-- generalize s, generalize q2, -- reverts (q2, s),
-- apply (path.rec_on r), intros (q2, s), -- apply (path.rec_on r), intros (q2, s),
-- apply (path.rec_on s), apply idp -- apply (path.rec_on s), apply idp
-- end -- end
/- A path between paths in a total space is commonly shown component wise. -/ /- A path between paths in a total space is commonly shown component wise. -/
definition path_path_sigma {p q : u ≈ v} (r : p..1 ≈ q..1) (s : r ▹ p..2 ≈ q..2) : p ≈ q := definition path_sigma_path {p q : u ≈ v} (r : p..1 ≈ q..1) (s : r ▹ p..2 ≈ q..2) : p ≈ q :=
begin begin
generalize s, generalize r, generalize q, reverts (q, r, s),
apply (path.rec_on p), apply (path.rec_on p),
apply (destruct u), intros (u1, u2, q, r, s), apply (destruct u), intros (u1, u2, q, r, s),
apply concat, rotate 1, apply concat, rotate 1,
apply eta_path_sigma, apply sigma_path_eta,
apply (path_path_sigma_path_sigma r s) apply (sigma_path_eq_sigma_path r s)
end end
/- In Coq they often have to give u and v explicitly when using the following definition -/ /- In Coq they often have to give u and v explicitly when using the following definition -/
definition path_path_sigma_uncurried {p q : u ≈ v} definition path_sigma_path_uncurried {p q : u ≈ v}
(rs : Σ(r : p..1 ≈ q..1), transport (λx, transport B x u.2 ≈ v.2) r p..2 ≈ q..2) : p ≈ q := (rs : Σ(r : p..1 ≈ q..1), transport (λx, transport B x u.2 ≈ v.2) r p..2 ≈ q..2) : p ≈ q :=
destruct rs path_path_sigma destruct rs path_sigma_path
/- Transport -/ /- Transport -/
@ -201,8 +202,8 @@ namespace sigma
In particular, this indicates why `transport` alone cannot be fully defined by induction on the structure of types, although Id-elim/transportD can be (cf. Observational Type Theory). A more thorough set of lemmas, along the lines of the present ones but dealing with Id-elim rather than just transport, might be nice to have eventually? -/ In particular, this indicates why `transport` alone cannot be fully defined by induction on the structure of types, although Id-elim/transportD can be (cf. Observational Type Theory). A more thorough set of lemmas, along the lines of the present ones but dealing with Id-elim rather than just transport, might be nice to have eventually? -/
definition transport_sigma (p : a ≈ a') (bc : Σ(b : B a), C a b) : p ▹ bc ≈ ⟨p ▹ bc.1, p ▹D bc.2⟩ definition transport_eq (p : a ≈ a') (bc : Σ(b : B a), C a b)
:= : p ▹ bc ≈ ⟨p ▹ bc.1, p ▹D bc.2⟩ :=
begin begin
apply (path.rec_on p), apply (path.rec_on p),
apply (destruct bc), intros (b, c), apply (destruct bc), intros (b, c),
@ -210,7 +211,7 @@ namespace sigma
end end
/- The special case when the second variable doesn't depend on the first is simpler. -/ /- The special case when the second variable doesn't depend on the first is simpler. -/
definition transport_sigma' {B : Type} {C : A → B → Type} (p : a ≈ a') (bc : Σ(b : B), C a b) definition transport_eq_deg {B : Type} {C : A → B → Type} (p : a ≈ a') (bc : Σ(b : B), C a b)
: p ▹ bc ≈ ⟨bc.1, p ▹ bc.2⟩ := : p ▹ bc ≈ ⟨bc.1, p ▹ bc.2⟩ :=
begin begin
apply (path.rec_on p), apply (path.rec_on p),
@ -218,12 +219,12 @@ namespace sigma
apply idp apply idp
end end
/- Or if the second variable contains a first component that doesn't depend on the first. Need to think about the naming of these. -/ /- Or if the second variable contains a first component that doesn't depend on the first. -/
definition transport_sigma_' {C : A → Type} {D : Π a:A, B a → C a → Type} (p : a ≈ a') definition transport_eq_4deg {C : A → Type} {D : Π a:A, B a → C a → Type} (p : a ≈ a')
(bcd : Σ(b : B a) (c : C a), D a b c) : p ▹ bcd ≈ ⟨p ▹ bcd.1, p ▹ bcd.2.1, p ▹D2 bcd.2.2⟩ := (bcd : Σ(b : B a) (c : C a), D a b c) : p ▹ bcd ≈ ⟨p ▹ bcd.1, p ▹ bcd.2.1, p ▹D2 bcd.2.2⟩ :=
begin begin
generalize bcd, revert bcd,
apply (path.rec_on p), intro bcd, apply (path.rec_on p), intro bcd,
apply (destruct bcd), intros (b, cd), apply (destruct bcd), intros (b, cd),
apply (destruct cd), intros (c, d), apply (destruct cd), intros (c, d),
@ -233,58 +234,117 @@ namespace sigma
/- Functorial action -/ /- Functorial action -/
variables (f : A → A') (g : Πa, B a → B' (f a)) variables (f : A → A') (g : Πa, B a → B' (f a))
definition functor_sigma (u : Σa, B a) : Σa', B' a' := protected definition functor (u : Σa, B a) : Σa', B' a' :=
⟨f u.1, g u.1 u.2⟩ ⟨f u.1, g u.1 u.2⟩
/- Equivalences -/ /- Equivalences -/
--remove explicit arguments of IsEquiv --TODO: remove explicit arguments of is_equiv
definition isequiv_functor_sigma [H1 : is_equiv f] [H2 : Π a, @is_equiv (B a) (B' (f a)) (g a)] definition is_equiv_functor [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)]
: is_equiv (functor_sigma f g) := : is_equiv (functor f g) :=
/-adjointify (functor_sigma f g) adjointify (functor f g)
(functor_sigma (f⁻¹) (λ(x : A') (y : B' x), ((g (f⁻¹ x))⁻¹ ((retr f x)⁻¹ ▹ y)))) (functor (f⁻¹) (λ(a' : A') (b' : B' a'),
sorry-/ ((g (f⁻¹ a'))⁻¹ (transport B' (retr f a'⁻¹) b'))))
sorry begin
intro u',
apply (destruct u'), intros (a', b'),
apply (sigma.path (retr f a')),
-- "rewrite retr (g (f⁻¹ a'))"
apply concat, apply (ap (λx, (transport B' (retr f a') x))), apply (retr (g (f⁻¹ a'))),
show retr f a' ▹ (((retr f a') ⁻¹) ▹ b') ≈ b',
from transport_pV B' (retr f a') b'
end
begin
intro u,
apply (destruct u), intros (a, b),
apply (sigma.path (sect f a)),
show transport B (sect f a) (g (f⁻¹ (f a))⁻¹ (transport B' (retr f (f a)⁻¹) (g a b))) ≈ b,
from calc
transport B (sect f a) (g (f⁻¹ (f a))⁻¹ (transport B' (retr f (f a)⁻¹) (g a b)))
≈ g a⁻¹ (transport (B' ∘ f) (sect f a) (transport B' (retr f (f a)⁻¹) (g a b)))
: ap_transport (sect f a) (λ a, g a⁻¹)
... ≈ g a⁻¹ (transport B' (ap f (sect f a)) (transport B' (retr f (f a)⁻¹) (g a b)))
: ap (g a⁻¹) !transport_compose
... ≈ g a⁻¹ (transport B' (ap f (sect f a)) (transport B' (ap f (sect f a)⁻¹) (g a b)))
: ap (λ x, g a⁻¹ (transport B' (ap f (sect f a)) (transport B' (x⁻¹) (g a b)))) (adj f a)
... ≈ g a⁻¹ (g a b) : transport_pV
... ≈ b : sect (g a) b
end
-- -- "rewrite ap_transport"
-- apply concat, apply inverse, apply (ap_transport (sect f a) (λ a, g a⁻¹)),
-- apply concat, apply (ap (g a⁻¹)),
-- -- "rewrite transport_compose"
-- apply concat, apply transport_compose,
-- -- "rewrite adj"
-- -- "rewrite transport_pV"
-- apply sect,
definition equiv_functor_sigma [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)] : (Σa, B a) ≃ (Σa', B' a') := definition equiv_functor_of_is_equiv [H1 : is_equiv f] [H2 : Π a, is_equiv (g a)]
equiv.mk (functor_sigma f g) !isequiv_functor_sigma : (Σa, B a) ≃ (Σa', B' a') :=
equiv.mk (functor f g) !is_equiv_functor
context --remove context --remove
irreducible inv function.compose --remove irreducible inv function.compose --remove
definition equiv_functor_sigma' (Hf : A ≃ A') (Hg : Π a, B a ≃ B' (to_fun Hf a)) : definition equiv_functor (Hf : A ≃ A') (Hg : Π a, B a ≃ B' (to_fun Hf a)) :
(Σa, B a) ≃ (Σa', B' a') := (Σa, B a) ≃ (Σa', B' a') :=
equiv_functor_sigma (to_fun Hf) (λ a, to_fun (Hg a)) equiv_functor_of_is_equiv (to_fun Hf) (λ a, to_fun (Hg a))
end --remove end --remove
definition equiv_functor_id {B' : A → Type} (Hg : Π a, B a ≃ B' a) : (Σa, B a) ≃ Σa, B' a :=
equiv_functor equiv.refl Hg
definition ap_functor_sigma_dpair (p : a ≈ a') (q : p ▹ b ≈ b')
: ap (sigma.functor f g) (sigma.path p q)
≈ sigma.path (ap f p)
(transport_compose _ f p (g a b)⁻¹ ⬝ ap_transport p g b⁻¹ ⬝ ap (g a') q) :=
begin
reverts (b', q),
apply (path.rec_on p),
intros (b', q), apply (path.rec_on q),
apply idp
end
definition ap_functor_sigma (p : u.1 ≈ v.1) (q : p ▹ u.2 ≈ v.2)
: ap (sigma.functor f g) (sigma.path p q)
≈ sigma.path (ap f p)
(transport_compose B' f p (g u.1 u.2)⁻¹ ⬝ ap_transport p g u.2⁻¹ ⬝ ap (g v.1) q) :=
begin
reverts (p, q),
apply (destruct u), intros (a, b),
apply (destruct v), intros (a', b', p, q),
apply ap_functor_sigma_dpair
end
/- definition 3.11.9(i): Summing up a contractible family of types does nothing. -/ /- definition 3.11.9(i): Summing up a contractible family of types does nothing. -/
open truncation open truncation
definition isequiv_pr1_contr [instance] (B : A → Type) [H : Π a, is_contr (B a)] definition is_equiv_dpr1 [instance] (B : A → Type) [H : Π a, is_contr (B a)]
: is_equiv (@dpr1 A B) := : is_equiv (@dpr1 A B) :=
adjointify dpr1 adjointify dpr1
(λa, ⟨a, !center⟩) (λa, ⟨a, !center⟩)
(λa, idp) (λa, idp)
(λu, path_sigma idp !contr) (λu, sigma.path idp !contr)
definition equiv_sigma_contr [H : Π a, is_contr (B a)] : (Σa, B a) ≃ A := definition equiv_of_all_contr [H : Π a, is_contr (B a)] : (Σa, B a) ≃ A :=
equiv.mk dpr1 _ equiv.mk dpr1 _
/- definition 3.11.9(ii): Dually, summing up over a contractible type does nothing. -/ /- definition 3.11.9(ii): Dually, summing up over a contractible type does nothing. -/
definition equiv_contr_sigma (B : A → Type) [H : is_contr A] : (Σa, B a) ≃ B (center A) := definition equiv_center_of_contr (B : A → Type) [H : is_contr A] : (Σa, B a) ≃ B (center A)
:=
equiv.mk _ (adjointify equiv.mk _ (adjointify
(λu, contr u.1⁻¹ ▹ u.2) (λu, contr u.1⁻¹ ▹ u.2)
(λb, ⟨!center, b⟩) (λb, ⟨!center, b⟩)
(λb, ap (λx, x ▹ b) !path2_contr) (λb, ap (λx, x ▹ b) !path2_contr)
(λu, path_sigma !contr !transport_pV)) (λu, sigma.path !contr !transport_pV))
/- Associativity -/ /- Associativity -/
--this proof is harder here than in Coq because we don't have eta definitionally for sigma --this proof is harder than in Coq because we don't have eta definitionally for sigma
definition equiv_sigma_assoc (C : (Σa, B a) → Type) : (Σa b, C ⟨a,b⟩) ≃ (Σu, C u) := protected definition assoc_equiv (C : (Σa, B a) → Type) : (Σa b, C ⟨a, b⟩) ≃ (Σu, C u) :=
-- begin -- begin
-- apply Equiv.mk, -- apply equiv.mk,
-- apply (adjointify (λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩) -- apply (adjointify (λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩)
-- (λuc, ⟨uc.1.1, uc.1.2, !eta_sigma⁻¹ ▹ uc.2⟩)), -- (λuc, ⟨uc.1.1, uc.1.2, !peta⁻¹ ▹ uc.2⟩)),
-- intro uc, apply (destruct uc), intro u, -- intro uc, apply (destruct uc), intro u,
-- apply (destruct u), intros (a, b, c), -- apply (destruct u), intros (a, b, c),
-- apply idp, -- apply idp,
@ -294,64 +354,121 @@ namespace sigma
-- end -- end
equiv.mk _ (adjointify equiv.mk _ (adjointify
(λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩) (λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩)
(λuc, ⟨uc.1.1, uc.1.2, !eta_sigma⁻¹ ▹ uc.2⟩) (λuc, ⟨uc.1.1, uc.1.2, !peta⁻¹ ▹ uc.2⟩)
proof (λuc, destruct uc (λu, destruct u (λa b c, idp))) qed proof (λuc, destruct uc (λu, destruct u (λa b c, idp))) qed
proof (λav, destruct av (λa v, destruct v (λb c, idp))) qed) proof (λav, destruct av (λa v, destruct v (λb c, idp))) qed)
open prod open prod
definition equiv_sigma_prod (C : (A × A') → Type) : (Σa a', C (a,a')) ≃ (Σu, C u) := definition assoc_equiv_prod (C : (A × A') → Type) : (Σa a', C (a,a')) ≃ (Σu, C u) :=
equiv.mk _ (adjointify equiv.mk _ (adjointify
(λav, ⟨(av.1, av.2.1), av.2.2⟩) (λav, ⟨(av.1, av.2.1), av.2.2⟩)
(λuc, ⟨pr₁ (uc.1), pr₂ (uc.1), !eta_prod⁻¹ ▹ uc.2⟩) (λuc, ⟨pr₁ (uc.1), pr₂ (uc.1), !prod.peta⁻¹ ▹ uc.2⟩)
proof (λuc, destruct uc (λu, prod.destruct u (λa b c, idp))) qed proof (λuc, destruct uc (λu, prod.destruct u (λa b c, idp))) qed
proof (λav, destruct av (λa v, destruct v (λb c, idp))) qed) proof (λav, destruct av (λa v, destruct v (λb c, idp))) qed)
/- Symmetry -/ /- Symmetry -/
-- if this breaks, replace "Equiv.id" by "proof Equiv.id qed" definition symm_equiv_uncurried (C : A × A' → Type) : (Σa a', C (a, a')) ≃ (Σa' a, C (a, a')) :=
definition equiv_sigma_symm_prod (C : A × A' → Type) : (Σa a', C (a, a')) ≃ (Σa' a, C (a, a')) :=
calc calc
(Σa a', C (a, a')) ≃ Σu, C u : equiv_sigma_prod (Σa a', C (a, a')) ≃ Σu, C u : assoc_equiv_prod
... ≃ Σv, C (flip v) : equiv_functor_sigma' !equiv_prod_symm ... ≃ Σv, C (flip v) : equiv_functor !prod.symm_equiv
(λu, prod.destruct u (λa a', equiv.refl)) (λu, prod.destruct u (λa a', equiv.refl))
... ≃ (Σa' a, C (a, a')) : equiv_sigma_prod ... ≃ (Σa' a, C (a, a')) : assoc_equiv_prod
definition equiv_sigma_symm (C : A → A' → Type) : (Σa a', C a a') ≃ (Σa' a, C a a') := definition symm_equiv (C : A → A' → Type) : (Σa a', C a a') ≃ (Σa' a, C a a') :=
sigma.equiv_sigma_symm_prod (λu, C (pr1 u) (pr2 u)) symm_equiv_uncurried (λu, C (pr1 u) (pr2 u))
definition equiv_sigma0_prod (A B : Type) : (Σ(a : A), B) ≃ A × B := definition equiv_prod (A B : Type) : (Σ(a : A), B) ≃ A × B :=
equiv.mk _ (adjointify equiv.mk _ (adjointify
(λs, (s.1, s.2)) (λs, (s.1, s.2))
(λp, ⟨pr₁ p, pr₂ p⟩) (λp, ⟨pr₁ p, pr₂ p⟩)
proof (λp, prod.destruct p (λa b, idp)) qed proof (λp, prod.destruct p (λa b, idp)) qed
proof (λs, destruct s (λa b, idp)) qed) proof (λs, destruct s (λa b, idp)) qed)
definition equiv_sigma_symm0 (A B : Type) : (Σ(a : A), B) ≃ Σ(b : B), A := definition symm_equiv_deg (A B : Type) : (Σ(a : A), B) ≃ Σ(b : B), A :=
calc calc
(Σ(a : A), B) ≃ A × B : equiv_sigma0_prod (Σ(a : A), B) ≃ A × B : equiv_prod
... ≃ B × A : equiv_prod_symm ... ≃ B × A : prod.symm_equiv
... ≃ Σ(b : B), A : equiv_sigma0_prod ... ≃ Σ(b : B), A : equiv_prod
/- ** Universal mapping properties -/
/- *** The positive universal property. -/
section
open funext
--in Coq this can be done without function extensionality
definition is_equiv_sigma_rec [instance] [FUN : funext] (C : (Σa, B a) → Type)
: is_equiv (@sigma.rec _ _ C) :=
adjointify _ (λ g a b, g ⟨a, b⟩)
(λ g, proof path_pi (λu, destruct u (λa b, idp)) qed)
(λ f, idpath f)
definition equiv_sigma_rec [FUN : funext] (C : (Σa, B a) → Type)
: (Π(a : A) (b: B a), C ⟨a, b⟩) ≃ (Πxy, C xy) :=
equiv.mk sigma.rec _
/- *** The negative universal property. -/
definition coind_uncurried (fg : Σ(f : Πa, B a), Πa, C a (f a)) (a : A) : Σ(b : B a), C a b
:= ⟨fg.1 a, fg.2 a⟩
definition coind (f : Π a, B a) (g : Π a, C a (f a)) (a : A) : Σ(b : B a), C a b :=
coind_uncurried ⟨f, g⟩ a
--is the instance below dangerous?
--in Coq this can be done without function extensionality
definition is_equiv_coind [instance] [FUN : funext] (C : Πa, B a → Type)
: is_equiv (@coind_uncurried _ _ C) :=
adjointify _ (λ h, ⟨λa, (h a).1, λa, (h a).2⟩)
(λ h, proof path_pi (λu, !peta) qed)
(λfg, destruct fg (λ(f : Π (a : A), B a) (g : Π (x : A), C x (f x)), proof idp qed))
definition equiv_coind [FUN : funext] : (Σ(f : Πa, B a), Πa, C a (f a)) ≃ (Πa, Σb, C a b) :=
equiv.mk coind_uncurried _
end
/- ** Subtypes (sigma types whose second components are hprops) -/
/- To prove equality in a subtype, we only need equality of the first component. -/
definition path_hprop [H : Πa, is_hprop (B a)] (u v : Σa, B a) : u.1 ≈ v.1 → u ≈ v :=
(sigma_path_uncurried ∘ (@inv _ _ dpr1 (@is_equiv_dpr1 _ _ (λp, !succ_is_trunc))))
definition is_equiv_path_hprop [instance] [H : Πa, is_hprop (B a)] (u v : Σa, B a)
: is_equiv (path_hprop u v) :=
!is_equiv.compose
definition equiv_path_hprop [H : Πa, is_hprop (B a)] (u v : Σa, B a) : (u.1 ≈ v.1) ≃ (u ≈ v)
:=
equiv.mk !path_hprop _
/- truncatedness -/ /- truncatedness -/
definition trunc_sigma [instance] (B : A → Type) (n : trunc_index) definition trunc_sigma [instance] (B : A → Type) (n : trunc_index)
[HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) := [HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) :=
begin begin
reverts (A, B, HA, HB), reverts (A, B, HA, HB),
apply (truncation.trunc_index.rec_on n), apply (trunc_index.rec_on n),
intros (A, B, HA, HB), intros (A, B, HA, HB),
fapply trunc_equiv', fapply trunc_equiv',
apply equiv.symm, apply equiv.symm,
apply equiv_contr_sigma, apply HA, --should be solved by term synthesis apply equiv_center_of_contr, apply HA, --should be solved by term synthesis
apply HB, apply HB,
intros (n, IH, A, B, HA, HB), intros (n, IH, A, B, HA, HB),
fapply is_trunc_succ, intros (u, v), fapply is_trunc_succ, intros (u, v),
fapply trunc_equiv', fapply trunc_equiv',
apply equiv_path_sigma, apply equiv_sigma_path,
apply IH, apply IH,
apply (succ_is_trunc n), apply succ_is_trunc, assumption,
intro p, intro p,
show is_trunc n (p ▹ u .2 ≈ v .2), from show is_trunc n (p ▹ u .2 ≈ v .2), from
succ_is_trunc n (p ▹ u.2) (v.2), succ_is_trunc n (p ▹ u.2) (v.2),
end end
end sigma end sigma
open truncation sigma
namespace prod
/- truncatedness -/
definition trunc_prod [instance] (A B : Type) (n : trunc_index)
[HA : is_trunc n A] [HB : is_trunc n B] : is_trunc n (A × B) :=
trunc_equiv' n !equiv_prod
end prod