feat(hott): make some arguments in init.path implicit and rename apD to apd
This commit is contained in:
parent
b98c109f73
commit
e769fdd9dc
21 changed files with 191 additions and 192 deletions
|
@ -189,7 +189,7 @@ namespace category
|
||||||
end)
|
end)
|
||||||
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = nat_trans.id :=
|
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = nat_trans.id :=
|
||||||
begin
|
begin
|
||||||
fapply (apD011 nat_trans.mk),
|
fapply (apd011 nat_trans.mk),
|
||||||
apply eq_of_homotopy, intro c, apply left_inverse,
|
apply eq_of_homotopy, intro c, apply left_inverse,
|
||||||
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
||||||
apply is_hset.elim
|
apply is_hset.elim
|
||||||
|
@ -197,7 +197,7 @@ namespace category
|
||||||
|
|
||||||
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = nat_trans.id :=
|
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = nat_trans.id :=
|
||||||
begin
|
begin
|
||||||
fapply (apD011 nat_trans.mk),
|
fapply (apd011 nat_trans.mk),
|
||||||
apply eq_of_homotopy, intro c, apply right_inverse,
|
apply eq_of_homotopy, intro c, apply right_inverse,
|
||||||
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
apply eq_of_homotopy, intros, apply eq_of_homotopy, intros, apply eq_of_homotopy, intros,
|
||||||
apply is_hset.elim
|
apply is_hset.elim
|
||||||
|
|
|
@ -50,7 +50,7 @@ namespace functor
|
||||||
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
|
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
|
||||||
(pF : F₁ = F₂) (pH : pF ▹ H₁ = H₂)
|
(pF : F₁ = F₂) (pH : pF ▹ H₁ = H₂)
|
||||||
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
||||||
apD01111 functor.mk pF pH !is_hprop.elim !is_hprop.elim
|
apd01111 functor.mk pF pH !is_hprop.elim !is_hprop.elim
|
||||||
|
|
||||||
definition functor_eq' {F₁ F₂ : C ⇒ D}
|
definition functor_eq' {F₁ F₂ : C ⇒ D}
|
||||||
: Π(p : to_fun_ob F₁ = to_fun_ob F₂),
|
: Π(p : to_fun_ob F₁ = to_fun_ob F₂),
|
||||||
|
@ -68,10 +68,10 @@ namespace functor
|
||||||
apply concat, rotate_left 1, apply transport_hom,
|
apply concat, rotate_left 1, apply transport_hom,
|
||||||
apply concat, rotate_left 1,
|
apply concat, rotate_left 1,
|
||||||
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c c') f),
|
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c c') f),
|
||||||
apply (apD10' f),
|
apply (apd10' f),
|
||||||
apply concat, rotate_left 1, esimp,
|
apply concat, rotate_left 1, esimp,
|
||||||
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c) c'),
|
exact (pi_transport_constant (eq_of_homotopy pF) (H₁ c) c'),
|
||||||
apply (apD10' c'),
|
apply (apd10' c'),
|
||||||
apply concat, rotate_left 1, esimp,
|
apply concat, rotate_left 1, esimp,
|
||||||
exact (pi_transport_constant (eq_of_homotopy pF) H₁ c),
|
exact (pi_transport_constant (eq_of_homotopy pF) H₁ c),
|
||||||
apply idp
|
apply idp
|
||||||
|
@ -169,7 +169,7 @@ namespace functor
|
||||||
definition functor_mk_eq'_idp (F : C → D) (H : Π(a b : C), hom a b → hom (F a) (F b))
|
definition functor_mk_eq'_idp (F : C → D) (H : Π(a b : C), hom a b → hom (F a) (F b))
|
||||||
(id comp) : functor_mk_eq' id id comp comp (idpath F) (idpath H) = idp :=
|
(id comp) : functor_mk_eq' id id comp comp (idpath F) (idpath H) = idp :=
|
||||||
begin
|
begin
|
||||||
fapply (apD011 (apD01111 functor.mk idp idp)),
|
fapply (apd011 (apd01111 functor.mk idp idp)),
|
||||||
apply is_hset.elim,
|
apply is_hset.elim,
|
||||||
apply is_hset.elim
|
apply is_hset.elim
|
||||||
end
|
end
|
||||||
|
@ -178,7 +178,7 @@ namespace functor
|
||||||
by (cases F; apply functor_mk_eq'_idp)
|
by (cases F; apply functor_mk_eq'_idp)
|
||||||
|
|
||||||
definition functor_eq_eta' {F₁ F₂ : C ⇒ D} (p : F₁ = F₂)
|
definition functor_eq_eta' {F₁ F₂ : C ⇒ D} (p : F₁ = F₂)
|
||||||
: functor_eq' (ap to_fun_ob p) (!transport_compose⁻¹ ⬝ apD to_fun_hom p) = p :=
|
: functor_eq' (ap to_fun_ob p) (!transport_compose⁻¹ ⬝ apd to_fun_hom p) = p :=
|
||||||
begin
|
begin
|
||||||
cases p, cases F₁,
|
cases p, cases F₁,
|
||||||
apply concat, rotate_left 1, apply functor_eq'_idp,
|
apply concat, rotate_left 1, apply functor_eq'_idp,
|
||||||
|
@ -203,7 +203,7 @@ namespace functor
|
||||||
|
|
||||||
-- definition ap010_functor_eq_mk' {F₁ F₂ : C ⇒ D} (p : to_fun_ob F₁ = to_fun_ob F₂)
|
-- definition ap010_functor_eq_mk' {F₁ F₂ : C ⇒ D} (p : to_fun_ob F₁ = to_fun_ob F₂)
|
||||||
-- (q : p ▹ F₁ = F₂) (c : C) :
|
-- (q : p ▹ F₁ = F₂) (c : C) :
|
||||||
-- ap to_fun_ob (functor_eq_mk (apD10 p) (λa b f, _)) = p := sorry
|
-- ap to_fun_ob (functor_eq_mk (apd10 p) (λa b f, _)) = p := sorry
|
||||||
-- begin
|
-- begin
|
||||||
-- cases F₂, revert q, apply (homotopy.rec_on p), clear p, esimp, intros (p, q),
|
-- cases F₂, revert q, apply (homotopy.rec_on p), clear p, esimp, intros (p, q),
|
||||||
-- cases p, clears (e_1, e_2),
|
-- cases p, clears (e_1, e_2),
|
||||||
|
|
|
@ -131,7 +131,7 @@ namespace iso
|
||||||
begin
|
begin
|
||||||
apply is_hprop.mk, intros [H, H'],
|
apply is_hprop.mk, intros [H, H'],
|
||||||
cases H with [g, li, ri], cases H' with [g', li', ri'],
|
cases H with [g, li, ri], cases H' with [g', li', ri'],
|
||||||
fapply (apD0111 (@is_iso.mk ob C a b f)),
|
fapply (apd0111 (@is_iso.mk ob C a b f)),
|
||||||
apply left_inverse_eq_right_inverse,
|
apply left_inverse_eq_right_inverse,
|
||||||
apply li,
|
apply li,
|
||||||
apply ri',
|
apply ri',
|
||||||
|
@ -167,7 +167,7 @@ namespace iso
|
||||||
|
|
||||||
protected definition eq_mk' {f f' : a ⟶ b} [H : is_iso f] [H' : is_iso f'] (p : f = f')
|
protected definition eq_mk' {f f' : a ⟶ b} [H : is_iso f] [H' : is_iso f'] (p : f = f')
|
||||||
: iso.mk f = iso.mk f' :=
|
: iso.mk f = iso.mk f' :=
|
||||||
apD011 iso.mk p !is_hprop.elim
|
apd011 iso.mk p !is_hprop.elim
|
||||||
|
|
||||||
protected definition eq_mk {f f' : a ≅ b} (p : to_hom f = to_hom f') : f = f' :=
|
protected definition eq_mk {f f' : a ≅ b} (p : to_hom f = to_hom f') : f = f' :=
|
||||||
by (cases f; cases f'; apply (iso.eq_mk' p))
|
by (cases f; cases f'; apply (iso.eq_mk' p))
|
||||||
|
@ -212,16 +212,16 @@ namespace iso
|
||||||
open funext
|
open funext
|
||||||
variables {X : Type} {x y : X} {F G : X → ob}
|
variables {X : Type} {x y : X} {F G : X → ob}
|
||||||
definition transport_hom_of_eq (p : F = G) (f : hom (F x) (F y))
|
definition transport_hom_of_eq (p : F = G) (f : hom (F x) (F y))
|
||||||
: p ▹ f = hom_of_eq (apD10 p y) ∘ f ∘ inv_of_eq (apD10 p x) :=
|
: p ▹ f = hom_of_eq (apd10 p y) ∘ f ∘ inv_of_eq (apd10 p x) :=
|
||||||
eq.rec_on p !id_leftright⁻¹
|
eq.rec_on p !id_leftright⁻¹
|
||||||
|
|
||||||
definition transport_hom (p : F ∼ G) (f : hom (F x) (F y))
|
definition transport_hom (p : F ∼ G) (f : hom (F x) (F y))
|
||||||
: eq_of_homotopy p ▹ f = hom_of_eq (p y) ∘ f ∘ inv_of_eq (p x) :=
|
: eq_of_homotopy p ▹ f = hom_of_eq (p y) ∘ f ∘ inv_of_eq (p x) :=
|
||||||
calc
|
calc
|
||||||
eq_of_homotopy p ▹ f =
|
eq_of_homotopy p ▹ f =
|
||||||
hom_of_eq (apD10 (eq_of_homotopy p) y) ∘ f ∘ inv_of_eq (apD10 (eq_of_homotopy p) x)
|
hom_of_eq (apd10 (eq_of_homotopy p) y) ∘ f ∘ inv_of_eq (apd10 (eq_of_homotopy p) x)
|
||||||
: transport_hom_of_eq
|
: transport_hom_of_eq
|
||||||
... = hom_of_eq (p y) ∘ f ∘ inv_of_eq (p x) : {retr apD10 p}
|
... = hom_of_eq (p y) ∘ f ∘ inv_of_eq (p x) : {retr apd10 p}
|
||||||
end
|
end
|
||||||
|
|
||||||
structure mono [class] (f : a ⟶ b) :=
|
structure mono [class] (f : a ⟶ b) :=
|
||||||
|
|
|
@ -43,7 +43,7 @@ namespace nat_trans
|
||||||
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)
|
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)
|
||||||
(p : η₁ ∼ η₂)
|
(p : η₁ ∼ η₂)
|
||||||
: nat_trans.mk η₁ nat₁ = nat_trans.mk η₂ nat₂ :=
|
: nat_trans.mk η₁ nat₁ = nat_trans.mk η₂ nat₂ :=
|
||||||
apD011 nat_trans.mk (eq_of_homotopy p) !is_hprop.elim
|
apd011 nat_trans.mk (eq_of_homotopy p) !is_hprop.elim
|
||||||
|
|
||||||
definition nat_trans_eq_mk {η₁ η₂ : F ⟹ G} : natural_map η₁ ∼ natural_map η₂ → η₁ = η₂ :=
|
definition nat_trans_eq_mk {η₁ η₂ : F ⟹ G} : natural_map η₁ ∼ natural_map η₂ → η₁ = η₂ :=
|
||||||
nat_trans.rec_on η₁ (λf₁ nat₁, nat_trans.rec_on η₂ (λf₂ nat₂ p, !nat_trans_eq_mk' p))
|
nat_trans.rec_on η₁ (λf₁ nat₁, nat_trans.rec_on η₂ (λf₂ nat₂ p, !nat_trans_eq_mk' p))
|
||||||
|
|
|
@ -29,10 +29,10 @@ namespace eq
|
||||||
transports in the theorem statement).
|
transports in the theorem statement).
|
||||||
For the fully-dependent versions (except that the conclusion doesn't contain a transport)
|
For the fully-dependent versions (except that the conclusion doesn't contain a transport)
|
||||||
we write
|
we write
|
||||||
apDi₀i₁...iₙ.
|
apdi₀i₁...iₙ.
|
||||||
|
|
||||||
For versions where only some arguments depend on some other arguments,
|
For versions where only some arguments depend on some other arguments,
|
||||||
or for versions with transport in the conclusion (like apD), we don't have a
|
or for versions with transport in the conclusion (like apd), we don't have a
|
||||||
consistent naming scheme (yet).
|
consistent naming scheme (yet).
|
||||||
|
|
||||||
We don't prove each theorem systematically, but prove only the ones which we actually need.
|
We don't prove each theorem systematically, but prove only the ones which we actually need.
|
||||||
|
@ -68,25 +68,25 @@ namespace eq
|
||||||
definition ap01000 (f : X → Πa b c, D a b c) (Hx : x = x') : f x ∼3 f x' :=
|
definition ap01000 (f : X → Πa b c, D a b c) (Hx : x = x') : f x ∼3 f x' :=
|
||||||
λa b c, eq.rec_on Hx idp
|
λa b c, eq.rec_on Hx idp
|
||||||
|
|
||||||
definition apD011 (f : Πa, B a → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd011 (f : Πa, B a → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
: f a b = f a' b' :=
|
: f a b = f a' b' :=
|
||||||
eq.rec_on Hb (eq.rec_on Ha idp)
|
eq.rec_on Hb (eq.rec_on Ha idp)
|
||||||
|
|
||||||
definition apD0111 (f : Πa b, C a b → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd0111 (f : Πa b, C a b → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
(Hc : apD011 C Ha Hb ▹ c = c')
|
(Hc : apd011 C Ha Hb ▹ c = c')
|
||||||
: f a b c = f a' b' c' :=
|
: f a b c = f a' b' c' :=
|
||||||
eq.rec_on Hc (eq.rec_on Hb (eq.rec_on Ha idp))
|
eq.rec_on Hc (eq.rec_on Hb (eq.rec_on Ha idp))
|
||||||
|
|
||||||
definition apD01111 (f : Πa b c, D a b c → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
definition apd01111 (f : Πa b c, D a b c → Z) (Ha : a = a') (Hb : (Ha ▹ b) = b')
|
||||||
(Hc : apD011 C Ha Hb ▹ c = c') (Hd : apD0111 D Ha Hb Hc ▹ d = d')
|
(Hc : apd011 C Ha Hb ▹ c = c') (Hd : apd0111 D Ha Hb Hc ▹ d = d')
|
||||||
: f a b c d = f a' b' c' d' :=
|
: f a b c d = f a' b' c' d' :=
|
||||||
eq.rec_on Hd (eq.rec_on Hc (eq.rec_on Hb (eq.rec_on Ha idp)))
|
eq.rec_on Hd (eq.rec_on Hc (eq.rec_on Hb (eq.rec_on Ha idp)))
|
||||||
|
|
||||||
definition apD100 {f g : Πa b, C a b} (p : f = g) : f ∼2 g :=
|
definition apd100 {f g : Πa b, C a b} (p : f = g) : f ∼2 g :=
|
||||||
λa b, apD10 (apD10 p a) b
|
λa b, apd10 (apd10 p a) b
|
||||||
|
|
||||||
definition apD1000 {f g : Πa b c, D a b c} (p : f = g) : f ∼3 g :=
|
definition apd1000 {f g : Πa b c, D a b c} (p : f = g) : f ∼3 g :=
|
||||||
λa b c, apD100 (apD10 p a) b c
|
λa b c, apd100 (apd10 p a) b c
|
||||||
|
|
||||||
/- some properties of these variants of ap -/
|
/- some properties of these variants of ap -/
|
||||||
|
|
||||||
|
@ -128,30 +128,30 @@ end eq
|
||||||
|
|
||||||
open eq is_equiv
|
open eq is_equiv
|
||||||
namespace funext
|
namespace funext
|
||||||
definition is_equiv_apD100 [instance] (f g : Πa b, C a b) : is_equiv (@apD100 A B C f g) :=
|
definition is_equiv_apd100 [instance] (f g : Πa b, C a b) : is_equiv (@apd100 A B C f g) :=
|
||||||
adjointify _
|
adjointify _
|
||||||
eq_of_homotopy2
|
eq_of_homotopy2
|
||||||
begin
|
begin
|
||||||
intro H, esimp [apD100, eq_of_homotopy2, function.compose],
|
intro H, esimp [apd100, eq_of_homotopy2, function.compose],
|
||||||
apply eq_of_homotopy, intro a,
|
apply eq_of_homotopy, intro a,
|
||||||
apply concat, apply (ap (λx, apD10 (x a))), apply (retr apD10),
|
apply concat, apply (ap (λx, apd10 (x a))), apply (retr apd10),
|
||||||
apply (retr apD10)
|
apply (retr apd10)
|
||||||
end
|
end
|
||||||
begin
|
begin
|
||||||
intro p, cases p, apply eq_of_homotopy2_id
|
intro p, cases p, apply eq_of_homotopy2_id
|
||||||
end
|
end
|
||||||
|
|
||||||
definition is_equiv_apD1000 [instance] (f g : Πa b c, D a b c)
|
definition is_equiv_apd1000 [instance] (f g : Πa b c, D a b c)
|
||||||
: is_equiv (@apD1000 A B C D f g) :=
|
: is_equiv (@apd1000 A B C D f g) :=
|
||||||
adjointify _
|
adjointify _
|
||||||
eq_of_homotopy3
|
eq_of_homotopy3
|
||||||
begin
|
begin
|
||||||
intro H, apply eq_of_homotopy, intro a,
|
intro H, apply eq_of_homotopy, intro a,
|
||||||
apply concat,
|
apply concat,
|
||||||
{apply (ap (λx, @apD100 _ _ (λ(b : B a)(c : C a b), _) _ _ (x a))),
|
{apply (ap (λx, @apd100 _ _ (λ(b : B a)(c : C a b), _) _ _ (x a))),
|
||||||
apply (retr apD10)},
|
apply (retr apd10)},
|
||||||
--TODO: remove implicit argument after #469 is closed
|
--TODO: remove implicit argument after #469 is closed
|
||||||
apply (@retr _ _ apD100 !is_equiv_apD100) --is explicit argument needed here?
|
apply (@retr _ _ apd100 !is_equiv_apd100) --is explicit argument needed here?
|
||||||
end
|
end
|
||||||
begin
|
begin
|
||||||
intro p, cases p, apply eq_of_homotopy3_id
|
intro p, cases p, apply eq_of_homotopy3_id
|
||||||
|
@ -160,22 +160,22 @@ end funext
|
||||||
|
|
||||||
namespace eq
|
namespace eq
|
||||||
open funext
|
open funext
|
||||||
local attribute funext.is_equiv_apD100 [instance]
|
local attribute funext.is_equiv_apd100 [instance]
|
||||||
protected definition homotopy2.rec_on {f g : Πa b, C a b} {P : (f ∼2 g) → Type}
|
protected definition homotopy2.rec_on {f g : Πa b, C a b} {P : (f ∼2 g) → Type}
|
||||||
(p : f ∼2 g) (H : Π(q : f = g), P (apD100 q)) : P p :=
|
(p : f ∼2 g) (H : Π(q : f = g), P (apd100 q)) : P p :=
|
||||||
retr apD100 p ▹ H (eq_of_homotopy2 p)
|
retr apd100 p ▹ H (eq_of_homotopy2 p)
|
||||||
|
|
||||||
protected definition homotopy3.rec_on {f g : Πa b c, D a b c} {P : (f ∼3 g) → Type}
|
protected definition homotopy3.rec_on {f g : Πa b c, D a b c} {P : (f ∼3 g) → Type}
|
||||||
(p : f ∼3 g) (H : Π(q : f = g), P (apD1000 q)) : P p :=
|
(p : f ∼3 g) (H : Π(q : f = g), P (apd1000 q)) : P p :=
|
||||||
retr apD1000 p ▹ H (eq_of_homotopy3 p)
|
retr apd1000 p ▹ H (eq_of_homotopy3 p)
|
||||||
|
|
||||||
definition apD10_ap (f : X → Πa, B a) (p : x = x')
|
definition apd10_ap (f : X → Πa, B a) (p : x = x')
|
||||||
: apD10 (ap f p) = ap010 f p :=
|
: apd10 (ap f p) = ap010 f p :=
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
definition eq_of_homotopy_ap010 (f : X → Πa, B a) (p : x = x')
|
definition eq_of_homotopy_ap010 (f : X → Πa, B a) (p : x = x')
|
||||||
: eq_of_homotopy (ap010 f p) = ap f p :=
|
: eq_of_homotopy (ap010 f p) = ap f p :=
|
||||||
inv_eq_of_eq !apD10_ap⁻¹
|
inv_eq_of_eq !apd10_ap⁻¹
|
||||||
|
|
||||||
definition ap_eq_ap_of_homotopy {f : X → Πa, B a} {p q : x = x'} (H : ap010 f p ∼ ap010 f q)
|
definition ap_eq_ap_of_homotopy {f : X → Πa, B a} {p q : x = x'} (H : ap010 f p ∼ ap010 f q)
|
||||||
: ap f p = ap f q :=
|
: ap f p = ap f q :=
|
||||||
|
|
|
@ -41,12 +41,12 @@ namespace circle
|
||||||
|
|
||||||
definition rec2_seg1 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2)
|
definition rec2_seg1 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2)
|
||||||
(Ps1 : seg1 ▹ Pb1 = Pb2) (Ps2 : seg2 ▹ Pb2 = Pb1)
|
(Ps1 : seg1 ▹ Pb1 = Pb2) (Ps2 : seg2 ▹ Pb2 = Pb1)
|
||||||
: apD (rec2 Pb1 Pb2 Ps1 Ps2) seg1 = sorry ⬝ Ps1 ⬝ sorry :=
|
: apd (rec2 Pb1 Pb2 Ps1 Ps2) seg1 = sorry ⬝ Ps1 ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
definition rec2_seg2 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2)
|
definition rec2_seg2 {P : circle → Type} (Pb1 : P base1) (Pb2 : P base2)
|
||||||
(Ps1 : seg1 ▹ Pb1 = Pb2) (Ps2 : seg2 ▹ Pb2 = Pb1)
|
(Ps1 : seg1 ▹ Pb1 = Pb2) (Ps2 : seg2 ▹ Pb2 = Pb1)
|
||||||
: apD (rec2 Pb1 Pb2 Ps1 Ps2) seg2 = sorry ⬝ Ps2 ⬝ sorry :=
|
: apd (rec2 Pb1 Pb2 Ps1 Ps2) seg2 = sorry ⬝ Ps2 ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
definition elim2 {P : Type} (Pb1 Pb2 : P) (Ps1 : Pb1 = Pb2) (Ps2 : Pb2 = Pb1) (x : circle) : P :=
|
definition elim2 {P : Type} (Pb1 Pb2 : P) (Ps1 : Pb1 = Pb2) (Ps2 : Pb2 = Pb1) (x : circle) : P :=
|
||||||
|
|
|
@ -46,7 +46,7 @@ parameters {A B : Type.{u}} (f g : A → B)
|
||||||
|
|
||||||
definition rec_cp {P : coeq → Type} (P_i : Π(x : B), P (coeq_i x))
|
definition rec_cp {P : coeq → Type} (P_i : Π(x : B), P (coeq_i x))
|
||||||
(Pcp : Π(x : A), cp x ▹ P_i (f x) = P_i (g x))
|
(Pcp : Π(x : A), cp x ▹ P_i (f x) = P_i (g x))
|
||||||
(x : A) : apD (rec P_i Pcp) (cp x) = sorry ⬝ Pcp x ⬝ sorry :=
|
(x : A) : apd (rec P_i Pcp) (cp x) = sorry ⬝ Pcp x ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
protected definition elim {P : Type} (P_i : B → P)
|
protected definition elim {P : Type} (P_i : B → P)
|
||||||
|
|
|
@ -52,7 +52,7 @@ section
|
||||||
definition rec_cglue [reducible] {P : colimit → Type}
|
definition rec_cglue [reducible] {P : colimit → Type}
|
||||||
(Pincl : Π⦃i : I⦄ (x : A i), P (ι x))
|
(Pincl : Π⦃i : I⦄ (x : A i), P (ι x))
|
||||||
(Pglue : Π(j : J) (x : A (dom j)), cglue j x ▹ Pincl (f j x) = Pincl x)
|
(Pglue : Π(j : J) (x : A (dom j)), cglue j x ▹ Pincl (f j x) = Pincl x)
|
||||||
{j : J} (x : A (dom j)) : apD (rec Pincl Pglue) (cglue j x) = Pglue j x :=
|
{j : J} (x : A (dom j)) : apd (rec Pincl Pglue) (cglue j x) = Pglue j x :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
protected definition elim {P : Type} (Pincl : Π⦃i : I⦄ (x : A i), P)
|
protected definition elim {P : Type} (Pincl : Π⦃i : I⦄ (x : A i), P)
|
||||||
|
@ -137,7 +137,7 @@ section
|
||||||
|
|
||||||
definition rec_glue {P : seq_colim → Type} (Pincl : Π⦃n : ℕ⦄ (a : A n), P (sι a))
|
definition rec_glue {P : seq_colim → Type} (Pincl : Π⦃n : ℕ⦄ (a : A n), P (sι a))
|
||||||
(Pglue : Π⦃n : ℕ⦄ (a : A n), glue a ▹ Pincl (f a) = Pincl a) {n : ℕ} (a : A n)
|
(Pglue : Π⦃n : ℕ⦄ (a : A n), glue a ▹ Pincl (f a) = Pincl a) {n : ℕ} (a : A n)
|
||||||
: apD (rec Pincl Pglue) (glue a) = sorry ⬝ Pglue a ⬝ sorry :=
|
: apd (rec Pincl Pglue) (glue a) = sorry ⬝ Pglue a ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
definition elim_glue {P : Type} (Pincl : Π⦃n : ℕ⦄ (a : A n), P)
|
definition elim_glue {P : Type} (Pincl : Π⦃n : ℕ⦄ (a : A n), P)
|
||||||
|
|
|
@ -54,7 +54,7 @@ parameters {A B : Type.{u}} (f : A → B)
|
||||||
definition rec_seg {P : cylinder → Type}
|
definition rec_seg {P : cylinder → Type}
|
||||||
(Pbase : Π(b : B), P (base b)) (Ptop : Π(a : A), P (top a))
|
(Pbase : Π(b : B), P (base b)) (Ptop : Π(a : A), P (top a))
|
||||||
(Pseg : Π(a : A), seg a ▹ Pbase (f a) = Ptop a)
|
(Pseg : Π(a : A), seg a ▹ Pbase (f a) = Ptop a)
|
||||||
(a : A) : apD (rec Pbase Ptop Pseg) (seg a) = sorry ⬝ Pseg a ⬝ sorry :=
|
(a : A) : apd (rec Pbase Ptop Pseg) (seg a) = sorry ⬝ Pseg a ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
protected definition elim {P : Type} (Pbase : B → P) (Ptop : A → P)
|
protected definition elim {P : Type} (Pbase : B → P) (Ptop : A → P)
|
||||||
|
|
|
@ -63,7 +63,7 @@ parameters {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
|
||||||
|
|
||||||
definition rec_glue {P : pushout → Type} (Pinl : Π(x : BL), P (inl x))
|
definition rec_glue {P : pushout → Type} (Pinl : Π(x : BL), P (inl x))
|
||||||
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x))
|
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x))
|
||||||
(x : TL) : apD (rec Pinl Pinr Pglue) (glue x) = Pglue x :=
|
(x : TL) : apd (rec Pinl Pinr Pglue) (glue x) = Pglue x :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
protected definition elim {P : Type} (Pinl : BL → P) (Pinr : TR → P)
|
protected definition elim {P : Type} (Pinl : BL → P) (Pinr : TR → P)
|
||||||
|
|
|
@ -46,7 +46,7 @@ parameters {A : Type} (R : A → A → hprop)
|
||||||
|
|
||||||
definition rec_eq_of_rel {P : quotient → Type} [Pt : Πaa, is_hset (P aa)]
|
definition rec_eq_of_rel {P : quotient → Type} [Pt : Πaa, is_hset (P aa)]
|
||||||
(Pc : Π(a : A), P (class_of a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), eq_of_rel H ▹ Pc a = Pc a')
|
(Pc : Π(a : A), P (class_of a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), eq_of_rel H ▹ Pc a = Pc a')
|
||||||
{a a' : A} (H : R a a') : apD (rec Pc Pp) (eq_of_rel H) = sorry ⬝ Pp H ⬝ sorry :=
|
{a a' : A} (H : R a a') : apd (rec Pc Pp) (eq_of_rel H) = sorry ⬝ Pp H ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
protected definition elim {P : Type} [Pt : is_hset P] (Pc : A → P)
|
protected definition elim {P : Type} [Pt : is_hset P] (Pc : A → P)
|
||||||
|
|
|
@ -41,7 +41,7 @@ namespace suspension
|
||||||
|
|
||||||
definition rec_merid {P : suspension A → Type} (PN : P !north) (PS : P !south)
|
definition rec_merid {P : suspension A → Type} (PN : P !north) (PS : P !south)
|
||||||
(Pm : Π(a : A), merid a ▹ PN = PS) (a : A)
|
(Pm : Π(a : A), merid a ▹ PN = PS) (a : A)
|
||||||
: apD (rec PN PS Pm) (merid a) = sorry ⬝ Pm a ⬝ sorry :=
|
: apd (rec PN PS Pm) (merid a) = sorry ⬝ Pm a ⬝ sorry :=
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
protected definition elim {P : Type} (PN : P) (PS : P) (Pm : A → PN = PS)
|
protected definition elim {P : Type} (PN : P) (PS : P) (Pm : A → PN = PS)
|
||||||
|
|
|
@ -138,25 +138,25 @@ definition funext_of_ua : funext :=
|
||||||
variables {A : Type} {P : A → Type} {f g : Π x, P x}
|
variables {A : Type} {P : A → Type} {f g : Π x, P x}
|
||||||
|
|
||||||
namespace funext
|
namespace funext
|
||||||
definition is_equiv_apD [instance] (f g : Π x, P x) : is_equiv (@apD10 A P f g) :=
|
definition is_equiv_apd [instance] (f g : Π x, P x) : is_equiv (@apd10 A P f g) :=
|
||||||
funext_of_ua f g
|
funext_of_ua f g
|
||||||
end funext
|
end funext
|
||||||
|
|
||||||
open funext
|
open funext
|
||||||
|
|
||||||
definition eq_equiv_homotopy : (f = g) ≃ (f ∼ g) :=
|
definition eq_equiv_homotopy : (f = g) ≃ (f ∼ g) :=
|
||||||
equiv.mk apD10 _
|
equiv.mk apd10 _
|
||||||
|
|
||||||
definition eq_of_homotopy [reducible] : f ∼ g → f = g :=
|
definition eq_of_homotopy [reducible] : f ∼ g → f = g :=
|
||||||
(@apD10 A P f g)⁻¹
|
(@apd10 A P f g)⁻¹
|
||||||
|
|
||||||
--TODO: rename to eq_of_homotopy_idp
|
--TODO: rename to eq_of_homotopy_idp
|
||||||
definition eq_of_homotopy_id (f : Π x, P x) : eq_of_homotopy (λx : A, idpath (f x)) = idpath f :=
|
definition eq_of_homotopy_id (f : Π x, P x) : eq_of_homotopy (λx : A, idpath (f x)) = idpath f :=
|
||||||
is_equiv.sect apD10 idp
|
is_equiv.sect apd10 idp
|
||||||
|
|
||||||
definition naive_funext_of_ua : naive_funext :=
|
definition naive_funext_of_ua : naive_funext :=
|
||||||
λ A P f g h, eq_of_homotopy h
|
λ A P f g h, eq_of_homotopy h
|
||||||
|
|
||||||
protected definition homotopy.rec_on {Q : (f ∼ g) → Type} (p : f ∼ g)
|
protected definition homotopy.rec_on {Q : (f ∼ g) → Type} (p : f ∼ g)
|
||||||
(H : Π(q : f = g), Q (apD10 q)) : Q p :=
|
(H : Π(q : f = g), Q (apd10 q)) : Q p :=
|
||||||
retr apD10 p ▹ H (eq_of_homotopy p)
|
retr apd10 p ▹ H (eq_of_homotopy p)
|
||||||
|
|
|
@ -14,7 +14,7 @@ import ..path ..trunc ..equiv
|
||||||
open eq is_trunc sigma function
|
open eq is_trunc sigma function
|
||||||
|
|
||||||
/- In init.axioms.funext, we defined function extensionality to be the assertion
|
/- In init.axioms.funext, we defined function extensionality to be the assertion
|
||||||
that the map apD10 is an equivalence. We now prove that this follows
|
that the map apd10 is an equivalence. We now prove that this follows
|
||||||
from a couple of weaker-looking forms of function extensionality. We do
|
from a couple of weaker-looking forms of function extensionality. We do
|
||||||
require eta conversion, which Coq 8.4+ has judgmentally.
|
require eta conversion, which Coq 8.4+ has judgmentally.
|
||||||
|
|
||||||
|
@ -22,7 +22,7 @@ open eq is_trunc sigma function
|
||||||
by Peter Lumsdaine and Michael Shulman. -/
|
by Peter Lumsdaine and Michael Shulman. -/
|
||||||
|
|
||||||
definition funext.{l k} :=
|
definition funext.{l k} :=
|
||||||
Π ⦃A : Type.{l}⦄ {P : A → Type.{k}} (f g : Π x, P x), is_equiv (@apD10 A P f g)
|
Π ⦃A : Type.{l}⦄ {P : A → Type.{k}} (f g : Π x, P x), is_equiv (@apd10 A P f g)
|
||||||
|
|
||||||
-- Naive funext is the simple assertion that pointwise equal functions are equal.
|
-- Naive funext is the simple assertion that pointwise equal functions are equal.
|
||||||
-- TODO think about universe levels
|
-- TODO think about universe levels
|
||||||
|
@ -98,13 +98,13 @@ theorem funext_of_weak_funext (wf : weak_funext.{l k}) : funext.{l k} :=
|
||||||
let sim2path := homotopy_ind f eq_to_f idp in
|
let sim2path := homotopy_ind f eq_to_f idp in
|
||||||
assert t1 : sim2path f (homotopy.refl f) = idp,
|
assert t1 : sim2path f (homotopy.refl f) = idp,
|
||||||
proof homotopy_ind_comp f eq_to_f idp qed,
|
proof homotopy_ind_comp f eq_to_f idp qed,
|
||||||
assert t2 : apD10 (sim2path f (homotopy.refl f)) = (homotopy.refl f),
|
assert t2 : apd10 (sim2path f (homotopy.refl f)) = (homotopy.refl f),
|
||||||
proof ap apD10 t1 qed,
|
proof ap apd10 t1 qed,
|
||||||
have sect : apD10 ∘ (sim2path g) ∼ id,
|
have sect : apd10 ∘ (sim2path g) ∼ id,
|
||||||
proof (homotopy_ind f (λ g' x, apD10 (sim2path g' x) = x) t2) g qed,
|
proof (homotopy_ind f (λ g' x, apd10 (sim2path g' x) = x) t2) g qed,
|
||||||
have retr : (sim2path g) ∘ apD10 ∼ id,
|
have retr : (sim2path g) ∘ apd10 ∼ id,
|
||||||
from (λ h, eq.rec_on h (homotopy_ind_comp f _ idp)),
|
from (λ h, eq.rec_on h (homotopy_ind_comp f _ idp)),
|
||||||
is_equiv.adjointify apD10 (sim2path g) sect retr
|
is_equiv.adjointify apd10 (sim2path g) sect retr
|
||||||
|
|
||||||
definition funext_from_naive_funext : naive_funext -> funext :=
|
definition funext_from_naive_funext : naive_funext -> funext :=
|
||||||
compose funext_of_weak_funext weak_funext_of_naive_funext
|
compose funext_of_weak_funext weak_funext_of_naive_funext
|
||||||
|
|
|
@ -78,7 +78,7 @@ namespace is_equiv
|
||||||
... = ap f (ap invf ff'a) ⬝ retrf'a : by rewrite ap_compose,
|
... = ap f (ap invf ff'a) ⬝ retrf'a : by rewrite ap_compose,
|
||||||
have eq2 : _ = _,
|
have eq2 : _ = _,
|
||||||
from calc retrf'a
|
from calc retrf'a
|
||||||
= (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ ff'a) : eq_inv_con_of_con_eq _ _ _ eq1⁻¹
|
= (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ ff'a) : eq_inv_con_of_con_eq eq1⁻¹
|
||||||
... = (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ Hty a) : ap_inv invf ff'a
|
... = (ap f (ap invf ff'a))⁻¹ ⬝ (ap f secta ⬝ Hty a) : ap_inv invf ff'a
|
||||||
... = (ap f (ap invf ff'a))⁻¹ ⬝ (Hty (invf (f a)) ⬝ ap f' secta) : by rewrite ap_con_eq_con_ap
|
... = (ap f (ap invf ff'a))⁻¹ ⬝ (Hty (invf (f a)) ⬝ ap f' secta) : by rewrite ap_con_eq_con_ap
|
||||||
... = ((ap f (ap invf ff'a))⁻¹ ⬝ Hty (invf (f a))) ⬝ ap f' secta : by rewrite con.assoc
|
... = ((ap f (ap invf ff'a))⁻¹ ⬝ Hty (invf (f a))) ⬝ ap f' secta : by rewrite con.assoc
|
||||||
|
@ -88,7 +88,7 @@ namespace is_equiv
|
||||||
... = Hty (invf (f' a)) ⬝ ((ap f' (ap invf ff'a))⁻¹ ⬝ ap f' secta) : by rewrite con.assoc,
|
... = Hty (invf (f' a)) ⬝ ((ap f' (ap invf ff'a))⁻¹ ⬝ ap f' secta) : by rewrite con.assoc,
|
||||||
have eq3 : _ = _,
|
have eq3 : _ = _,
|
||||||
from calc (Hty (invf (f' a)))⁻¹ ⬝ retrf'a
|
from calc (Hty (invf (f' a)))⁻¹ ⬝ retrf'a
|
||||||
= (ap f' (ap invf ff'a))⁻¹ ⬝ ap f' secta : inv_con_eq_of_eq_con _ _ _ eq2
|
= (ap f' (ap invf ff'a))⁻¹ ⬝ ap f' secta : inv_con_eq_of_eq_con eq2
|
||||||
... = (ap f' (ap invf ff'a)⁻¹) ⬝ ap f' secta : by rewrite ap_inv
|
... = (ap f' (ap invf ff'a)⁻¹) ⬝ ap f' secta : by rewrite ap_inv
|
||||||
... = ap f' ((ap invf ff'a)⁻¹ ⬝ secta) : by rewrite ap_con,
|
... = ap f' ((ap invf ff'a)⁻¹ ⬝ secta) : by rewrite ap_con,
|
||||||
eq3) in
|
eq3) in
|
||||||
|
@ -119,7 +119,7 @@ namespace is_equiv
|
||||||
from !con_idp ⬝ eq1,
|
from !con_idp ⬝ eq1,
|
||||||
have eq3 : idp = _,
|
have eq3 : idp = _,
|
||||||
from calc idp
|
from calc idp
|
||||||
= (ap f (sec a))⁻¹ ⬝ ((ret fgfa)⁻¹ ⬝ (fgretrfa ⬝ ap f (sec a))) : eq_inv_con_of_con_eq _ _ _ eq2
|
= (ap f (sec a))⁻¹ ⬝ ((ret fgfa)⁻¹ ⬝ (fgretrfa ⬝ ap f (sec a))) : eq_inv_con_of_con_eq eq2
|
||||||
... = ((ap f (sec a))⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : by rewrite con.assoc'
|
... = ((ap f (sec a))⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : by rewrite con.assoc'
|
||||||
... = (ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : by rewrite ap_inv
|
... = (ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ (fgretrfa ⬝ ap f (sec a)) : by rewrite ap_inv
|
||||||
... = ((ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ fgretrfa) ⬝ ap f (sec a) : by rewrite con.assoc'
|
... = ((ap f (sec a)⁻¹ ⬝ (ret fgfa)⁻¹) ⬝ fgretrfa) ⬝ ap f (sec a) : by rewrite con.assoc'
|
||||||
|
@ -130,7 +130,7 @@ namespace is_equiv
|
||||||
... = retrfa⁻¹ ⬝ (ap f (ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ ap f (sec a)) : by rewrite con.assoc'
|
... = retrfa⁻¹ ⬝ (ap f (ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ ap f (sec a)) : by rewrite con.assoc'
|
||||||
... = retrfa⁻¹ ⬝ ap f ((ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ sec a) : by rewrite -ap_con,
|
... = retrfa⁻¹ ⬝ ap f ((ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ sec a) : by rewrite -ap_con,
|
||||||
have eq4 : ret (f a) = ap f ((ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ sec a),
|
have eq4 : ret (f a) = ap f ((ap g (ap f (sec a)⁻¹) ⬝ ap g (ret (f a))) ⬝ sec a),
|
||||||
from eq_of_idp_eq_inv_con _ _ eq3,
|
from eq_of_idp_eq_inv_con eq3,
|
||||||
eq4)
|
eq4)
|
||||||
|
|
||||||
definition adjointify : is_equiv f :=
|
definition adjointify : is_equiv f :=
|
||||||
|
@ -188,7 +188,7 @@ namespace is_equiv
|
||||||
= transport P (retr f (f x)) (df (f⁻¹ (f x))) : by esimp
|
= transport P (retr f (f x)) (df (f⁻¹ (f x))) : by esimp
|
||||||
... = transport P (eq.ap f (sect f x)) (df (f⁻¹ (f x))) : by rewrite adj
|
... = transport P (eq.ap f (sect f x)) (df (f⁻¹ (f x))) : by rewrite adj
|
||||||
... = transport (P ∘ f) (sect f x) (df (f⁻¹ (f x))) : by rewrite -transport_compose
|
... = transport (P ∘ f) (sect f x) (df (f⁻¹ (f x))) : by rewrite -transport_compose
|
||||||
... = df x : by rewrite (apD df (sect f x))
|
... = df x : by rewrite (apd df (sect f x))
|
||||||
|
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -82,5 +82,5 @@ namespace type_quotient
|
||||||
|
|
||||||
constant rec_eq_of_rel {A : Type} {R : A → A → Type} {P : type_quotient R → Type}
|
constant rec_eq_of_rel {A : Type} {R : A → A → Type} {P : type_quotient R → Type}
|
||||||
(Pc : Π(a : A), P (class_of R a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), eq_of_rel H ▹ Pc a = Pc a')
|
(Pc : Π(a : A), P (class_of R a)) (Pp : Π⦃a a' : A⦄ (H : R a a'), eq_of_rel H ▹ Pc a = Pc a')
|
||||||
{a a' : A} (H : R a a') : apD (type_quotient.rec Pc Pp) (eq_of_rel H) = Pp H
|
{a a' : A} (H : R a a') : apd (type_quotient.rec Pc Pp) (eq_of_rel H) = Pp H
|
||||||
end type_quotient
|
end type_quotient
|
||||||
|
|
|
@ -112,69 +112,69 @@ namespace eq
|
||||||
|
|
||||||
/- Theorems for moving things around in equations -/
|
/- Theorems for moving things around in equations -/
|
||||||
|
|
||||||
definition con_eq_of_eq_inv_con (p : x = z) (q : y = z) (r : y = x) :
|
definition con_eq_of_eq_inv_con {p : x = z} {q : y = z} {r : y = x} :
|
||||||
p = r⁻¹ ⬝ q → r ⬝ p = q :=
|
p = r⁻¹ ⬝ q → r ⬝ p = q :=
|
||||||
eq.rec_on r (take p h, idp_con _ ⬝ h ⬝ idp_con _) p
|
eq.rec_on r (take p h, !idp_con ⬝ h ⬝ !idp_con) p
|
||||||
|
|
||||||
definition con_eq_of_eq_con_inv (p : x = z) (q : y = z) (r : y = x) :
|
definition con_eq_of_eq_con_inv {p : x = z} {q : y = z} {r : y = x} :
|
||||||
r = q ⬝ p⁻¹ → r ⬝ p = q :=
|
r = q ⬝ p⁻¹ → r ⬝ p = q :=
|
||||||
eq.rec_on p (take q h, (con_idp _ ⬝ h ⬝ con_idp _)) q
|
eq.rec_on p (take q h, (!con_idp ⬝ h ⬝ !con_idp)) q
|
||||||
|
|
||||||
definition inv_con_eq_of_eq_con (p : x = z) (q : y = z) (r : x = y) :
|
definition inv_con_eq_of_eq_con {p : x = z} {q : y = z} {r : x = y} :
|
||||||
p = r ⬝ q → r⁻¹ ⬝ p = q :=
|
p = r ⬝ q → r⁻¹ ⬝ p = q :=
|
||||||
eq.rec_on r (take q h, idp_con _ ⬝ h ⬝ idp_con _) q
|
eq.rec_on r (take q h, !idp_con ⬝ h ⬝ !idp_con) q
|
||||||
|
|
||||||
definition con_inv_eq_of_eq_con (p : z = x) (q : y = z) (r : y = x) :
|
definition con_inv_eq_of_eq_con {p : z = x} {q : y = z} {r : y = x} :
|
||||||
r = q ⬝ p → r ⬝ p⁻¹ = q :=
|
r = q ⬝ p → r ⬝ p⁻¹ = q :=
|
||||||
eq.rec_on p (take r h, con_idp _ ⬝ h ⬝ con_idp _) r
|
eq.rec_on p (take r h, !con_idp ⬝ h ⬝ !con_idp) r
|
||||||
|
|
||||||
definition eq_con_of_inv_con_eq (p : x = z) (q : y = z) (r : y = x) :
|
definition eq_con_of_inv_con_eq {p : x = z} {q : y = z} {r : y = x} :
|
||||||
r⁻¹ ⬝ q = p → q = r ⬝ p :=
|
r⁻¹ ⬝ q = p → q = r ⬝ p :=
|
||||||
eq.rec_on r (take p h, (idp_con _)⁻¹ ⬝ h ⬝ (idp_con _)⁻¹) p
|
eq.rec_on r (take p h, !idp_con⁻¹ ⬝ h ⬝ !idp_con⁻¹) p
|
||||||
|
|
||||||
definition eq_con_of_con_inv_eq (p : x = z) (q : y = z) (r : y = x) :
|
definition eq_con_of_con_inv_eq {p : x = z} {q : y = z} {r : y = x} :
|
||||||
q ⬝ p⁻¹ = r → q = r ⬝ p :=
|
q ⬝ p⁻¹ = r → q = r ⬝ p :=
|
||||||
eq.rec_on p (take q h, (con_idp _)⁻¹ ⬝ h ⬝ (con_idp _)⁻¹) q
|
eq.rec_on p (take q h, !con_idp⁻¹ ⬝ h ⬝ !con_idp⁻¹) q
|
||||||
|
|
||||||
definition eq_inv_con_of_con_eq (p : x = z) (q : y = z) (r : x = y) :
|
definition eq_inv_con_of_con_eq {p : x = z} {q : y = z} {r : x = y} :
|
||||||
r ⬝ q = p → q = r⁻¹ ⬝ p :=
|
r ⬝ q = p → q = r⁻¹ ⬝ p :=
|
||||||
eq.rec_on r (take q h, (idp_con _)⁻¹ ⬝ h ⬝ (idp_con _)⁻¹) q
|
eq.rec_on r (take q h, !idp_con⁻¹ ⬝ h ⬝ !idp_con⁻¹) q
|
||||||
|
|
||||||
definition eq_con_inv_of_con_eq (p : z = x) (q : y = z) (r : y = x) :
|
definition eq_con_inv_of_con_eq {p : z = x} {q : y = z} {r : y = x} :
|
||||||
q ⬝ p = r → q = r ⬝ p⁻¹ :=
|
q ⬝ p = r → q = r ⬝ p⁻¹ :=
|
||||||
eq.rec_on p (take r h, (con_idp _)⁻¹ ⬝ h ⬝ (con_idp _)⁻¹) r
|
eq.rec_on p (take r h, !con_idp⁻¹ ⬝ h ⬝ !con_idp⁻¹) r
|
||||||
|
|
||||||
definition eq_of_con_inv_eq_idp (p q : x = y) :
|
definition eq_of_con_inv_eq_idp {p q : x = y} :
|
||||||
p ⬝ q⁻¹ = idp → p = q :=
|
p ⬝ q⁻¹ = idp → p = q :=
|
||||||
eq.rec_on q (take p h, (con_idp _)⁻¹ ⬝ h) p
|
eq.rec_on q (take p h, !con_idp⁻¹ ⬝ h) p
|
||||||
|
|
||||||
definition eq_of_inv_con_eq_idp (p q : x = y) :
|
definition eq_of_inv_con_eq_idp {p q : x = y} :
|
||||||
q⁻¹ ⬝ p = idp → p = q :=
|
q⁻¹ ⬝ p = idp → p = q :=
|
||||||
eq.rec_on q (take p h, (idp_con _)⁻¹ ⬝ h) p
|
eq.rec_on q (take p h, !idp_con⁻¹ ⬝ h) p
|
||||||
|
|
||||||
definition eq_inv_of_con_eq_idp' (p : x = y) (q : y = x) :
|
definition eq_inv_of_con_eq_idp' {p : x = y} {q : y = x} :
|
||||||
p ⬝ q = idp → p = q⁻¹ :=
|
p ⬝ q = idp → p = q⁻¹ :=
|
||||||
eq.rec_on q (take p h, (con_idp _)⁻¹ ⬝ h) p
|
eq.rec_on q (take p h, !con_idp⁻¹ ⬝ h) p
|
||||||
|
|
||||||
definition eq_inv_of_con_eq_idp (p : x = y) (q : y = x) :
|
definition eq_inv_of_con_eq_idp {p : x = y} {q : y = x} :
|
||||||
q ⬝ p = idp → p = q⁻¹ :=
|
q ⬝ p = idp → p = q⁻¹ :=
|
||||||
eq.rec_on q (take p h, (idp_con _)⁻¹ ⬝ h) p
|
eq.rec_on q (take p h, !idp_con⁻¹ ⬝ h) p
|
||||||
|
|
||||||
definition eq_of_idp_eq_inv_con (p q : x = y) :
|
definition eq_of_idp_eq_inv_con {p q : x = y} :
|
||||||
idp = p⁻¹ ⬝ q → p = q :=
|
idp = p⁻¹ ⬝ q → p = q :=
|
||||||
eq.rec_on p (take q h, h ⬝ (idp_con _)) q
|
eq.rec_on p (take q h, h ⬝ !idp_con) q
|
||||||
|
|
||||||
definition eq_of_idp_eq_con_inv (p q : x = y) :
|
definition eq_of_idp_eq_con_inv {p q : x = y} :
|
||||||
idp = q ⬝ p⁻¹ → p = q :=
|
idp = q ⬝ p⁻¹ → p = q :=
|
||||||
eq.rec_on p (take q h, h ⬝ (con_idp _)) q
|
eq.rec_on p (take q h, h ⬝ !con_idp) q
|
||||||
|
|
||||||
definition inv_eq_of_idp_eq_con (p : x = y) (q : y = x) :
|
definition inv_eq_of_idp_eq_con {p : x = y} {q : y = x} :
|
||||||
idp = q ⬝ p → p⁻¹ = q :=
|
idp = q ⬝ p → p⁻¹ = q :=
|
||||||
eq.rec_on p (take q h, h ⬝ (con_idp _)) q
|
eq.rec_on p (take q h, h ⬝ !con_idp) q
|
||||||
|
|
||||||
definition inv_eq_of_idp_eq_con' (p : x = y) (q : y = x) :
|
definition inv_eq_of_idp_eq_con' {p : x = y} {q : y = x} :
|
||||||
idp = p ⬝ q → p⁻¹ = q :=
|
idp = p ⬝ q → p⁻¹ = q :=
|
||||||
eq.rec_on p (take q h, h ⬝ (idp_con _)) q
|
eq.rec_on p (take q h, h ⬝ !idp_con) q
|
||||||
|
|
||||||
/- Transport -/
|
/- Transport -/
|
||||||
|
|
||||||
|
@ -184,13 +184,16 @@ namespace eq
|
||||||
-- This idiom makes the operation right associative.
|
-- This idiom makes the operation right associative.
|
||||||
notation p `▹`:65 x:64 := transport _ p x
|
notation p `▹`:65 x:64 := transport _ p x
|
||||||
|
|
||||||
|
definition cast [reducible] {A B : Type} (p : A = B) (a : A) : B :=
|
||||||
|
p ▹ a
|
||||||
|
|
||||||
definition tr_inv [reducible] (P : A → Type) {x y : A} (p : x = y) (u : P y) : P x :=
|
definition tr_inv [reducible] (P : A → Type) {x y : A} (p : x = y) (u : P y) : P x :=
|
||||||
p⁻¹ ▹ u
|
p⁻¹ ▹ u
|
||||||
|
|
||||||
definition ap ⦃A B : Type⦄ (f : A → B) {x y:A} (p : x = y) : f x = f y :=
|
definition ap ⦃A B : Type⦄ (f : A → B) {x y:A} (p : x = y) : f x = f y :=
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
definition ap01 [reducible] := ap
|
abbreviation ap01 [parsing-only] := ap
|
||||||
|
|
||||||
definition homotopy [reducible] (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
|
||||||
|
@ -212,19 +215,19 @@ namespace eq
|
||||||
calc_trans trans
|
calc_trans trans
|
||||||
end homotopy
|
end homotopy
|
||||||
|
|
||||||
definition apD10 {f g : Πx, P x} (H : f = g) : f ∼ g :=
|
definition apd10 {f g : Πx, P x} (H : f = g) : f ∼ g :=
|
||||||
λx, eq.rec_on H idp
|
λx, eq.rec_on H idp
|
||||||
|
|
||||||
--the next theorem is useful if you want to write "apply (apD10' a)"
|
--the next theorem is useful if you want to write "apply (apd10' a)"
|
||||||
definition apD10' {f g : Πx, P x} (a : A) (H : f = g) : f a = g a :=
|
definition apd10' {f g : Πx, P x} (a : A) (H : f = g) : f a = g a :=
|
||||||
eq.rec_on H idp
|
eq.rec_on H idp
|
||||||
|
|
||||||
definition ap10 {f g : A → B} (H : f = g) : f ∼ g := apD10 H
|
definition ap10 [reducible] {f g : A → B} (H : f = g) : f ∼ g := apd10 H
|
||||||
|
|
||||||
definition ap11 {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 :=
|
||||||
eq.rec_on H (eq.rec_on p idp)
|
eq.rec_on H (eq.rec_on p idp)
|
||||||
|
|
||||||
definition apD (f : Πa:A, P 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 :=
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
/- calc enviroment -/
|
/- calc enviroment -/
|
||||||
|
@ -236,19 +239,19 @@ namespace eq
|
||||||
|
|
||||||
/- More theorems for moving things around in equations -/
|
/- More theorems for moving things around in equations -/
|
||||||
|
|
||||||
definition tr_eq_of_eq_inv_tr (P : A → Type) {x y : A} (p : x = y) (u : P x) (v : P y) :
|
definition tr_eq_of_eq_inv_tr (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 :=
|
||||||
eq.rec_on p (take v, id) v
|
eq.rec_on p (take v, id) v
|
||||||
|
|
||||||
definition inv_tr_eq_of_eq_tr (P : A → Type) {x y : A} (p : y = x) (u : P x) (v : P y) :
|
definition inv_tr_eq_of_eq_tr (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 :=
|
||||||
eq.rec_on p (take u, id) u
|
eq.rec_on p (take u, id) u
|
||||||
|
|
||||||
definition eq_inv_tr_of_tr_eq (P : A → Type) {x y : A} (p : x = y) (u : P x) (v : P y) :
|
definition eq_inv_tr_of_tr_eq (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 :=
|
||||||
eq.rec_on p (take v, id) v
|
eq.rec_on p (take v, id) v
|
||||||
|
|
||||||
definition eq_tr_of_inv_tr_eq (P : A → Type) {x y : A} (p : y = x) (u : P x) (v : P y) :
|
definition eq_tr_of_inv_tr_eq (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 :=
|
||||||
eq.rec_on p (take u, id) u
|
eq.rec_on p (take u, id) u
|
||||||
|
|
||||||
|
@ -258,9 +261,8 @@ namespace eq
|
||||||
-- functorial.
|
-- functorial.
|
||||||
|
|
||||||
-- Functions take identity paths to identity paths
|
-- Functions take identity paths to identity paths
|
||||||
definition ap_idp (x : A) (f : A → B) : (ap f idp) = idp :> (f x = f x) := idp
|
definition ap_idp (x : A) (f : A → B) : ap f idp = idp :> (f x = f x) := idp
|
||||||
|
definition apd_idp (x : A) (f : Πx, P x) : apd f idp = idp :> (f x = f x) := idp
|
||||||
definition apD_idp (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_con (f : A → B) {x y z : A} (p : x = y) (q : y = z) :
|
definition ap_con (f : A → B) {x y z : A} (p : x = y) (q : y = z) :
|
||||||
|
@ -297,23 +299,22 @@ namespace eq
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
-- The action of constant maps.
|
-- The action of constant maps.
|
||||||
definition ap_constant (p : x = y) (z : B) :
|
definition ap_constant (p : x = y) (z : B) : ap (λu, z) p = idp :=
|
||||||
ap (λu, z) p = idp :=
|
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
-- Naturality of [ap].
|
-- Naturality of [ap].
|
||||||
definition ap_con_eq_con_ap {f g : A → B} (p : Π x, f x = g x) {x y : A} (q : x = y) :
|
definition ap_con_eq_con_ap {f g : A → B} (p : f ∼ g) {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 :=
|
||||||
eq.rec_on q (idp_con _ ⬝ (con_idp _)⁻¹)
|
eq.rec_on q (!idp_con ⬝ !con_idp⁻¹)
|
||||||
|
|
||||||
-- Naturality of [ap] at identity.
|
-- Naturality of [ap] at identity.
|
||||||
definition ap_con_eq_con {f : A → A} (p : Πx, f x = x) {x y : A} (q : x = y) :
|
definition ap_con_eq_con {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 :=
|
||||||
eq.rec_on q (idp_con _ ⬝ (con_idp _)⁻¹)
|
eq.rec_on q (!idp_con ⬝ !con_idp⁻¹)
|
||||||
|
|
||||||
definition con_ap_eq_con {f : A → A} (p : Πx, x = f x) {x y : A} (q : x = y) :
|
definition con_ap_eq_con {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 :=
|
||||||
eq.rec_on q (con_idp _ ⬝ (idp_con _)⁻¹)
|
eq.rec_on q (!con_idp ⬝ !idp_con⁻¹)
|
||||||
|
|
||||||
-- Naturality with other paths hanging around.
|
-- Naturality with other paths hanging around.
|
||||||
|
|
||||||
|
@ -330,14 +331,14 @@ namespace eq
|
||||||
-- TODO: try this using the simplifier, and compare proofs
|
-- TODO: try this using the simplifier, and compare proofs
|
||||||
definition ap_con_con_eq_con_ap_con {f g : A → B} (p : f ∼ g) {x y : A} (q : x = y)
|
definition ap_con_con_eq_con_ap_con {f g : A → B} (p : f ∼ g) {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) :=
|
||||||
eq.rec_on s (eq.rec_on q
|
eq.rec_on s (eq.rec_on q
|
||||||
(calc
|
(calc
|
||||||
(ap f idp) ⬝ (p x ⬝ idp) = idp ⬝ p x : idp
|
(ap f idp) ⬝ (p x ⬝ idp) = idp ⬝ p x : idp
|
||||||
... = p x : idp_con _
|
... = p x : !idp_con
|
||||||
... = (p x) ⬝ (ap g idp ⬝ idp) : idp))
|
... = (p x) ⬝ (ap g idp ⬝ idp) : idp))
|
||||||
-- This also works:
|
-- This also works:
|
||||||
-- eq.rec_on s (eq.rec_on q (idp_con _ ▹ idp))
|
-- eq.rec_on s (eq.rec_on q (!idp_con ▹ idp))
|
||||||
|
|
||||||
definition con_ap_con_con_eq_con_con_con {f : A → A} (p : f ∼ id) {x y : A} (q : x = y)
|
definition con_ap_con_con_eq_con_con_con {f : A → A} (p : f ∼ id) {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) :
|
||||||
|
@ -356,8 +357,8 @@ namespace eq
|
||||||
|
|
||||||
definition ap_con_con_eq_con_con {f : A → A} (p : f ∼ id) {x y : A} (q : x = y)
|
definition ap_con_con_eq_con_con {f : A → A} (p : f ∼ id) {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) :=
|
||||||
eq.rec_on s (eq.rec_on q (idp_con _ ▹ idp))
|
eq.rec_on s (eq.rec_on q (!idp_con ▹ idp))
|
||||||
|
|
||||||
definition con_con_ap_eq_con_con {g : A → A} (p : id ∼ g) {x y : A} (q : x = y)
|
definition con_con_ap_eq_con_con {g : A → A} (p : id ∼ g) {x y : A} (q : x = y)
|
||||||
{w : A} (r : w = x) :
|
{w : A} (r : w = x) :
|
||||||
|
@ -373,27 +374,27 @@ namespace eq
|
||||||
apply (idp_con (p x) ▹ idp)
|
apply (idp_con (p x) ▹ idp)
|
||||||
end
|
end
|
||||||
|
|
||||||
/- Action of [apD10] and [ap10] on paths -/
|
/- Action of [apd10] and [ap10] on paths -/
|
||||||
|
|
||||||
-- Application of paths between functions preserves the groupoid structure
|
-- Application of paths between functions preserves the groupoid structure
|
||||||
|
|
||||||
definition apD10_idp (f : Πx, P x) (x : A) : apD10 (refl f) x = idp := idp
|
definition apd10_idp (f : Πx, P x) (x : A) : apd10 (refl f) x = idp := idp
|
||||||
|
|
||||||
definition apD10_con {f f' f'' : Πx, P x} (h : f = f') (h' : f' = f'') (x : A) :
|
definition apd10_con {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 :=
|
||||||
eq.rec_on h (take h', eq.rec_on h' idp) h'
|
eq.rec_on h (take h', eq.rec_on h' idp) h'
|
||||||
|
|
||||||
definition apD10_inv {f g : Πx : A, P x} (h : f = g) (x : A) :
|
definition apd10_inv {f g : Πx : A, P x} (h : f = g) (x : A) :
|
||||||
apD10 h⁻¹ x = (apD10 h x)⁻¹ :=
|
apd10 h⁻¹ x = (apd10 h x)⁻¹ :=
|
||||||
eq.rec_on h idp
|
eq.rec_on h idp
|
||||||
|
|
||||||
definition ap10_idp {f : A → B} (x : A) : ap10 (refl f) x = idp := idp
|
definition ap10_idp {f : A → B} (x : A) : ap10 (refl f) x = idp := idp
|
||||||
|
|
||||||
definition ap10_con {f f' f'' : A → B} (h : f = f') (h' : f' = f'') (x : A) :
|
definition ap10_con {f f' f'' : A → B} (h : f = f') (h' : f' = f'') (x : A) :
|
||||||
ap10 (h ⬝ h') x = ap10 h x ⬝ ap10 h' x := apD10_con h h' x
|
ap10 (h ⬝ h') x = ap10 h x ⬝ ap10 h' x := apd10_con h h' x
|
||||||
|
|
||||||
definition ap10_inv {f g : A → B} (h : f = g) (x : A) : ap10 h⁻¹ x = (ap10 h x)⁻¹ :=
|
definition ap10_inv {f g : A → B} (h : f = g) (x : A) : ap10 h⁻¹ x = (ap10 h x)⁻¹ :=
|
||||||
apD10_inv h x
|
apd10_inv h x
|
||||||
|
|
||||||
-- [ap10] also behaves nicely on paths produced by [ap]
|
-- [ap10] also behaves nicely on paths produced by [ap]
|
||||||
definition ap_ap10 (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) :
|
||||||
|
@ -403,8 +404,7 @@ namespace eq
|
||||||
|
|
||||||
/- Transport and the groupoid structure of paths -/
|
/- Transport and the groupoid structure of paths -/
|
||||||
|
|
||||||
definition tr_idp (P : A → Type) {x : A} (u : P x) :
|
definition tr_idp (P : A → Type) {x : A} (u : P x) : idp ▹ u = u := idp
|
||||||
idp ▹ u = u := idp
|
|
||||||
|
|
||||||
definition tr_con (P : A → Type) {x y z : A} (p : x = y) (q : y = z) (u : P x) :
|
definition tr_con (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 :=
|
||||||
|
@ -432,12 +432,12 @@ namespace eq
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
-- Dependent transport in a doubly dependent type.
|
-- Dependent transport in a doubly dependent type.
|
||||||
-- should P, Q and y all be explicit here?
|
definition transportD {P : A → Type} (Q : Π a : A, P a → Type)
|
||||||
definition transportD (P : A → Type) (Q : Π a : A, P a → Type)
|
|
||||||
{a a' : A} (p : a = a') (b : P a) (z : Q a b) : Q a' (p ▹ b) :=
|
{a a' : A} (p : a = a') (b : P a) (z : Q a b) : Q a' (p ▹ b) :=
|
||||||
eq.rec_on p z
|
eq.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
|
|
||||||
notation p `▹D`:65 x:64 := transportD _ _ p _ x
|
-- In Coq the variables P, Q and b are explicit, but in Lean we can probably have them implicit using the following notation
|
||||||
|
notation p `▹D`:65 x:64 := transportD _ p _ x
|
||||||
|
|
||||||
-- Transporting along higher-dimensional paths
|
-- Transporting along higher-dimensional paths
|
||||||
definition transport2 (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) :
|
||||||
|
@ -470,7 +470,7 @@ namespace eq
|
||||||
definition ap_tr_con_tr2 (P : A → Type) {x y : A} {p q : x = y} {z w : P x} (r : p = q)
|
definition ap_tr_con_tr2 (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 :=
|
||||||
eq.rec_on r (con_idp _ ⬝ (idp_con _)⁻¹)
|
eq.rec_on r (!con_idp ⬝ !idp_con⁻¹)
|
||||||
|
|
||||||
|
|
||||||
definition fn_tr_eq_tr_fn {P Q : A → Type} {x y : A} (p : x = y) (f : Πx, P x → Q x) (z : P x) :
|
definition fn_tr_eq_tr_fn {P Q : A → Type} {x y : A} (p : x = y) (f : Πx, P x → Q x) (z : P x) :
|
||||||
|
@ -494,7 +494,7 @@ namespace eq
|
||||||
|
|
||||||
definition tr2_constant {p q : x = y} (r : p = q) (z : B) :
|
definition tr2_constant {p q : x = y} (r : p = q) (z : B) :
|
||||||
tr_constant p z = transport2 (λu, B) r z ⬝ tr_constant q z :=
|
tr_constant p z = transport2 (λu, B) r z ⬝ tr_constant q z :=
|
||||||
eq.rec_on r (idp_con _)⁻¹
|
eq.rec_on r !idp_con⁻¹
|
||||||
|
|
||||||
-- Transporting in a pulled back fibration.
|
-- Transporting in a pulled back fibration.
|
||||||
-- rename: tr_compose
|
-- rename: tr_compose
|
||||||
|
@ -506,33 +506,31 @@ namespace eq
|
||||||
ap (λh, h ∘ f) p = 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 :=
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
definition apD10_ap_precompose (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) :=
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
definition apD10_ap_postcompose (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) :=
|
||||||
eq.rec_on p idp
|
eq.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 tr_eq_tr_id_ap (P : A → Type) x y (p : x = y) (u : P x) :
|
definition tr_eq_cast_ap (P : A → Type) {x y} (p : x = y) (u : P x) : p ▹ u = cast (ap P p) u :=
|
||||||
transport P p u = transport id (ap P p) u :=
|
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
|
|
||||||
/- The behavior of [ap] and [apD] -/
|
/- The behavior of [ap] and [apd] -/
|
||||||
|
|
||||||
-- In a constant fibration, [apD] reduces to [ap], modulo [transport_const].
|
-- In a constant fibration, [apd] reduces to [ap], modulo [transport_const].
|
||||||
definition apD_eq_tr_constant_con_ap (f : A → B) (p: x = y) :
|
definition apd_eq_tr_constant_con_ap (f : A → B) (p : x = y) :
|
||||||
apD f p = tr_constant p (f x) ⬝ ap f p :=
|
apd f p = tr_constant p (f x) ⬝ ap f p :=
|
||||||
eq.rec_on p idp
|
eq.rec_on p idp
|
||||||
|
|
||||||
|
|
||||||
/- The 2-dimensional groupoid structure -/
|
/- The 2-dimensional groupoid structure -/
|
||||||
|
|
||||||
-- Horizontal composition of 2-dimensional paths.
|
-- Horizontal composition of 2-dimensional paths.
|
||||||
definition concat2 {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' :=
|
|
||||||
eq.rec_on h (eq.rec_on h' idp)
|
eq.rec_on h (eq.rec_on h' idp)
|
||||||
|
|
||||||
infixl `◾`:75 := concat2
|
infixl `◾`:75 := concat2
|
||||||
|
@ -552,10 +550,10 @@ namespace eq
|
||||||
|
|
||||||
-- Unwhiskering, a.k.a. cancelling
|
-- Unwhiskering, a.k.a. cancelling
|
||||||
|
|
||||||
definition cancel_left {x y z : A} (p : x = y) (q r : y = z) : (p ⬝ q = p ⬝ r) → (q = r) :=
|
definition cancel_left {x y z : A} {p : x = y} {q r : y = z} : (p ⬝ q = p ⬝ r) → (q = r) :=
|
||||||
λs, !inv_con_cancel_left⁻¹ ⬝ whisker_left p⁻¹ s ⬝ !inv_con_cancel_left
|
λs, !inv_con_cancel_left⁻¹ ⬝ whisker_left p⁻¹ s ⬝ !inv_con_cancel_left
|
||||||
|
|
||||||
definition cancel_right {x y z : A} (p q : x = y) (r : y = z) : (p ⬝ r = q ⬝ r) → (p = q) :=
|
definition cancel_right {x y z : A} {p q : x = y} {r : y = z} : (p ⬝ r = q ⬝ r) → (p = q) :=
|
||||||
λs, !con_inv_cancel_right⁻¹ ⬝ whisker_right s r⁻¹ ⬝ !con_inv_cancel_right
|
λs, !con_inv_cancel_right⁻¹ ⬝ whisker_right s r⁻¹ ⬝ !con_inv_cancel_right
|
||||||
|
|
||||||
-- Whiskering and identity paths.
|
-- Whiskering and identity paths.
|
||||||
|
@ -590,9 +588,10 @@ namespace eq
|
||||||
(a ◾ c) ⬝ (b ◾ d) = (a ⬝ b) ◾ (c ⬝ d) :=
|
(a ◾ c) ⬝ (b ◾ d) = (a ⬝ b) ◾ (c ⬝ d) :=
|
||||||
eq.rec_on d (eq.rec_on c (eq.rec_on b (eq.rec_on a idp)))
|
eq.rec_on d (eq.rec_on c (eq.rec_on b (eq.rec_on a idp)))
|
||||||
|
|
||||||
definition whisker_right_con_whisker_left {x y z : A} (p p' : x = y) (q q' : y = z) (a : p = p') (b : q = q') :
|
definition whisker_right_con_whisker_left {x y z : A} {p p' : x = y} {q q' : y = z}
|
||||||
|
(a : p = p') (b : q = q') :
|
||||||
(whisker_right a q) ⬝ (whisker_left p' b) = (whisker_left p b) ⬝ (whisker_right a q') :=
|
(whisker_right a q) ⬝ (whisker_left p' b) = (whisker_left p b) ⬝ (whisker_right a q') :=
|
||||||
eq.rec_on b (eq.rec_on a (idp_con _)⁻¹)
|
eq.rec_on b (eq.rec_on a !idp_con⁻¹)
|
||||||
|
|
||||||
-- Structure corresponding to the coherence equations of a bicategory.
|
-- Structure corresponding to the coherence equations of a bicategory.
|
||||||
|
|
||||||
|
@ -609,7 +608,7 @@ namespace eq
|
||||||
con.assoc' p idp q ⬝ whisker_right (con_idp p) q = whisker_left p (idp_con q) :=
|
con.assoc' p idp q ⬝ whisker_right (con_idp p) q = whisker_left p (idp_con q) :=
|
||||||
eq.rec_on q (eq.rec_on p idp)
|
eq.rec_on q (eq.rec_on p idp)
|
||||||
|
|
||||||
definition eckmann_hilton {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 :=
|
||||||
(!whisker_right_idp_right ◾ !whisker_left_idp_left)⁻¹
|
(!whisker_right_idp_right ◾ !whisker_left_idp_left)⁻¹
|
||||||
⬝ (!con_idp ◾ !con_idp)
|
⬝ (!con_idp ◾ !con_idp)
|
||||||
⬝ (!idp_con ◾ !idp_con)
|
⬝ (!idp_con ◾ !idp_con)
|
||||||
|
@ -619,7 +618,7 @@ namespace eq
|
||||||
⬝ (!whisker_left_idp_left ◾ !whisker_right_idp_right)
|
⬝ (!whisker_left_idp_left ◾ !whisker_right_idp_right)
|
||||||
|
|
||||||
-- The action of functions on 2-dimensional paths
|
-- The action of functions on 2-dimensional paths
|
||||||
definition ap02 (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 :=
|
||||||
eq.rec_on r idp
|
eq.rec_on r idp
|
||||||
|
|
||||||
definition ap02_con (f : A → B) {x y : A} {p p' p'' : x = y} (r : p = p') (r' : p' = p'') :
|
definition ap02_con (f : A → B) {x y : A} {p p' p'' : x = y} (r : p = p') (r' : p' = p'') :
|
||||||
|
@ -633,16 +632,16 @@ namespace eq
|
||||||
⬝ (ap_con f p' q')⁻¹ :=
|
⬝ (ap_con f p' q')⁻¹ :=
|
||||||
eq.rec_on r (eq.rec_on s (eq.rec_on q (eq.rec_on p idp)))
|
eq.rec_on r (eq.rec_on s (eq.rec_on q (eq.rec_on p idp)))
|
||||||
|
|
||||||
definition apD02 {p q : x = y} (f : Π x, P x) (r : p = q) :
|
definition apd02 {p q : x = y} (f : Π x, P x) (r : p = q) :
|
||||||
apD f p = transport2 P r (f x) ⬝ apD f q :=
|
apd f p = transport2 P r (f x) ⬝ apd f q :=
|
||||||
eq.rec_on r (idp_con _)⁻¹
|
eq.rec_on r !idp_con⁻¹
|
||||||
|
|
||||||
-- 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_con (P : A → Type) (f : Π x:A, P x) {x y : A}
|
definition apd02_con (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
|
||||||
⬝ whisker_left (transport2 P r1 (f x)) (apD02 f r2)
|
⬝ whisker_left (transport2 P r1 (f x)) (apd02 f r2)
|
||||||
⬝ con.assoc' _ _ _
|
⬝ con.assoc' _ _ _
|
||||||
⬝ (whisker_right (tr2_con P r1 r2 (f x))⁻¹ (apD f p3)) :=
|
⬝ (whisker_right (tr2_con P r1 r2 (f x))⁻¹ (apd f p3)) :=
|
||||||
eq.rec_on r2 (eq.rec_on r1 (eq.rec_on p1 idp))
|
eq.rec_on r2 (eq.rec_on r1 (eq.rec_on p1 idp))
|
||||||
end eq
|
end eq
|
||||||
|
|
|
@ -96,7 +96,7 @@ namespace eq
|
||||||
|
|
||||||
definition transport_eq_FlFr_D {B : A → Type} {f g : Πa, B a}
|
definition transport_eq_FlFr_D {B : A → Type} {f g : Πa, B a}
|
||||||
(p : a1 = a2) (q : f a1 = g a1)
|
(p : a1 = a2) (q : f a1 = g a1)
|
||||||
: transport (λx, f x = g x) p q = (apD f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apD g p) :=
|
: transport (λx, f x = g x) p q = (apd f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apd g p) :=
|
||||||
begin
|
begin
|
||||||
apply (eq.rec_on p),
|
apply (eq.rec_on p),
|
||||||
apply inverse,
|
apply inverse,
|
||||||
|
@ -201,7 +201,7 @@ namespace eq
|
||||||
equiv.mk _ !is_equiv_whisker_right
|
equiv.mk _ !is_equiv_whisker_right
|
||||||
|
|
||||||
definition is_equiv_con_eq_of_eq_inv_con (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
definition is_equiv_con_eq_of_eq_inv_con (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||||
: is_equiv (con_eq_of_eq_inv_con p q r) :=
|
: is_equiv (con_eq_of_eq_inv_con : p = r⁻¹ ⬝ q → r ⬝ p = q) :=
|
||||||
begin
|
begin
|
||||||
cases r,
|
cases r,
|
||||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right),
|
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right),
|
||||||
|
@ -212,7 +212,7 @@ namespace eq
|
||||||
equiv.mk _ !is_equiv_con_eq_of_eq_inv_con
|
equiv.mk _ !is_equiv_con_eq_of_eq_inv_con
|
||||||
|
|
||||||
definition is_equiv_con_eq_of_eq_con_inv (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
definition is_equiv_con_eq_of_eq_con_inv (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||||
: is_equiv (con_eq_of_eq_con_inv p q r) :=
|
: is_equiv (con_eq_of_eq_con_inv : r = q ⬝ p⁻¹ → r ⬝ p = q) :=
|
||||||
begin
|
begin
|
||||||
cases p,
|
cases p,
|
||||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||||
|
@ -223,7 +223,7 @@ namespace eq
|
||||||
equiv.mk _ !is_equiv_con_eq_of_eq_con_inv
|
equiv.mk _ !is_equiv_con_eq_of_eq_con_inv
|
||||||
|
|
||||||
definition is_equiv_inv_con_eq_of_eq_con (p : a1 = a3) (q : a2 = a3) (r : a1 = a2)
|
definition is_equiv_inv_con_eq_of_eq_con (p : a1 = a3) (q : a2 = a3) (r : a1 = a2)
|
||||||
: is_equiv (inv_con_eq_of_eq_con p q r) :=
|
: is_equiv (inv_con_eq_of_eq_con : p = r ⬝ q → r⁻¹ ⬝ p = q) :=
|
||||||
begin
|
begin
|
||||||
cases r,
|
cases r,
|
||||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||||
|
@ -234,7 +234,7 @@ namespace eq
|
||||||
equiv.mk _ !is_equiv_inv_con_eq_of_eq_con
|
equiv.mk _ !is_equiv_inv_con_eq_of_eq_con
|
||||||
|
|
||||||
definition is_equiv_con_inv_eq_of_eq_con (p : a3 = a1) (q : a2 = a3) (r : a2 = a1)
|
definition is_equiv_con_inv_eq_of_eq_con (p : a3 = a1) (q : a2 = a3) (r : a2 = a1)
|
||||||
: is_equiv (con_inv_eq_of_eq_con p q r) :=
|
: is_equiv (con_inv_eq_of_eq_con : r = q ⬝ p → r ⬝ p⁻¹ = q) :=
|
||||||
begin
|
begin
|
||||||
cases p,
|
cases p,
|
||||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||||
|
@ -245,7 +245,7 @@ namespace eq
|
||||||
equiv.mk _ !is_equiv_con_inv_eq_of_eq_con
|
equiv.mk _ !is_equiv_con_inv_eq_of_eq_con
|
||||||
|
|
||||||
definition is_equiv_eq_con_of_inv_con_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
definition is_equiv_eq_con_of_inv_con_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||||
: is_equiv (eq_con_of_inv_con_eq p q r) :=
|
: is_equiv (eq_con_of_inv_con_eq : r⁻¹ ⬝ q = p → q = r ⬝ p) :=
|
||||||
begin
|
begin
|
||||||
cases r,
|
cases r,
|
||||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||||
|
@ -256,7 +256,7 @@ namespace eq
|
||||||
equiv.mk _ !is_equiv_eq_con_of_inv_con_eq
|
equiv.mk _ !is_equiv_eq_con_of_inv_con_eq
|
||||||
|
|
||||||
definition is_equiv_eq_con_of_con_inv_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
definition is_equiv_eq_con_of_con_inv_eq (p : a1 = a3) (q : a2 = a3) (r : a2 = a1)
|
||||||
: is_equiv (eq_con_of_con_inv_eq p q r) :=
|
: is_equiv (eq_con_of_con_inv_eq : q ⬝ p⁻¹ = r → q = r ⬝ p) :=
|
||||||
begin
|
begin
|
||||||
cases p,
|
cases p,
|
||||||
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
apply (@is_equiv_compose _ _ _ _ _ !is_equiv_concat_left !is_equiv_concat_right)
|
||||||
|
|
|
@ -91,7 +91,7 @@ namespace equiv
|
||||||
|
|
||||||
protected definition eq_mk' {f f' : A → B} [H : is_equiv f] [H' : is_equiv f'] (p : f = f')
|
protected definition eq_mk' {f f' : A → B} [H : is_equiv f] [H' : is_equiv f'] (p : f = f')
|
||||||
: equiv.mk f H = equiv.mk f' H' :=
|
: equiv.mk f H = equiv.mk f' H' :=
|
||||||
apD011 equiv.mk p !is_hprop.elim
|
apd011 equiv.mk p !is_hprop.elim
|
||||||
|
|
||||||
protected definition eq_mk {f f' : A ≃ B} (p : to_fun f = to_fun f') : f = f' :=
|
protected definition eq_mk {f f' : A ≃ B} (p : to_fun f = to_fun f') : f = f' :=
|
||||||
by (cases f; cases f'; apply (equiv.eq_mk' p))
|
by (cases f; cases f'; apply (equiv.eq_mk' p))
|
||||||
|
|
|
@ -22,15 +22,15 @@ namespace pi
|
||||||
|
|
||||||
/- 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 apD10_eq_of_homotopy (h : f ∼ g) : apD10 (eq_of_homotopy h) ∼ h :=
|
definition apd10_eq_of_homotopy (h : f ∼ g) : apd10 (eq_of_homotopy h) ∼ h :=
|
||||||
apD10 (retr apD10 h)
|
apd10 (retr apd10 h)
|
||||||
|
|
||||||
definition eq_of_homotopy_eta (p : f = g) : eq_of_homotopy (apD10 p) = p :=
|
definition eq_of_homotopy_eta (p : f = g) : eq_of_homotopy (apd10 p) = p :=
|
||||||
sect apD10 p
|
sect apd10 p
|
||||||
|
|
||||||
definition eq_of_homotopy_idp (f : Πa, B a) : eq_of_homotopy (λx : A, refl (f x)) = refl f :=
|
definition eq_of_homotopy_idp (f : Πa, B a) : eq_of_homotopy (λx : A, refl (f x)) = refl f :=
|
||||||
!eq_of_homotopy_eta
|
!eq_of_homotopy_eta
|
||||||
|
@ -38,11 +38,11 @@ namespace pi
|
||||||
/- The identification of the path space of a dependent function space, up to equivalence, is of course just funext. -/
|
/- The identification of the path space of a dependent function space, up to equivalence, is of course just funext. -/
|
||||||
|
|
||||||
definition eq_equiv_homotopy (f g : Πx, B x) : (f = g) ≃ (f ∼ g) :=
|
definition eq_equiv_homotopy (f g : Πx, B x) : (f = g) ≃ (f ∼ g) :=
|
||||||
equiv.mk _ !is_equiv_apD
|
equiv.mk _ !is_equiv_apd
|
||||||
|
|
||||||
definition is_equiv_eq_of_homotopy [instance] (f g : Πx, B x)
|
definition is_equiv_eq_of_homotopy [instance] (f g : Πx, B x)
|
||||||
: is_equiv (@eq_of_homotopy _ _ f g) :=
|
: is_equiv (@eq_of_homotopy _ _ f g) :=
|
||||||
is_equiv_inv apD10
|
is_equiv_inv apd10
|
||||||
|
|
||||||
definition homotopy_equiv_eq (f g : Πx, B x) : (f ∼ g) ≃ (f = g) :=
|
definition homotopy_equiv_eq (f g : Πx, B x) : (f ∼ g) ≃ (f = g) :=
|
||||||
equiv.mk _ !is_equiv_eq_of_homotopy
|
equiv.mk _ !is_equiv_eq_of_homotopy
|
||||||
|
@ -52,7 +52,7 @@ namespace pi
|
||||||
|
|
||||||
definition pi_transport (p : a = a') (f : Π(b : B a), C a b)
|
definition pi_transport (p : a = a') (f : Π(b : B a), C a b)
|
||||||
: (transport (λa, Π(b : B a), C a b) p f)
|
: (transport (λa, Π(b : B a), C a b) p f)
|
||||||
∼ (λb, transport (C a') !tr_inv_tr (transportD _ _ p _ (f (p⁻¹ ▹ b)))) :=
|
∼ (λb, transport (C a') !tr_inv_tr (transportD _ p _ (f (p⁻¹ ▹ b)))) :=
|
||||||
eq.rec_on p (λx, idp)
|
eq.rec_on p (λx, idp)
|
||||||
|
|
||||||
/- A special case of [transport_pi] where the type [B] does not depend on [A],
|
/- A special case of [transport_pi] where the type [B] does not depend on [A],
|
||||||
|
@ -107,7 +107,7 @@ namespace pi
|
||||||
definition ap_pi_functor {g g' : Π(a:A), B a} (h : g ∼ g')
|
definition ap_pi_functor {g g' : Π(a:A), B a} (h : g ∼ g')
|
||||||
: ap (pi_functor f0 f1) (eq_of_homotopy h) = eq_of_homotopy (λa':A', (ap (f1 a') (h (f0 a')))) :=
|
: ap (pi_functor f0 f1) (eq_of_homotopy h) = eq_of_homotopy (λa':A', (ap (f1 a') (h (f0 a')))) :=
|
||||||
begin
|
begin
|
||||||
apply (equiv_rect (@apD10 A B g g')), intro p, clear h,
|
apply (equiv_rect (@apd10 A B g g')), intro p, clear h,
|
||||||
cases p,
|
cases p,
|
||||||
apply concat,
|
apply concat,
|
||||||
exact (ap (ap (pi_functor f0 f1)) (eq_of_homotopy_idp g)),
|
exact (ap (ap (pi_functor f0 f1)) (eq_of_homotopy_idp g)),
|
||||||
|
@ -130,13 +130,13 @@ namespace pi
|
||||||
apply (transport (λx, f1 a' x = h a') (transport_compose B f0 (sect f0 a') _)),
|
apply (transport (λx, f1 a' x = h a') (transport_compose B f0 (sect f0 a') _)),
|
||||||
apply (tr_inv (λx, x = h a') (fn_tr_eq_tr_fn _ f1 _)), unfold function.compose,
|
apply (tr_inv (λx, x = h a') (fn_tr_eq_tr_fn _ f1 _)), unfold function.compose,
|
||||||
apply (tr_inv (λx, sect f0 a' ▹ x = h a') (retr (f1 _) _)), unfold function.id,
|
apply (tr_inv (λx, sect f0 a' ▹ x = h a') (retr (f1 _) _)), unfold function.id,
|
||||||
apply apD
|
apply apd
|
||||||
end,
|
end,
|
||||||
begin
|
begin
|
||||||
intro h,
|
intro h,
|
||||||
apply eq_of_homotopy, intro a,
|
apply eq_of_homotopy, intro a,
|
||||||
apply (tr_inv (λx, retr f0 a ▹ x = h a) (sect (f1 _) _)), unfold function.id,
|
apply (tr_inv (λx, retr f0 a ▹ x = h a) (sect (f1 _) _)), unfold function.id,
|
||||||
apply apD
|
apply apd
|
||||||
end
|
end
|
||||||
end
|
end
|
||||||
|
|
||||||
|
|
|
@ -80,7 +80,7 @@ namespace is_trunc
|
||||||
(imp : Π{a b : A}, R a b → a = b) : is_hset A :=
|
(imp : Π{a b : A}, R a b → a = b) : is_hset A :=
|
||||||
is_hset_of_axiom_K
|
is_hset_of_axiom_K
|
||||||
(λa p,
|
(λa p,
|
||||||
have H2 : transport (λx, R a x → a = x) p (@imp a a) = @imp a a, from !apD,
|
have H2 : transport (λx, R a x → a = x) p (@imp a a) = @imp a a, from !apd,
|
||||||
have H3 : Π(r : R a a), transport (λx, a = x) p (imp r)
|
have H3 : Π(r : R a a), transport (λx, a = x) p (imp r)
|
||||||
= imp (transport (λx, R a x) p r), from
|
= imp (transport (λx, R a x) p r), from
|
||||||
to_fun (equiv.symm !heq_pi) H2,
|
to_fun (equiv.symm !heq_pi) H2,
|
||||||
|
@ -89,7 +89,7 @@ namespace is_trunc
|
||||||
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_eq_r
|
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_eq_r
|
||||||
... = imp (transport (λx, R a x) p (refl a)) : H3
|
... = imp (transport (λx, R a x) p (refl a)) : H3
|
||||||
... = imp (refl a) : is_hprop.elim,
|
... = imp (refl a) : is_hprop.elim,
|
||||||
cancel_left (imp (refl a)) _ _ H4)
|
cancel_left H4)
|
||||||
|
|
||||||
definition relation_equiv_eq {A : Type} (R : A → A → Type)
|
definition relation_equiv_eq {A : Type} (R : A → A → Type)
|
||||||
(mere : Π(a b : A), is_hprop (R a b)) (refl : Π(a : A), R a a)
|
(mere : Π(a b : A), is_hprop (R a b)) (refl : Π(a : A), R a a)
|
||||||
|
|
Loading…
Reference in a new issue