feat(hott): functoriality of pushout; connectedness in is_conn namespace

other changes:
- move result about connectedness of susp to homotopy.susp
- improved definition of circle multiplication
- improved the interface to join
This commit is contained in:
Ulrik Buchholtz 2016-02-08 12:07:53 +01:00 committed by Leonardo de Moura
parent a8db8bc61a
commit bd9e47c82c
15 changed files with 764 additions and 198 deletions

View file

@ -514,6 +514,10 @@ namespace eq
definition is_set.elims [H : is_set A] : square p₁₀ p₁₂ p₀₁ p₂₁ :=
square_of_eq !is_set.elim
definition is_trunc_square [instance] (n : trunc_index) [H : is_trunc n .+2 A]
: is_trunc n (square p₁₀ p₁₂ p₀₁ p₂₁) :=
is_trunc_equiv_closed_rev n !square_equiv_eq
-- definition square_of_con_inv_hsquare {p₁ p₂ p₃ p₄ : a₁ = a₂}
-- {t : p₁ = p₂} {b : p₃ = p₄} {l : p₁ = p₃} {r : p₂ = p₄}
-- (s : square (con_inv_eq_idp t) (con_inv_eq_idp b) (l ◾ r⁻²) idp)

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn
Declaration of the pushout
-/
import .quotient cubical.square types.sigma
import .quotient types.sigma types.arrow_2
open quotient eq sum equiv is_trunc
@ -120,6 +120,11 @@ namespace pushout
apply eq_pathover, apply hdeg_square, esimp, apply elim_glue},
end
/- glue squares -/
protected definition glue_square {x x' : TL} (p : x = x')
: square (glue x) (glue x') (ap inl (ap f p)) (ap inr (ap g p)) :=
by cases p; apply vrefl
end pushout
open function sigma.ops
@ -231,4 +236,117 @@ namespace pushout
end
-- Functoriality of pushouts
section
section lemmas
variables {X : Type} {x₀ x₁ x₂ x₃ : X}
(p : x₀ = x₁) (q : x₁ = x₂) (r : x₂ = x₃)
private definition is_equiv_functor_lemma₁
: (r ⬝ ((p ⬝ q ⬝ r)⁻¹ ⬝ p)) = q⁻¹ :=
by cases p; cases r; cases q; reflexivity
private definition is_equiv_functor_lemma₂
: (p ⬝ q ⬝ r)⁻¹ ⬝ (p ⬝ q) = r⁻¹ :=
by cases p; cases r; cases q; reflexivity
end lemmas
variables {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
{TL' BL' TR' : Type} (f' : TL' → BL') (g' : TL' → TR')
(tl : TL → TL') (bl : BL → BL') (tr : TR → TR')
(fh : bl ∘ f ~ f' ∘ tl) (gh : tr ∘ g ~ g' ∘ tl)
include fh gh
protected definition functor [reducible] : pushout f g → pushout f' g' :=
begin
intro x, induction x with a b z,
{ exact inl (bl a) },
{ exact inr (tr b) },
{ exact (ap inl (fh z)) ⬝ glue (tl z) ⬝ (ap inr (gh z)⁻¹) }
end
protected definition ap_functor_inl [reducible] {x x' : BL} (p : x = x')
: ap (pushout.functor f g f' g' tl bl tr fh gh) (ap inl p) = ap inl (ap bl p) :=
by cases p; reflexivity
protected definition ap_functor_inr [reducible] {x x' : TR} (p : x = x')
: ap (pushout.functor f g f' g' tl bl tr fh gh) (ap inr p) = ap inr (ap tr p) :=
by cases p; reflexivity
variables [ietl : is_equiv tl] [iebl : is_equiv bl] [ietr : is_equiv tr]
include ietl iebl ietr
open equiv is_equiv arrow
protected definition is_equiv_functor [instance]
: is_equiv (pushout.functor f g f' g' tl bl tr fh gh) :=
adjointify
(pushout.functor f g f' g' tl bl tr fh gh)
(pushout.functor f' g' f g tl⁻¹ bl⁻¹ tr⁻¹
(inv_commute_of_commute tl bl f f' fh)
(inv_commute_of_commute tl tr g g' gh))
abstract begin
intro x', induction x' with a' b' z',
{ apply ap inl, apply right_inv },
{ apply ap inr, apply right_inv },
{ apply eq_pathover,
rewrite [ap_id,ap_compose' (pushout.functor f g f' g' tl bl tr fh gh)],
krewrite elim_glue,
rewrite [ap_inv,ap_con,ap_inv],
krewrite [pushout.ap_functor_inr], rewrite ap_con,
krewrite [pushout.ap_functor_inl,elim_glue],
apply transpose,
apply move_top_of_right, apply move_top_of_left',
krewrite [-(ap_inv inl),-ap_con,-(ap_inv inr),-ap_con],
apply move_top_of_right, apply move_top_of_left',
krewrite [-ap_con,-(ap_inv inl),-ap_con],
rewrite ap_bot_inv_commute_of_commute,
apply eq_hconcat (ap02 inl
(is_equiv_functor_lemma₁
(right_inv bl (f' z'))
(ap f' (right_inv tl z')⁻¹)
(fh (tl⁻¹ z'))⁻¹)),
rewrite [ap_inv f',inv_inv],
rewrite ap_bot_inv_commute_of_commute,
refine hconcat_eq _ (ap02 inr
(is_equiv_functor_lemma₁
(right_inv tr (g' z'))
(ap g' (right_inv tl z')⁻¹)
(gh (tl⁻¹ z'))⁻¹))⁻¹,
rewrite [ap_inv g',inv_inv],
apply pushout.glue_square }
end end
abstract begin
intro x, induction x with a b z,
{ apply ap inl, apply left_inv },
{ apply ap inr, apply left_inv },
{ apply eq_pathover,
rewrite [ap_id,ap_compose'
(pushout.functor f' g' f g tl⁻¹ bl⁻¹ tr⁻¹ _ _)
(pushout.functor f g f' g' tl bl tr _ _)],
krewrite elim_glue,
rewrite [ap_inv,ap_con,ap_inv],
krewrite [pushout.ap_functor_inr], rewrite ap_con,
krewrite [pushout.ap_functor_inl,elim_glue],
apply transpose,
apply move_top_of_right, apply move_top_of_left',
krewrite [-(ap_inv inl),-ap_con,-(ap_inv inr),-ap_con],
apply move_top_of_right, apply move_top_of_left',
krewrite [-ap_con,-(ap_inv inl),-ap_con],
rewrite inv_commute_of_commute_top,
apply eq_hconcat (ap02 inl
(is_equiv_functor_lemma₂
(ap bl⁻¹ (fh z))⁻¹
(left_inv bl (f z))
(ap f (left_inv tl z)⁻¹))),
rewrite [ap_inv f,inv_inv],
rewrite inv_commute_of_commute_top,
refine hconcat_eq _ (ap02 inr
(is_equiv_functor_lemma₂
(ap tr⁻¹ (gh z))⁻¹
(left_inv tr (g z))
(ap g (left_inv tl z)⁻¹)))⁻¹,
rewrite [ap_inv g,inv_inv],
apply pushout.glue_square }
end end
end
end pushout

View file

@ -10,7 +10,7 @@ import .sphere
import types.bool types.int.hott types.equiv
import algebra.homotopy_group algebra.hott .connectedness
open eq susp bool sphere_index is_equiv equiv is_trunc pi algebra homotopy
open eq susp bool sphere_index is_equiv equiv is_trunc is_conn pi algebra
definition circle : Type₀ := sphere 1
@ -130,6 +130,30 @@ namespace circle
rewrite [▸*,-apdo_eq_pathover_of_eq_ap,↑circle.elim,rec_loop],
end
theorem elim_seg1 {P : Type} (Pbase : P) (Ploop : Pbase = Pbase)
: ap (circle.elim Pbase Ploop) seg1 = (tr_constant seg1 Pbase)⁻¹ :=
begin
apply eq_of_fn_eq_fn_inv !(pathover_constant seg1),
rewrite [▸*,-apdo_eq_pathover_of_eq_ap,↑circle.elim,↑circle.rec],
rewrite [↑circle.rec2_on,rec2_seg1], apply inverse,
apply pathover_of_eq_tr_constant_inv
end
theorem elim_seg2 {P : Type} (Pbase : P) (Ploop : Pbase = Pbase)
: ap (circle.elim Pbase Ploop) seg2 = Ploop ⬝ (tr_constant seg1 Pbase)⁻¹ :=
begin
apply eq_of_fn_eq_fn_inv !(pathover_constant seg2),
rewrite [▸*,-apdo_eq_pathover_of_eq_ap,↑circle.elim,↑circle.rec],
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')⁻¹)
:> b =[p] p' ▸ b',
{ intros, cases q, cases p', cases p, reflexivity },
apply l
end
protected definition elim_type (Pbase : Type) (Ploop : Pbase ≃ Pbase)
(x : circle) : Type :=
circle.elim Pbase (ua Ploop) x
@ -269,34 +293,25 @@ namespace circle
induction H', reflexivity}
end
definition is_trunc_circle [instance] : is_trunc 1 S¹ :=
proposition is_trunc_circle [instance] : is_trunc 1 S¹ :=
begin
apply is_trunc_succ_of_is_trunc_loop,
{ apply trunc_index.minus_one_le_succ},
{ intro x, apply is_trunc_equiv_closed_rev, apply eq_equiv_Z}
end
definition is_conn_circle [instance] : is_conn 0 S¹ :=
proposition is_conn_circle [instance] : is_conn 0 circle :=
sphere.is_conn_sphere -1.+2
definition circle_turn [reducible] (x : S¹) : x = x :=
begin
fapply is_contr.mk,
{ exact tr base},
{ intro x, induction x with x,
induction x,
{ reflexivity},
{ apply is_prop.elimo}}
induction x,
{ exact loop },
{ apply eq_pathover, apply square_of_eq, rewrite ap_id }
end
definition circle_mul [reducible] (x y : S¹) : S¹ :=
begin
induction x,
{ induction y,
{ exact base },
{ exact loop } },
{ induction y,
{ exact loop },
{ apply eq_pathover, rewrite elim_loop,
apply square_of_eq, reflexivity } }
end
circle.elim y (circle_turn y) x
definition circle_mul_base (x : S¹) : circle_mul x base = x :=
begin
@ -305,11 +320,8 @@ namespace circle
{ apply eq_pathover, krewrite [elim_loop,ap_id], apply hrefl }
end
definition circle_base_mul (x : S¹) : circle_mul base x = x :=
begin
induction x,
{ reflexivity },
{ apply eq_pathover, krewrite [elim_loop,ap_id], apply hrefl }
end
definition circle_base_mul [reducible] (x : S¹)
: circle_mul base x = x :=
idp
end circle

View file

@ -3,11 +3,11 @@ Copyright (c) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz, Floris van Doorn
-/
import types.trunc types.arrow_2 .sphere
import types.trunc types.arrow_2 types.fiber
open eq is_trunc is_equiv nat equiv trunc function fiber funext pi
namespace homotopy
namespace is_conn
definition is_conn [reducible] (n : ℕ₋₂) (A : Type) : Type :=
is_contr (trunc n A)
@ -248,39 +248,7 @@ namespace homotopy
-- all types are -2-connected
definition is_conn_minus_two (A : Type) : is_conn -2 A :=
_
-- Theorem 8.2.1
open susp
theorem is_conn_susp [instance] (n : ℕ₋₂) (A : Type)
[H : is_conn n A] : is_conn (n .+1) (susp A) :=
is_contr.mk (tr north)
begin
apply trunc.rec,
fapply susp.rec,
{ reflexivity },
{ exact (trunc.rec (λa, ap tr (merid a)) (center (trunc n A))) },
{ intro a,
generalize (center (trunc n A)),
apply trunc.rec,
intro a',
apply pathover_of_tr_eq,
rewrite [transport_eq_Fr,idp_con],
revert H, induction n with [n, IH],
{ intro H, apply is_prop.elim },
{ intros H,
change ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid a'),
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),
reflexivity
}
}
end
is_trunc_trunc -2 A
-- Lemma 7.5.14
theorem is_equiv_trunc_functor_of_is_conn_fun {A B : Type} (n : ℕ₋₂) (f : A → B)
@ -297,19 +265,4 @@ namespace homotopy
[H : is_conn_fun n f] : trunc n A ≃ trunc n B :=
equiv.mk (trunc_functor n f) (is_equiv_trunc_functor_of_is_conn_fun n f)
open trunc_index pointed sphere.ops
-- Corollary 8.2.2
theorem is_conn_sphere [instance] (n : ℕ₋₁) : is_conn (n..-1) (S n) :=
begin
induction n with n IH,
{ apply is_conn_minus_two},
{ rewrite [succ_sub_one, sphere.sphere_succ], apply is_conn_susp}
end
section
open sphere_index
theorem is_conn_psphere [instance] (n : ) : is_conn (n.-1) (S. n) :=
transport (λx, is_conn x (sphere n)) (of_nat_sub_one n) (is_conn_sphere n)
end
end homotopy
end is_conn

View file

@ -5,9 +5,9 @@ Authors: Floris van Doorn, Clive Newstead
-/
import algebra.homotopy_group .connectedness
import algebra.homotopy_group .sphere
open eq is_trunc trunc_index pointed algebra trunc nat homotopy fiber pointed
open eq is_trunc trunc_index pointed algebra trunc nat is_conn fiber pointed
namespace is_trunc
-- Lemma 8.3.1
@ -33,13 +33,13 @@ namespace is_trunc
-- Corollary 8.3.3
section
open sphere.ops sphere_index
open sphere sphere.ops sphere_index
theorem homotopy_group_sphere_le (n k : ) (H : k < n) : is_contr (π[k] (S. n)) :=
begin
cases n with n,
{ exfalso, apply not_lt_zero, exact H},
{ have H2 : k ≤ n, from le_of_lt_succ H,
apply @(trivial_homotopy_group_of_is_conn _ H2)}
apply @(trivial_homotopy_group_of_is_conn _ H2) }
end
end

View file

@ -1,94 +1,267 @@
/-
Copyright (c) 2015 Jakob von Raumer. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jakob von Raumer
Authors: Jakob von Raumer, Ulrik Buchholtz
Declaration of a join as a special case of a pushout
-/
import hit.pushout .susp cubical.cube cubical.squareover
import hit.pushout .sphere cubical.cube
open eq function prod equiv pushout is_trunc bool sigma.ops function
open eq function prod equiv is_trunc bool sigma.ops
definition join (A B : Type) : Type := @pushout.pushout (A × B) A B pr1 pr2
namespace join
section
variables {A B : Type}
definition inl (a : A) : join A B := @pushout.inl (A × B) A B pr1 pr2 a
definition inr (b : B) : join A B := @pushout.inr (A × B) A B pr1 pr2 b
definition glue (a : A) (b : B) : inl a = inr b :=
@pushout.glue (A × B) A B pr1 pr2 (a, b)
protected definition rec {P : join A B → Type}
(Pinl : Π(x : A), P (inl x))
(Pinr : Π(y : B), P (inr y))
(Pglue : Π(x : A)(y : B), Pinl x =[glue x y] Pinr y)
(z : join A B) : P z :=
pushout.rec Pinl Pinr (prod.rec Pglue) z
protected definition rec_glue {P : join A B → Type}
(Pinl : Π(x : A), P (inl x))
(Pinr : Π(y : B), P (inr y))
(Pglue : Π(x : A)(y : B), Pinl x =[glue x y] Pinr y)
(x : A) (y : B)
: apdo (join.rec Pinl Pinr Pglue) (glue x y) = Pglue x y :=
!quotient.rec_eq_of_rel
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
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)
: ap (join.elim Pinl Pinr Pglue) (glue x y) = Pglue x y :=
begin
apply equiv.eq_of_fn_eq_fn_inv !(pathover_constant (glue x y)),
rewrite [▸*,-apdo_eq_pathover_of_eq_ap,↑join.elim],
apply join.rec_glue
end
protected definition elim_ap_inl {P : Type} (Pinl : A → P) (Pinr : B → P)
(Pglue : Π(x : A)(y : B), Pinl x = Pinr y) {a a' : A} (p : a = a')
: ap (join.elim Pinl Pinr Pglue) (ap inl p) = ap Pinl p :=
by cases p; reflexivity
protected definition hsquare {a a' : A} {b b' : B} (p : a = a') (q : b = b') :
square (ap inl p) (ap inr q) (glue a b) (glue a' b') :=
eq.rec_on p (eq.rec_on q hrfl)
protected definition vsquare {a a' : A} {b b' : B} (p : a = a') (q : b = b') :
square (glue a b) (glue a' b') (ap inl p) (ap inr q) :=
eq.rec_on p (eq.rec_on q vrfl)
end
end join
attribute join.inl join.inr [constructor]
attribute join.rec [recursor]
attribute join.elim [recursor 7]
attribute join.rec join.elim [unfold 7]
/- Diamonds in joins -/
namespace join
variables {A B : Type}
protected definition diamond (a a' : A) (b b' : B) :=
square (glue a b) (glue a' b')⁻¹ (glue a b') (glue a' b)⁻¹
protected definition hdiamond {a a' : A} (b b' : B) (p : a = a')
: join.diamond a a' b b' :=
begin
cases p, unfold join.diamond,
assert H : (glue a b' ⬝ (glue a b')⁻¹ ⬝ (glue a b)⁻¹⁻¹) = glue a b,
{ rewrite [con.right_inv,inv_inv,idp_con] },
exact H ▸ top_deg_square (glue a b') (glue a b')⁻¹ (glue a b)⁻¹,
end
protected definition vdiamond (a a' : A) {b b' : B} (q : b = b')
: join.diamond a a' b b' :=
begin
cases q, unfold join.diamond,
assert H : (glue a b ⬝ (glue a' b)⁻¹ ⬝ (glue a' b)⁻¹⁻¹) = glue a b,
{ rewrite [con.assoc,con.right_inv] },
exact H ▸ top_deg_square (glue a b) (glue a' b)⁻¹ (glue a' b)⁻¹
end
protected definition symm_diamond (a : A) (b : B)
: join.vdiamond a a idp = join.hdiamond b b idp :=
begin
unfold join.hdiamond, unfold join.vdiamond,
assert H : Π{X : Type} ⦃x y : X⦄ (p : x = y),
eq.rec (eq.rec (refl p) (symm (con.right_inv p⁻¹)))
(symm (con.assoc p p⁻¹ p⁻¹⁻¹)) ▸ top_deg_square p p⁻¹ p⁻¹
= eq.rec (eq.rec (eq.rec (refl p) (symm (idp_con p))) (symm (inv_inv p)))
(symm (con.right_inv p)) ▸ top_deg_square p p⁻¹ p⁻¹
:> square p p⁻¹ p p⁻¹,
{ intros X x y p, cases p, reflexivity },
apply H (glue a b)
end
end join
namespace join
variables {A₁ A₂ B₁ B₂ : Type}
protected definition functor [reducible]
(f : A₁ → A₂) (g : B₁ → B₂) : join A₁ B₁ → join A₂ B₂ :=
begin
intro x, induction x with a b a b,
{ exact inl (f a) }, { exact inr (g b) }, { apply glue }
end
protected definition ap_diamond (f : A₁ → A₂) (g : B₁ → B₂)
{a a' : A₁} {b b' : B₁}
: join.diamond a a' b b' → join.diamond (f a) (f a') (g b) (g b') :=
begin
unfold join.diamond, intro s,
note s' := aps (join.functor f g) s,
do 2 rewrite eq.ap_inv at s',
do 4 rewrite join.elim_glue at s', exact s'
end
protected definition equiv_closed
: A₁ ≃ A₂ → B₁ ≃ B₂ → join A₁ B₁ ≃ join A₂ B₂ :=
begin
intros H K,
fapply equiv.MK,
{ intro x, induction x with a b a b,
{ exact inl (to_fun H a) }, { exact inr (to_fun K b) },
{ apply glue } },
{ intro y, induction y with a b a b,
{ exact inl (to_inv H a) }, { exact inr (to_inv K b) },
{ apply glue } },
{ intro y, induction y with a b a b,
{ apply ap inl, apply to_right_inv },
{ apply ap inr, apply to_right_inv },
{ apply eq_pathover, rewrite ap_id,
rewrite (ap_compose' (join.elim _ _ _)),
do 2 krewrite join.elim_glue, apply join.hsquare } },
{ intro x, induction x with a b a b,
{ apply ap inl, apply to_left_inv },
{ apply ap inr, apply to_left_inv },
{ apply eq_pathover, rewrite ap_id,
rewrite (ap_compose' (join.elim _ _ _)),
do 2 krewrite join.elim_glue, apply join.hsquare } }
end
protected definition twist_diamond {A : Type} {a a' : A} (p : a = a')
: pathover (λx, join.diamond a' x a x)
(join.vdiamond a' a idp) p
(join.hdiamond a a' idp) :=
begin
cases p, apply pathover_idp_of_eq, apply join.symm_diamond
end
protected definition empty (A : Type) : join empty A ≃ A :=
begin
fapply equiv.MK,
{ intro x, induction x with z a z a,
{ induction z },
{ exact a },
{ induction z } },
{ intro a, exact inr a },
{ intro a, reflexivity },
{ intro x, induction x with z a z a,
{ induction z },
{ reflexivity },
{ induction z } }
end
protected definition bool (A : Type) : join bool A ≃ susp A :=
begin
fapply equiv.MK,
{ intro ba, induction ba with [b, a, b, a],
{ induction b, exact susp.south, exact susp.north },
{ exact susp.north },
{ induction b, esimp,
{ apply inverse, apply susp.merid, exact a },
{ reflexivity } } },
{ intro s, induction s with a,
{ exact inl tt },
{ exact inl ff },
{ exact (glue tt a) ⬝ (glue ff a)⁻¹ } },
{ intro s, induction s with a,
{ reflexivity },
{ reflexivity },
{ esimp, apply eq_pathover, rewrite ap_id,
rewrite (ap_compose' (join.elim _ _ _)),
rewrite [susp.elim_merid,ap_con,ap_inv],
krewrite [join.elim_glue,join.elim_glue],
esimp, rewrite [inv_inv,idp_con],
apply hdeg_square, reflexivity } },
{ intro ba, induction ba with [b, a, b, a], esimp,
{ induction b, do 2 reflexivity },
{ apply glue },
{ induction b,
{ esimp, apply eq_pathover, rewrite ap_id,
rewrite (ap_compose' (susp.elim _ _ _)),
krewrite join.elim_glue, rewrite ap_inv,
krewrite susp.elim_merid,
apply square_of_eq_top, apply inverse,
rewrite con.assoc, apply con.left_inv },
{ esimp, apply eq_pathover, rewrite ap_id,
rewrite (ap_compose' (susp.elim _ _ _)),
krewrite join.elim_glue, esimp,
apply square_of_eq_top,
rewrite [idp_con,con.right_inv] } } }
end
end join
namespace join
variables (A B C : Type)
definition join : Type := @pushout (A × B) A B pr1 pr2
definition jglue {A B : Type} (a : A) (b : B) := @glue (A × B) A B pr1 pr2 (a, b)
protected definition is_contr [HA : is_contr A] :
is_contr (join A B) :=
begin
fapply is_contr.mk, exact inl (center A),
intro x, induction x with a b, apply ap inl, apply center_eq,
apply jglue, induction x with a b, apply pathover_of_tr_eq,
intro x, induction x with a b a b, apply ap inl, apply center_eq,
apply glue, apply pathover_of_tr_eq,
apply concat, apply transport_eq_Fr, esimp, rewrite ap_id,
generalize center_eq a, intro p, cases p, apply idp_con,
end
protected definition bool : join bool A ≃ susp A :=
begin
fapply equiv.MK, intro ba, induction ba with b a,
induction b, exact susp.south, exact susp.north, exact susp.north,
induction x with b a, esimp,
induction b, apply inverse, apply susp.merid, exact a, reflexivity,
intro s, induction s with m,
exact inl tt, exact inl ff, exact (jglue tt m) ⬝ (jglue ff m)⁻¹,
intros, induction b with m, do 2 reflexivity, esimp,
apply eq_pathover, apply hconcat, apply hdeg_square, apply concat,
apply ap_compose' (pushout.elim _ _ _), apply concat,
apply ap (ap (pushout.elim _ _ _)), apply susp.elim_merid, apply ap_con,
apply hconcat, apply vconcat, apply hdeg_square, apply elim_glue,
apply hdeg_square, apply ap_inv, esimp,
apply hconcat, apply hdeg_square, apply concat, apply idp_con,
apply concat, apply ap inverse, apply elim_glue, apply inv_inv,
apply hinverse, apply hdeg_square, apply ap_id,
intro x, induction x with b a, induction b, do 2 reflexivity,
esimp, apply jglue, induction x with b a, induction b, esimp,
apply eq_pathover, rewrite ap_id,
apply eq_hconcat, apply concat, apply ap_compose' (susp.elim _ _ _),
apply concat, apply ap (ap _) !elim_glue,
apply concat, apply ap_inv,
apply concat, apply ap inverse !susp.elim_merid,
apply concat, apply con_inv, apply ap (λ x, x ⬝ _) !inv_inv,
apply square_of_eq_top, apply inverse,
apply concat, apply ap (λ x, x ⬝ _) !con.assoc,
rewrite [con.left_inv, con_idp], apply con.right_inv,
esimp, apply eq_pathover, rewrite ap_id,
apply eq_hconcat, apply concat, apply ap_compose' (susp.elim _ _ _),
apply concat, apply ap (ap _) !elim_glue, esimp, reflexivity,
apply square_of_eq_top, rewrite idp_con, apply !con.right_inv⁻¹,
end
protected definition swap : join A B → join B A :=
begin
intro x, induction x with a b, exact inr a, exact inl b,
apply !jglue⁻¹
intro x, induction x with a b a b, exact inr a, exact inl b,
apply !glue⁻¹
end
protected definition swap_involutive (x : join A B) :
join.swap B A (join.swap A B x) = x :=
begin
induction x with a b, do 2 reflexivity,
induction x with a b, esimp,
induction x with a b a b, do 2 reflexivity,
apply eq_pathover, rewrite ap_id,
apply hdeg_square, esimp[join.swap],
apply concat, apply ap_compose' (pushout.elim _ _ _),
krewrite [elim_glue, ap_inv, elim_glue], apply inv_inv,
apply concat, apply ap_compose' (join.elim _ _ _),
krewrite [join.elim_glue, ap_inv, join.elim_glue], apply inv_inv,
end
protected definition symm : join A B ≃ join B A :=
by fapply equiv.MK; do 2 apply join.swap; do 2 apply join.swap_involutive
end
end join
/- This proves that the join operator is associative.
The proof is more or less ported from Evan Cavallo's agda version:
https://github.com/HoTT/HoTT-Agda/blob/master/homotopy/JoinAssocCubical.agda -/
/- This proves that the join operator is associative.
The proof is more or less ported from Evan Cavallo's agda version:
https://github.com/HoTT/HoTT-Agda/blob/master/homotopy/JoinAssocCubical.agda -/
namespace join
section join_switch
section join_switch
private definition massage_sq' {A : Type} {a₀₀ a₂₀ a₀₂ a₂₂ : A}
{p₁₀ : a₀₀ = a₂₀} {p₁₂ : a₀₂ = a₂₂} {p₀₁ : a₀₀ = a₀₂} {p₂₁ : a₂₀ = a₂₂}
@ -140,111 +313,122 @@ namespace join
private definition square_Flr_ap_idp_cube {A B : Type} {b : B} {f : A → B}
{p₁ p₂ : Π a, f a = b} (α : Π a, p₁ a = p₂ a) {a₁ a₂ : A} (q : a₁ = a₂) :
cube hrfl hrfl (square_Flr_ap_idp p₁ q) (square_Flr_ap_idp p₂ q)
cube hrfl hrfl (square_Flr_ap_idp p₁ q) (square_Flr_ap_idp p₂ q)
(hdeg_square (α _)) (hdeg_square (α _)) :=
by cases q; esimp[square_Flr_ap_idp]; apply deg3_cube; esimp
variables {A B C : Type}
private definition switch_left [reducible] : join A B → join (join C B) A :=
definition switch_left [reducible] : join A B → join (join C B) A :=
begin
intro x, induction x with a b, exact inr a, exact inl (inr b), apply !jglue⁻¹,
intro x, induction x with a b a b, exact inr a, exact inl (inr b), apply !glue⁻¹,
end
private definition switch_coh_fill (a : A) (b : B) (c : C) :
Σ sq : square (jglue (inl c) a)⁻¹ (ap inl (jglue c b))⁻¹ (ap switch_left (jglue a b)) idp,
cube (hdeg_square !elim_glue) ids sq (massage_sq !square_Flr_ap_idp) hrfl hrfl :=
private definition switch_coh_fill_square (a : A) (b : B) (c : C) :=
square (glue (inl c) a)⁻¹ (ap inl (glue c b))⁻¹ (ap switch_left (glue a b)) idp
private definition switch_coh_fill_cube (a : A) (b : B) (c : C)
(sq : switch_coh_fill_square a b c) :=
cube (hdeg_square !join.elim_glue) ids
sq (massage_sq !square_Flr_ap_idp)
hrfl hrfl
private definition switch_coh_fill_type (a : A) (b : B) (c : C) :=
Σ sq : switch_coh_fill_square a b c, switch_coh_fill_cube a b c sq
private definition switch_coh_fill (a : A) (b : B) (c : C)
: switch_coh_fill_type a b c :=
by esimp; apply cube_fill101
private definition switch_coh (ab : join A B) (c : C) : switch_left ab = inl (inl c) :=
begin
induction ab with a b, apply !jglue⁻¹, apply (ap inl !jglue)⁻¹, induction x with a b,
induction ab with a b a b, apply !glue⁻¹, apply (ap inl !glue)⁻¹,
apply eq_pathover, refine _ ⬝hp !ap_constant⁻¹,
apply !switch_coh_fill.1,
end
protected definition switch [reducible] : join (join A B) C → join (join C B) A :=
begin
intro x, induction x with ab c, exact switch_left ab, exact inl (inl c),
induction x with ab c, exact switch_coh ab c,
intro x, induction x with ab c ab c, exact switch_left ab, exact inl (inl c),
exact switch_coh ab c,
end
private definition switch_inv_left_square (a : A) (b : B) :
square idp idp (ap (!(@join.switch C) ∘ switch_left) (jglue a b)) (ap inl (jglue a b)) :=
square idp idp (ap (!(@join.switch C) ∘ switch_left) (glue a b)) (ap inl (glue a b)) :=
begin
refine hdeg_square !ap_compose ⬝h _,
refine aps join.switch (hdeg_square !elim_glue) ⬝h _, esimp,
refine aps join.switch (hdeg_square !join.elim_glue) ⬝h _, esimp,
refine hdeg_square !(ap_inv join.switch) ⬝h _,
refine hrfl⁻¹ʰ⁻¹ᵛ ⬝h _, esimp[join.switch,switch_left,switch_coh],
refine (hdeg_square !elim_glue)⁻¹ᵛ ⬝h _, esimp,
refine (hdeg_square !join.elim_glue)⁻¹ᵛ ⬝h _, esimp,
refine hrfl⁻¹ᵛ ⬝h _, apply hdeg_square !inv_inv,
end
private definition switch_inv_coh_left (c : C) (a : A) :
square idp idp (ap !(@join.switch C B) (switch_coh (inl a) c)) (jglue (inl a) c) :=
square idp idp (ap !(@join.switch C B) (switch_coh (inl a) c)) (glue (inl a) c) :=
begin
refine hrfl ⬝h _,
refine aps join.switch hrfl ⬝h _, esimp[switch_coh],
refine hdeg_square !ap_inv ⬝h _,
refine hrfl⁻¹ʰ⁻¹ᵛ ⬝h _, esimp[join.switch,switch_left],
refine (hdeg_square !elim_glue)⁻¹ᵛ ⬝h _,
refine (hdeg_square !join.elim_glue)⁻¹ᵛ ⬝h _,
refine hrfl⁻¹ᵛ ⬝h _, apply hdeg_square !inv_inv,
end
private definition switch_inv_coh_right (c : C) (b : B) :
square idp idp (ap !(@join.switch _ _ A) (switch_coh (inr b) c)) (jglue (inr b) c) :=
square idp idp (ap !(@join.switch _ _ A) (switch_coh (inr b) c)) (glue (inr b) c) :=
begin
refine hrfl ⬝h _,
refine aps join.switch hrfl ⬝h _, esimp[switch_coh],
refine hdeg_square !ap_inv ⬝h _,
refine (hdeg_square !ap_compose)⁻¹ʰ⁻¹ᵛ ⬝h _,
refine hrfl⁻¹ᵛ ⬝h _, esimp[join.switch,switch_left],
refine (hdeg_square !elim_glue)⁻¹ᵛ ⬝h _, apply hdeg_square !inv_inv,
refine (hdeg_square !join.elim_glue)⁻¹ᵛ ⬝h _, apply hdeg_square !inv_inv,
end
private definition switch_inv_left (ab : join A B) :
!(@join.switch C) (join.switch (inl ab)) = inl ab :=
begin
induction ab with a b, do 2 reflexivity,
induction x with a b, apply eq_pathover, exact !switch_inv_left_square,
induction ab with a b a b, do 2 reflexivity,
apply eq_pathover, exact !switch_inv_left_square,
end
section
variables (a : A) (b : B) (c : C)
variables (a : A) (b : B) (c : C)
private definition switch_inv_cube_aux1 {A B C : Type} {b : B} {f : A → B} (h : B → C)
(g : Π a, f a = b) {x y : A} (p : x = y) :
cube (hdeg_square (ap_compose h f p)) ids (square_Flr_ap_idp (λ a, ap h (g a)) p)
(aps h (square_Flr_ap_idp _ _)) hrfl hrfl :=
by cases p; esimp[square_Flr_ap_idp]; apply deg2_cube; cases (g x); esimp
private definition switch_inv_cube_aux1 {A B C : Type} {b : B} {f : A → B} (h : B → C)
(g : Π a, f a = b) {x y : A} (p : x = y) :
cube (hdeg_square (ap_compose h f p)) ids (square_Flr_ap_idp (λ a, ap h (g a)) p)
(aps h (square_Flr_ap_idp _ _)) hrfl hrfl :=
by cases p; esimp[square_Flr_ap_idp]; apply deg2_cube; cases (g x); esimp
private definition switch_inv_cube_aux2 {A B : Type} {b : B} {f : A → B}
(g : Π a, f a = b) {x y : A} (p : x = y) {sq : square (g x) (g y) (ap f p) idp}
(q : apdo g p = eq_pathover (sq ⬝hp !ap_constant⁻¹)) : square_Flr_ap_idp _ _ = sq :=
begin
cases p, esimp at *, apply concat, apply inverse, apply vdeg_square_idp,
apply concat, apply ap vdeg_square, exact ap eq_of_pathover_idp q,
krewrite (is_equiv.right_inv (equiv.to_fun !pathover_idp)),
exact is_equiv.left_inv (equiv.to_fun (vdeg_square_equiv _ _)) sq,
end
private definition switch_inv_cube_aux2 {A B : Type} {b : B} {f : A → B}
(g : Π a, f a = b) {x y : A} (p : x = y) {sq : square (g x) (g y) (ap f p) idp}
(q : apdo g p = eq_pathover (sq ⬝hp !ap_constant⁻¹)) : square_Flr_ap_idp _ _ = sq :=
begin
cases p, esimp at *, apply concat, apply inverse, apply vdeg_square_idp,
apply concat, apply ap vdeg_square, exact ap eq_of_pathover_idp q,
krewrite (is_equiv.right_inv (equiv.to_fun !pathover_idp)),
exact is_equiv.left_inv (equiv.to_fun (vdeg_square_equiv _ _)) sq,
end
private definition switch_inv_cube (a : A) (b : B) (c : C) :
cube (switch_inv_left_square a b) ids (square_Flr_ap_idp _ _)
(square_Flr_ap_idp _ _) (switch_inv_coh_left c a) (switch_inv_coh_right c b) :=
begin
esimp [switch_inv_coh_left, switch_inv_coh_right, switch_inv_left_square],
apply cube_concat2, apply switch_inv_cube_aux1,
apply cube_concat2, apply cube_transport101, apply inverse,
apply ap (λ x, aps join.switch x), apply switch_inv_cube_aux2, apply rec_glue,
apply apc, apply (switch_coh_fill a b c).2,
apply cube_concat2, esimp, apply ap_square_massage,
apply cube_concat2, apply massage_cube, apply cube_inverse2, apply switch_inv_cube_aux1,
apply cube_concat2, apply massage_cube, apply square_Flr_ap_idp_cube,
apply cube_concat2, apply massage_cube, apply cube_transport101,
apply inverse, apply switch_inv_cube_aux2,
esimp[switch_coh], apply rec_glue, apply (switch_coh_fill c b a).2,
apply massage_massage,
end
private definition switch_inv_cube (a : A) (b : B) (c : C) :
cube (switch_inv_left_square a b) ids (square_Flr_ap_idp _ _)
(square_Flr_ap_idp _ _) (switch_inv_coh_left c a) (switch_inv_coh_right c b) :=
begin
esimp [switch_inv_coh_left, switch_inv_coh_right, switch_inv_left_square],
apply cube_concat2, apply switch_inv_cube_aux1,
apply cube_concat2, apply cube_transport101, apply inverse,
apply ap (λ x, aps join.switch x), apply switch_inv_cube_aux2, apply join.rec_glue,
apply apc, apply (switch_coh_fill a b c).2,
apply cube_concat2, esimp, apply ap_square_massage,
apply cube_concat2, apply massage_cube, apply cube_inverse2, apply switch_inv_cube_aux1,
apply cube_concat2, apply massage_cube, apply square_Flr_ap_idp_cube,
apply cube_concat2, apply massage_cube, apply cube_transport101,
apply inverse, apply switch_inv_cube_aux2,
esimp[switch_coh], apply join.rec_glue, apply (switch_coh_fill c b a).2,
apply massage_massage,
end
end
@ -274,25 +458,25 @@ namespace join
end
private definition switch_inv_coh (c : C) (k : join A B) :
square (switch_inv_left k) idp (ap join.switch (switch_coh k c)) (jglue k c) :=
square (switch_inv_left k) idp (ap join.switch (switch_coh k c)) (glue k c) :=
begin
induction k, apply switch_inv_coh_left, apply switch_inv_coh_right,
induction k with a b a b, apply switch_inv_coh_left, apply switch_inv_coh_right,
refine pathover_of_triangle_cube _,
induction x with [a, b], esimp, apply cube_transport011,
esimp, apply cube_transport011,
apply inverse, rotate 1, apply switch_inv_cube,
apply natural_square_tr_beta, apply rec_glue,
apply natural_square_tr_beta, apply join.rec_glue,
end
protected definition switch_involutive (x : join (join A B) C) :
join.switch (join.switch x) = x :=
begin
induction x, apply switch_inv_left, reflexivity,
induction x with ab c ab c, apply switch_inv_left, reflexivity,
apply pathover_of_ap_ap_square join.switch join.switch,
induction x with [k, c], krewrite elim_glue, esimp,
krewrite join.elim_glue, esimp,
apply transpose, exact !switch_inv_coh,
end
end join_switch
end join_switch
protected definition switch_equiv (A B C : Type) : join (join A B) C ≃ join (join C B) A :=
by apply equiv.MK; do 2 apply join.switch_involutive
@ -300,6 +484,43 @@ namespace join
protected definition assoc (A B C : Type) : join (join A B) C ≃ join A (join B C) :=
calc join (join A B) C ≃ join (join C B) A : join.switch_equiv
... ≃ join A (join C B) : join.symm
... ≃ join A (join B C) : join.symm
... ≃ join A (join B C) : join.equiv_closed erfl (join.symm C B)
protected definition ap_assoc_inv_glue_inl {A B : Type} (C : Type) (a : A) (b : B)
: ap (to_inv (join.assoc A B C)) (glue a (inl b)) = ap inl (glue a b) :=
begin
unfold join.assoc, unfold equiv.trans, rewrite ap_compose, krewrite join.elim_glue,
rewrite ap_compose, krewrite join.elim_glue, rewrite ap_inv, krewrite join.elim_glue,
unfold switch_coh, unfold join.symm, unfold join.swap, esimp, rewrite eq.inv_inv
end
protected definition ap_assoc_inv_glue_inr {A C : Type} (B : Type) (a : A) (c : C)
: ap (to_inv (join.assoc A B C)) (glue a (inr c)) = glue (inl a) c :=
begin
unfold join.assoc, unfold equiv.trans, rewrite ap_compose, krewrite join.elim_glue,
rewrite ap_compose, krewrite join.elim_glue, rewrite ap_inv, krewrite join.elim_glue,
unfold switch_coh, unfold join.symm, unfold join.swap, esimp, rewrite eq.inv_inv
end
end join
namespace join
open sphere sphere_index sphere.ops
protected definition spheres (n m : ℕ₋₁) : join (S n) (S m) ≃ S (n+1+m) :=
begin
apply equiv.trans (join.symm (S n) (S m)),
induction m with m IH,
{ exact join.empty (S n) },
{ calc join (S m.+1) (S n)
≃ join (join bool (S m)) (S n)
: join.equiv_closed (equiv.symm (join.bool (S m))) erfl
... ≃ join bool (join (S m) (S n))
: join.assoc
... ≃ join bool (S (n+1+m))
: join.equiv_closed erfl IH
... ≃ sphere (n+1+m.+1)
: join.bool (S (n+1+m)) }
end
end join

View file

@ -272,6 +272,23 @@ namespace sphere
end sphere
namespace sphere
open is_conn trunc_index sphere_index sphere.ops
-- Corollary 8.2.2
theorem is_conn_sphere [instance] (n : ℕ₋₁) : is_conn (n..-1) (S n) :=
begin
induction n with n IH,
{ apply is_conn_minus_two },
{ rewrite [trunc_index.succ_sub_one n, sphere.sphere_succ],
apply is_conn_susp }
end
theorem is_conn_psphere [instance] (n : ) : is_conn (n.-1) (S. n) :=
transport (λx, is_conn x (sphere n)) (of_nat_sub_one n) (is_conn_sphere n)
end sphere
open sphere sphere.ops
namespace is_trunc

View file

@ -1,12 +1,12 @@
/-
Copyright (c) 2015 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
Authors: Floris van Doorn, Ulrik Buchholtz
Declaration of suspension
-/
import hit.pushout types.pointed cubical.square
import hit.pushout types.pointed cubical.square .connectedness
open pushout unit eq equiv
@ -77,6 +77,72 @@ attribute susp.elim_type [unfold 5]
attribute susp.rec_on susp.elim_on [unfold 3]
attribute susp.elim_type_on [unfold 2]
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
apply trunc.rec,
fapply susp.rec,
{ reflexivity },
{ exact (trunc.rec (λa, ap tr (merid a)) (center (trunc n A))) },
{ intro a,
generalize (center (trunc n A)),
apply trunc.rec,
intro a',
apply pathover_of_tr_eq,
rewrite [transport_eq_Fr,idp_con],
revert H, induction n with [n, IH],
{ intro H, apply is_prop.elim },
{ intros H,
change ap (@tr n .+2 (susp A)) (merid a) = ap tr (merid a'),
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),
reflexivity
}
}
end
end susp
/-
namespace susp
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
/-
sigma P ≃ sigma P' (P' := pushout.elim_type (λx, PN) (λx, PS) Pm) : foo
≃ pushout F' G' : pushout.flattening
≃ pushout F G : pushout_functor_is_equiv
-/
exact sorry
end
end
end susp
-/
namespace susp
open pointed

View file

@ -30,7 +30,8 @@ namespace wedge
end
end wedge
open trunc is_trunc function homotopy
open trunc is_trunc is_conn function
namespace wedge_extension
section
-- The wedge connectivity lemma (Lemma 8.6.2)

View file

@ -253,7 +253,6 @@ namespace is_equiv
!ap_eq_of_fn_eq_fn'
end
end is_equiv
open is_equiv

View file

@ -138,6 +138,10 @@ namespace eq
{ 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' :=
by cases p; constructor
definition eq_of_pathover_idp [unfold 6] {b' : B a} (q : b =[idpath a] b') : b = b' :=
tr_eq_of_pathover q

View file

@ -109,3 +109,77 @@ namespace arrow
(λa, con_eq_of_eq_inv_con (ap_id _))
end arrow
namespace arrow
/-
equivalences in the arrow category; could be packaged into structures.
cannot be moved to types.pi because of the dependence on types.equiv.
-/
variables {A A' B B' : Type} (f : A → B) (f' : A' → B')
(α : A → A') (β : B → B')
[Hf : is_equiv f] [Hf' : is_equiv f']
include Hf Hf'
open function
definition inv_commute_of_commute (p : f' ∘ α ~ β ∘ f) : f'⁻¹ ∘ β ~ α ∘ f⁻¹ :=
begin
apply homotopy_inv_of_homotopy_post f' β (α ∘ f⁻¹),
apply homotopy.symm,
apply homotopy_inv_of_homotopy_pre f (f' ∘ α) β,
apply p
end
definition inv_commute_of_commute_top (p : f' ∘ α ~ β ∘ f) (a : A)
: inv_commute_of_commute f f' α β p (f a)
= (ap f'⁻¹ (p a))⁻¹ ⬝ left_inv f' (α a) ⬝ ap α (left_inv f a)⁻¹ :=
begin
unfold inv_commute_of_commute,
unfold homotopy_inv_of_homotopy_post,
unfold homotopy_inv_of_homotopy_pre,
rewrite [adj f a,-(ap_compose β f)],
rewrite [eq_of_square (natural_square_tr p (left_inv f a))],
rewrite [ap_inv f'⁻¹,ap_con f'⁻¹,con_inv,con.assoc,con.assoc],
apply whisker_left (ap f'⁻¹ (p a))⁻¹,
apply eq_of_square, rewrite [ap_inv α,-(ap_compose f'⁻¹ (f' ∘ α))],
apply hinverse, rewrite [ap_compose (f'⁻¹ ∘ f') α],
refine vconcat_eq _ (ap_id (ap α (left_inv f a))),
apply natural_square (left_inv f') (ap α (left_inv f a))
end
definition ap_bot_inv_commute_of_commute (p : f' ∘ α ~ β ∘ f) (b : B)
: ap f' (inv_commute_of_commute f f' α β p b)
= right_inv f' (β b) ⬝ ap β (right_inv f b)⁻¹ ⬝ (p (f⁻¹ b))⁻¹ :=
begin
unfold inv_commute_of_commute,
unfold homotopy_inv_of_homotopy_post,
unfold homotopy_inv_of_homotopy_pre,
rewrite [ap_con,-(ap_compose f' f'⁻¹),-(adj f' (α (f⁻¹ b)))],
rewrite [con.assoc (right_inv f' (β b)) (ap β (right_inv f b)⁻¹)
(p (f⁻¹ b))⁻¹],
apply eq_of_square,
refine vconcat_eq _
(whisker_right (ap_inv β (right_inv f b)) (p (f⁻¹ b))⁻¹)⁻¹,
refine vconcat_eq _
(con_inv (p (f⁻¹ b)) (ap β (right_inv f b))),
refine vconcat_eq _
(ap_id (p (f⁻¹ b) ⬝ ap β (right_inv f b))⁻¹),
apply natural_square (right_inv f')
(p (f⁻¹ b) ⬝ ap β (right_inv f b))⁻¹
end
definition is_equiv_inv_commute_of_commute
: is_equiv (inv_commute_of_commute f f' α β) :=
begin
unfold inv_commute_of_commute,
apply @is_equiv_compose _ _ _
(homotopy.symm ∘ (homotopy_inv_of_homotopy_pre f (f' ∘ α) β))
(homotopy_inv_of_homotopy_post f' β (α ∘ f⁻¹)),
{ apply @is_equiv_compose _ _ _
(homotopy_inv_of_homotopy_pre f (f' ∘ α) β) homotopy.symm,
{ apply homotopy_inv_of_homotopy_pre.is_equiv },
{ apply pi.is_equiv_homotopy_symm }
},
{ apply homotopy_inv_of_homotopy_post.is_equiv }
end
end arrow

View file

@ -7,7 +7,7 @@ Ported from Coq HoTT
Theorems about the types equiv and is_equiv
-/
import .fiber .arrow arity ..prop_trunc
import .fiber .arrow arity ..prop_trunc cubical.square
open eq is_trunc sigma sigma.ops pi fiber function equiv
@ -102,6 +102,83 @@ namespace is_equiv
end is_equiv
/- Moving equivalences around in homotopies -/
namespace is_equiv
variables {A B C : Type} (f : A → B) [Hf : is_equiv f]
include Hf
section pre_compose
variables (α : A → C) (β : B → C)
definition homotopy_inv_of_homotopy_pre (p : α ~ β ∘ f) : α ∘ f⁻¹ ~ β :=
λb, p (f⁻¹ b) ⬝ ap β (right_inv f b)
protected definition homotopy_inv_of_homotopy_pre.is_equiv
: is_equiv (homotopy_inv_of_homotopy_pre f α β) :=
adjointify _ (λq a, (ap α (left_inv f a))⁻¹ ⬝ q (f a))
abstract begin
intro q, apply eq_of_homotopy, intro b,
unfold homotopy_inv_of_homotopy_pre,
apply inverse, apply eq_bot_of_square,
apply eq_hconcat (ap02 α (adj_inv f b)),
apply eq_hconcat (ap_compose α f⁻¹ (right_inv f b))⁻¹,
apply natural_square_tr q (right_inv f b)
end end
abstract begin
intro p, apply eq_of_homotopy, intro a,
unfold homotopy_inv_of_homotopy_pre,
apply trans (con.assoc
(ap α (left_inv f a))⁻¹
(p (f⁻¹ (f a)))
(ap β (right_inv f (f a))))⁻¹,
apply inverse, apply eq_bot_of_square,
refine hconcat_eq _ (ap02 β (adj f a))⁻¹,
refine hconcat_eq _ (ap_compose β f (left_inv f a)),
apply natural_square_tr p (left_inv f a)
end end
end pre_compose
section post_compose
variables (β : C → B) (α : C → A)
definition homotopy_inv_of_homotopy_post (p : β ~ f ∘ α) : f⁻¹ ∘ β ~ α :=
λc, ap f⁻¹ (p c) ⬝ (left_inv f (α c))
protected definition homotopy_inv_of_homotopy_post.is_equiv
: is_equiv (homotopy_inv_of_homotopy_post f β α) :=
adjointify _ (λq c, (right_inv f (β c))⁻¹ ⬝ ap f (q c))
abstract begin
intro q, apply eq_of_homotopy, intro c,
unfold homotopy_inv_of_homotopy_post,
apply trans (whisker_right
(ap_con f⁻¹ (right_inv f (β c))⁻¹ (ap f (q c))
⬝ whisker_right (ap_inv f⁻¹ (right_inv f (β c)))
(ap f⁻¹ (ap f (q c)))) (left_inv f (α c))),
apply inverse, apply eq_bot_of_square,
apply eq_hconcat (adj_inv f (β c))⁻¹,
apply eq_vconcat (ap_compose f⁻¹ f (q c))⁻¹,
refine vconcat_eq _ (ap_id (q c)),
apply natural_square (left_inv f) (q c)
end end
abstract begin
intro p, apply eq_of_homotopy, intro c,
unfold homotopy_inv_of_homotopy_post,
apply trans (whisker_left (right_inv f (β c))⁻¹
(ap_con f (ap f⁻¹ (p c)) (left_inv f (α c)))),
apply trans (con.assoc (right_inv f (β c))⁻¹ (ap f (ap f⁻¹ (p c)))
(ap f (left_inv f (α c))))⁻¹,
apply inverse, apply eq_bot_of_square,
refine hconcat_eq _ (adj f (α c)),
apply eq_vconcat (ap_compose f f⁻¹ (p c))⁻¹,
refine vconcat_eq _ (ap_id (p c)),
apply natural_square (right_inv f) (p c)
end end
end post_compose
end is_equiv
namespace is_equiv
/- Theorem 4.7.7 -/

View file

@ -37,6 +37,16 @@ namespace pi
definition eq_of_homotopy_idp (f : Πa, B a) : eq_of_homotopy (λx : A, refl (f x)) = refl f :=
!eq_of_homotopy_eta
/- homotopy.symm is an equivalence -/
definition is_equiv_homotopy_symm : is_equiv (homotopy.symm : f ~ g → g ~ f) :=
begin
fapply adjointify homotopy.symm homotopy.symm,
{ intro p, apply eq_of_homotopy, intro a,
unfold homotopy.symm, apply inv_inv },
{ intro p, apply eq_of_homotopy, intro a,
unfold homotopy.symm, apply inv_inv }
end
/-
The identification of the path space of a dependent function space,
up to equivalence, is of course just funext.

View file

@ -127,6 +127,16 @@ namespace prod
: ap (prod_functor f g) (prod_eq p q) = prod_eq (ap f p) (ap g q) :=
by induction u; induction v; esimp at *; induction p; induction q; reflexivity
/- Helpers for functions of two arguments -/
definition ap_diagonal {a a' : A} (p : a = a')
: ap (λx : A, (x,x)) p = prod_eq p p :=
by cases p; constructor
definition ap_binary (m : A → B → C) (p : a = a') (q : b = b')
: ap (λz : A × B, m z.1 z.2) (prod_eq p q)
= (ap (m a) q) ⬝ (ap (λx : A, m x b') p) :=
by cases p; cases q; constructor
/- Equivalences -/
definition is_equiv_prod_functor [instance] [constructor] [H : is_equiv f] [H : is_equiv g]