2016-01-25 16:54:24 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2016 Jakob von Raumer. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Jakob von Raumer
|
|
|
|
|
|
|
|
|
|
The Cofiber Type
|
|
|
|
|
-/
|
2016-11-24 02:36:25 +00:00
|
|
|
|
import hit.pushout function .susp types.unit
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
2016-11-24 02:36:25 +00:00
|
|
|
|
open eq pushout unit pointed is_trunc is_equiv susp unit equiv
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
2017-02-17 04:31:58 +00:00
|
|
|
|
definition cofiber {A B : Type} (f : A → B) := pushout f (λ (a : A), ⋆)
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
|
|
|
|
namespace cofiber
|
|
|
|
|
section
|
|
|
|
|
parameters {A B : Type} (f : A → B)
|
|
|
|
|
|
2017-02-17 04:31:58 +00:00
|
|
|
|
definition cod : B → cofiber f := inl
|
|
|
|
|
definition base : cofiber f := inr ⋆
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
2016-11-23 22:59:13 +00:00
|
|
|
|
parameter {f}
|
2017-02-17 04:31:58 +00:00
|
|
|
|
protected definition glue (a : A) : cofiber.cod f (f a) = cofiber.base f :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
pushout.glue a
|
|
|
|
|
|
2017-02-17 04:31:58 +00:00
|
|
|
|
protected definition rec {P : cofiber f → Type} (Pcod : Π (b : B), P (cod b)) (Pbase : P base)
|
|
|
|
|
(Pglue : Π (a : A), pathover P (Pcod (f a)) (glue a) Pbase) :
|
2016-01-27 14:59:54 +00:00
|
|
|
|
(Π y, P y) :=
|
|
|
|
|
begin
|
2017-02-17 04:31:58 +00:00
|
|
|
|
intro y, induction y, exact Pcod x, induction x, exact Pbase, exact Pglue x
|
2016-01-27 14:59:54 +00:00
|
|
|
|
end
|
|
|
|
|
|
2016-11-23 22:59:13 +00:00
|
|
|
|
protected definition rec_on {P : cofiber f → Type} (y : cofiber f)
|
2017-02-17 04:31:58 +00:00
|
|
|
|
(Pcod : Π (b : B), P (cod b)) (Pbase : P base)
|
|
|
|
|
(Pglue : Π (a : A), pathover P (Pcod (f a)) (glue a) Pbase) : P y :=
|
|
|
|
|
cofiber.rec Pcod Pbase Pglue y
|
|
|
|
|
|
|
|
|
|
protected theorem rec_glue {P : cofiber f → Type} (Pcod : Π (b : B), P (cod b)) (Pbase : P base)
|
|
|
|
|
(Pglue : Π (a : A), pathover P (Pcod (f a)) (glue a) Pbase) (a : A)
|
|
|
|
|
: apd (cofiber.rec Pcod Pbase Pglue) (cofiber.glue a) = Pglue a :=
|
|
|
|
|
!pushout.rec_glue
|
|
|
|
|
|
|
|
|
|
protected definition elim {P : Type} (Pcod : B → P) (Pbase : P)
|
|
|
|
|
(Pglue : Π (x : A), Pcod (f x) = Pbase) (y : cofiber f) : P :=
|
|
|
|
|
pushout.elim Pcod (λu, Pbase) Pglue y
|
|
|
|
|
|
|
|
|
|
protected definition elim_on {P : Type} (y : cofiber f) (Pcod : B → P) (Pbase : P)
|
|
|
|
|
(Pglue : Π (x : A), Pcod (f x) = Pbase) : P :=
|
|
|
|
|
cofiber.elim Pcod Pbase Pglue y
|
|
|
|
|
|
|
|
|
|
protected theorem elim_glue {P : Type} (Pcod : B → P) (Pbase : P)
|
|
|
|
|
(Pglue : Π (x : A), Pcod (f x) = Pbase) (a : A)
|
|
|
|
|
: ap (cofiber.elim Pcod Pbase Pglue) (cofiber.glue a) = Pglue a :=
|
2016-11-24 02:36:25 +00:00
|
|
|
|
!pushout.elim_glue
|
|
|
|
|
|
2016-01-25 16:54:24 +00:00
|
|
|
|
end
|
2017-02-17 04:31:58 +00:00
|
|
|
|
|
2016-01-25 16:54:24 +00:00
|
|
|
|
end cofiber
|
|
|
|
|
|
2016-11-23 22:59:13 +00:00
|
|
|
|
attribute cofiber.base cofiber.cod [constructor]
|
|
|
|
|
attribute cofiber.rec cofiber.elim [recursor 8] [unfold 8]
|
|
|
|
|
attribute cofiber.rec_on cofiber.elim_on [unfold 5]
|
|
|
|
|
|
2016-02-15 23:23:28 +00:00
|
|
|
|
-- pointed version
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
2016-11-23 22:59:13 +00:00
|
|
|
|
definition pcofiber [constructor] {A B : Type*} (f : A →* B) : Type* :=
|
|
|
|
|
pointed.MK (cofiber f) !cofiber.base
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
2017-01-08 21:47:48 +00:00
|
|
|
|
notation `ℂ` := pcofiber
|
|
|
|
|
|
2016-01-27 11:48:32 +00:00
|
|
|
|
namespace cofiber
|
2016-01-27 14:59:54 +00:00
|
|
|
|
|
2017-02-17 04:31:58 +00:00
|
|
|
|
variables {A B : Type*} (f : A →* B)
|
|
|
|
|
|
|
|
|
|
definition is_contr_cofiber_of_equiv [H : is_equiv f] : is_contr (cofiber f) :=
|
|
|
|
|
begin
|
|
|
|
|
fapply is_contr.mk, exact cofiber.base f,
|
|
|
|
|
intro a, induction a with b a,
|
|
|
|
|
{ exact !glue⁻¹ ⬝ ap inl (right_inv f b) },
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ apply eq_pathover_constant_left_id_right, apply move_top_of_left,
|
|
|
|
|
refine _ ⬝pv natural_square_tr cofiber.glue (left_inv f a) ⬝vp !ap_constant,
|
|
|
|
|
refine ap02 inl _ ⬝ !ap_compose⁻¹, exact adj f a },
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition pcod [constructor] (f : A →* B) : B →* pcofiber f :=
|
|
|
|
|
pmap.mk (cofiber.cod f) (ap inl (respect_pt f)⁻¹ ⬝ cofiber.glue pt)
|
|
|
|
|
|
|
|
|
|
definition pcod_pcompose [constructor] (f : A →* B) : pcod f ∘* f ~* pconst A (ℂ f) :=
|
|
|
|
|
begin
|
|
|
|
|
fapply phomotopy.mk,
|
|
|
|
|
{ intro a, exact cofiber.glue a },
|
|
|
|
|
{ exact !con_inv_cancel_left⁻¹ ⬝ idp ◾ (!ap_inv⁻¹ ◾ idp) }
|
|
|
|
|
end
|
2016-01-25 16:54:24 +00:00
|
|
|
|
|
2017-02-17 04:31:58 +00:00
|
|
|
|
definition pcofiber_punit (A : Type*) : pcofiber (pconst A punit) ≃* psusp A :=
|
2016-01-27 11:48:32 +00:00
|
|
|
|
begin
|
2016-02-15 19:40:25 +00:00
|
|
|
|
fapply pequiv_of_pmap,
|
2016-01-27 11:48:32 +00:00
|
|
|
|
{ fconstructor, intro x, induction x, exact north, exact south, exact merid x,
|
2017-02-17 04:31:58 +00:00
|
|
|
|
exact (merid pt)⁻¹ },
|
2016-01-27 11:48:32 +00:00
|
|
|
|
{ esimp, fapply adjointify,
|
2016-11-24 02:36:25 +00:00
|
|
|
|
{ intro s, induction s, exact inl ⋆, exact inr ⋆, apply glue a },
|
|
|
|
|
{ intro s, induction s, do 2 reflexivity, esimp,
|
|
|
|
|
apply eq_pathover, refine _ ⬝hp !ap_id⁻¹, apply hdeg_square,
|
|
|
|
|
refine !(ap_compose (pushout.elim _ _ _)) ⬝ _,
|
|
|
|
|
refine ap _ !elim_merid ⬝ _, apply elim_glue },
|
2017-02-17 04:31:58 +00:00
|
|
|
|
{ intro c, induction c with u, induction u, reflexivity,
|
|
|
|
|
reflexivity, esimp, apply eq_pathover, apply hdeg_square,
|
2016-11-24 02:36:25 +00:00
|
|
|
|
refine _ ⬝ !ap_id⁻¹, refine !(ap_compose (pushout.elim _ _ _)) ⬝ _,
|
|
|
|
|
refine ap02 _ !elim_glue ⬝ _, apply elim_merid }},
|
2016-01-27 11:48:32 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end cofiber
|