move to lib and older things
This commit is contained in:
parent
ec5b9dba12
commit
be0d5977f6
5 changed files with 408 additions and 785 deletions
|
@ -326,6 +326,18 @@ namespace EM
|
|||
EMadd1_pequiv_succ_natural f n !isomorphism.refl !isomorphism.refl (π→g[n+2] f)
|
||||
proof λa, idp qed
|
||||
|
||||
definition EMadd1_pmap_equiv (n : ℕ) (X Y : Type*) [is_conn (n+1) X] [is_trunc (n+1+1) X]
|
||||
[is_conn (n+1) Y] [is_trunc (n+1+1) Y] : (X →* Y) ≃ πag[n+2] X →g πag[n+2] Y :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ exact π→g[n+2] },
|
||||
{ exact λφ, (EMadd1_pequiv_type Y n ∘* EMadd1_functor φ (n+1)) ∘* (EMadd1_pequiv_type X n)⁻¹ᵉ* },
|
||||
{ intro φ, apply homomorphism_eq,
|
||||
refine homotopy_group_functor_compose (n+2) _ _ ⬝hty _, exact sorry }, -- easy but tedious
|
||||
{ intro f, apply eq_of_phomotopy, refine (phomotopy_pinv_right_of_phomotopy _)⁻¹*,
|
||||
apply EMadd1_pequiv_type_natural }
|
||||
end
|
||||
|
||||
/- The Eilenberg-MacLane space K(G,n) with the same homotopy group as X on level n.
|
||||
On paper this is written K(πₙ(X), n). The problem is that for
|
||||
* n = 0 the expression π₀(X) is a pointed set, and K(X,0) needs X to be a pointed set
|
||||
|
|
|
@ -700,6 +700,19 @@ end
|
|||
apply eq_inv_con_of_con_eq, exact (to_homotopy_pt p)⁻¹ }
|
||||
end
|
||||
|
||||
definition pcofiber_elim_unique {X Y Z : Type*} {f : X →* Y}
|
||||
{g₁ g₂ : pcofiber f →* Z} (h : g₁ ∘* pcod f ~* g₂ ∘* pcod f)
|
||||
(sq : Πx, square (h (f x)) (respect_pt g₁ ⬝ (respect_pt g₂)⁻¹)
|
||||
(ap g₁ (cofiber.glue x)) (ap g₂ (cofiber.glue x))) : g₁ ~* g₂ :=
|
||||
begin
|
||||
fapply phomotopy.mk,
|
||||
{ intro x, induction x with y x,
|
||||
{ exact h y },
|
||||
{ exact respect_pt g₁ ⬝ (respect_pt g₂)⁻¹ },
|
||||
{ apply eq_pathover, exact sq x }},
|
||||
{ apply inv_con_cancel_right }
|
||||
end
|
||||
|
||||
/-
|
||||
The maps Z^{C_f} --> Z^Y --> Z^X are exact at Z^Y.
|
||||
Here Y^X means pointed maps from X to Y and C_f is the cofiber of f.
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
-- Authors: Floris van Doorn
|
||||
|
||||
import homotopy.wedge
|
||||
import homotopy.wedge homotopy.cofiber ..move_to_lib .pushout
|
||||
|
||||
open wedge pushout eq prod sum pointed equiv is_equiv unit lift bool option
|
||||
|
||||
|
@ -107,4 +107,164 @@ equiv.MK add_point_of_wedge_pbool wedge_pbool_of_add_point
|
|||
calc plift.{u v} (A ∨ B) ≃* A ∨ B : by exact !pequiv_plift⁻¹ᵉ*
|
||||
... ≃* plift.{u v} A ∨ plift.{u v} B : by exact wedge_pequiv !pequiv_plift !pequiv_plift
|
||||
|
||||
protected definition pelim [constructor] {X Y Z : Type*} (f : X →* Z) (g : Y →* Z) : X ∨ Y →* Z :=
|
||||
pmap.mk (wedge.elim f g (respect_pt f ⬝ (respect_pt g)⁻¹)) (respect_pt f)
|
||||
|
||||
definition wedge_pr1 [constructor] (X Y : Type*) : X ∨ Y →* X := wedge.pelim (pid X) (pconst Y X)
|
||||
definition wedge_pr2 [constructor] (X Y : Type*) : X ∨ Y →* Y := wedge.pelim (pconst X Y) (pid Y)
|
||||
|
||||
open fiber prod cofiber pi
|
||||
|
||||
variables {X Y : Type*}
|
||||
definition pcofiber_pprod_incl1_of_pfiber_wedge_pr2' [unfold 3]
|
||||
(x : pfiber (wedge_pr2 X Y)) : pcofiber (pprod_incl1 (Ω Y) X) :=
|
||||
begin
|
||||
induction x with x p, induction x with x y,
|
||||
{ exact cod _ (p, x) },
|
||||
{ exact pt },
|
||||
{ apply arrow_pathover_constant_right, intro p, apply cofiber.glue }
|
||||
end
|
||||
|
||||
--set_option pp.all true
|
||||
/-
|
||||
X : Type* has a nondegenerate basepoint iff
|
||||
it has the homotopy extension property iff
|
||||
Π(f : X → Y) (y : Y) (p : f pt = y), ∃(g : X → Y) (h : f ~ g) (q : y = g pt), h pt = p ⬝ q
|
||||
(or Σ?)
|
||||
-/
|
||||
|
||||
definition pfiber_wedge_pr2_of_pcofiber_pprod_incl1' [unfold 3]
|
||||
(x : pcofiber (pprod_incl1 (Ω Y) X)) : pfiber (wedge_pr2 X Y) :=
|
||||
begin
|
||||
induction x with v p,
|
||||
{ induction v with p x, exact fiber.mk (inl x) p },
|
||||
{ exact fiber.mk (inr pt) idp },
|
||||
{ esimp, apply fiber_eq (wedge.glue ⬝ ap inr p), symmetry,
|
||||
refine !ap_con ⬝ !wedge.elim_glue ◾ (!ap_compose'⁻¹ ⬝ !ap_id) ⬝ !idp_con }
|
||||
end
|
||||
|
||||
variables (X Y)
|
||||
definition pcofiber_pprod_incl1_of_pfiber_wedge_pr2 [constructor] :
|
||||
pfiber (wedge_pr2 X Y) →* pcofiber (pprod_incl1 (Ω Y) X) :=
|
||||
pmap.mk pcofiber_pprod_incl1_of_pfiber_wedge_pr2' (cofiber.glue idp)
|
||||
|
||||
-- definition pfiber_wedge_pr2_of_pprod [constructor] :
|
||||
-- Ω Y ×* X →* pfiber (wedge_pr2 X Y) :=
|
||||
-- begin
|
||||
-- fapply pmap.mk,
|
||||
-- { intro v, induction v with p x, exact fiber.mk (inl x) p },
|
||||
-- { reflexivity }
|
||||
-- end
|
||||
|
||||
definition pfiber_wedge_pr2_of_pcofiber_pprod_incl1 [constructor] :
|
||||
pcofiber (pprod_incl1 (Ω Y) X) →* pfiber (wedge_pr2 X Y) :=
|
||||
pmap.mk pfiber_wedge_pr2_of_pcofiber_pprod_incl1'
|
||||
begin refine (fiber_eq wedge.glue _)⁻¹, exact !wedge.elim_glue⁻¹ end
|
||||
-- pcofiber.elim (pfiber_wedge_pr2_of_pprod X Y)
|
||||
-- begin
|
||||
-- fapply phomotopy.mk,
|
||||
-- { intro p, apply fiber_eq (wedge.glue ⬝ ap inr p ⬝ wedge.glue⁻¹), symmetry,
|
||||
-- refine !ap_con ⬝ (!ap_con ⬝ !wedge.elim_glue ◾ (!ap_compose'⁻¹ ⬝ !ap_id)) ◾
|
||||
-- (!ap_inv ⬝ !wedge.elim_glue⁻²) ⬝ _, exact idp_con p },
|
||||
-- { esimp, refine fiber_eq2 (con.right_inv wedge.glue) _ ⬝ !fiber_eq_eta⁻¹,
|
||||
-- rewrite [idp_con, ↑fiber_eq_pr2, con2_idp, whisker_right_idp, whisker_right_idp],
|
||||
-- -- refine _ ⬝ (eq_bot_of_square (transpose (ap_con_right_inv_sq
|
||||
-- -- (wedge.elim (λx : X, Point Y) (@id Y) idp) wedge.glue)))⁻¹,
|
||||
-- -- refine whisker_right _ !con_inv ⬝ _,
|
||||
-- exact sorry
|
||||
-- }
|
||||
-- end
|
||||
|
||||
--set_option pp.notation false
|
||||
set_option pp.binder_types true
|
||||
open sigma.ops
|
||||
definition pfiber_wedge_pr2_pequiv_pcofiber_pprod_incl1 [constructor] :
|
||||
pfiber (wedge_pr2 X Y) ≃* pcofiber (pprod_incl1 (Ω Y) X) :=
|
||||
pequiv.MK (pcofiber_pprod_incl1_of_pfiber_wedge_pr2 _ _)
|
||||
(pfiber_wedge_pr2_of_pcofiber_pprod_incl1 _ _)
|
||||
abstract begin
|
||||
fapply phomotopy.mk,
|
||||
{ intro x, esimp, induction x with x p, induction x with x y,
|
||||
{ reflexivity },
|
||||
{ refine (fiber_eq (ap inr p) _)⁻¹, refine !ap_id⁻¹ ⬝ !ap_compose' },
|
||||
{ apply @pi_pathover_right' _ _ _ _ (λ(xp : Σ(x : X ∨ Y), pppi.to_fun (wedge_pr2 X Y) x = pt),
|
||||
pfiber_wedge_pr2_of_pcofiber_pprod_incl1'
|
||||
(pcofiber_pprod_incl1_of_pfiber_wedge_pr2' (mk xp.1 xp.2)) = mk xp.1 xp.2),
|
||||
intro p, apply eq_pathover, exact sorry }},
|
||||
{ symmetry, refine !cofiber.elim_glue ◾ idp ⬝ _, apply con_inv_eq_idp,
|
||||
apply ap (fiber_eq wedge.glue), esimp, rewrite [idp_con], refine !whisker_right_idp⁻² }
|
||||
end end
|
||||
abstract begin
|
||||
exact sorry
|
||||
end end
|
||||
-- apply eq_pathover_id_right, refine ap_compose pcofiber_pprod_incl1_of_pfiber_wedge_pr2 _ _ ⬝ ap02 _ !elim_glue ⬝ph _
|
||||
-- apply eq_pathover_id_right, refine ap_compose pfiber_wedge_pr2_of_pcofiber_pprod_incl1 _ _ ⬝ ap02 _ !elim_glue ⬝ph _
|
||||
|
||||
/- move -/
|
||||
definition ap1_idp {A B : Type*} (f : A →* B) : Ω→ f idp = idp :=
|
||||
respect_pt (Ω→ f)
|
||||
|
||||
definition ap1_phomotopy_idp {A B : Type*} {f g : A →* B} (h : f ~* g) :
|
||||
Ω⇒ h idp = !respect_pt ⬝ !respect_pt⁻¹ :=
|
||||
sorry
|
||||
|
||||
|
||||
variables {A} {B : Type*} {f : A →* B} {g : B →* A} (h : f ∘* g ~* pid B)
|
||||
include h
|
||||
definition nar_of_noo' (p : Ω A) : Ω (pfiber f) ×* Ω B :=
|
||||
begin
|
||||
refine (_, Ω→ f p),
|
||||
have z : Ω A →* Ω A, from pmap.mk (λp, Ω→ (g ∘* f) p ⬝ p⁻¹) proof (respect_pt (Ω→ (g ∘* f))) qed,
|
||||
refine fiber_eq ((Ω→ g ∘* Ω→ f) p ⬝ p⁻¹) (!idp_con⁻¹ ⬝ whisker_right (respect_pt f) _⁻¹),
|
||||
induction B with B b₀, induction f with f f₀, esimp at * ⊢, induction f₀,
|
||||
refine !idp_con⁻¹ ⬝ ap1_con (pmap_of_map f pt) _ _ ⬝
|
||||
((ap1_pcompose (pmap_of_map f pt) g _)⁻¹ ⬝ Ω⇒ h _ ⬝ ap1_pid _) ◾ ap1_inv _ _ ⬝ !con.right_inv
|
||||
end
|
||||
|
||||
definition noo_of_nar' (x : Ω (pfiber f) ×* Ω B) : Ω A :=
|
||||
begin
|
||||
induction x with p q, exact Ω→ (ppoint f) p ⬝ Ω→ g q
|
||||
end
|
||||
|
||||
variables (f g)
|
||||
definition nar_of_noo [constructor] :
|
||||
Ω A →* Ω (pfiber f) ×* Ω B :=
|
||||
begin
|
||||
refine pmap.mk (nar_of_noo' h) (prod_eq _ (ap1_gen_idp f (respect_pt f))),
|
||||
esimp [nar_of_noo'],
|
||||
refine fiber_eq2 (ap (ap1_gen _ _ _) (ap1_gen_idp f _) ⬝ !ap1_gen_idp) _ ⬝ !fiber_eq_eta⁻¹,
|
||||
induction B with B b₀, induction f with f f₀, esimp at * ⊢, induction f₀, esimp,
|
||||
refine (!idp_con ⬝ !whisker_right_idp) ◾ !whisker_right_idp ⬝ _, esimp [fiber_eq_pr2],
|
||||
apply inv_con_eq_idp,
|
||||
refine ap (ap02 f) !idp_con ⬝ _,
|
||||
esimp [ap1_con, ap1_gen_con, ap1_inv, ap1_gen_inv],
|
||||
refine _ ⬝ whisker_left _ (!con2_idp ⬝ !whisker_right_idp ⬝ idp ◾ ap1_phomotopy_idp h)⁻¹ᵖ,
|
||||
esimp, exact sorry
|
||||
end
|
||||
|
||||
definition noo_of_nar [constructor] :
|
||||
Ω (pfiber f) ×* Ω B →* Ω A :=
|
||||
pmap.mk (noo_of_nar' h) (respect_pt (Ω→ (ppoint f)) ◾ respect_pt (Ω→ g))
|
||||
|
||||
definition noo_pequiv_nar [constructor] :
|
||||
Ω A ≃* Ω (pfiber f) ×* Ω B :=
|
||||
pequiv.MK (nar_of_noo f g h) (noo_of_nar f g h)
|
||||
abstract begin
|
||||
exact sorry
|
||||
end end
|
||||
abstract begin
|
||||
exact sorry
|
||||
end end
|
||||
-- apply eq_pathover_id_right, refine ap_compose nar_of_noo _ _ ⬝ ap02 _ !elim_glue ⬝ph _
|
||||
-- apply eq_pathover_id_right, refine ap_compose noo_of_nar _ _ ⬝ ap02 _ !elim_glue ⬝ph _
|
||||
|
||||
definition loop_pequiv_of_cross_section {A B : Type*} (f : A →* B) (g : B →* A)
|
||||
(h : f ∘* g ~* pid B) : Ω A ≃* Ω (pfiber f) ×* Ω B :=
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
end wedge
|
||||
|
|
File diff suppressed because it is too large
Load diff
|
@ -606,5 +606,31 @@ definition ap1_gen_idp_eq {A B : Type} (f : A → B) {a : A} (q : f a = f a) (r
|
|||
ap1_gen_idp f q = ap (λx, ap1_gen f x x idp) r :=
|
||||
begin cases r, reflexivity end
|
||||
|
||||
definition pointed_change_point [constructor] (A : Type*) {a : A} (p : a = pt) :
|
||||
pointed.MK A a ≃* A :=
|
||||
pequiv_of_eq_pt p ⬝e* (pointed_eta_pequiv A)⁻¹ᵉ*
|
||||
|
||||
definition change_path_psquare {A B : Type*} (f : A →* B)
|
||||
{a' : A} {b' : B} (p : a' = pt) (q : pt = b') :
|
||||
psquare (pointed_change_point _ p)
|
||||
(pointed_change_point _ q⁻¹)
|
||||
(pmap.mk f (ap f p ⬝ respect_pt f ⬝ q)) f :=
|
||||
begin
|
||||
fapply phomotopy.mk, exact homotopy.rfl,
|
||||
exact !idp_con ⬝ !ap_id ◾ !ap_id ⬝ !con_inv_cancel_right ⬝ whisker_right _ (ap02 f !ap_id⁻¹)
|
||||
end
|
||||
|
||||
definition change_path_psquare_cod {A B : Type*} (f : A →* B) {b' : B} (p : pt = b') :
|
||||
f ~* pointed_change_point _ p⁻¹ ∘* pmap.mk f (respect_pt f ⬝ p) :=
|
||||
begin
|
||||
fapply phomotopy.mk, exact homotopy.rfl, exact !idp_con ⬝ !ap_id ◾ !ap_id ⬝ !con_inv_cancel_right
|
||||
end
|
||||
|
||||
definition change_path_psquare_cod' {A B : Type} (f : A → B) (a : A) {b' : B} (p : f a = b') :
|
||||
pointed_change_point _ p ∘* pmap_of_map f a ~* pmap.mk f p :=
|
||||
begin
|
||||
fapply phomotopy.mk, exact homotopy.rfl, refine whisker_left idp (ap_id p)⁻¹
|
||||
end
|
||||
|
||||
|
||||
end pointed
|
||||
|
|
Loading…
Reference in a new issue