2015-06-17 19:58:58 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2016-02-08 11:07:53 +00:00
|
|
|
|
Authors: Floris van Doorn, Ulrik Buchholtz
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
|
|
|
|
Declaration of suspension
|
|
|
|
|
-/
|
|
|
|
|
|
2017-06-15 01:56:23 +00:00
|
|
|
|
import hit.pushout types.pointed2 cubical.square .connectedness
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2017-07-20 17:55:31 +00:00
|
|
|
|
open pushout unit eq equiv pointed is_equiv
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp' (A : Type) : Type := pushout (λ(a : A), star) (λ(a : A), star)
|
|
|
|
|
|
|
|
|
|
namespace susp
|
|
|
|
|
|
|
|
|
|
definition north' {A : Type} : susp' A :=
|
|
|
|
|
inl star
|
|
|
|
|
|
|
|
|
|
definition pointed_susp [instance] [constructor] (X : Type)
|
|
|
|
|
: pointed (susp' X) :=
|
|
|
|
|
pointed.mk north'
|
|
|
|
|
|
|
|
|
|
end susp open susp
|
|
|
|
|
|
|
|
|
|
definition susp [constructor] (X : Type) : Type* :=
|
|
|
|
|
pointed.MK (susp' X) north'
|
|
|
|
|
|
|
|
|
|
notation `⅀` := susp
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
|
|
|
|
namespace susp
|
|
|
|
|
variable {A : Type}
|
|
|
|
|
|
|
|
|
|
definition north {A : Type} : susp A :=
|
2017-07-20 14:01:40 +00:00
|
|
|
|
north'
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
|
|
|
|
definition south {A : Type} : susp A :=
|
2015-09-22 16:01:55 +00:00
|
|
|
|
inr star
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
|
|
|
|
definition merid (a : A) : @north A = @south A :=
|
2015-09-22 16:01:55 +00:00
|
|
|
|
glue a
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
|
|
|
|
protected definition rec {P : susp A → Type} (PN : P north) (PS : P south)
|
2017-07-20 14:01:40 +00:00
|
|
|
|
(Pm : Π(a : A), PN =[merid a] PS) (x : susp' A) : P x :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
begin
|
2015-06-17 23:31:05 +00:00
|
|
|
|
induction x with u u,
|
|
|
|
|
{ cases u, exact PN},
|
|
|
|
|
{ cases u, exact PS},
|
|
|
|
|
{ apply Pm},
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
protected definition rec_on [reducible] {P : susp A → Type} (y : susp' A)
|
2015-06-17 19:58:58 +00:00
|
|
|
|
(PN : P north) (PS : P south) (Pm : Π(a : A), PN =[merid a] PS) : P y :=
|
|
|
|
|
susp.rec PN PS Pm y
|
|
|
|
|
|
|
|
|
|
theorem rec_merid {P : susp A → Type} (PN : P north) (PS : P south)
|
|
|
|
|
(Pm : Π(a : A), PN =[merid a] PS) (a : A)
|
2016-03-19 15:25:08 +00:00
|
|
|
|
: apd (susp.rec PN PS Pm) (merid a) = Pm a :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
!rec_glue
|
|
|
|
|
|
|
|
|
|
protected definition elim {P : Type} (PN : P) (PS : P) (Pm : A → PN = PS)
|
2017-07-20 14:01:40 +00:00
|
|
|
|
(x : susp' A) : P :=
|
2016-06-23 20:49:54 +00:00
|
|
|
|
susp.rec PN PS (λa, pathover_of_eq _ (Pm a)) x
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
protected definition elim_on [reducible] {P : Type} (x : susp' A)
|
2015-06-17 19:58:58 +00:00
|
|
|
|
(PN : P) (PS : P) (Pm : A → PN = PS) : P :=
|
|
|
|
|
susp.elim PN PS Pm x
|
|
|
|
|
|
|
|
|
|
theorem elim_merid {P : Type} {PN PS : P} (Pm : A → PN = PS) (a : A)
|
|
|
|
|
: ap (susp.elim PN PS Pm) (merid a) = Pm a :=
|
|
|
|
|
begin
|
2018-09-07 14:30:58 +00:00
|
|
|
|
apply inj_inv !(pathover_constant (merid a)),
|
2016-03-19 15:25:08 +00:00
|
|
|
|
rewrite [▸*,-apd_eq_pathover_of_eq_ap,↑susp.elim,rec_merid],
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
protected definition elim_type (PN : Type) (PS : Type) (Pm : A → PN ≃ PS)
|
2017-07-20 14:01:40 +00:00
|
|
|
|
(x : susp' A) : Type :=
|
2016-04-14 21:09:38 +00:00
|
|
|
|
pushout.elim_type (λx, PN) (λx, PS) Pm x
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
protected definition elim_type_on [reducible] (x : susp' A)
|
2015-06-17 19:58:58 +00:00
|
|
|
|
(PN : Type) (PS : Type) (Pm : A → PN ≃ PS) : Type :=
|
|
|
|
|
susp.elim_type PN PS Pm x
|
|
|
|
|
|
|
|
|
|
theorem elim_type_merid (PN : Type) (PS : Type) (Pm : A → PN ≃ PS)
|
|
|
|
|
(a : A) : transport (susp.elim_type PN PS Pm) (merid a) = Pm a :=
|
2016-04-14 21:09:38 +00:00
|
|
|
|
!elim_type_glue
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
|
theorem elim_type_merid_inv {A : Type} (PN : Type) (PS : Type) (Pm : A → PN ≃ PS)
|
|
|
|
|
(a : A) : transport (susp.elim_type PN PS Pm) (merid a)⁻¹ = to_inv (Pm a) :=
|
|
|
|
|
!elim_type_glue_inv
|
|
|
|
|
|
2016-03-06 18:59:48 +00:00
|
|
|
|
protected definition merid_square {a a' : A} (p : a = a')
|
|
|
|
|
: square (merid a) (merid a') idp idp :=
|
|
|
|
|
by cases p; apply vrefl
|
|
|
|
|
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end susp
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
attribute susp.north' susp.north susp.south [constructor]
|
2015-07-07 23:37:06 +00:00
|
|
|
|
attribute susp.rec susp.elim [unfold 6] [recursor 6]
|
|
|
|
|
attribute susp.elim_type [unfold 5]
|
|
|
|
|
attribute susp.rec_on susp.elim_on [unfold 3]
|
|
|
|
|
attribute susp.elim_type_on [unfold 2]
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2016-02-08 11:07:53 +00:00
|
|
|
|
namespace susp
|
|
|
|
|
|
|
|
|
|
open is_trunc is_conn trunc
|
|
|
|
|
|
|
|
|
|
-- Theorem 8.2.1
|
|
|
|
|
definition is_conn_susp [instance] (n : trunc_index) (A : Type)
|
|
|
|
|
[H : is_conn n A] : is_conn (n .+1) (susp A) :=
|
|
|
|
|
is_contr.mk (tr north)
|
|
|
|
|
begin
|
2017-07-20 14:01:40 +00:00
|
|
|
|
intro x, induction x with x, induction x,
|
2016-02-08 11:07:53 +00:00
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ exact (trunc.rec (λa, ap tr (merid a)) (center (trunc n A))) },
|
2017-07-20 14:01:40 +00:00
|
|
|
|
{ generalize (center (trunc n A)),
|
|
|
|
|
intro x, induction x with a',
|
2016-02-08 11:07:53 +00:00
|
|
|
|
apply pathover_of_tr_eq,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
rewrite [eq_transport_Fr,idp_con],
|
2017-07-20 14:01:40 +00:00
|
|
|
|
revert H, induction n with n IH: intro H,
|
|
|
|
|
{ apply is_prop.elim },
|
|
|
|
|
{ change ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid a'),
|
2016-02-08 11:07:53 +00:00
|
|
|
|
generalize a',
|
|
|
|
|
apply is_conn_fun.elim n
|
|
|
|
|
(is_conn_fun_from_unit n A a)
|
|
|
|
|
(λx : A, trunctype.mk' n (ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid x))),
|
|
|
|
|
intros,
|
|
|
|
|
change ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid a),
|
2017-07-20 14:01:40 +00:00
|
|
|
|
reflexivity }
|
2016-02-08 11:07:53 +00:00
|
|
|
|
}
|
|
|
|
|
end
|
|
|
|
|
|
2016-11-23 22:59:13 +00:00
|
|
|
|
/- Flattening lemma -/
|
2016-03-19 15:25:08 +00:00
|
|
|
|
|
2016-02-08 11:07:53 +00:00
|
|
|
|
open prod prod.ops
|
|
|
|
|
section
|
|
|
|
|
universe variable u
|
|
|
|
|
parameters (A : Type) (PN PS : Type.{u}) (Pm : A → PN ≃ PS)
|
|
|
|
|
include Pm
|
|
|
|
|
|
|
|
|
|
local abbreviation P [unfold 5] := susp.elim_type PN PS Pm
|
|
|
|
|
|
|
|
|
|
local abbreviation F : A × PN → PN := λz, z.2
|
|
|
|
|
|
|
|
|
|
local abbreviation G : A × PN → PS := λz, Pm z.1 z.2
|
|
|
|
|
|
|
|
|
|
protected definition flattening : sigma P ≃ pushout F G :=
|
|
|
|
|
begin
|
2016-04-14 21:09:38 +00:00
|
|
|
|
apply equiv.trans !pushout.flattening,
|
2016-03-06 18:59:48 +00:00
|
|
|
|
fapply pushout.equiv,
|
|
|
|
|
{ exact sigma.equiv_prod A PN },
|
|
|
|
|
{ apply sigma.sigma_unit_left },
|
|
|
|
|
{ apply sigma.sigma_unit_left },
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ reflexivity }
|
2016-02-08 11:07:53 +00:00
|
|
|
|
end
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end susp
|
2016-03-06 18:59:48 +00:00
|
|
|
|
|
|
|
|
|
/- Functoriality and equivalence -/
|
|
|
|
|
namespace susp
|
|
|
|
|
variables {A B : Type} (f : A → B)
|
|
|
|
|
include f
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_functor' [unfold 4] : susp A → susp B :=
|
2016-03-06 18:59:48 +00:00
|
|
|
|
begin
|
|
|
|
|
intro x, induction x with a,
|
|
|
|
|
{ exact north },
|
|
|
|
|
{ exact south },
|
|
|
|
|
{ exact merid (f a) }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
variable [Hf : is_equiv f]
|
|
|
|
|
include Hf
|
|
|
|
|
|
|
|
|
|
open is_equiv
|
2017-07-20 14:01:40 +00:00
|
|
|
|
protected definition is_equiv_functor [instance] [constructor] : is_equiv (susp_functor' f) :=
|
|
|
|
|
adjointify (susp_functor' f) (susp_functor' f⁻¹)
|
2016-03-06 18:59:48 +00:00
|
|
|
|
abstract begin
|
|
|
|
|
intro sb, induction sb with b, do 2 reflexivity,
|
|
|
|
|
apply eq_pathover,
|
2018-09-07 13:57:43 +00:00
|
|
|
|
rewrite [ap_id,-ap_compose' (susp_functor' f) (susp_functor' f⁻¹)],
|
2016-03-06 18:59:48 +00:00
|
|
|
|
krewrite [susp.elim_merid,susp.elim_merid], apply transpose,
|
|
|
|
|
apply susp.merid_square (right_inv f b)
|
|
|
|
|
end end
|
|
|
|
|
abstract begin
|
|
|
|
|
intro sa, induction sa with a, do 2 reflexivity,
|
|
|
|
|
apply eq_pathover,
|
2018-09-07 13:57:43 +00:00
|
|
|
|
rewrite [ap_id,-ap_compose' (susp_functor' f⁻¹) (susp_functor' f)],
|
2016-03-06 18:59:48 +00:00
|
|
|
|
krewrite [susp.elim_merid,susp.elim_merid], apply transpose,
|
|
|
|
|
apply susp.merid_square (left_inv f a)
|
|
|
|
|
end end
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end susp
|
|
|
|
|
|
|
|
|
|
namespace susp
|
|
|
|
|
variables {A B : Type} (f : A ≃ B)
|
|
|
|
|
|
|
|
|
|
protected definition equiv : susp A ≃ susp B :=
|
2017-07-20 14:01:40 +00:00
|
|
|
|
equiv.mk (susp_functor' f) _
|
2016-03-06 18:59:48 +00:00
|
|
|
|
end susp
|
2016-02-08 11:07:53 +00:00
|
|
|
|
|
2016-03-11 00:53:42 +00:00
|
|
|
|
namespace susp
|
2016-11-23 22:59:13 +00:00
|
|
|
|
open pointed is_trunc
|
2017-07-20 17:55:31 +00:00
|
|
|
|
variables {X X' Y Y' Z : Type*}
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_functor [constructor] (f : X →* Y) : susp X →* susp Y :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
begin
|
2015-06-23 16:47:52 +00:00
|
|
|
|
fconstructor,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
{ exact susp_functor' f },
|
2016-03-11 00:53:42 +00:00
|
|
|
|
{ reflexivity }
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
2018-09-10 12:29:06 +00:00
|
|
|
|
notation `⅀→`:(max+5) := susp_functor
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition is_equiv_susp_functor [constructor] (f : X →* Y) [Hf : is_equiv f]
|
|
|
|
|
: is_equiv (susp_functor f) :=
|
2016-03-11 00:53:42 +00:00
|
|
|
|
susp.is_equiv_functor f
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_pequiv [constructor] (f : X ≃* Y) : susp X ≃* susp Y :=
|
2016-03-11 00:53:42 +00:00
|
|
|
|
pequiv_of_equiv (susp.equiv f) idp
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_functor_pcompose (g : Y →* Z) (f : X →* Y) :
|
|
|
|
|
susp_functor (g ∘* f) ~* susp_functor g ∘* susp_functor f :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
begin
|
2017-06-02 16:13:20 +00:00
|
|
|
|
fapply phomotopy.mk,
|
|
|
|
|
{ intro x, induction x,
|
2016-03-14 23:11:21 +00:00
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ reflexivity },
|
2017-07-20 14:01:40 +00:00
|
|
|
|
{ apply eq_pathover, apply hdeg_square,
|
|
|
|
|
refine !elim_merid ⬝ _ ⬝ (ap_compose (susp_functor g) _ _)⁻¹ᵖ,
|
2017-06-02 16:13:20 +00:00
|
|
|
|
refine _ ⬝ ap02 _ !elim_merid⁻¹, exact !elim_merid⁻¹ }},
|
|
|
|
|
{ reflexivity },
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_functor_phomotopy {f g : X →* Y} (p : f ~* g) :
|
|
|
|
|
susp_functor f ~* susp_functor g :=
|
2017-06-02 16:13:20 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply phomotopy.mk,
|
|
|
|
|
{ intro x, induction x,
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ apply eq_pathover, apply hdeg_square, esimp, refine !elim_merid ⬝ _ ⬝ !elim_merid⁻¹ᵖ,
|
|
|
|
|
exact ap merid (p a), }},
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
end
|
|
|
|
|
|
2018-09-10 12:29:06 +00:00
|
|
|
|
notation `⅀⇒`:(max+5) := susp_functor_phomotopy
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_functor_pid (A : Type*) : susp_functor (pid A) ~* pid (susp A) :=
|
2017-06-02 16:13:20 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply phomotopy.mk,
|
|
|
|
|
{ intro x, induction x,
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ apply eq_pathover_id_right, apply hdeg_square, apply elim_merid }},
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
/- adjunction originally ported from Coq-HoTT,
|
|
|
|
|
but we proved some additional naturality conditions -/
|
2015-06-17 19:58:58 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_unit [constructor] (X : Type*) : X →* Ω(susp X) :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
begin
|
2015-06-23 16:47:52 +00:00
|
|
|
|
fconstructor,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
{ intro x, exact merid x ⬝ (merid pt)⁻¹ },
|
|
|
|
|
{ apply con.right_inv },
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_unit_natural (f : X →* Y)
|
2018-09-10 12:29:06 +00:00
|
|
|
|
: psquare (loop_susp_unit X) (loop_susp_unit Y) f (Ω→ (susp_functor f)) :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
begin
|
2018-09-10 12:29:06 +00:00
|
|
|
|
apply ptranspose,
|
2015-06-17 19:58:58 +00:00
|
|
|
|
induction X with X x, induction Y with Y y, induction f with f pf, esimp at *, induction pf,
|
2017-06-17 21:20:04 +00:00
|
|
|
|
fapply phomotopy.mk,
|
2017-03-22 22:35:35 +00:00
|
|
|
|
{ intro x', symmetry,
|
2015-06-17 23:31:05 +00:00
|
|
|
|
exact
|
2017-03-22 22:35:35 +00:00
|
|
|
|
!ap1_gen_idp_left ⬝
|
2015-06-17 23:31:05 +00:00
|
|
|
|
(!ap_con ⬝
|
|
|
|
|
whisker_left _ !ap_inv) ⬝
|
2016-11-23 22:59:13 +00:00
|
|
|
|
(!elim_merid ◾ (inverse2 !elim_merid)) },
|
2017-03-22 22:35:35 +00:00
|
|
|
|
{ rewrite [▸*, idp_con (con.right_inv _)],
|
2015-06-17 23:31:05 +00:00
|
|
|
|
apply inv_con_eq_of_eq_con,
|
|
|
|
|
refine _ ⬝ !con.assoc',
|
|
|
|
|
rewrite inverse2_right_inv,
|
|
|
|
|
refine _ ⬝ !con.assoc',
|
2016-02-15 17:57:51 +00:00
|
|
|
|
rewrite [ap_con_right_inv],
|
2017-03-22 22:35:35 +00:00
|
|
|
|
rewrite [ap1_gen_idp_left_con],
|
|
|
|
|
rewrite [-ap_compose (concat idp)] },
|
2015-06-17 23:31:05 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_counit [constructor] (X : Type*) : susp (Ω X) →* X :=
|
2015-06-17 23:31:05 +00:00
|
|
|
|
begin
|
2017-07-21 12:35:23 +00:00
|
|
|
|
fapply pmap.mk,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
{ intro x, induction x, exact pt, exact pt, exact a },
|
|
|
|
|
{ reflexivity },
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_counit_natural (f : X →* Y)
|
2018-09-10 12:29:06 +00:00
|
|
|
|
: psquare (loop_susp_counit X) (loop_susp_counit Y) (⅀→ (Ω→ f)) f :=
|
2015-06-17 23:31:05 +00:00
|
|
|
|
begin
|
|
|
|
|
induction X with X x, induction Y with Y y, induction f with f pf, esimp at *, induction pf,
|
2015-06-23 16:47:52 +00:00
|
|
|
|
fconstructor,
|
2015-06-17 23:31:05 +00:00
|
|
|
|
{ intro x', induction x' with p,
|
2017-01-17 19:32:55 +00:00
|
|
|
|
{ reflexivity },
|
2016-11-23 22:59:13 +00:00
|
|
|
|
{ reflexivity },
|
2015-07-29 12:17:16 +00:00
|
|
|
|
{ esimp, apply eq_pathover, apply hdeg_square,
|
2018-09-07 13:57:43 +00:00
|
|
|
|
xrewrite [-ap_compose' f, -ap_compose' (susp.elim (f x) (f x) (λ (a : f x = f x), a)),▸*],
|
2017-03-22 22:35:35 +00:00
|
|
|
|
xrewrite [+elim_merid, ap1_gen_idp_left] }},
|
2016-11-23 22:59:13 +00:00
|
|
|
|
{ reflexivity }
|
2015-06-17 23:31:05 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_counit_unit (X : Type*)
|
|
|
|
|
: ap1 (loop_susp_counit X) ∘* loop_susp_unit (Ω X) ~* pid (Ω X) :=
|
2015-06-17 23:31:05 +00:00
|
|
|
|
begin
|
2015-06-23 16:47:52 +00:00
|
|
|
|
induction X with X x, fconstructor,
|
2015-06-17 23:31:05 +00:00
|
|
|
|
{ intro p, esimp,
|
2017-03-22 22:35:35 +00:00
|
|
|
|
refine !ap1_gen_idp_left ⬝
|
2015-06-17 23:31:05 +00:00
|
|
|
|
(!ap_con ⬝
|
|
|
|
|
whisker_left _ !ap_inv) ⬝
|
2016-11-23 22:59:13 +00:00
|
|
|
|
(!elim_merid ◾ inverse2 !elim_merid) },
|
2015-12-09 05:02:05 +00:00
|
|
|
|
{ rewrite [▸*,inverse2_right_inv (elim_merid id idp)],
|
2015-06-17 23:31:05 +00:00
|
|
|
|
refine !con.assoc ⬝ _,
|
2017-03-22 22:35:35 +00:00
|
|
|
|
xrewrite [ap_con_right_inv (susp.elim x x (λa, a)) (merid idp),ap1_gen_idp_left_con,
|
|
|
|
|
-ap_compose] }
|
2015-06-17 23:31:05 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_unit_counit (X : Type*)
|
|
|
|
|
: loop_susp_counit (susp X) ∘* susp_functor (loop_susp_unit X) ~* pid (susp X) :=
|
2015-06-17 23:31:05 +00:00
|
|
|
|
begin
|
2015-06-23 16:47:52 +00:00
|
|
|
|
induction X with X x, fconstructor,
|
2015-06-17 23:31:05 +00:00
|
|
|
|
{ intro x', induction x',
|
2016-11-23 22:59:13 +00:00
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ exact merid pt },
|
2015-07-29 12:17:16 +00:00
|
|
|
|
{ apply eq_pathover,
|
2018-09-07 13:57:43 +00:00
|
|
|
|
xrewrite [▸*, ap_id, -ap_compose' (susp.elim north north (λa, a)), +elim_merid,▸*],
|
2016-11-23 22:59:13 +00:00
|
|
|
|
apply square_of_eq, exact !idp_con ⬝ !inv_con_cancel_right⁻¹ }},
|
|
|
|
|
{ reflexivity }
|
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_elim [constructor] {X Y : Type*} (f : X →* Ω Y) : susp X →* Y :=
|
|
|
|
|
loop_susp_counit Y ∘* susp_functor f
|
2016-11-23 22:59:13 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_intro [constructor] {X Y : Type*} (f : susp X →* Y) : X →* Ω Y :=
|
|
|
|
|
ap1 f ∘* loop_susp_unit X
|
2016-11-23 22:59:13 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_elim_susp_functor {A B C : Type*} (g : B →* Ω C) (f : A →* B) :
|
|
|
|
|
susp_elim g ∘* susp_functor f ~* susp_elim (g ∘* f) :=
|
2017-06-02 16:13:20 +00:00
|
|
|
|
begin
|
2017-07-20 14:01:40 +00:00
|
|
|
|
refine !passoc ⬝* _, exact pwhisker_left _ !susp_functor_pcompose⁻¹*
|
2017-06-02 16:13:20 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_elim_phomotopy {A B : Type*} {f g : A →* Ω B} (p : f ~* g) : susp_elim f ~* susp_elim g :=
|
|
|
|
|
pwhisker_left _ (susp_functor_phomotopy p)
|
2017-06-02 16:13:20 +00:00
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_elim_natural {X Y Z : Type*} (g : Y →* Z) (f : X →* Ω Y)
|
|
|
|
|
: g ∘* susp_elim f ~* susp_elim (Ω→ g ∘* f) :=
|
2017-06-02 16:13:20 +00:00
|
|
|
|
begin
|
2017-07-20 14:01:40 +00:00
|
|
|
|
refine _ ⬝* pwhisker_left _ !susp_functor_pcompose⁻¹*,
|
2017-06-02 16:13:20 +00:00
|
|
|
|
refine !passoc⁻¹* ⬝* _ ⬝* !passoc,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
exact pwhisker_right _ !loop_susp_counit_natural
|
2017-06-02 16:13:20 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition loop_susp_intro_natural {X Y Z : Type*} (g : susp Y →* Z) (f : X →* Y) :
|
|
|
|
|
loop_susp_intro (g ∘* susp_functor f) ~* loop_susp_intro g ∘* f :=
|
2018-09-10 12:29:06 +00:00
|
|
|
|
pwhisker_right _ !ap1_pcompose ⬝* !passoc ⬝* pwhisker_left _ !loop_susp_unit_natural ⬝*
|
2017-06-02 16:13:20 +00:00
|
|
|
|
!passoc⁻¹*
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_adjoint_loop_right_inv {X Y : Type*} (g : X →* Ω Y) :
|
|
|
|
|
loop_susp_intro (susp_elim g) ~* g :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
|
|
|
|
refine !pwhisker_right !ap1_pcompose ⬝* _,
|
|
|
|
|
refine !passoc ⬝* _,
|
2018-09-10 12:29:06 +00:00
|
|
|
|
refine !pwhisker_left !loop_susp_unit_natural ⬝* _,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
refine !passoc⁻¹* ⬝* _,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
refine !pwhisker_right !loop_susp_counit_unit ⬝* _,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
apply pid_pcompose
|
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_adjoint_loop_left_inv {X Y : Type*} (f : susp X →* Y) :
|
|
|
|
|
susp_elim (loop_susp_intro f) ~* f :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
2017-07-20 14:01:40 +00:00
|
|
|
|
refine !pwhisker_left !susp_functor_pcompose ⬝* _,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
refine !passoc⁻¹* ⬝* _,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
refine !pwhisker_right !loop_susp_counit_natural⁻¹* ⬝* _,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
refine !passoc ⬝* _,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
refine !pwhisker_left !loop_susp_unit_counit ⬝* _,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
apply pcompose_pid
|
2015-06-17 23:31:05 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_adjoint_loop_unpointed [constructor] (X Y : Type*) : susp X →* Y ≃ X →* Ω Y :=
|
2015-06-17 19:58:58 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply equiv.MK,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
{ exact loop_susp_intro },
|
|
|
|
|
{ exact susp_elim },
|
|
|
|
|
{ intro g, apply eq_of_phomotopy, exact susp_adjoint_loop_right_inv g },
|
|
|
|
|
{ intro f, apply eq_of_phomotopy, exact susp_adjoint_loop_left_inv f }
|
2016-11-23 22:59:13 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 17:55:31 +00:00
|
|
|
|
definition susp_functor_pconst_homotopy [unfold 3] {X Y : Type*} (x : susp X) :
|
|
|
|
|
susp_functor (pconst X Y) x = pt :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
2017-07-20 17:55:31 +00:00
|
|
|
|
induction x,
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ exact (merid pt)⁻¹ },
|
|
|
|
|
{ apply eq_pathover, refine !elim_merid ⬝ph _ ⬝hp !ap_constant⁻¹, exact square_of_eq !con.right_inv⁻¹ }
|
2016-11-23 22:59:13 +00:00
|
|
|
|
end
|
|
|
|
|
|
2018-09-10 12:29:06 +00:00
|
|
|
|
definition susp_functor_pconst [constructor] (X Y : Type*) :
|
|
|
|
|
susp_functor (pconst X Y) ~* pconst (susp X) (susp Y) :=
|
2017-07-20 17:55:31 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply phomotopy.mk,
|
|
|
|
|
{ exact susp_functor_pconst_homotopy },
|
|
|
|
|
{ reflexivity }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition susp_pfunctor [constructor] (X Y : Type*) : ppmap X Y →* ppmap (susp X) (susp Y) :=
|
|
|
|
|
pmap.mk susp_functor (eq_of_phomotopy !susp_functor_pconst)
|
|
|
|
|
|
|
|
|
|
definition susp_pelim [constructor] (X Y : Type*) : ppmap X (Ω Y) →* ppmap (susp X) Y :=
|
|
|
|
|
ppcompose_left (loop_susp_counit Y) ∘* susp_pfunctor X (Ω Y)
|
|
|
|
|
|
|
|
|
|
definition loop_susp_pintro [constructor] (X Y : Type*) : ppmap (susp X) Y →* ppmap X (Ω Y) :=
|
|
|
|
|
ppcompose_right (loop_susp_unit X) ∘* pap1 (susp X) Y
|
|
|
|
|
|
|
|
|
|
definition loop_susp_pintro_natural_left (f : X' →* X) :
|
|
|
|
|
psquare (loop_susp_pintro X Y) (loop_susp_pintro X' Y)
|
|
|
|
|
(ppcompose_right (susp_functor f)) (ppcompose_right f) :=
|
2018-09-10 12:29:06 +00:00
|
|
|
|
!pap1_natural_left ⬝h* ppcompose_right_psquare (loop_susp_unit_natural f)
|
2017-07-20 17:55:31 +00:00
|
|
|
|
|
|
|
|
|
definition loop_susp_pintro_natural_right (f : Y →* Y') :
|
|
|
|
|
psquare (loop_susp_pintro X Y) (loop_susp_pintro X Y')
|
|
|
|
|
(ppcompose_left f) (ppcompose_left (Ω→ f)) :=
|
|
|
|
|
!pap1_natural_right ⬝h* !ppcompose_left_ppcompose_right⁻¹*
|
|
|
|
|
|
|
|
|
|
definition is_equiv_loop_susp_pintro [constructor] (X Y : Type*) :
|
|
|
|
|
is_equiv (loop_susp_pintro X Y) :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
2017-07-20 17:55:31 +00:00
|
|
|
|
fapply adjointify,
|
|
|
|
|
{ exact susp_pelim X Y },
|
|
|
|
|
{ intro g, apply eq_of_phomotopy, exact susp_adjoint_loop_right_inv g },
|
|
|
|
|
{ intro f, apply eq_of_phomotopy, exact susp_adjoint_loop_left_inv f }
|
2015-06-17 23:31:05 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 17:55:31 +00:00
|
|
|
|
definition susp_adjoint_loop [constructor] (X Y : Type*) : ppmap (susp X) Y ≃* ppmap X (Ω Y) :=
|
|
|
|
|
pequiv_of_pmap (loop_susp_pintro X Y) (is_equiv_loop_susp_pintro X Y)
|
|
|
|
|
|
|
|
|
|
definition susp_adjoint_loop_natural_right (f : Y →* Y') :
|
|
|
|
|
psquare (susp_adjoint_loop X Y) (susp_adjoint_loop X Y')
|
|
|
|
|
(ppcompose_left f) (ppcompose_left (Ω→ f)) :=
|
|
|
|
|
loop_susp_pintro_natural_right f
|
|
|
|
|
|
|
|
|
|
definition susp_adjoint_loop_natural_left (f : X' →* X) :
|
|
|
|
|
psquare (susp_adjoint_loop X Y) (susp_adjoint_loop X' Y)
|
|
|
|
|
(ppcompose_right (susp_functor f)) (ppcompose_right f) :=
|
|
|
|
|
loop_susp_pintro_natural_left f
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition ap1_susp_elim {A : Type*} {X : Type*} (p : A →* Ω X) :
|
|
|
|
|
Ω→(susp_elim p) ∘* loop_susp_unit A ~* p :=
|
|
|
|
|
susp_adjoint_loop_right_inv p
|
2016-11-23 22:59:13 +00:00
|
|
|
|
|
2017-07-20 17:55:31 +00:00
|
|
|
|
/- the underlying homotopies of susp_adjoint_loop_natural_* -/
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_adjoint_loop_nat_right (f : susp X →* Y) (g : Y →* Z)
|
|
|
|
|
: susp_adjoint_loop X Z (g ∘* f) ~* ap1 g ∘* susp_adjoint_loop X Y f :=
|
2015-06-17 23:31:05 +00:00
|
|
|
|
begin
|
2017-07-20 14:01:40 +00:00
|
|
|
|
esimp [susp_adjoint_loop],
|
2015-06-17 23:31:05 +00:00
|
|
|
|
refine _ ⬝* !passoc,
|
|
|
|
|
apply pwhisker_right,
|
2016-09-22 19:42:46 +00:00
|
|
|
|
apply ap1_pcompose
|
2015-06-17 23:31:05 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition susp_adjoint_loop_nat_left (f : Y →* Ω Z) (g : X →* Y)
|
|
|
|
|
: (susp_adjoint_loop X Z)⁻¹ᵉ (f ∘* g) ~* (susp_adjoint_loop Y Z)⁻¹ᵉ f ∘* susp_functor g :=
|
2015-06-17 23:31:05 +00:00
|
|
|
|
begin
|
2017-07-20 14:01:40 +00:00
|
|
|
|
esimp [susp_adjoint_loop],
|
2015-06-17 23:31:05 +00:00
|
|
|
|
refine _ ⬝* !passoc⁻¹*,
|
|
|
|
|
apply pwhisker_left,
|
2017-07-20 14:01:40 +00:00
|
|
|
|
apply susp_functor_pcompose
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end
|
|
|
|
|
|
2016-11-23 22:59:13 +00:00
|
|
|
|
/- iterated suspension -/
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition iterate_susp (n : ℕ) (A : Type*) : Type* := iterate (λX, susp X) n A
|
2016-04-22 19:12:25 +00:00
|
|
|
|
|
|
|
|
|
open is_conn trunc_index nat
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition iterate_susp_succ (n : ℕ) (A : Type*) :
|
2016-04-22 19:12:25 +00:00
|
|
|
|
iterate_susp (succ n) A = susp (iterate_susp n A) :=
|
|
|
|
|
idp
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition is_conn_iterate_susp [instance] (n : ℕ₋₂) (m : ℕ) (A : Type*)
|
2016-04-22 19:12:25 +00:00
|
|
|
|
[H : is_conn n A] : is_conn (n + m) (iterate_susp m A) :=
|
|
|
|
|
begin induction m with m IH, exact H, exact @is_conn_susp _ _ IH end
|
|
|
|
|
|
|
|
|
|
-- Separate cases for n = 0, which comes up often
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition is_conn_iterate_susp_zero [instance] (m : ℕ) (A : Type*)
|
2016-04-22 19:12:25 +00:00
|
|
|
|
[H : is_conn 0 A] : is_conn m (iterate_susp m A) :=
|
|
|
|
|
begin induction m with m IH, exact H, exact @is_conn_susp _ _ IH end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition iterate_susp_functor (n : ℕ) {A B : Type*} (f : A →* B) :
|
|
|
|
|
iterate_susp n A →* iterate_susp n B :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
|
|
|
|
induction n with n g,
|
|
|
|
|
{ exact f },
|
2017-07-20 14:01:40 +00:00
|
|
|
|
{ exact susp_functor g }
|
2016-11-23 22:59:13 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition iterate_susp_succ_in (n : ℕ) (A : Type*) :
|
|
|
|
|
iterate_susp (succ n) A ≃* iterate_susp n (susp A) :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
|
|
|
|
induction n with n IH,
|
|
|
|
|
{ reflexivity},
|
2017-07-20 14:01:40 +00:00
|
|
|
|
{ exact susp_pequiv IH}
|
2016-11-23 22:59:13 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-07-20 14:01:40 +00:00
|
|
|
|
definition iterate_susp_adjoint_loopn [constructor] (X Y : Type*) (n : ℕ) :
|
|
|
|
|
ppmap (iterate_susp n X) Y ≃* ppmap X (Ω[n] Y) :=
|
2016-11-23 22:59:13 +00:00
|
|
|
|
begin
|
|
|
|
|
revert X Y, induction n with n IH: intro X Y,
|
|
|
|
|
{ reflexivity },
|
2018-09-07 14:30:58 +00:00
|
|
|
|
{ refine !susp_adjoint_loop ⬝e* !IH ⬝e* _, apply ppmap_pequiv_ppmap_right,
|
2016-11-23 22:59:13 +00:00
|
|
|
|
symmetry, apply loopn_succ_in }
|
|
|
|
|
end
|
|
|
|
|
|
2015-06-17 19:58:58 +00:00
|
|
|
|
end susp
|