feat(hott): move basic theorems from colimit development to library.
Most notable changes: rename apo011 -> apd011 and apd011 -> apdt011 make an argument of pathover_of_eq explicit
This commit is contained in:
parent
ae1b2e854c
commit
17ccc283a9
37 changed files with 426 additions and 159 deletions
|
@ -44,18 +44,16 @@ namespace functor
|
|||
|
||||
definition nat_trans_left_inverse : nat_trans_inverse η ∘n η = 1 :=
|
||||
begin
|
||||
fapply (apd011 nat_trans.mk),
|
||||
fapply (apdt011 nat_trans.mk),
|
||||
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 is_set.elim
|
||||
apply eq_of_homotopy3, intros, apply is_set.elim
|
||||
end
|
||||
|
||||
definition nat_trans_right_inverse : η ∘n nat_trans_inverse η = 1 :=
|
||||
begin
|
||||
fapply (apd011 nat_trans.mk),
|
||||
fapply (apdt011 nat_trans.mk),
|
||||
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 is_set.elim
|
||||
apply eq_of_homotopy3, intros, apply is_set.elim
|
||||
end
|
||||
|
||||
definition is_natural_iso [constructor] : is_iso η :=
|
||||
|
|
|
@ -52,7 +52,7 @@ namespace category
|
|||
→ is_left_adjoint.mk G η ε H K = is_left_adjoint.mk G' η' ε' H' K',
|
||||
begin
|
||||
intros p q r, induction p, induction q, induction r, esimp,
|
||||
apply apd011 (is_left_adjoint.mk G η ε) !is_prop.elim !is_prop.elim
|
||||
apply apd011 (is_left_adjoint.mk G η ε) !is_prop.elim !is_prop.elimo
|
||||
end,
|
||||
have lem₂ : Π (d : carrier D),
|
||||
(to_fun_hom G (natural_map ε' d) ∘
|
||||
|
|
|
@ -56,7 +56,7 @@ namespace functor
|
|||
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} (id₁ id₂ comp₁ comp₂)
|
||||
(pF : F₁ = F₂) (pH : pF ▸ H₁ = H₂)
|
||||
: functor.mk F₁ H₁ id₁ comp₁ = functor.mk F₂ H₂ id₂ comp₂ :=
|
||||
apd01111 functor.mk pF pH !is_prop.elim !is_prop.elim
|
||||
apdt01111 functor.mk pF pH !is_prop.elim !is_prop.elim
|
||||
|
||||
definition functor_eq' {F₁ F₂ : C ⇒ D} : Π(p : to_fun_ob F₁ = to_fun_ob F₂),
|
||||
(transport (λx, Πa b f, hom (x a) (x b)) p @(to_fun_hom F₁) = @(to_fun_hom F₂)) → F₁ = F₂ :=
|
||||
|
@ -192,9 +192,9 @@ namespace functor
|
|||
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 :=
|
||||
begin
|
||||
fapply apd011 (apd01111 functor.mk idp idp),
|
||||
apply is_set.elim,
|
||||
apply is_set.elim
|
||||
fapply apd011 (apdt01111 functor.mk idp idp),
|
||||
apply is_prop.elim,
|
||||
apply is_prop.elimo
|
||||
end
|
||||
|
||||
definition functor_eq'_idp (F : C ⇒ D) : functor_eq' idp idp = (idpath F) :=
|
||||
|
@ -225,9 +225,9 @@ namespace functor
|
|||
|
||||
theorem ap010_apd01111_functor {F₁ F₂ : C → D} {H₁ : Π(a b : C), hom a b → hom (F₁ a) (F₁ b)}
|
||||
{H₂ : Π(a b : C), hom a b → hom (F₂ a) (F₂ b)} {id₁ id₂ comp₁ comp₂}
|
||||
(pF : F₁ = F₂) (pH : pF ▸ H₁ = H₂) (pid : cast (apd011 _ pF pH) id₁ = id₂)
|
||||
(pcomp : cast (apd0111 _ pF pH pid) comp₁ = comp₂) (c : C)
|
||||
: ap010 to_fun_ob (apd01111 functor.mk pF pH pid pcomp) c = ap10 pF c :=
|
||||
(pF : F₁ = F₂) (pH : pF ▸ H₁ = H₂) (pid : cast (apdt011 _ pF pH) id₁ = id₂)
|
||||
(pcomp : cast (apdt0111 _ pF pH pid) comp₁ = comp₂) (c : C)
|
||||
: ap010 to_fun_ob (apdt01111 functor.mk pF pH pid pcomp) c = ap10 pF c :=
|
||||
by induction pF; induction pH; induction pid; induction pcomp; reflexivity
|
||||
|
||||
definition ap010_functor_eq {F₁ F₂ : C ⇒ D} (p : to_fun_ob F₁ ~ to_fun_ob F₂)
|
||||
|
|
|
@ -334,13 +334,13 @@ namespace category
|
|||
definition equivalence_eq {C : Category} {D : Precategory} {F F' : C ≃c D}
|
||||
(p : equivalence.to_functor F = equivalence.to_functor F') : F = F' :=
|
||||
begin
|
||||
induction F, induction F', exact apd011 equivalence.mk p !is_prop.elim
|
||||
induction F, induction F', exact apd011 equivalence.mk p !is_prop.elimo
|
||||
end
|
||||
|
||||
definition isomorphism_eq {F F' : C ≅c D}
|
||||
(p : isomorphism.to_functor F = isomorphism.to_functor F') : F = F' :=
|
||||
begin
|
||||
induction F, induction F', exact apd011 isomorphism.mk p !is_prop.elim
|
||||
induction F, induction F', exact apd011 isomorphism.mk p !is_prop.elimo
|
||||
end
|
||||
|
||||
definition is_equiv_isomorphism_of_equivalence [constructor] (C D : Category)
|
||||
|
|
|
@ -121,8 +121,8 @@ namespace iso
|
|||
apply left_inverse_eq_right_inverse,
|
||||
apply li,
|
||||
apply ri',
|
||||
apply is_prop.elim,
|
||||
apply is_prop.elim,
|
||||
apply is_prop.elimo,
|
||||
apply is_prop.elimo,
|
||||
end
|
||||
end iso open iso
|
||||
|
||||
|
@ -174,7 +174,7 @@ namespace iso
|
|||
|
||||
definition iso_mk_eq {f f' : a ⟶ b} [H : is_iso f] [H' : is_iso f'] (p : f = f')
|
||||
: iso.mk f _ = iso.mk f' _ :=
|
||||
apd011 iso.mk p !is_prop.elim
|
||||
apd011 iso.mk p !is_prop.elimo
|
||||
|
||||
variable {C}
|
||||
definition iso_eq {f f' : a ≅ b} (p : to_hom f = to_hom f') : f = f' :=
|
||||
|
|
|
@ -52,7 +52,7 @@ namespace nat_trans
|
|||
(nat₂ : Π (a b : C) (f : hom a b), G f ∘ η₂ a = η₂ b ∘ F f)
|
||||
(p : η₁ ~ η₂)
|
||||
: nat_trans.mk η₁ nat₁ = nat_trans.mk η₂ nat₂ :=
|
||||
apd011 nat_trans.mk (eq_of_homotopy p) !is_prop.elim
|
||||
apd011 nat_trans.mk (eq_of_homotopy p) !is_prop.elimo
|
||||
|
||||
definition nat_trans_eq {η₁ η₂ : F ⟹ G} : natural_map η₁ ~ natural_map η₂ → η₁ = η₂ :=
|
||||
by induction η₁; induction η₂; apply nat_trans_mk_eq
|
||||
|
|
|
@ -69,7 +69,7 @@ namespace algebra
|
|||
(resp_mul : Π(g h : G), cast p (g * h) = cast p g * cast p h) : G = H :=
|
||||
begin
|
||||
cases G with Gc G, cases H with Hc H,
|
||||
apply (apo011 mk p),
|
||||
apply (apd011 mk p),
|
||||
exact group_pathover resp_mul
|
||||
end
|
||||
|
||||
|
|
|
@ -79,29 +79,29 @@ namespace eq
|
|||
definition ap01000 (f : X → Πa b c, D a b c) (Hx : x = x') : f x ~3 f x' :=
|
||||
by intros; cases Hx; reflexivity
|
||||
|
||||
definition apd011 (f : Πa, B a → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
definition apdt011 (f : Πa, B a → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
: f a b = f a' b' :=
|
||||
by cases Ha; cases Hb; reflexivity
|
||||
|
||||
definition apd0111 (f : Πa b, C a b → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apd011 C Ha Hb) c = c')
|
||||
definition apdt0111 (f : Πa b, C a b → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apdt011 C Ha Hb) c = c')
|
||||
: f a b c = f a' b' c' :=
|
||||
by cases Ha; cases Hb; cases Hc; reflexivity
|
||||
|
||||
definition apd01111 (f : Πa b c, D a b c → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apd011 C Ha Hb) c = c') (Hd : cast (apd0111 D Ha Hb Hc) d = d')
|
||||
definition apdt01111 (f : Πa b c, D a b c → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apdt011 C Ha Hb) c = c') (Hd : cast (apdt0111 D Ha Hb Hc) d = d')
|
||||
: f a b c d = f a' b' c' d' :=
|
||||
by cases Ha; cases Hb; cases Hc; cases Hd; reflexivity
|
||||
|
||||
definition apd011111 (f : Πa b c d, E a b c d → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apd011 C Ha Hb) c = c') (Hd : cast (apd0111 D Ha Hb Hc) d = d')
|
||||
(He : cast (apd01111 E Ha Hb Hc Hd) e = e')
|
||||
definition apdt011111 (f : Πa b c d, E a b c d → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apdt011 C Ha Hb) c = c') (Hd : cast (apdt0111 D Ha Hb Hc) d = d')
|
||||
(He : cast (apdt01111 E Ha Hb Hc Hd) e = e')
|
||||
: f a b c d e = f a' b' c' d' e' :=
|
||||
by cases Ha; cases Hb; cases Hc; cases Hd; cases He; reflexivity
|
||||
|
||||
definition apd0111111 (f : Πa b c d e, F a b c d e → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apd011 C Ha Hb) c = c') (Hd : cast (apd0111 D Ha Hb Hc) d = d')
|
||||
(He : cast (apd01111 E Ha Hb Hc Hd) e = e') (Hf : cast (apd011111 F Ha Hb Hc Hd He) ff = f')
|
||||
definition apdt0111111 (f : Πa b c d e, F a b c d e → Z) (Ha : a = a') (Hb : transport B Ha b = b')
|
||||
(Hc : cast (apdt011 C Ha Hb) c = c') (Hd : cast (apdt0111 D Ha Hb Hc) d = d')
|
||||
(He : cast (apdt01111 E Ha Hb Hc Hd) e = e') (Hf : cast (apdt011111 F Ha Hb Hc Hd He) ff = f')
|
||||
: f a b c d e ff = f a' b' c' d' e' f' :=
|
||||
begin cases Ha, cases Hb, cases Hc, cases Hd, cases He, cases Hf, reflexivity end
|
||||
|
||||
|
|
|
@ -37,8 +37,8 @@ namespace eq
|
|||
: pathover_of_tr_eq r = pathover_tr p b ⬝o pathover_idp_of_eq r :=
|
||||
by induction r; induction p; reflexivity
|
||||
|
||||
definition apo011_eq_apo11_apd (f : Πa, B a → A') (p : a = a₂) (q : b =[p] b₂)
|
||||
: apo011 f p q = eq_of_pathover (apo11 (apd f p) q) :=
|
||||
definition apd011_eq_apo11_apd (f : Πa, B a → A') (p : a = a₂) (q : b =[p] b₂)
|
||||
: apd011 f p q = apo11_constant_right (apd f p) q :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition change_path_con (q : p = p') (q' : p' = p'') (r : b =[p] b₂) :
|
||||
|
@ -125,4 +125,9 @@ namespace eq
|
|||
: change_path (con.right_inv p) (q ⬝o q'⁻¹ᵒ) = idpo :=
|
||||
by induction r; induction q; reflexivity
|
||||
|
||||
definition tr_eq_of_pathover_concato_eq {A : Type} {B : A → Type} {a a' : A} {p : a = a'}
|
||||
{b : B a} {b' b'' : B a'} (q : b =[p] b') (r : b' = b'') :
|
||||
tr_eq_of_pathover (q ⬝op r) = tr_eq_of_pathover q ⬝ r :=
|
||||
by induction r; reflexivity
|
||||
|
||||
end eq
|
||||
|
|
|
@ -60,6 +60,10 @@ namespace eq
|
|||
: square p₁₀ p₁₄ (p₀₁ ⬝ p₀₃) (p₂₁ ⬝ p₂₃) :=
|
||||
by induction s₁₃; exact s₁₁
|
||||
|
||||
definition dconcat [unfold 14] {p₀₀ : a₀₀ = a} {p₂₂ : a = a₂₂}
|
||||
(s₂₁ : square p₀₀ p₁₂ p₀₁ p₂₂) (s₁₂ : square p₁₀ p₂₂ p₀₀ p₂₁) : square p₁₀ p₁₂ p₀₁ p₂₁ :=
|
||||
by induction s₁₂; exact s₂₁
|
||||
|
||||
definition hinverse [unfold 10] (s₁₁ : square p₁₀ p₁₂ p₀₁ p₂₁) : square p₁₀⁻¹ p₁₂⁻¹ p₂₁ p₀₁ :=
|
||||
by induction s₁₁;exact ids
|
||||
|
||||
|
@ -94,7 +98,7 @@ namespace eq
|
|||
definition transpose [unfold 10] (s₁₁ : square p₁₀ p₁₂ p₀₁ p₂₁) : square p₀₁ p₂₁ p₁₀ p₁₂ :=
|
||||
by induction s₁₁;exact ids
|
||||
|
||||
definition aps [unfold 12] {B : Type} (f : A → B) (s₁₁ : square p₁₀ p₁₂ p₀₁ p₂₁)
|
||||
definition aps [unfold 12] (f : A → B) (s₁₁ : square p₁₀ p₁₂ p₀₁ p₂₁)
|
||||
: square (ap f p₁₀) (ap f p₁₂) (ap f p₀₁) (ap f p₂₁) :=
|
||||
by induction s₁₁;exact ids
|
||||
|
||||
|
@ -541,15 +545,15 @@ namespace eq
|
|||
|
||||
/- Squares having an 'ap' term on one face -/
|
||||
--TODO find better names
|
||||
definition square_Flr_ap_idp {A B : Type} {c : B} {f : A → B} (p : Π a, f a = c)
|
||||
definition square_Flr_ap_idp {c : B} {f : A → B} (p : Π a, f a = c)
|
||||
{a b : A} (q : a = b) : square (p a) (p b) (ap f q) idp :=
|
||||
by induction q; apply vrfl
|
||||
|
||||
definition square_Flr_idp_ap {A B : Type} {c : B} {f : A → B} (p : Π a, c = f a)
|
||||
definition square_Flr_idp_ap {c : B} {f : A → B} (p : Π a, c = f a)
|
||||
{a b : A} (q : a = b) : square (p a) (p b) idp (ap f q) :=
|
||||
by induction q; apply vrfl
|
||||
|
||||
definition square_ap_idp_Flr {A B : Type} {b : B} {f : A → B} (p : Π a, f a = b)
|
||||
definition square_ap_idp_Flr {b : B} {f : A → B} (p : Π a, f a = b)
|
||||
{a b : A} (q : a = b) : square (ap f q) idp (p a) (p b) :=
|
||||
by induction q; apply hrfl
|
||||
|
||||
|
@ -580,4 +584,35 @@ namespace eq
|
|||
s₁₁ ⬝vp r = s₁₁ ⬝v vdeg_square r :=
|
||||
by cases r; cases s₁₁; esimp
|
||||
|
||||
definition natural_square011 {A A' : Type} {B : A → Type}
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} (q : b =[p] b')
|
||||
{l r : Π⦃a⦄, B a → A'} (g : Π⦃a⦄ (b : B a), l b = r b)
|
||||
: square (apd011 l p q) (apd011 r p q) (g b) (g b') :=
|
||||
begin
|
||||
induction q, exact hrfl
|
||||
end
|
||||
|
||||
definition natural_square0111' {A A' : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} {q : b =[p] b'}
|
||||
{c : C b} {c' : C b'} (s : c =[apd011 C p q] c')
|
||||
{l r : Π⦃a⦄ {b : B a}, C b → A'}
|
||||
(g : Π⦃a⦄ {b : B a} (c : C b), l c = r c)
|
||||
: square (apd0111 l p q s) (apd0111 r p q s) (g c) (g c') :=
|
||||
begin
|
||||
induction q, esimp at s, induction s using idp_rec_on, exact hrfl
|
||||
end
|
||||
|
||||
-- this can be generalized a bit, by making the domain and codomain of k different, and also have
|
||||
-- a function at the RHS of s (similar to m)
|
||||
definition natural_square0111 {A A' : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} {q : b =[p] b'}
|
||||
{c : C b} {c' : C b'} (r : c =[apd011 C p q] c')
|
||||
{k : A → A} {l : Π⦃a⦄, B a → B (k a)} (m : Π⦃a⦄ {b : B a}, C b → C (l b))
|
||||
{f : Π⦃a⦄ {b : B a}, C b → A'}
|
||||
(s : Π⦃a⦄ {b : B a} (c : C b), f (m c) = f c)
|
||||
: square (apd0111 (λa b c, f (m c)) p q r) (apd0111 f p q r) (s c) (s c') :=
|
||||
begin
|
||||
induction q, esimp at r, induction r using idp_rec_on, exact hrfl
|
||||
end
|
||||
|
||||
end eq
|
||||
|
|
|
@ -8,7 +8,7 @@ Squareovers
|
|||
|
||||
import .square
|
||||
|
||||
open eq equiv is_equiv
|
||||
open eq equiv is_equiv sigma
|
||||
|
||||
namespace eq
|
||||
|
||||
|
@ -277,4 +277,24 @@ namespace eq
|
|||
: squareover B s₁₁ q₁₀ q₁₂ q₀₁ (change_path r⁻¹ q₂₁) :=
|
||||
by induction r; exact t
|
||||
|
||||
/- You can construct a square in a sigma-type by giving a squareover -/
|
||||
definition square_dpair_eq_dpair {a₀₀ a₂₀ a₀₂ a₂₂ : A}
|
||||
{p₁₀ : a₀₀ = a₂₀} {p₀₁ : a₀₀ = a₀₂} {p₂₁ : a₂₀ = a₂₂} {p₁₂ : a₀₂ = a₂₂}
|
||||
(s₁₁ : square p₁₀ p₁₂ p₀₁ p₂₁) {b₀₀ : B a₀₀} {b₂₀ : B a₂₀} {b₀₂ : B a₀₂} {b₂₂ : B a₂₂}
|
||||
{q₁₀ : b₀₀ =[p₁₀] b₂₀} {q₀₁ : b₀₀ =[p₀₁] b₀₂} {q₂₁ : b₂₀ =[p₂₁] b₂₂} {q₁₂ : b₀₂ =[p₁₂] b₂₂}
|
||||
(t₁₁ : squareover B s₁₁ q₁₀ q₁₂ q₀₁ q₂₁) :
|
||||
square (dpair_eq_dpair p₁₀ q₁₀) (dpair_eq_dpair p₁₂ q₁₂)
|
||||
(dpair_eq_dpair p₀₁ q₀₁) (dpair_eq_dpair p₂₁ q₂₁) :=
|
||||
by induction t₁₁; constructor
|
||||
|
||||
definition sigma_square {v₀₀ v₂₀ v₀₂ v₂₂ : Σa, B a}
|
||||
{p₁₀ : v₀₀ = v₂₀} {p₀₁ : v₀₀ = v₀₂} {p₂₁ : v₂₀ = v₂₂} {p₁₂ : v₀₂ = v₂₂}
|
||||
(s₁₁ : square p₁₀..1 p₁₂..1 p₀₁..1 p₂₁..1)
|
||||
(t₁₁ : squareover B s₁₁ p₁₀..2 p₁₂..2 p₀₁..2 p₂₁..2) : square p₁₀ p₁₂ p₀₁ p₂₁ :=
|
||||
begin
|
||||
induction v₀₀, induction v₂₀, induction v₀₂, induction v₂₂,
|
||||
rewrite [▸* at *, -sigma_eq_eta p₁₀, -sigma_eq_eta p₁₂, -sigma_eq_eta p₀₁, -sigma_eq_eta p₂₁],
|
||||
exact square_dpair_eq_dpair s₁₁ t₁₁
|
||||
end
|
||||
|
||||
end eq
|
||||
|
|
|
@ -49,7 +49,7 @@ parameters {A B : Type.{u}} (f g : A → B)
|
|||
|
||||
protected definition elim {P : Type} (P_i : B → P)
|
||||
(Pcp : Π(x : A), P_i (f x) = P_i (g x)) (y : coeq) : P :=
|
||||
rec P_i (λx, pathover_of_eq (Pcp x)) y
|
||||
rec P_i (λx, pathover_of_eq _ (Pcp x)) y
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (y : coeq) (P_i : B → P)
|
||||
(Pcp : Π(x : A), P_i (f x) = P_i (g x)) : P :=
|
||||
|
@ -123,12 +123,12 @@ section
|
|||
begin
|
||||
change
|
||||
quotient.rec P_i
|
||||
(λb b' r, coeq_rel.cases_on r (λx, pathover_of_eq (ua (Pcp x))))
|
||||
(λb b' r, coeq_rel.cases_on r (λx, pathover_of_eq _ (ua (Pcp x))))
|
||||
= quotient.rec P_i
|
||||
(λb b' r, pathover_of_eq (ua (coeq_rel.cases_on r Pcp))),
|
||||
(λb b' r, pathover_of_eq _ (ua (coeq_rel.cases_on r Pcp))),
|
||||
have H2 : Π⦃b b' : B⦄ (r : coeq_rel f g b b'),
|
||||
coeq_rel.cases_on r (λx, pathover_of_eq (ua (Pcp x)))
|
||||
= pathover_of_eq (ua (coeq_rel.cases_on r Pcp))
|
||||
coeq_rel.cases_on r (λx, pathover_of_eq _ (ua (Pcp x)))
|
||||
= pathover_of_eq _ (ua (coeq_rel.cases_on r Pcp))
|
||||
:> P_i b =[eq_of_rel (coeq_rel f g) r] P_i b',
|
||||
begin intros b b' r, cases r, reflexivity end,
|
||||
rewrite (eq_of_homotopy3 H2)
|
||||
|
|
|
@ -55,7 +55,7 @@ section
|
|||
|
||||
protected definition elim {P : Type} (Pincl : Π⦃i : I⦄ (x : A i), P)
|
||||
(Pglue : Π(j : J) (x : A (dom j)), Pincl (f j x) = Pincl x) (y : colimit) : P :=
|
||||
rec Pincl (λj a, pathover_of_eq (Pglue j a)) y
|
||||
rec Pincl (λj a, pathover_of_eq _ (Pglue j a)) y
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (y : colimit)
|
||||
(Pincl : Π⦃i : I⦄ (x : A i), P)
|
||||
|
@ -146,7 +146,7 @@ section
|
|||
|
||||
protected definition elim {P : Type} (Pincl : Π⦃n : ℕ⦄ (a : A n), P)
|
||||
(Pglue : Π⦃n : ℕ⦄ (a : A n), Pincl (f a) = Pincl a) : seq_colim → P :=
|
||||
rec Pincl (λn a, pathover_of_eq (Pglue a))
|
||||
rec Pincl (λn a, pathover_of_eq _ (Pglue a))
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (aa : seq_colim)
|
||||
(Pincl : Π⦃n : ℕ⦄ (a : A n), P)
|
||||
|
@ -175,6 +175,11 @@ section
|
|||
: transport (elim_type Pincl Pglue) (glue a) = Pglue a :=
|
||||
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_glue]; apply cast_ua_fn
|
||||
|
||||
theorem elim_type_glue_inv (Pincl : Π⦃n : ℕ⦄ (a : A n), Type)
|
||||
(Pglue : Π⦃n : ℕ⦄ (a : A n), Pincl (f a) ≃ Pincl a) {n : ℕ} (a : A n)
|
||||
: transport (seq_colim.elim_type f Pincl Pglue) (glue a)⁻¹ = to_inv (Pglue a) :=
|
||||
by rewrite [tr_eq_cast_ap_fn, ↑seq_colim.elim_type, ap_inv, elim_glue]; apply cast_ua_inv_fn
|
||||
|
||||
protected definition rec_prop {P : seq_colim → Type} [H : Πx, is_prop (P x)]
|
||||
(Pincl : Π⦃n : ℕ⦄ (a : A n), P (sι a)) (aa : seq_colim) : P aa :=
|
||||
rec Pincl (λa b, !is_prop.elimo) aa
|
||||
|
|
|
@ -63,7 +63,7 @@ namespace one_step_tr
|
|||
|
||||
protected definition elim {P : Type} (Pt : A → P)
|
||||
(Pe : Π(a a' : A), Pt a = Pt a') (x : one_step_tr) : P :=
|
||||
rec Pt (λa a', pathover_of_eq (Pe a a')) x
|
||||
rec Pt (λa a', pathover_of_eq _ (Pe a a')) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : one_step_tr) (Pt : A → P)
|
||||
(Pe : Π(a a' : A), Pt a = Pt a') : P :=
|
||||
|
|
|
@ -56,7 +56,7 @@ parameters {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
|
|||
|
||||
protected definition elim {P : Type} (Pinl : BL → P) (Pinr : TR → P)
|
||||
(Pglue : Π(x : TL), Pinl (f x) = Pinr (g x)) (y : pushout) : P :=
|
||||
rec Pinl Pinr (λx, pathover_of_eq (Pglue x)) y
|
||||
rec Pinl Pinr (λx, pathover_of_eq _ (Pglue x)) y
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (y : pushout) (Pinl : BL → P)
|
||||
(Pinr : TR → P) (Pglue : Π(x : TL), Pinl (f x) = Pinr (g x)) : P :=
|
||||
|
|
|
@ -24,7 +24,7 @@ namespace quotient
|
|||
|
||||
protected definition elim {P : Type} (Pc : A → P) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a')
|
||||
(x : quotient R) : P :=
|
||||
quotient.rec Pc (λa a' H, pathover_of_eq (Pp H)) x
|
||||
quotient.rec Pc (λa a' H, pathover_of_eq _ (Pp H)) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : quotient R)
|
||||
(Pc : A → P) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') : P :=
|
||||
|
@ -178,8 +178,10 @@ namespace quotient
|
|||
(c : C a) : ap (elim @Qpt Qeq) (Peq r c) = Qeq r c :=
|
||||
begin
|
||||
refine !ap_dpair_eq_dpair ⬝ _,
|
||||
rewrite [apo011_eq_apo11_apd, rec_eq_of_rel, ▸*, apo011_arrow_pathover_constant_right,
|
||||
↑elim_type_eq_of_rel', to_right_inv !pathover_equiv_tr_eq, ap_inv],
|
||||
refine !apd011_eq_apo11_apd ⬝ _,
|
||||
rewrite [rec_eq_of_rel, ▸*],
|
||||
refine !apo11_arrow_pathover_constant_right ⬝ _,
|
||||
rewrite [↑elim_type_eq_of_rel', to_right_inv !pathover_equiv_tr_eq, ap_inv],
|
||||
apply inv_con_cancel_right
|
||||
end
|
||||
|
||||
|
|
|
@ -49,7 +49,7 @@ section
|
|||
|
||||
protected definition elim {P : Type} [Pt : is_set P] (Pc : A → P)
|
||||
(Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') (x : set_quotient) : P :=
|
||||
rec Pc (λa a' H, pathover_of_eq (Pp H)) x
|
||||
rec Pc (λa a' H, pathover_of_eq _ (Pp H)) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : set_quotient) [Pt : is_set P]
|
||||
(Pc : A → P) (Pp : Π⦃a a' : A⦄ (H : R a a'), Pc a = Pc a') : P :=
|
||||
|
|
|
@ -80,7 +80,7 @@ namespace simple_two_quotient
|
|||
protected definition pre_elim [unfold 8] {P : Type} (Pj : A → P)
|
||||
(Pa : Π⦃a : A⦄ ⦃r : T a a⦄, Q r → P) (Pe : Π⦃a a' : A⦄ (s : R a a'), Pj a = Pj a') (x : C)
|
||||
: P :=
|
||||
pre_rec Pj Pa (λa a' s, pathover_of_eq (Pe s)) x
|
||||
pre_rec Pj Pa (λa a' s, pathover_of_eq _ (Pe s)) x
|
||||
|
||||
protected theorem rec_e {P : C → Type}
|
||||
(Pj : Πa, P (j a)) (Pa : Π⦃a : A⦄ ⦃r : T a a⦄ (q : Q r), P (pre_aux q))
|
||||
|
|
|
@ -51,7 +51,7 @@ namespace circle
|
|||
!rec_merid
|
||||
|
||||
definition elim2 {P : Type} (Pb1 Pb2 : P) (Ps1 Ps2 : Pb1 = Pb2) (x : S¹) : P :=
|
||||
rec2 Pb1 Pb2 (pathover_of_eq Ps1) (pathover_of_eq Ps2) x
|
||||
rec2 Pb1 Pb2 (pathover_of_eq _ Ps1) (pathover_of_eq _ Ps2) x
|
||||
|
||||
definition elim2_on [reducible] {P : Type} (x : S¹) (Pb1 Pb2 : P)
|
||||
(Ps1 : Pb1 = Pb2) (Ps2 : Pb1 = Pb2) : P :=
|
||||
|
@ -117,7 +117,7 @@ namespace circle
|
|||
|
||||
protected definition elim {P : Type} (Pbase : P) (Ploop : Pbase = Pbase)
|
||||
(x : S¹) : P :=
|
||||
circle.rec Pbase (pathover_of_eq Ploop) x
|
||||
circle.rec Pbase (pathover_of_eq _ Ploop) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : S¹) (Pbase : P)
|
||||
(Ploop : Pbase = Pbase) : P :=
|
||||
|
@ -147,8 +147,8 @@ namespace circle
|
|||
rewrite [↑circle.rec2_on,rec2_seg2],
|
||||
assert l : Π(A B : Type)(a a₂ a₂' : A)(b b' : B)(p : a = a₂)(p' : a₂' = a₂)
|
||||
(q : b = b'),
|
||||
pathover_tr_of_pathover (pathover_of_eq q)
|
||||
= pathover_of_eq (q ⬝ (tr_constant p' b')⁻¹)
|
||||
pathover_tr_of_pathover (pathover_of_eq _ q)
|
||||
= pathover_of_eq _ (q ⬝ (tr_constant p' b')⁻¹)
|
||||
:> b =[p] p' ▸ b',
|
||||
{ intros, cases q, cases p', cases p, reflexivity },
|
||||
apply l
|
||||
|
|
|
@ -57,7 +57,7 @@ section
|
|||
|
||||
protected definition elim {P : Type} (Pbase : B → P) (Ptop : A → P)
|
||||
(Pseg : Π(a : A), Pbase (f a) = Ptop a) (x : cylinder) : P :=
|
||||
rec Pbase Ptop (λa, pathover_of_eq (Pseg a)) x
|
||||
rec Pbase Ptop (λa, pathover_of_eq _ (Pseg a)) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : cylinder) (Pbase : B → P) (Ptop : A → P)
|
||||
(Pseg : Π(a : A), Pbase (f a) = Ptop a) : P :=
|
||||
|
|
|
@ -35,7 +35,7 @@ namespace interval
|
|||
!rec_merid
|
||||
|
||||
protected definition elim {P : Type} (P0 P1 : P) (Ps : P0 = P1) (x : interval) : P :=
|
||||
interval.rec P0 P1 (pathover_of_eq Ps) x
|
||||
interval.rec P0 P1 (pathover_of_eq _ Ps) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : interval) (P0 P1 : P)
|
||||
(Ps : P0 = P1) : P :=
|
||||
|
|
|
@ -39,7 +39,7 @@ namespace join
|
|||
|
||||
protected definition elim {P : Type} (Pinl : A → P) (Pinr : B → P)
|
||||
(Pglue : Π(x : A)(y : B), Pinl x = Pinr y) (z : join A B) : P :=
|
||||
join.rec Pinl Pinr (λx y, pathover_of_eq (Pglue x y)) z
|
||||
join.rec Pinl Pinr (λx y, pathover_of_eq _ (Pglue x y)) z
|
||||
|
||||
protected definition elim_glue {P : Type} (Pinl : A → P) (Pinr : B → P)
|
||||
(Pglue : Π(x : A)(y : B), Pinl x = Pinr y) (x : A) (y : B)
|
||||
|
|
|
@ -44,7 +44,7 @@ namespace susp
|
|||
|
||||
protected definition elim {P : Type} (PN : P) (PS : P) (Pm : A → PN = PS)
|
||||
(x : susp A) : P :=
|
||||
susp.rec PN PS (λa, pathover_of_eq (Pm a)) x
|
||||
susp.rec PN PS (λa, pathover_of_eq _ (Pm a)) x
|
||||
|
||||
protected definition elim_on [reducible] {P : Type} (x : susp A)
|
||||
(PN : P) (PS : P) (Pm : A → PN = PS) : P :=
|
||||
|
|
|
@ -68,7 +68,7 @@ namespace is_equiv
|
|||
parameters {A B : Type} (f : A → B) (g : B → A)
|
||||
(ret : Πb, f (g b) = b) (sec : Πa, g (f a) = a)
|
||||
|
||||
definition adjointify_left_inv' (a : A) : g (f a) = a :=
|
||||
definition adjointify_left_inv' [unfold_full] (a : A) : g (f a) = a :=
|
||||
ap g (ap f (inverse (sec a))) ⬝ ap g (ret (f a)) ⬝ sec a
|
||||
|
||||
theorem adjointify_adj' (a : A) : ret (f a) = ap f (adjointify_left_inv' a) :=
|
||||
|
@ -276,6 +276,7 @@ namespace is_equiv
|
|||
= right_inv f (h' c) ⬝ ap h' (right_inv f c)⁻¹ ⬝ (p (f⁻¹ c))⁻¹ :=
|
||||
!ap_eq_of_fn_eq_fn'
|
||||
|
||||
-- inv_commute'_fn is in types.equiv
|
||||
end
|
||||
|
||||
-- This is inv_commute' for A ≡ unit
|
||||
|
|
|
@ -238,42 +238,65 @@ end funext
|
|||
|
||||
open funext
|
||||
|
||||
definition eq_equiv_homotopy : (f = g) ≃ (f ~ g) :=
|
||||
equiv.mk apd10 _
|
||||
namespace eq
|
||||
definition eq_equiv_homotopy : (f = g) ≃ (f ~ g) :=
|
||||
equiv.mk apd10 _
|
||||
|
||||
definition eq_of_homotopy [reducible] : f ~ g → f = g :=
|
||||
(@apd10 A P f g)⁻¹
|
||||
definition eq_of_homotopy [reducible] : f ~ g → f = g :=
|
||||
(@apd10 A P f g)⁻¹
|
||||
|
||||
definition apd10_eq_of_homotopy (p : f ~ g) : apd10 (eq_of_homotopy p) = p :=
|
||||
right_inv apd10 p
|
||||
definition apd10_eq_of_homotopy (p : f ~ g) : apd10 (eq_of_homotopy p) = p :=
|
||||
right_inv apd10 p
|
||||
|
||||
definition eq_of_homotopy_apd10 (p : f = g) : eq_of_homotopy (apd10 p) = p :=
|
||||
left_inv apd10 p
|
||||
definition eq_of_homotopy_apd10 (p : f = g) : eq_of_homotopy (apd10 p) = p :=
|
||||
left_inv apd10 p
|
||||
|
||||
definition eq_of_homotopy_idp (f : Π x, P x) : eq_of_homotopy (λx : A, idpath (f x)) = idpath f :=
|
||||
is_equiv.left_inv apd10 idp
|
||||
definition eq_of_homotopy_idp (f : Π x, P x) : eq_of_homotopy (λx : A, idpath (f x)) = idpath f :=
|
||||
is_equiv.left_inv apd10 idp
|
||||
|
||||
definition naive_funext_of_ua : naive_funext :=
|
||||
λ A P f g h, eq_of_homotopy h
|
||||
definition naive_funext_of_ua : naive_funext :=
|
||||
λ A P f g h, eq_of_homotopy h
|
||||
|
||||
protected definition homotopy.rec_on [recursor] {Q : (f ~ g) → Type} (p : f ~ g)
|
||||
(H : Π(q : f = g), Q (apd10 q)) : Q p :=
|
||||
right_inv apd10 p ▸ H (eq_of_homotopy p)
|
||||
protected definition homotopy.rec_on [recursor] {Q : (f ~ g) → Type} (p : f ~ g)
|
||||
(H : Π(q : f = g), Q (apd10 q)) : Q p :=
|
||||
right_inv apd10 p ▸ H (eq_of_homotopy p)
|
||||
|
||||
protected definition homotopy.rec_on_idp [recursor] {Q : Π{g}, (f ~ g) → Type} {g : Π x, P x} (p : f ~ g) (H : Q (homotopy.refl f)) : Q p :=
|
||||
homotopy.rec_on p (λq, eq.rec_on q H)
|
||||
protected definition homotopy.rec_on_idp [recursor] {Q : Π{g}, (f ~ g) → Type} {g : Π x, P x}
|
||||
(p : f ~ g) (H : Q (homotopy.refl f)) : Q p :=
|
||||
homotopy.rec_on p (λq, eq.rec_on q H)
|
||||
|
||||
definition eq_of_homotopy_inv {f g : Π x, P x} (H : f ~ g)
|
||||
: eq_of_homotopy (λx, (H x)⁻¹) = (eq_of_homotopy H)⁻¹ :=
|
||||
begin
|
||||
apply homotopy.rec_on_idp H,
|
||||
rewrite [+eq_of_homotopy_idp]
|
||||
end
|
||||
protected definition homotopy.rec_on' {f f' : Πa, P a} {Q : (f ~ f') → (f = f') → Type}
|
||||
(p : f ~ f') (H : Π(q : f = f'), Q (apd10 q) q) : Q p (eq_of_homotopy p) :=
|
||||
begin
|
||||
refine homotopy.rec_on p _,
|
||||
intro q, exact (left_inv (apd10) q)⁻¹ ▸ H q
|
||||
end
|
||||
|
||||
definition eq_of_homotopy_con {f g h : Π x, P x} (H1 : f ~ g) (H2 : g ~ h)
|
||||
: eq_of_homotopy (λx, H1 x ⬝ H2 x) = eq_of_homotopy H1 ⬝ eq_of_homotopy H2 :=
|
||||
begin
|
||||
apply homotopy.rec_on_idp H1,
|
||||
apply homotopy.rec_on_idp H2,
|
||||
rewrite [+eq_of_homotopy_idp]
|
||||
end
|
||||
protected definition homotopy.rec_on_idp' {f : Πa, P a} {Q : Π{g}, (f ~ g) → (f = g) → Type}
|
||||
{g : Πa, P a} (p : f ~ g) (H : Q (homotopy.refl f) idp) : Q p (eq_of_homotopy p) :=
|
||||
begin
|
||||
refine homotopy.rec_on' p _, intro q, induction q, exact H
|
||||
end
|
||||
|
||||
definition eq_of_homotopy_inv {f g : Π x, P x} (H : f ~ g)
|
||||
: eq_of_homotopy (λx, (H x)⁻¹) = (eq_of_homotopy H)⁻¹ :=
|
||||
begin
|
||||
apply homotopy.rec_on_idp H,
|
||||
rewrite [+eq_of_homotopy_idp]
|
||||
end
|
||||
|
||||
definition eq_of_homotopy_con {f g h : Π x, P x} (H1 : f ~ g) (H2 : g ~ h)
|
||||
: eq_of_homotopy (λx, H1 x ⬝ H2 x) = eq_of_homotopy H1 ⬝ eq_of_homotopy H2 :=
|
||||
begin
|
||||
apply homotopy.rec_on_idp H1,
|
||||
apply homotopy.rec_on_idp H2,
|
||||
rewrite [+eq_of_homotopy_idp]
|
||||
end
|
||||
|
||||
definition compose_eq_of_homotopy {A B C : Type} (g : B → C) {f f' : A → B} (H : f ~ f') :
|
||||
ap (λf, g ∘ f) (eq_of_homotopy H) = eq_of_homotopy (hwhisker_left g H) :=
|
||||
begin
|
||||
refine homotopy.rec_on_idp' H _, exact !eq_of_homotopy_idp⁻¹
|
||||
end
|
||||
|
||||
end eq
|
||||
|
|
|
@ -104,11 +104,12 @@ namespace eq
|
|||
by induction q; induction p; reflexivity
|
||||
|
||||
-- Inverse is an involution.
|
||||
definition inv_inv (p : x = y) : p⁻¹⁻¹ = p :=
|
||||
definition inv_inv [unfold 4] (p : x = y) : p⁻¹⁻¹ = p :=
|
||||
by induction p; reflexivity
|
||||
|
||||
-- auxiliary definition used by 'cases' tactic
|
||||
definition elim_inv_inv {A : Type} {a b : A} {C : a = b → Type} (H₁ : a = b) (H₂ : C (H₁⁻¹⁻¹)) : C H₁ :=
|
||||
definition elim_inv_inv [unfold 5] {A : Type} {a b : A} {C : a = b → Type}
|
||||
(H₁ : a = b) (H₂ : C (H₁⁻¹⁻¹)) : C H₁ :=
|
||||
eq.rec_on (inv_inv H₁) H₂
|
||||
|
||||
/- Theorems for moving things around in equations -/
|
||||
|
@ -241,6 +242,14 @@ namespace eq
|
|||
(H1 : f ~ g) (H2 : g ~ h) : f ~ h :=
|
||||
λ x, H1 x ⬝ H2 x
|
||||
|
||||
definition hwhisker_left [unfold_full] (g : B → C) {f f' : A → B} (H : f ~ f') :
|
||||
g ∘ f ~ g ∘ f' :=
|
||||
λa, ap g (H a)
|
||||
|
||||
definition hwhisker_right [unfold_full] (f : A → B) {g g' : B → C} (H : g ~ g') :
|
||||
g ∘ f ~ g' ∘ f :=
|
||||
λa, H (f a)
|
||||
|
||||
definition homotopy_of_eq {f g : Πx, P x} (H1 : f = g) : f ~ g :=
|
||||
H1 ▸ homotopy.refl f
|
||||
|
||||
|
|
|
@ -70,9 +70,6 @@ namespace eq
|
|||
definition inverseo [unfold 8] (r : b =[p] b₂) : b₂ =[p⁻¹] b :=
|
||||
by induction r; constructor
|
||||
|
||||
definition apd [unfold 6] (f : Πa, B a) (p : a = a₂) : f a =[p] f a₂ :=
|
||||
by induction p; constructor
|
||||
|
||||
definition concato_eq [unfold 10] (r : b =[p] b₂) (q : b₂ = b₂') : b =[p] b₂' :=
|
||||
by induction q; exact r
|
||||
|
||||
|
@ -126,20 +123,20 @@ namespace eq
|
|||
definition eq_of_pathover {a' a₂' : A'} (q : a' =[p] a₂') : a' = a₂' :=
|
||||
by cases q;reflexivity
|
||||
|
||||
definition pathover_of_eq [unfold 5 8] {a' a₂' : A'} (q : a' = a₂') : a' =[p] a₂' :=
|
||||
definition pathover_of_eq [unfold 5 8] (p : a = a₂) {a' a₂' : A'} (q : a' = a₂') : a' =[p] a₂' :=
|
||||
by cases p;cases q;constructor
|
||||
|
||||
definition pathover_constant [constructor] (p : a = a₂) (a' a₂' : A') : a' =[p] a₂' ≃ a' = a₂' :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ exact eq_of_pathover},
|
||||
{ exact pathover_of_eq},
|
||||
{ exact pathover_of_eq p},
|
||||
{ intro r, cases p, cases r, reflexivity},
|
||||
{ intro r, cases r, reflexivity},
|
||||
end
|
||||
|
||||
definition pathover_of_eq_tr_constant_inv (p : a = a₂) (a' : A')
|
||||
: pathover_of_eq (tr_constant p a')⁻¹ = pathover_tr p a' :=
|
||||
: pathover_of_eq p (tr_constant p a')⁻¹ = pathover_tr p a' :=
|
||||
by cases p; constructor
|
||||
|
||||
definition eq_of_pathover_idp [unfold 6] {b' : B a} (q : b =[idpath a] b') : b = b' :=
|
||||
|
@ -155,6 +152,9 @@ namespace eq
|
|||
(to_right_inv !pathover_equiv_tr_eq)
|
||||
(to_left_inv !pathover_equiv_tr_eq)
|
||||
|
||||
definition eq_of_pathover_idp_pathover_of_eq {A X : Type} (x : X) {a a' : A} (p : a = a') :
|
||||
eq_of_pathover_idp (pathover_of_eq (idpath x) p) = p :=
|
||||
by induction p; reflexivity
|
||||
|
||||
-- definition pathover_idp (b : B a) (b' : B a) : b =[idpath a] b' ≃ b = b' :=
|
||||
-- pathover_equiv_tr_eq idp b b'
|
||||
|
@ -165,7 +165,7 @@ namespace eq
|
|||
-- definition pathover_idp_of_eq [reducible] {b' : B a} (q : b = b') : b =[idpath a] b' :=
|
||||
-- to_inv !pathover_idp q
|
||||
|
||||
definition idp_rec_on [recursor] {P : Π⦃b₂ : B a⦄, b =[idpath a] b₂ → Type}
|
||||
definition idp_rec_on [recursor] [unfold 7] {P : Π⦃b₂ : B a⦄, b =[idpath a] b₂ → Type}
|
||||
{b₂ : B a} (r : b =[idpath a] b₂) (H : P idpo) : P r :=
|
||||
have H2 : P (pathover_idp_of_eq (eq_of_pathover_idp r)), from
|
||||
eq.rec_on (eq_of_pathover_idp r) H,
|
||||
|
@ -198,17 +198,6 @@ namespace eq
|
|||
{ intro q, cases q, reflexivity},
|
||||
end
|
||||
|
||||
definition apd_con (f : Πa, B a) (p : a = a₂) (q : a₂ = a₃)
|
||||
: apd f (p ⬝ q) = apd f p ⬝o apd f q :=
|
||||
by cases p; cases q; reflexivity
|
||||
|
||||
definition apd_inv (f : Πa, B a) (p : a = a₂) : apd f p⁻¹ = (apd f p)⁻¹ᵒ :=
|
||||
by cases p; reflexivity
|
||||
|
||||
definition apd_eq_pathover_of_eq_ap (f : A → A') (p : a = a₂) :
|
||||
apd f p = pathover_of_eq (ap f p) :=
|
||||
eq.rec_on p idp
|
||||
|
||||
definition pathover_of_pathover_tr (q : b =[p ⬝ p₂] p₂ ▸ b₂) : b =[p] b₂ :=
|
||||
pathover_cancel_right q !pathover_tr⁻¹ᵒ
|
||||
|
||||
|
@ -232,29 +221,33 @@ namespace eq
|
|||
by induction r;exact c
|
||||
infix ` ▸o `:75 := transporto _
|
||||
|
||||
definition fn_tro_eq_tro_fn (C' : Π ⦃a : A⦄, B a → Type) (q : b =[p] b₂)
|
||||
(f : Π(b : B a), C b → C' b) (c : C b) : f b (q ▸o c) = (q ▸o (f b c)) :=
|
||||
by induction q;reflexivity
|
||||
definition fn_tro_eq_tro_fn {C' : Π ⦃a : A⦄, B a → Type} (q : b =[p] b₂)
|
||||
(f : Π⦃a : A⦄ (b : B a), C b → C' b) (c : C b) : f b₂ (q ▸o c) = q ▸o (f b c) :=
|
||||
by induction q; reflexivity
|
||||
variable {C}
|
||||
|
||||
definition apo {f : A → A'} (g : Πa, B a → B'' (f a))
|
||||
(q : b =[p] b₂) : g a b =[p] g a₂ b₂ :=
|
||||
/- various variants of ap for pathovers -/
|
||||
definition apd [unfold 6] (f : Πa, B a) (p : a = a₂) : f a =[p] f a₂ :=
|
||||
by induction p; constructor
|
||||
|
||||
definition apo [unfold 12] {f : A → A'} (g : Πa, B a → B'' (f a)) (q : b =[p] b₂) :
|
||||
g a b =[p] g a₂ b₂ :=
|
||||
by induction q; constructor
|
||||
|
||||
definition apo011 [unfold 10] (f : Πa, B a → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
||||
definition apd011 [unfold 10] (f : Πa, B a → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
||||
: f a b = f a₂ b₂ :=
|
||||
by cases Hb; reflexivity
|
||||
|
||||
definition apo0111 (f : Πa b, C b → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
||||
(Hc : c =[apo011 C Ha Hb] c₂) : f a b c = f a₂ b₂ c₂ :=
|
||||
definition apd0111 [unfold 13 14] (f : Πa b, C b → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
||||
(Hc : c =[apd011 C Ha Hb] c₂) : f a b c = f a₂ b₂ c₂ :=
|
||||
by cases Hb; apply (idp_rec_on Hc); apply idp
|
||||
|
||||
definition apod11 {f : Πb, C b} {g : Πb₂, C b₂} (r : f =[p] g)
|
||||
{b : B a} {b₂ : B a₂} (q : b =[p] b₂) : f b =[apo011 C p q] g b₂ :=
|
||||
{b : B a} {b₂ : B a₂} (q : b =[p] b₂) : f b =[apd011 C p q] g b₂ :=
|
||||
by cases r; apply (idp_rec_on q); constructor
|
||||
|
||||
definition apdo10 {f : Πb, C b} {g : Πb₂, C b₂} (r : f =[p] g)
|
||||
(b : B a) : f b =[apo011 C p !pathover_tr] g (p ▸ b) :=
|
||||
(b : B a) : f b =[apd011 C p !pathover_tr] g (p ▸ b) :=
|
||||
by cases r; constructor
|
||||
|
||||
definition apo10 [unfold 9] {f : B a → B' a} {g : B a₂ → B' a₂} (r : f =[p] g)
|
||||
|
@ -273,6 +266,46 @@ namespace eq
|
|||
(q : b =[p] b₂) : f b =[p] g b₂ :=
|
||||
by induction q; exact apo10 r b
|
||||
|
||||
definition apdo011 {A : Type} {B : A → Type} {C : Π⦃a⦄, B a → Type}
|
||||
(f : Π⦃a⦄ (b : B a), C b) {a a' : A} (p : a = a') {b : B a} {b' : B a'} (q : b =[p] b')
|
||||
: f b =[apd011 C p q] f b' :=
|
||||
by cases q; constructor
|
||||
|
||||
definition apdo0111 {A : Type} {B : A → Type} {C C' : Π⦃a⦄, B a → Type}
|
||||
(f : Π⦃a⦄ {b : B a}, C b → C' b) {a a' : A} (p : a = a') {b : B a} {b' : B a'} (q : b =[p] b')
|
||||
{c : C b} {c' : C b'} (r : c =[apd011 C p q] c')
|
||||
: f c =[apd011 C' p q] f c' :=
|
||||
begin
|
||||
induction q, esimp at r, induction r using idp_rec_on, constructor
|
||||
end
|
||||
|
||||
definition apo11_constant_right [unfold 12] {f : B a → A'} {g : B a₂ → A'}
|
||||
(q : f =[p] g) (r : b =[p] b₂) : f b = g b₂ :=
|
||||
eq_of_pathover (apo11 q r)
|
||||
|
||||
/- properties about these "ap"s, transporto and pathover_ap -/
|
||||
definition apd_con (f : Πa, B a) (p : a = a₂) (q : a₂ = a₃)
|
||||
: apd f (p ⬝ q) = apd f p ⬝o apd f q :=
|
||||
by cases p; cases q; reflexivity
|
||||
|
||||
definition apd_inv (f : Πa, B a) (p : a = a₂) : apd f p⁻¹ = (apd f p)⁻¹ᵒ :=
|
||||
by cases p; reflexivity
|
||||
|
||||
definition apd_eq_pathover_of_eq_ap (f : A → A') (p : a = a₂) :
|
||||
apd f p = pathover_of_eq p (ap f p) :=
|
||||
eq.rec_on p idp
|
||||
|
||||
definition apo_invo (f : Πa, B a → B' a) {Ha : a = a₂} (Hb : b =[Ha] b₂)
|
||||
: (apo f Hb)⁻¹ᵒ = apo f Hb⁻¹ᵒ :=
|
||||
by induction Hb; reflexivity
|
||||
|
||||
definition apd011_inv (f : Πa, B a → A') (Ha : a = a₂) (Hb : b =[Ha] b₂)
|
||||
: (apd011 f Ha Hb)⁻¹ = (apd011 f Ha⁻¹ Hb⁻¹ᵒ) :=
|
||||
by induction Hb; reflexivity
|
||||
|
||||
definition cast_apd011 (q : b =[p] b₂) (c : C b) : cast (apd011 C p q) c = q ▸o c :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition apd_compose1 (g : Πa, B a → B' a) (f : Πa, B a) (p : a = a₂)
|
||||
: apd (g ∘' f) p = apo g (apd f p) :=
|
||||
by induction p; reflexivity
|
||||
|
@ -281,25 +314,79 @@ namespace eq
|
|||
: apd (λa, g (f a)) p = pathover_of_pathover_ap B'' f (apd g (ap f p)) :=
|
||||
by induction p; reflexivity
|
||||
|
||||
definition cono.right_inv_eq (q : b = b')
|
||||
: concato_eq (pathover_idp_of_eq q) q⁻¹ = (idpo : b =[refl a] b) :=
|
||||
definition apo_tro (C : Π⦃a⦄, B' a → Type) (f : Π⦃a⦄, B a → B' a) (q : b =[p] b₂)
|
||||
(c : C (f b)) : apo f q ▸o c = q ▸o c :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition pathover_ap_tro {B' : A' → Type} (C : Π⦃a'⦄, B' a' → Type) (f : A → A')
|
||||
{b : B' (f a)} {b₂ : B' (f a₂)} (q : b =[p] b₂) (c : C b) : pathover_ap B' f q ▸o c = q ▸o c :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition pathover_ap_invo_tro {B' : A' → Type} (C : Π⦃a'⦄, B' a' → Type) (f : A → A')
|
||||
{b : B' (f a)} {b₂ : B' (f a₂)} (q : b =[p] b₂) (c : C b₂)
|
||||
: (pathover_ap B' f q)⁻¹ᵒ ▸o c = q⁻¹ᵒ ▸o c :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition pathover_tro (q : b =[p] b₂) (c : C b) : c =[apd011 C p q] q ▸o c :=
|
||||
by induction q; constructor
|
||||
|
||||
definition pathover_ap_invo {B' : A' → Type} (f : A → A') {p : a = a₂}
|
||||
{b : B' (f a)} {b₂ : B' (f a₂)} (q : b =[p] b₂)
|
||||
: pathover_ap B' f q⁻¹ᵒ =[ap_inv f p] (pathover_ap B' f q)⁻¹ᵒ :=
|
||||
by induction q; exact idpo
|
||||
|
||||
definition tro_invo_tro {A : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} (q : b =[p] b') (c : C b') :
|
||||
q ▸o (q⁻¹ᵒ ▸o c) = c :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition invo_tro_tro {A : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} (q : b =[p] b') (c : C b) :
|
||||
q⁻¹ᵒ ▸o (q ▸o c) = c :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition cono_tro {A : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a₁ a₂ a₃ : A} {p₁ : a₁ = a₂} {p₂ : a₂ = a₃} {b₁ : B a₁} {b₂ : B a₂} {b₃ : B a₃}
|
||||
(q₁ : b₁ =[p₁] b₂) (q₂ : b₂ =[p₂] b₃) (c : C b₁) :
|
||||
transporto C (q₁ ⬝o q₂) c = transporto C q₂ (transporto C q₁ c) :=
|
||||
by induction q₂; reflexivity
|
||||
|
||||
definition is_equiv_transporto [constructor] {A : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} (q : b =[p] b') : is_equiv (transporto C q) :=
|
||||
begin
|
||||
fapply adjointify,
|
||||
{ exact transporto C q⁻¹ᵒ},
|
||||
{ exact tro_invo_tro C q},
|
||||
{ exact invo_tro_tro C q}
|
||||
end
|
||||
|
||||
definition equiv_apd011 [constructor] {A : Type} {B : A → Type} (C : Π⦃a⦄, B a → Type)
|
||||
{a a' : A} {p : a = a'} {b : B a} {b' : B a'} (q : b =[p] b') : C b ≃ C b' :=
|
||||
equiv.mk (transporto C q) !is_equiv_transporto
|
||||
|
||||
/- some cancellation laws for concato_eq an variants -/
|
||||
|
||||
definition cono.right_inv_eq (q : b = b') :
|
||||
pathover_idp_of_eq q ⬝op q⁻¹ = (idpo : b =[refl a] b) :=
|
||||
by induction q;constructor
|
||||
|
||||
definition cono.right_inv_eq' (q : b = b')
|
||||
: eq_concato q (pathover_idp_of_eq q⁻¹) = (idpo : b =[refl a] b) :=
|
||||
definition cono.right_inv_eq' (q : b = b') :
|
||||
q ⬝po (pathover_idp_of_eq q⁻¹) = (idpo : b =[refl a] b) :=
|
||||
by induction q;constructor
|
||||
|
||||
definition cono.left_inv_eq (q : b = b')
|
||||
: concato_eq (pathover_idp_of_eq q⁻¹) q = (idpo : b' =[refl a] b') :=
|
||||
definition cono.left_inv_eq (q : b = b') :
|
||||
pathover_idp_of_eq q⁻¹ ⬝op q = (idpo : b' =[refl a] b') :=
|
||||
by induction q;constructor
|
||||
|
||||
definition cono.left_inv_eq' (q : b = b')
|
||||
: eq_concato q⁻¹ (pathover_idp_of_eq q) = (idpo : b' =[refl a] b') :=
|
||||
definition cono.left_inv_eq' (q : b = b') :
|
||||
q⁻¹ ⬝po pathover_idp_of_eq q = (idpo : b' =[refl a] b') :=
|
||||
by induction q;constructor
|
||||
|
||||
definition pathover_of_fn_pathover_fn (f : Π{a}, B a ≃ B' a) (r : f b =[p] f b₂) : b =[p] b₂ :=
|
||||
(left_inv f b)⁻¹ ⬝po apo (λa, f⁻¹ᵉ) r ⬝op left_inv f b₂
|
||||
|
||||
/- a pathover in a pathover type where the only thing which varies is the path is the same as
|
||||
an equality with a change_path -/
|
||||
definition change_path_of_pathover (s : p = p') (r : b =[p] b₂) (r' : b =[p'] b₂)
|
||||
(q : r =[s] r') : change_path s r = r' :=
|
||||
by induction s; eapply idp_rec_on q; reflexivity
|
||||
|
@ -318,6 +405,7 @@ namespace eq
|
|||
{ intro q, induction s, eapply idp_rec_on q, reflexivity},
|
||||
end
|
||||
|
||||
/- variants of inverse2 and concat2 -/
|
||||
definition inverseo2 [unfold 10] {r r' : b =[p] b₂} (s : r = r') : r⁻¹ᵒ = r'⁻¹ᵒ :=
|
||||
by induction s; reflexivity
|
||||
|
||||
|
@ -328,4 +416,19 @@ namespace eq
|
|||
infixl ` ◾o `:75 := concato2
|
||||
postfix [parsing_only] `⁻²ᵒ`:(max+10) := inverseo2 --this notation is abusive, should we use it?
|
||||
|
||||
-- find a better name for this
|
||||
definition fn_tro_eq_tro_fn2 (q : b =[p] b₂)
|
||||
{k : A → A} {l : Π⦃a⦄, B a → B (k a)} (m : Π⦃a⦄ {b : B a}, C b → C (l b))
|
||||
(c : C b) :
|
||||
m (q ▸o c) = (pathover_ap B k (apo l q)) ▸o (m c) :=
|
||||
by induction q; reflexivity
|
||||
|
||||
definition apd0111_precompose (f : Π⦃a⦄ {b : B a}, C b → A')
|
||||
{k : A → A} {l : Π⦃a⦄, B a → B (k a)} (m : Π⦃a⦄ {b : B a}, C b → C (l b))
|
||||
{q : b =[p] b₂} (c : C b)
|
||||
: apd0111 (λa b c, f (m c)) p q (pathover_tro q c) ⬝ ap (@f _ _) (fn_tro_eq_tro_fn2 q m c) =
|
||||
apd0111 f (ap k p) (pathover_ap B k (apo l q)) (pathover_tro _ (m c)) :=
|
||||
by induction q; reflexivity
|
||||
|
||||
|
||||
end eq
|
||||
|
|
|
@ -122,16 +122,20 @@ namespace pi
|
|||
definition arrow_pathover_constant_right' {B : A → Type} {C : Type}
|
||||
{f : B a → C} {g : B a' → C} {p : a = a'}
|
||||
(r : Π⦃b : B a⦄ ⦃b' : B a'⦄ (q : b =[p] b'), f b = g b') : f =[p] g :=
|
||||
arrow_pathover (λb b' q, pathover_of_eq (r q))
|
||||
arrow_pathover (λb b' q, pathover_of_eq p (r q))
|
||||
|
||||
definition arrow_pathover_constant_right {B : A → Type} {C : Type} {f : B a → C}
|
||||
{g : B a' → C} {p : a = a'} (r : Π(b : B a), f b = g (p ▸ b)) : f =[p] g :=
|
||||
arrow_pathover_left (λb, pathover_of_eq (r b))
|
||||
arrow_pathover_left (λb, pathover_of_eq p (r b))
|
||||
|
||||
/- a lemma used for the flattening lemma -/
|
||||
definition apo011_arrow_pathover_constant_right {f : D a → A'} {g : D a' → A'} {p : a = a'}
|
||||
definition arrow_pathover_constant_right_rev {A : Type} {B : A → Type} {C : Type} {a a' : A}
|
||||
{f : B a → C} {g : B a' → C} {p : a = a'} (r : Π(b : B a'), f (p⁻¹ ▸ b) = g b) : f =[p] g :=
|
||||
arrow_pathover_right (λb, pathover_of_eq p (r b))
|
||||
|
||||
/- a lemma used for the flattening lemma, and is also used in the colimits file -/
|
||||
definition apo11_arrow_pathover_constant_right {f : D a → A'} {g : D a' → A'} {p : a = a'}
|
||||
{q : d =[p] d'} (r : Π(d : D a), f d = g (p ▸ d))
|
||||
: eq_of_pathover (apo11 (arrow_pathover_constant_right r) q) = r d ⬝ ap g (tr_eq_of_pathover q)
|
||||
: apo11_constant_right (arrow_pathover_constant_right r) q = r d ⬝ ap g (tr_eq_of_pathover q)
|
||||
:=
|
||||
begin
|
||||
induction q, esimp at r,
|
||||
|
@ -140,7 +144,6 @@ namespace pi
|
|||
rewrite [eq_of_homotopy_idp]
|
||||
end
|
||||
|
||||
|
||||
/-
|
||||
The fact that the arrow type preserves truncation level is a direct consequence
|
||||
of the fact that pi's preserve truncation level
|
||||
|
|
|
@ -61,8 +61,9 @@ namespace eq
|
|||
: whisker_right r⁻¹ q = (whisker_right r q)⁻¹ :=
|
||||
by induction r; reflexivity
|
||||
|
||||
theorem ap_eq_ap10 {f g : A → B} (p : f = g) (a : A) : ap (λh, h a) p = ap10 p a :=
|
||||
by induction p;reflexivity
|
||||
theorem ap_eq_apd10 {B : A → Type} {f g : Πa, B a} (p : f = g) (a : A) :
|
||||
ap (λh, h a) p = apd10 p a :=
|
||||
by induction p; reflexivity
|
||||
|
||||
theorem inverse2_right_inv (r : p = p') : r ◾ inverse2 r ⬝ con.right_inv p' = con.right_inv p :=
|
||||
by induction r;induction p;reflexivity
|
||||
|
|
|
@ -81,7 +81,7 @@ namespace is_equiv
|
|||
|
||||
definition inv_eq_inv {A B : Type} {f f' : A → B} {Hf : is_equiv f} {Hf' : is_equiv f'}
|
||||
(p : f = f') : f⁻¹ = f'⁻¹ :=
|
||||
apd011 inv p !is_prop.elim
|
||||
apd011 inv p !is_prop.elimo
|
||||
|
||||
/- contractible fibers -/
|
||||
definition is_contr_fun_of_is_equiv [H : is_equiv f] : is_contr_fun f :=
|
||||
|
@ -100,6 +100,20 @@ namespace is_equiv
|
|||
definition is_equiv_equiv_is_contr_fun : is_equiv f ≃ is_contr_fun f :=
|
||||
equiv_of_is_prop _ (λH, !is_equiv_of_is_contr_fun)
|
||||
|
||||
theorem inv_commute'_fn {A : Type} {B C : A → Type} (f : Π{a}, B a → C a) [H : Πa, is_equiv (@f a)]
|
||||
{g : A → A} (h : Π{a}, B a → B (g a)) (h' : Π{a}, C a → C (g a))
|
||||
(p : Π⦃a : A⦄ (b : B a), f (h b) = h' (f b)) {a : A} (b : B a) :
|
||||
inv_commute' @f @h @h' p (f b)
|
||||
= (ap f⁻¹ (p b))⁻¹ ⬝ left_inv f (h b) ⬝ (ap h (left_inv f b))⁻¹ :=
|
||||
begin
|
||||
rewrite [↑[inv_commute',eq_of_fn_eq_fn'],+ap_con,-adj_inv f,+con.assoc,inv_con_cancel_left,
|
||||
adj f,+ap_inv,-+ap_compose,
|
||||
eq_bot_of_square (natural_square (λb, (left_inv f (h b))⁻¹ ⬝ ap f⁻¹ (p b)) (left_inv f b))⁻¹ʰ,
|
||||
con_inv,inv_inv,+con.assoc],
|
||||
do 3 apply whisker_left,
|
||||
rewrite [con_inv_cancel_left,con.left_inv]
|
||||
end
|
||||
|
||||
end is_equiv
|
||||
|
||||
/- Moving equivalences around in homotopies -/
|
||||
|
@ -206,7 +220,7 @@ namespace equiv
|
|||
|
||||
definition equiv_mk_eq {f f' : A → B} [H : is_equiv f] [H' : is_equiv f'] (p : f = f')
|
||||
: equiv.mk f H = equiv.mk f' H' :=
|
||||
apd011 equiv.mk p !is_prop.elim
|
||||
apd011 equiv.mk p !is_prop.elimo
|
||||
|
||||
definition equiv_eq' {f f' : A ≃ B} (p : to_fun f = to_fun f') : f = f' :=
|
||||
by (cases f; cases f'; apply (equiv_mk_eq p))
|
||||
|
|
|
@ -122,7 +122,7 @@ definition mk_mod [reducible] (n i : nat) : fin (succ n) :=
|
|||
mk (i % (succ n)) (mod_lt _ !zero_lt_succ)
|
||||
|
||||
theorem mk_mod_zero_eq (n : nat) : mk_mod n 0 = 0 :=
|
||||
apd011 fin.mk rfl !is_prop.elim
|
||||
apd011 fin.mk rfl !is_prop.elimo
|
||||
|
||||
variable {n : nat}
|
||||
|
||||
|
|
|
@ -137,7 +137,7 @@ rfl
|
|||
|
||||
theorem last_congr {l₁ l₂ : list T} (h₁ : l₁ ≠ []) (h₂ : l₂ ≠ []) (h₃ : l₁ = l₂)
|
||||
: last l₁ h₁ = last l₂ h₂ :=
|
||||
apd011 last h₃ !is_prop.elim
|
||||
apd011 last h₃ !is_prop.elimo
|
||||
|
||||
theorem last_concat {x : T} : ∀ {l : list T} (h : concat x l ≠ []), last (concat x l) h = x
|
||||
| [] h := rfl
|
||||
|
|
|
@ -7,9 +7,9 @@ Partially ported from Coq HoTT
|
|||
Theorems about pi-types (dependent function spaces)
|
||||
-/
|
||||
|
||||
import types.sigma arity
|
||||
import types.sigma arity cubical.square
|
||||
|
||||
open eq equiv is_equiv funext sigma unit bool is_trunc prod
|
||||
open eq equiv is_equiv funext sigma unit bool is_trunc prod function sigma.ops
|
||||
|
||||
namespace pi
|
||||
variables {A A' : Type} {B : A → Type} {B' : A' → Type} {C : Πa, B a → Type}
|
||||
|
@ -80,7 +80,7 @@ namespace pi
|
|||
/- Pathovers -/
|
||||
|
||||
definition pi_pathover {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'}
|
||||
(r : Π(b : B a) (b' : B a') (q : b =[p] b'), f b =[apo011 C p q] g b') : f =[p] g :=
|
||||
(r : Π(b : B a) (b' : B a') (q : b =[p] b'), f b =[apd011 C p q] g b') : f =[p] g :=
|
||||
begin
|
||||
cases p, apply pathover_idp_of_eq,
|
||||
apply eq_of_homotopy, intro b,
|
||||
|
@ -88,7 +88,7 @@ namespace pi
|
|||
end
|
||||
|
||||
definition pi_pathover_left {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'}
|
||||
(r : Π(b : B a), f b =[apo011 C p !pathover_tr] g (p ▸ b)) : f =[p] g :=
|
||||
(r : Π(b : B a), f b =[apd011 C p !pathover_tr] g (p ▸ b)) : f =[p] g :=
|
||||
begin
|
||||
cases p, apply pathover_idp_of_eq,
|
||||
apply eq_of_homotopy, intro b,
|
||||
|
@ -96,7 +96,7 @@ namespace pi
|
|||
end
|
||||
|
||||
definition pi_pathover_right {f : Πb, C a b} {g : Πb', C a' b'} {p : a = a'}
|
||||
(r : Π(b' : B a'), f (p⁻¹ ▸ b') =[apo011 C p !tr_pathover] g b') : f =[p] g :=
|
||||
(r : Π(b' : B a'), f (p⁻¹ ▸ b') =[apd011 C p !tr_pathover] g b') : f =[p] g :=
|
||||
begin
|
||||
cases p, apply pathover_idp_of_eq,
|
||||
apply eq_of_homotopy, intro b,
|
||||
|
@ -338,6 +338,39 @@ namespace pi
|
|||
definition pi_comm_equiv {P : A → A' → Type} : (Πa b, P a b) ≃ (Πb a, P a b) :=
|
||||
equiv.mk (@function.flip _ _ P) _
|
||||
|
||||
/- Dependent functions are equivalent to nondependent functions into the total space together
|
||||
with a homotopy -/
|
||||
definition pi_equiv_arrow_sigma_right [constructor] {A : Type} {B : A → Type} (f : Πa, B a) :
|
||||
Σ(f : A → Σa, B a), pr1 ∘ f ~ id :=
|
||||
⟨λa, ⟨a, f a⟩, λa, idp⟩
|
||||
|
||||
definition pi_equiv_arrow_sigma_left.{u v} [unfold 3] {A : Type.{u}} {B : A → Type.{v}}
|
||||
(v : Σ(f : A → Σa, B a), pr1 ∘ f ~ id) (a : A) : B a :=
|
||||
transport B (v.2 a) (v.1 a).2
|
||||
|
||||
open funext
|
||||
definition pi_equiv_arrow_sigma [constructor] {A : Type} (B : A → Type) :
|
||||
(Πa, B a) ≃ Σ(f : A → Σa, B a), pr1 ∘ f ~ id :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ exact pi_equiv_arrow_sigma_right},
|
||||
{ exact pi_equiv_arrow_sigma_left},
|
||||
{ intro v, induction v with f p, fapply sigma_eq: esimp,
|
||||
{ apply eq_of_homotopy, intro a, fapply sigma_eq: esimp,
|
||||
{ exact (p a)⁻¹},
|
||||
{ apply inverseo, apply pathover_tr}},
|
||||
{ apply pi_pathover_constant, intro a, apply eq_pathover_constant_right,
|
||||
refine ap_compose (λf, f a) _ _ ⬝ph _,
|
||||
refine ap02 _ !compose_eq_of_homotopy ⬝ph _,
|
||||
refine !ap_eq_apd10 ⬝ph _,
|
||||
refine apd10 (right_inv apd10 _) a ⬝ph _,
|
||||
esimp, refine !sigma_eq_pr1 ⬝ph _, apply square_of_eq, exact !con.left_inv⁻¹}},
|
||||
{ intro a, reflexivity}
|
||||
end
|
||||
|
||||
|
||||
|
||||
|
||||
end pi
|
||||
|
||||
attribute pi.is_trunc_pi [instance] [priority 1520]
|
||||
|
|
|
@ -52,9 +52,9 @@ namespace pointed
|
|||
definition pType_eq {A B : Type*} (f : A ≃ B) (p : f pt = pt) : A = B :=
|
||||
begin
|
||||
cases A with A a, cases B with B b, esimp at *,
|
||||
fapply apd011 @pType.mk,
|
||||
fapply apdt011 @pType.mk,
|
||||
{ apply ua f},
|
||||
{ rewrite [cast_ua,p]},
|
||||
{ rewrite [cast_ua, p]},
|
||||
end
|
||||
|
||||
definition pType_eq_elim {A B : Type*} (p : A = B :> Type*)
|
||||
|
@ -193,10 +193,10 @@ namespace pointed
|
|||
begin
|
||||
cases f with f p, cases g with g q,
|
||||
esimp at *,
|
||||
fapply apo011 pmap.mk,
|
||||
fapply apd011 pmap.mk,
|
||||
{ exact eq_of_homotopy r},
|
||||
{ apply concato_eq, apply pathover_eq_Fl, apply inv_con_eq_of_eq_con,
|
||||
rewrite [ap_eq_ap10,↑ap10,apd10_eq_of_homotopy,s]}
|
||||
rewrite [ap_eq_apd10, apd10_eq_of_homotopy, s]}
|
||||
end
|
||||
|
||||
definition pmap_equiv_left (A : Type) (B : Type*) : A₊ →* B ≃ (A → B) :=
|
||||
|
@ -554,7 +554,7 @@ namespace pointed
|
|||
(gf : g ∘* f ~* !pid) (fg : f ∘* g ~* !pid) : to_pinv (pequiv.MK2 f g gf fg) ~* g :=
|
||||
phomotopy.mk (λb, idp)
|
||||
abstract [irreducible] begin
|
||||
esimp, unfold [adjointify_left_inv'],
|
||||
esimp,
|
||||
note H := to_homotopy_pt gf, note H2 := to_homotopy_pt fg,
|
||||
note H3 := eq_top_of_square (natural_square_tr (to_homotopy fg) (respect_pt f)),
|
||||
rewrite [▸* at *, H, H3, H2, ap_id, - +con.assoc, ap_compose' f g, con_inv,
|
||||
|
@ -614,7 +614,7 @@ namespace pointed
|
|||
definition pequiv_eq {p q : A ≃* B} (H : p = q :> (A →* B)) : p = q :=
|
||||
begin
|
||||
cases p with f Hf, cases q with g Hg, esimp at *,
|
||||
exact apd011 pequiv_of_pmap H !is_prop.elim
|
||||
exact apd011 pequiv_of_pmap H !is_prop.elimo
|
||||
end
|
||||
|
||||
infix ` ⬝e*p `:75 := peconcat_eq
|
||||
|
|
|
@ -212,7 +212,7 @@ namespace prod
|
|||
definition equiv_prod_rec [constructor] (P : A × B → Type) : (Πa b, P (a, b)) ≃ (Πu, P u) :=
|
||||
equiv.mk prod.rec _
|
||||
|
||||
definition imp_imp_equiv_prod_imp (A B C : Type) : (A → B → C) ≃ (A × B → C) :=
|
||||
definition imp_imp_equiv_prod_imp [constructor] (A B C : Type) : (A → B → C) ≃ (A × B → C) :=
|
||||
!equiv_prod_rec
|
||||
|
||||
definition prod_corec_unc [unfold 4] {P Q : A → Type} (u : (Πa, P a) × (Πa, Q a)) (a : A)
|
||||
|
|
|
@ -30,7 +30,7 @@ namespace sigma
|
|||
| eta3 ⟨u₁, u₂, u₃, u₄⟩ := idp
|
||||
|
||||
definition dpair_eq_dpair [unfold 8] (p : a = a') (q : b =[p] b') : ⟨a, b⟩ = ⟨a', b'⟩ :=
|
||||
apo011 sigma.mk p q
|
||||
apd011 sigma.mk p q
|
||||
|
||||
definition sigma_eq [unfold 3 4] (p : u.1 = v.1) (q : u.2 =[p] v.2) : u = v :=
|
||||
by induction u; induction v; exact (dpair_eq_dpair p q)
|
||||
|
@ -159,7 +159,7 @@ namespace sigma
|
|||
destruct rs sigma_eq2
|
||||
|
||||
definition ap_dpair_eq_dpair (f : Πa, B a → A') (p : a = a') (q : b =[p] b')
|
||||
: ap (sigma.rec f) (dpair_eq_dpair p q) = apo011 f p q :=
|
||||
: ap (sigma.rec f) (dpair_eq_dpair p q) = apd011 f p q :=
|
||||
by induction q; reflexivity
|
||||
|
||||
/- Transport -/
|
||||
|
@ -192,10 +192,10 @@ namespace sigma
|
|||
by induction p; induction bc; apply idpo
|
||||
|
||||
definition sigma_pathover (p : a = a') (u : Σ(b : B a), C a b) (v : Σ(b : B a'), C a' b)
|
||||
(r : u.1 =[p] v.1) (s : u.2 =[apo011 C p r] v.2) : u =[p] v :=
|
||||
(r : u.1 =[p] v.1) (s : u.2 =[apd011 C p r] v.2) : u =[p] v :=
|
||||
begin
|
||||
induction u, induction v, esimp at *, induction r,
|
||||
esimp [apo011] at s, induction s using idp_rec_on, apply idpo
|
||||
esimp [apd011] at s, induction s using idp_rec_on, apply idpo
|
||||
end
|
||||
|
||||
/-
|
||||
|
@ -203,6 +203,7 @@ namespace sigma
|
|||
* define the projections from the type u =[p] v
|
||||
* show that the uncurried version of sigma_pathover is an equivalence
|
||||
-/
|
||||
/- Squares in a sigma type are characterized in cubical.squareover (to avoid circular imports) -/
|
||||
|
||||
/- Functorial action -/
|
||||
variables (f : A → A') (g : Πa, B a → B' (f a))
|
||||
|
|
|
@ -6,7 +6,7 @@ Authors: Floris van Doorn
|
|||
Theorems about the unit type
|
||||
-/
|
||||
|
||||
open equiv option eq pointed is_trunc
|
||||
open is_equiv equiv option eq pointed is_trunc function
|
||||
|
||||
namespace unit
|
||||
|
||||
|
@ -32,6 +32,20 @@ namespace unit
|
|||
|
||||
notation `unit*` := punit
|
||||
|
||||
definition unit_arrow_eq {X : Type} (f : unit → X) : (λx, f ⋆) = f :=
|
||||
by apply eq_of_homotopy; intro u; induction u; reflexivity
|
||||
|
||||
open funext
|
||||
definition unit_arrow_eq_compose {X Y : Type} (g : X → Y) (f : unit → X) :
|
||||
unit_arrow_eq (g ∘ f) = ap (λf, g ∘ f) (unit_arrow_eq f) :=
|
||||
begin
|
||||
apply eq_of_fn_eq_fn' apd10,
|
||||
refine right_inv apd10 _ ⬝ _,
|
||||
refine _ ⬝ ap apd10 (!compose_eq_of_homotopy)⁻¹,
|
||||
refine _ ⬝ (right_inv apd10 _)⁻¹,
|
||||
apply eq_of_homotopy, intro u, induction u, reflexivity
|
||||
end
|
||||
|
||||
end unit
|
||||
|
||||
open unit is_trunc
|
||||
|
|
Loading…
Reference in a new issue