feat(hott) move many lemmas to library, and cleanup various parts

This commit is contained in:
Floris van Doorn 2016-11-23 17:59:13 -05:00 committed by Leonardo de Moura
parent ecbe4af3c7
commit 9342fe2716
47 changed files with 1512 additions and 835 deletions

View file

@ -18,7 +18,6 @@ The following files are [ported](../port.md) from the standard library. If anyth
Files which are not ported from the standard library:
* [hott](hott.hlean) : Basic theorems about the algebraic hierarchy specific to HoTT
* [group_theory](group_theory.hlean) : Basic theorems about group homomorphisms and isomorphisms
* [trunc_group](trunc_group.hlean) : truncate an infinity-group to a group
* [homotopy_group](homotopy_group.hlean) : homotopy groups of a pointed type

View file

@ -6,7 +6,7 @@ Authors: Jeremy Avigad
Bundled structures
-/
import algebra.group homotopy.interval
open algebra
open algebra pointed is_trunc
namespace algebra
structure Semigroup :=
@ -37,6 +37,21 @@ structure Group :=
(carrier : Type) (struct : group carrier)
attribute Group.carrier [coercion]
attribute Group.struct [instance]
section
local attribute Group.struct [instance]
definition pSet_of_Group [constructor] [reducible] [coercion] (G : Group) : Set* :=
ptrunctype.mk G !semigroup.is_set_carrier 1
end
attribute algebra._trans_of_pSet_of_Group [unfold 1]
attribute algebra._trans_of_pSet_of_Group_1 algebra._trans_of_pSet_of_Group_2 [constructor]
definition pType_of_Group [reducible] [constructor] : Group → Type* :=
algebra._trans_of_pSet_of_Group_1
definition Set_of_Group [reducible] [constructor] : Group → Set :=
algebra._trans_of_pSet_of_Group_2
definition AddGroup : Type := Group
@ -64,6 +79,14 @@ CommGroup.struct G
attribute AddCommGroup.struct CommGroup.struct [instance] [priority 2000]
definition Group_of_CommGroup [coercion] [constructor] (G : CommGroup) : Group :=
Group.mk G _
attribute algebra._trans_of_Group_of_CommGroup_1
algebra._trans_of_Group_of_CommGroup
algebra._trans_of_Group_of_CommGroup_3 [constructor]
attribute algebra._trans_of_Group_of_CommGroup_2 [unfold 1]
-- structure AddSemigroup :=
-- (carrier : Type) (struct : add_semigroup carrier)

View file

@ -201,6 +201,9 @@ section group
theorem inv_eq_one_iff_eq_one (a : A) : a⁻¹ = 1 ↔ a = 1 :=
one_inv ▸ inv_eq_inv_iff_eq a 1
theorem inv_eq_one {a : A} (H : a = 1) : a⁻¹ = 1 :=
iff.mpr (inv_eq_one_iff_eq_one a) H
theorem eq_one_of_inv_eq_one (a : A) : a⁻¹ = 1 → a = 1 :=
iff.mp !inv_eq_one_iff_eq_one

View file

@ -9,27 +9,27 @@ This file will be rewritten in the future, when we develop are more systematic n
describing homomorphisms
-/
import algebra.category.category algebra.hott
import algebra.category.category algebra.bundled
open eq algebra pointed function is_trunc pi category equiv is_equiv
open eq algebra pointed function is_trunc pi equiv is_equiv
set_option class.force_new true
namespace group
definition pointed_Group [instance] [constructor] (G : Group) : pointed G :=
pointed.mk 1
definition pType_of_Group [constructor] [reducible] (G : Group) : Type* :=
pointed.MK G 1
definition Set_of_Group [constructor] (G : Group) : Set := trunctype.mk G _
definition Group_of_CommGroup [coercion] [constructor] (G : CommGroup) : Group :=
Group.mk G _
definition Group.struct' [instance] [reducible] (G : Group) : group G :=
Group.struct G
definition comm_group_Group_of_CommGroup [instance] [constructor] [priority 900]
(G : CommGroup) : comm_group (Group_of_CommGroup G) :=
begin esimp, exact _ end
definition group_pType_of_Group [instance] [priority 900] (G : Group) :
group (pType_of_Group G) :=
definition comm_group_pSet_of_Group [instance] (G : CommGroup) : comm_group (pSet_of_Group G) :=
CommGroup.struct G
definition group_pSet_of_Group [instance] [priority 900] (G : Group) :
group (pSet_of_Group G) :=
Group.struct G
/- group homomorphisms -/
@ -43,7 +43,8 @@ namespace group
[group G] [group G₁] [group G₂] [group G₃]
[is_homomorphism ψ] [is_homomorphism φ₁] [is_homomorphism φ₂] [is_homomorphism φ]
definition respect_mul /- φ -/ : Π(g h : G₁), φ (g * h) = φ g * φ h :=
definition respect_mul {G₁ G₂ : Type} [has_mul G₁] [has_mul G₂] (φ : G₁ → G₂)
[is_homomorphism φ] : Π(g h : G₁), φ (g * h) = φ g * φ h :=
by assumption
theorem respect_one /- φ -/ : φ 1 = 1 :=
@ -78,7 +79,7 @@ namespace group
section additive
definition is_add_homomorphism [class] [reducible] {G₁ G₂ : Type} [add_group G₁] [add_group G₂]
definition is_add_homomorphism [class] [reducible] {G₁ G₂ : Type} [has_add G₁] [has_add G₂]
(φ : G₁ → G₂) : Type :=
Π(g h : G₁), φ (g + h) = φ g + φ h
@ -106,11 +107,11 @@ namespace group
infix ` →g `:55 := homomorphism
definition group_fun [unfold 3] [coercion] := @homomorphism.φ
definition homomorphism.struct [instance] [priority 900] {G₁ : Group} {G₂ : Group} (φ : G₁ →g G₂)
: is_homomorphism φ :=
definition homomorphism.struct [unfold 3] [instance] [priority 900] {G₁ G₂ : Group}
(φ : G₁ →g G₂) : is_homomorphism φ :=
homomorphism.p φ
definition homomorphism.mulstruct [instance] [priority 2000] {G₁ G₂ : Group } (φ : G₁ →g G₂)
definition homomorphism.mulstruct [instance] [priority 2000] {G₁ G₂ : Group} (φ : G₁ →g G₂)
: is_homomorphism φ :=
homomorphism.p φ
@ -147,9 +148,13 @@ namespace group
end
variables {G₁ G₂}
definition pmap_of_homomorphism [constructor] /- φ -/ : pType_of_Group G₁ →* pType_of_Group G₂ :=
definition pmap_of_homomorphism [constructor] /- φ -/ : G₁ →* G₂ :=
pmap.mk φ begin esimp, exact respect_one φ end
definition homomorphism_change_fun [constructor] {G₁ G₂ : Group}
(φ : G₁ →g G₂) (f : G₁ → G₂) (p : φ ~ f) : G₁ →g G₂ :=
homomorphism.mk f (λg h, (p (g * h))⁻¹ ⬝ to_respect_mul φ g h ⬝ ap011 mul (p g) (p h))
definition homomorphism_eq (p : group_fun φ₁ ~ group_fun φ₂) : φ₁ = φ₂ :=
begin
induction φ₁ with φ₁ q₁, induction φ₂ with φ₂ q₂, esimp at p, induction p,
@ -205,6 +210,7 @@ namespace group
homomorphism.mk (@id G) (is_homomorphism_id G)
variable {G}
abbreviation gid [constructor] := @homomorphism_id
infixr ` ∘g `:75 := homomorphism_compose
notation 1 := homomorphism_id _
@ -215,25 +221,33 @@ namespace group
infix ` ≃g `:25 := isomorphism
attribute isomorphism.to_hom [coercion]
attribute isomorphism.is_equiv_to_hom [instance]
attribute isomorphism._trans_of_to_hom [unfold 3]
definition equiv_of_isomorphism [constructor] (φ : G₁ ≃g G₂) : G₁ ≃ G₂ :=
equiv.mk φ _
definition pequiv_of_isomorphism [constructor] (φ : G₁ ≃g G₂) :
pType_of_Group G₁ ≃* pType_of_Group G₂ :=
G₁ ≃* G₂ :=
pequiv.mk φ begin esimp, exact _ end begin esimp, exact respect_one φ end
definition isomorphism_of_equiv [constructor] (φ : G₁ ≃ G₂)
(p : Πg₁ g₂, φ (g₁ * g₂) = φ g₁ * φ g₂) : G₁ ≃g G₂ :=
isomorphism.mk (homomorphism.mk φ p) !to_is_equiv
definition eq_of_isomorphism {G₁ G₂ : Group} (φ : G₁ ≃g G₂) : G₁ = G₂ :=
Group_eq (equiv_of_isomorphism φ) (respect_mul φ)
definition isomorphism_of_eq {G₁ G₂ : Group} (φ : G₁ = G₂) : G₁ ≃g G₂ :=
definition isomorphism_of_eq [constructor] {G₁ G₂ : Group} (φ : G₁ = G₂) : G₁ ≃g G₂ :=
isomorphism_of_equiv (equiv_of_eq (ap Group.carrier φ))
begin intros, induction φ, reflexivity end
definition pequiv_of_isomorphism_of_eq {G₁ G₂ : Group} (p : G₁ = G₂) :
pequiv_of_isomorphism (isomorphism_of_eq p) = pequiv_of_eq (ap pType_of_Group p) :=
begin
induction p,
apply pequiv_eq,
fapply pmap_eq,
{ intro g, reflexivity},
{ apply is_prop.elim}
end
definition to_ginv [constructor] (φ : G₁ ≃g G₂) : G₂ →g G₁ :=
homomorphism.mk φ⁻¹
abstract begin
@ -265,18 +279,14 @@ namespace group
infixl ` ⬝gp `:75 := isomorphism.trans_eq
infixl ` ⬝pg `:75 := isomorphism.eq_trans
-- TODO
-- definition Group_univalence (G₁ G₂ : Group) : (G₁ ≃g G₂) ≃ (G₁ = G₂) :=
-- begin
-- fapply equiv.MK,
-- { exact eq_of_isomorphism},
-- { intro p, apply transport _ p, reflexivity},
-- { intro p, induction p, esimp, },
-- { }
-- end
definition pmap_of_isomorphism [constructor] (φ : G₁ ≃g G₂) :
G₁ →* G₂ :=
pequiv_of_isomorphism φ
/- category of groups -/
section
open category
definition precategory_group [constructor] : precategory Group :=
precategory.mk homomorphism
@homomorphism_compose
@ -284,6 +294,7 @@ namespace group
(λG₁ G₂ G₃ G₄ φ₃ φ₂ φ₁, homomorphism_eq (λg, idp))
(λG₁ G₂ φ, homomorphism_eq (λg, idp))
(λG₁ G₂ φ, homomorphism_eq (λg, idp))
end
-- TODO
-- definition category_group : category Group :=
@ -339,6 +350,17 @@ namespace group
end
variable (G)
/- the trivial group -/
open unit
definition trivial_group [constructor] : group unit :=
group.mk (λx y, star) _ (λx y z, idp) star (unit.rec idp) (unit.rec idp) (λx, star) (λx, idp)
definition Trivial_group [constructor] : Group :=
Group.mk _ trivial_group
abbreviation G0 := Trivial_group
definition trivial_group_of_is_contr [H : is_contr G] : G ≃g G0 :=
begin
fapply isomorphism_of_equiv,
@ -346,8 +368,6 @@ namespace group
{ intros, reflexivity}
end
definition trivial_group_of_is_contr' (G : Group) [H : is_contr G] : G = G0 :=
eq_of_isomorphism (trivial_group_of_is_contr G)
variable {G}
/-
@ -377,4 +397,113 @@ namespace group
mul_left_inv_pt := mul.left_inv⦄
end
definition Group_of_pgroup (G : Type*) [pgroup G] : Group :=
Group.mk G _
definition pgroup_Group [instance] (G : Group) : pgroup G :=
⦃ pgroup, Group.struct G,
pt_mul := one_mul,
mul_pt := mul_one,
mul_left_inv_pt := mul.left_inv ⦄
/- equality of groups and abelian groups -/
definition group.to_has_mul {A : Type} (H : group A) : has_mul A := _
definition group.to_has_inv {A : Type} (H : group A) : has_inv A := _
definition group.to_has_one {A : Type} (H : group A) : has_one A := _
local attribute group.to_has_mul group.to_has_inv [coercion]
universe variable l
variables {A B : Type.{l}}
definition group_eq {G H : group A} (same_mul' : Π(g h : A), @mul A G g h = @mul A H g h)
: G = H :=
begin
have foo : Π(g : A), @inv A G g = (@inv A G g * g) * @inv A H g,
from λg, !mul_inv_cancel_right⁻¹,
cases G with Gm Gs Gh1 G1 Gh2 Gh3 Gi Gh4,
cases H with Hm Hs Hh1 H1 Hh2 Hh3 Hi Hh4,
rewrite [↑[semigroup.to_has_mul,group.to_has_inv] at (same_mul',foo)],
have same_mul : Gm = Hm, from eq_of_homotopy2 same_mul',
cases same_mul,
have same_one : G1 = H1, from calc
G1 = Hm G1 H1 : Hh3
... = H1 : Gh2,
have same_inv : Gi = Hi, from eq_of_homotopy (take g, calc
Gi g = Hm (Hm (Gi g) g) (Hi g) : foo
... = Hm G1 (Hi g) : by rewrite Gh4
... = Hi g : Gh2),
cases same_one, cases same_inv,
have ps : Gs = Hs, from !is_prop.elim,
have ph1 : Gh1 = Hh1, from !is_prop.elim,
have ph2 : Gh2 = Hh2, from !is_prop.elim,
have ph3 : Gh3 = Hh3, from !is_prop.elim,
have ph4 : Gh4 = Hh4, from !is_prop.elim,
cases ps, cases ph1, cases ph2, cases ph3, cases ph4, reflexivity
end
definition group_pathover {G : group A} {H : group B} {p : A = B}
(resp_mul : Π(g h : A), cast p (g * h) = cast p g * cast p h) : G =[p] H :=
begin
induction p,
apply pathover_idp_of_eq, exact group_eq (resp_mul)
end
definition Group_eq_of_eq {G H : Group} (p : Group.carrier G = Group.carrier H)
(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 (apd011 Group.mk p),
exact group_pathover resp_mul
end
definition Group_eq {G H : Group} (f : Group.carrier G ≃ Group.carrier H)
(resp_mul : Π(g h : G), f (g * h) = f g * f h) : G = H :=
Group_eq_of_eq (ua f) (λg h, !cast_ua ⬝ resp_mul g h ⬝ ap011 mul !cast_ua⁻¹ !cast_ua⁻¹)
definition eq_of_isomorphism {G₁ G₂ : Group} (φ : G₁ ≃g G₂) : G₁ = G₂ :=
Group_eq (equiv_of_isomorphism φ) (respect_mul φ)
definition comm_group.to_has_mul {A : Type} (H : comm_group A) : has_mul A := _
local attribute comm_group.to_has_mul [coercion]
definition comm_group_eq {A : Type} {G H : comm_group A}
(same_mul : Π(g h : A), @mul A G g h = @mul A H g h)
: G = H :=
begin
have g_eq : @comm_group.to_group A G = @comm_group.to_group A H, from group_eq same_mul,
cases G with Gm Gs Gh1 G1 Gh2 Gh3 Gi Gh4 Gh5,
cases H with Hm Hs Hh1 H1 Hh2 Hh3 Hi Hh4 Hh5,
have pm : Gm = Hm, from ap (@mul _ ∘ group.to_has_mul) g_eq,
have pi : Gi = Hi, from ap (@inv _ ∘ group.to_has_inv) g_eq,
have p1 : G1 = H1, from ap (@one _ ∘ group.to_has_one) g_eq,
induction pm, induction pi, induction p1,
have ps : Gs = Hs, from !is_prop.elim,
have ph1 : Gh1 = Hh1, from !is_prop.elim,
have ph2 : Gh2 = Hh2, from !is_prop.elim,
have ph3 : Gh3 = Hh3, from !is_prop.elim,
have ph4 : Gh4 = Hh4, from !is_prop.elim,
have ph5 : Gh5 = Hh5, from !is_prop.elim,
induction ps, induction ph1, induction ph2, induction ph3, induction ph4, induction ph5,
reflexivity
end
definition comm_group_pathover {A B : Type} {G : comm_group A} {H : comm_group B} {p : A = B}
(resp_mul : Π(g h : A), cast p (g * h) = cast p g * cast p h) : G =[p] H :=
begin
induction p,
apply pathover_idp_of_eq, exact comm_group_eq (resp_mul)
end
definition CommGroup_eq_of_isomorphism {G₁ G₂ : CommGroup} (φ : G₁ ≃g G₂) : G₁ = G₂ :=
begin
induction G₁, induction G₂,
apply apd011 CommGroup.mk (ua (equiv_of_isomorphism φ)),
apply comm_group_pathover,
intro g h, exact !cast_ua ⬝ respect_mul φ g h ⬝ ap011 mul !cast_ua⁻¹ !cast_ua⁻¹
end
definition trivial_group_of_is_contr' (G : Group) [H : is_contr G] : G = G0 :=
eq_of_isomorphism (trivial_group_of_is_contr G)
end group

View file

@ -103,7 +103,8 @@ namespace eq
apply ap tr, apply loopn_succ_in_con
end
definition ghomotopy_group_succ_in (A : Type*) (n : ) : πg[n + 2] A ≃g πg[n + 1] (Ω A) :=
definition ghomotopy_group_succ_in [constructor] (A : Type*) (n : ) :
πg[n + 2] A ≃g πg[n + 1] (Ω A) :=
begin
fapply isomorphism_of_equiv,
{ exact homotopy_group_succ_in A (succ n)},
@ -120,6 +121,9 @@ namespace eq
(p : f ~* g) : π→[n] f ~* π→[n] g :=
ptrunc_functor_phomotopy 0 (apn_phomotopy n p)
definition homotopy_group_functor_pid (n : ) (A : Type*) : π→[n] (pid A) ~* pid (π[n] A) :=
ptrunc_functor_phomotopy 0 !apn_pid ⬝* !ptrunc_functor_pid
definition homotopy_group_functor_compose [constructor] (n : ) {A B C : Type*} (g : B →* C)
(f : A →* B) : π→[n] (g ∘* f) ~* π→[n] g ∘* π→[n] f :=
ptrunc_functor_phomotopy 0 !apn_pcompose ⬝* !ptrunc_functor_pcompose
@ -227,7 +231,7 @@ namespace eq
begin
fapply isomorphism_of_equiv,
{ exact homotopy_group_ptrunc (k+1) A},
{ intro g₁ g₂, esimp,
{ intro g₁ g₂,
refine _ ⬝ !homotopy_group_pequiv_loop_ptrunc_inv_con,
apply ap ((homotopy_group_pequiv_loop_ptrunc (k+1) A)⁻¹ᵉ*),
refine _ ⬝ !loopn_pequiv_loopn_con ,

View file

@ -1,80 +0,0 @@
/-
Copyright (c) 2015 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Floris van Doorn
Theorems about algebra specific to HoTT
-/
import .group arity types.pi prop_trunc types.unit .bundled
open equiv eq is_trunc unit
namespace algebra
definition trivial_group [constructor] : group unit :=
group.mk (λx y, star) _ (λx y z, idp) star (unit.rec idp) (unit.rec idp) (λx, star) (λx, idp)
definition Trivial_group [constructor] : Group :=
Group.mk _ trivial_group
abbreviation G0 := Trivial_group
open Group has_mul has_inv
-- we prove under which conditions two groups are equal
-- group and has_mul are classes. So, lean does not automatically generate
-- coercions between them anymore.
-- So, an application such as (@mul A G g h) in the following definition
-- is type incorrect if the coercion is not added (explicitly or implicitly).
definition group.to_has_mul {A : Type} (H : group A) : has_mul A := _
local attribute group.to_has_mul group.to_has_inv [coercion]
universe variable l
variables {A B : Type.{l}}
definition group_eq {G H : group A} (same_mul' : Π(g h : A), @mul A G g h = @mul A H g h)
: G = H :=
begin
have foo : Π(g : A), @inv A G g = (@inv A G g * g) * @inv A H g,
from λg, !mul_inv_cancel_right⁻¹,
cases G with Gm Gs Gh1 G1 Gh2 Gh3 Gi Gh4,
cases H with Hm Hs Hh1 H1 Hh2 Hh3 Hi Hh4,
rewrite [↑[semigroup.to_has_mul,group.to_has_inv] at (same_mul',foo)],
have same_mul : Gm = Hm, from eq_of_homotopy2 same_mul',
cases same_mul,
have same_one : G1 = H1, from calc
G1 = Hm G1 H1 : Hh3
... = H1 : Gh2,
have same_inv : Gi = Hi, from eq_of_homotopy (take g, calc
Gi g = Hm (Hm (Gi g) g) (Hi g) : foo
... = Hm G1 (Hi g) : by rewrite Gh4
... = Hi g : Gh2),
cases same_one, cases same_inv,
have ps : Gs = Hs, from !is_prop.elim,
have ph1 : Gh1 = Hh1, from !is_prop.elim,
have ph2 : Gh2 = Hh2, from !is_prop.elim,
have ph3 : Gh3 = Hh3, from !is_prop.elim,
have ph4 : Gh4 = Hh4, from !is_prop.elim,
cases ps, cases ph1, cases ph2, cases ph3, cases ph4, reflexivity
end
definition group_pathover {G : group A} {H : group B} {p : A = B}
(resp_mul : Π(g h : A), cast p (g * h) = cast p g * cast p h) : G =[p] H :=
begin
induction p,
apply pathover_idp_of_eq, exact group_eq (resp_mul)
end
definition Group_eq_of_eq {G H : Group} (p : carrier G = carrier H)
(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 (apd011 Group.mk p),
exact group_pathover resp_mul
end
definition Group_eq {G H : Group} (f : carrier G ≃ carrier H)
(resp_mul : Π(g h : G), f (g * h) = f g * f h) : G = H :=
Group_eq_of_eq (ua f) (λg h, !cast_ua ⬝ resp_mul g h ⬝ ap011 mul !cast_ua⁻¹ !cast_ua⁻¹)
end algebra

View file

@ -106,7 +106,16 @@ namespace eq
(c : cube vrfl vrfl vrfl vrfl s s') : s = s' :=
by induction s; exact eq_of_cube c
definition square_pathover [unfold 7]
definition square_pathover {A B : Type} {a a' : A} {b₁ b₂ b₃ b₄ : A → B}
{f₁ : b₁ ~ b₂} {f₂ : b₃ ~ b₄} {f₃ : b₁ ~ b₃} {f₄ : b₂ ~ b₄} {p : a = a'}
{q : square (f₁ a) (f₂ a) (f₃ a) (f₄ a)}
{r : square (f₁ a') (f₂ a') (f₃ a') (f₄ a')}
(s : cube (natural_square f₁ p) (natural_square f₂ p)
(natural_square f₃ p) (natural_square f₄ p) q r) : q =[p] r :=
by induction p; apply pathover_idp_of_eq; exact eq_of_deg12_cube s
-- a special case where the endpoints do not depend on `p`
definition square_pathover'
{f₁ : A → b₁ = b₂} {f₂ : A → b₃ = b₄} {f₃ : A → b₁ = b₃} {f₄ : A → b₂ = b₄}
{p : a = a'}
{q : square (f₁ a) (f₂ a) (f₃ a) (f₄ a)}
@ -298,7 +307,43 @@ namespace eq
definition cube_inverse3 : cube s₀₁₁⁻¹ᵛ s₂₁₁⁻¹ᵛ s₁₀₁⁻¹ᵛ s₁₂₁⁻¹ᵛ s₁₁₂ s₁₁₀ :=
by cases c; exact idc
omit c
definition eq_concat1 {s₀₁₁' : square p₀₁₀ p₀₁₂ p₀₀₁ p₀₂₁} (r : s₀₁₁' = s₀₁₁)
(c : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂) : cube s₀₁₁' s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂ :=
by induction r; exact c
definition concat1_eq {s₂₁₁' : square p₂₁₀ p₂₁₂ p₂₀₁ p₂₂₁}
(c : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂) (r : s₂₁₁ = s₂₁₁')
: cube s₀₁₁ s₂₁₁' s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂ :=
by induction r; exact c
definition eq_concat2 {s₁₀₁' : square p₁₀₀ p₁₀₂ p₀₀₁ p₂₀₁} (r : s₁₀₁' = s₁₀₁)
(c : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂) : cube s₀₁₁ s₂₁₁ s₁₀₁' s₁₂₁ s₁₁₀ s₁₁₂ :=
by induction r; exact c
definition concat2_eq {s₁₂₁' : square p₁₂₀ p₁₂₂ p₀₂₁ p₂₂₁}
(c : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂) (r : s₁₂₁ = s₁₂₁')
: cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁' s₁₁₀ s₁₁₂ :=
by induction r; exact c
definition eq_concat3 {s₁₁₀' : square p₀₁₀ p₂₁₀ p₁₀₀ p₁₂₀} (r : s₁₁₀' = s₁₁₀)
(c : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂) : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀' s₁₁₂ :=
by induction r; exact c
definition concat3_eq {s₁₁₂' : square p₀₁₂ p₂₁₂ p₁₀₂ p₁₂₂}
(c : cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂) (r : s₁₁₂ = s₁₁₂')
: cube s₀₁₁ s₂₁₁ s₁₀₁ s₁₂₁ s₁₁₀ s₁₁₂' :=
by induction r; exact c
infix ` ⬝1 `:75 := cube_concat1
infix ` ⬝2 `:75 := cube_concat2
infix ` ⬝3 `:75 := cube_concat3
infix ` ⬝p1 `:75 := eq_concat1
infix ` ⬝1p `:75 := concat1_eq
infix ` ⬝p2 `:75 := eq_concat3
infix ` ⬝2p `:75 := concat2_eq
infix ` ⬝p3 `:75 := eq_concat3
infix ` ⬝3p `:75 := concat3_eq
end eq

View file

@ -102,14 +102,6 @@ namespace eq
: square (ap f p₁₀) (ap f p₁₂) (ap f p₀₁) (ap f p₂₁) :=
by induction s₁₁;exact ids
definition natural_square [unfold 8] {f g : A → B} (p : f ~ g) (q : a = a') :
square (ap f q) (ap g q) (p a) (p a') :=
eq.rec_on q hrfl
definition natural_square_tr [unfold 8] {f g : A → B} (p : f ~ g) (q : a = a') :
square (p a) (p a') (ap f q) (ap g q) :=
eq.rec_on q vrfl
/- canceling, whiskering and moving thinks along the sides of the square -/
definition whisker_tl (p : a = a₀₀) (s₁₁ : square p₁₀ p₁₂ p₀₁ p₂₁)
: square (p ⬝ p₁₀) p₁₂ (p ⬝ p₀₁) p₂₁ :=
@ -584,6 +576,14 @@ namespace eq
s₁₁ ⬝vp r = s₁₁ ⬝v vdeg_square r :=
by cases r; cases s₁₁; esimp
definition natural_square [unfold 8] {f g : A → B} (p : f ~ g) (q : a = a') :
square (p a) (p a') (ap f q) (ap g q) :=
square_of_pathover (apd p q)
definition natural_square_tr [unfold 8] {f g : A → B} (p : f ~ g) (q : a = a') :
square (ap f q) (ap g q) (p a) (p a') :=
transpose (natural_square p q)
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)

View file

@ -25,7 +25,7 @@ namespace eq
by induction p; reflexivity
theorem ap_is_constant_natural_square {g : B → C} {f : A → B} (H : Πa, g (f a) = c) (p : a = a') :
(ap_is_constant H p)⁻¹ ⬝ph natural_square_tr H p ⬝hp ap_constant p c =
(ap_is_constant H p)⁻¹ ⬝ph natural_square H p ⬝hp ap_constant p c =
whisker_bl (H a') (whisker_tl (H a) ids) :=
begin induction p, esimp, rewrite inv_inv, rewrite whisker_bl_whisker_tl_eq end

View file

@ -255,7 +255,7 @@ namespace eq
{b : Πa, B (f a)} {b₂ : Πa, B (g a)} {q : Π(a' : A'), f a' = g a'}
(r : pathover B (b a') (q a') (b₂ a'))
(r₂ : pathover B (b a₂') (q a₂') (b₂ a₂'))
(s : squareover B (natural_square_tr q p) r r₂
(s : squareover B (natural_square q p) r r₂
(pathover_ap B f (apd b p)) (pathover_ap B g (apd b₂ p)))
: pathover (λa, pathover B (b a) (q a) (b₂ a)) r p r₂ :=
begin

View file

@ -88,7 +88,7 @@ namespace eq
(ap (ap g ∘ ap f) r)
(ap_compose g f p)
(ap_compose g f q) :=
natural_square (ap_compose g f) r
natural_square_tr (ap_compose g f) r
theorem whisker_right_eq_of_con_inv_eq_idp {p q : a₁ = a₂} (r : p ⬝ q⁻¹ = idp) :
whisker_right (eq_of_con_inv_eq_idp r) q⁻¹ ⬝ con.right_inv q = r :=
@ -118,8 +118,8 @@ namespace eq
con_tr idp q u = ap (λp, p ▸ u) (idp_con q) :=
by induction q;reflexivity
definition transport_eq_Fl_idp_left {A B : Type} {a : A} {b : B} (f : A → B) (q : f a = b)
: transport_eq_Fl idp q = !idp_con⁻¹ :=
definition eq_transport_Fl_idp_left {A B : Type} {a : A} {b : B} (f : A → B) (q : f a = b)
: eq_transport_Fl idp q = !idp_con⁻¹ :=
by induction q; reflexivity
definition whisker_left_idp_con_eq_assoc

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn
Declaration of the coequalizer
-/
import .quotient_functor types.equiv
import types.equiv .quotient
open quotient eq equiv is_trunc sigma sigma.ops

View file

@ -1,55 +0,0 @@
/-
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, Floris van Doorn
Pointed Pushouts
-/
import .pushout types.pointed
open eq pushout
namespace pointed
definition pointed_pushout [instance] [constructor] {TL BL TR : Type} [HTL : pointed TL]
[HBL : pointed BL] [HTR : pointed TR] (f : TL → BL) (g : TL → TR) : pointed (pushout f g) :=
pointed.mk (inl (point _))
end pointed
open pointed pType
namespace pushout
section
parameters {TL BL TR : Type*} (f : TL →* BL) (g : TL →* TR)
definition ppushout [constructor] : Type* :=
pointed.mk' (pushout f g)
parameters {f g}
definition pinl [constructor] : BL →* ppushout :=
pmap.mk inl idp
definition pinr [constructor] : TR →* ppushout :=
pmap.mk inr ((ap inr (respect_pt g))⁻¹ ⬝ !glue⁻¹ ⬝ (ap inl (respect_pt f)))
definition pglue (x : TL) : pinl (f x) = pinr (g x) := -- TODO do we need this?
!glue
definition prec {P : ppushout → Type} (Pinl : Π x, P (pinl x)) (Pinr : Π x, P (pinr x))
(H : Π x, Pinl (f x) =[pglue x] Pinr (g x)) : (Π y, P y) :=
pushout.rec Pinl Pinr H
end
section
variables {TL BL TR : Type*} (f : TL →* BL) (g : TL →* TR)
protected definition psymm [constructor] : ppushout f g ≃* ppushout g f :=
begin
fapply pequiv_of_equiv,
{ apply pushout.symm},
{ exact ap inr (respect_pt f)⁻¹ ⬝ !glue⁻¹ ⬝ ap inl (respect_pt g)}
end
end
end pushout

View file

@ -296,7 +296,7 @@ namespace hide
intro b,
induction b with n b n b,
{ apply eq_constructors},
{ apply (equiv.to_inv !pathover_eq_equiv_r), apply eq_constructors_comp}
{ apply (equiv.to_inv !eq_pathover_equiv_r), apply eq_constructors_comp}
end
end

View file

@ -1,14 +1,14 @@
/-
Copyright (c) 2015-16 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn, Ulrik Buchholtz
Authors: Floris van Doorn, Ulrik Buchholtz, Jakob von Raumer
Declaration and properties of the pushout
-/
import .quotient types.sigma types.arrow_2
open quotient eq sum equiv is_trunc
open quotient eq sum equiv is_trunc pointed
namespace pushout
section
@ -339,4 +339,40 @@ namespace pushout
protected definition equiv : pushout f g ≃ pushout f' g' :=
equiv.mk (pushout.functor f g f' g' tl bl tr fh gh) _
end
definition pointed_pushout [instance] [constructor] {TL BL TR : Type} [HTL : pointed TL]
[HBL : pointed BL] [HTR : pointed TR] (f : TL → BL) (g : TL → TR) : pointed (pushout f g) :=
pointed.mk (inl (point _))
end pushout open pushout
definition ppushout [constructor] {TL BL TR : Type*} (f : TL →* BL) (g : TL →* TR) : Type* :=
pointed.mk' (pushout f g)
namespace pushout
section
parameters {TL BL TR : Type*} (f : TL →* BL) (g : TL →* TR)
parameters {f g}
definition pinl [constructor] : BL →* ppushout f g :=
pmap.mk inl idp
definition pinr [constructor] : TR →* ppushout f g :=
pmap.mk inr ((ap inr (respect_pt g))⁻¹ ⬝ !glue⁻¹ ⬝ (ap inl (respect_pt f)))
definition pglue (x : TL) : pinl (f x) = pinr (g x) := -- TODO do we need this?
!glue
end
section
variables {TL BL TR : Type*} (f : TL →* BL) (g : TL →* TR)
protected definition psymm [constructor] : ppushout f g ≃* ppushout g f :=
begin
fapply pequiv_of_equiv,
{ apply pushout.symm},
{ exact ap inr (respect_pt f)⁻¹ ⬝ !glue⁻¹ ⬝ ap inl (respect_pt g)}
end
end
end pushout

View file

@ -1,7 +1,7 @@
/-
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
Quotients. This is a quotient without truncation for an arbitrary type-valued binary relation.
See also .set_quotient
@ -211,4 +211,110 @@ namespace quotient
end
end flattening
section
open is_equiv equiv prod prod.ops
variables {A : Type} (R : A → A → Type)
{B : Type} (Q : B → B → Type)
(f : A → B) (k : Πa a' : A, R a a' → Q (f a) (f a'))
include f k
protected definition functor [reducible] : quotient R → quotient Q :=
quotient.elim (λa, class_of Q (f a)) (λa a' r, eq_of_rel Q (k a a' r))
variables [F : is_equiv f] [K : Πa a', is_equiv (k a a')]
include F K
protected definition functor_inv [reducible] : quotient Q → quotient R :=
quotient.elim (λb, class_of R (f⁻¹ b))
(λb b' q, eq_of_rel R ((k (f⁻¹ b) (f⁻¹ b'))⁻¹
((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q)))
protected definition is_equiv [instance]
: is_equiv (quotient.functor R Q f k):=
begin
fapply adjointify _ (quotient.functor_inv R Q f k),
{ intro qb, induction qb with b b b' q,
{ apply ap (class_of Q), apply right_inv },
{ apply eq_pathover, rewrite [ap_id,ap_compose' (quotient.elim _ _)],
do 2 krewrite elim_eq_of_rel, rewrite (right_inv (k (f⁻¹ b) (f⁻¹ b'))),
have H1 : pathover (λz : B × B, Q z.1 z.2)
((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q)
(prod_eq (right_inv f b) (right_inv f b')) q,
begin apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] end,
have H2 : square
(ap (λx : (Σz : B × B, Q z.1 z.2), class_of Q x.1.1)
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1))
(ap (λx : (Σz : B × B, Q z.1 z.2), class_of Q x.1.2)
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1))
(eq_of_rel Q ((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q))
(eq_of_rel Q q),
from
natural_square_tr (λw : (Σz : B × B, Q z.1 z.2), eq_of_rel Q w.2)
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1),
krewrite (ap_compose' (class_of Q)) at H2,
krewrite (ap_compose' (λz : B × B, z.1)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr1 at H2, krewrite prod_eq_pr1 at H2,
krewrite (ap_compose' (class_of Q) (λx : (Σz : B × B, Q z.1 z.2), x.1.2)) at H2,
krewrite (ap_compose' (λz : B × B, z.2)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr2 at H2, krewrite prod_eq_pr2 at H2,
apply H2 } },
{ intro qa, induction qa with a a a' r,
{ apply ap (class_of R), apply left_inv },
{ apply eq_pathover, rewrite [ap_id,(ap_compose' (quotient.elim _ _))],
do 2 krewrite elim_eq_of_rel,
have H1 : pathover (λz : A × A, R z.1 z.2)
((left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r)
(prod_eq (left_inv f a) (left_inv f a')) r,
begin apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] end,
have H2 : square
(ap (λx : (Σz : A × A, R z.1 z.2), class_of R x.1.1)
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1))
(ap (λx : (Σz : A × A, R z.1 z.2), class_of R x.1.2)
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1))
(eq_of_rel R ((left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r))
(eq_of_rel R r),
begin
exact
natural_square_tr (λw : (Σz : A × A, R z.1 z.2), eq_of_rel R w.2)
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1)
end,
krewrite (ap_compose' (class_of R)) at H2,
krewrite (ap_compose' (λz : A × A, z.1)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr1 at H2, krewrite prod_eq_pr1 at H2,
krewrite (ap_compose' (class_of R) (λx : (Σz : A × A, R z.1 z.2), x.1.2)) at H2,
krewrite (ap_compose' (λz : A × A, z.2)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr2 at H2, krewrite prod_eq_pr2 at H2,
have H3 :
(k (f⁻¹ (f a)) (f⁻¹ (f a')))⁻¹
((right_inv f (f a))⁻¹ ▸ (right_inv f (f a'))⁻¹ ▸ k a a' r)
= (left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r,
begin
rewrite [adj f a,adj f a',ap_inv',ap_inv'],
rewrite [-(tr_compose _ f (left_inv f a')⁻¹ (k a a' r)),
-(tr_compose _ f (left_inv f a)⁻¹)],
rewrite [-(fn_tr_eq_tr_fn (left_inv f a')⁻¹ (λx, k a x) r),
-(fn_tr_eq_tr_fn (left_inv f a)⁻¹
(λx, k x (f⁻¹ (f a')))),
left_inv (k _ _)]
end,
rewrite H3, apply H2 } }
end
end
section
variables {A : Type} (R : A → A → Type)
{B : Type} (Q : B → B → Type)
(f : A ≃ B) (k : Πa a' : A, R a a' ≃ Q (f a) (f a'))
include f k
/- This could also be proved using ua, but then it wouldn't compute -/
protected definition equiv : quotient R ≃ quotient Q :=
equiv.mk (quotient.functor R Q f k) _
end
end quotient

View file

@ -1,116 +0,0 @@
/-
Copyright (c) 2015 Ulrik Buchholtz. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Ulrik Buchholtz
Functoriality of quotients and a condition for when an equivalence is induced.
-/
import types.sigma .quotient
open eq is_equiv equiv prod prod.ops sigma sigma.ops
namespace quotient
section
variables {A : Type} (R : A → A → Type)
{B : Type} (Q : B → B → Type)
(f : A → B) (k : Πa a' : A, R a a' → Q (f a) (f a'))
include f k
protected definition functor [reducible] : quotient R → quotient Q :=
quotient.elim (λa, class_of Q (f a)) (λa a' r, eq_of_rel Q (k a a' r))
variables [F : is_equiv f] [K : Πa a', is_equiv (k a a')]
include F K
protected definition functor_inv [reducible] : quotient Q → quotient R :=
quotient.elim (λb, class_of R (f⁻¹ b))
(λb b' q, eq_of_rel R ((k (f⁻¹ b) (f⁻¹ b'))⁻¹
((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q)))
protected definition is_equiv [instance]
: is_equiv (quotient.functor R Q f k):=
begin
fapply adjointify _ (quotient.functor_inv R Q f k),
{ intro qb, induction qb with b b b' q,
{ apply ap (class_of Q), apply right_inv },
{ apply eq_pathover, rewrite [ap_id,ap_compose' (quotient.elim _ _)],
do 2 krewrite elim_eq_of_rel, rewrite (right_inv (k (f⁻¹ b) (f⁻¹ b'))),
have H1 : pathover (λz : B × B, Q z.1 z.2)
((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q)
(prod_eq (right_inv f b) (right_inv f b')) q,
begin apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] end,
have H2 : square
(ap (λx : (Σz : B × B, Q z.1 z.2), class_of Q x.1.1)
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1))
(ap (λx : (Σz : B × B, Q z.1 z.2), class_of Q x.1.2)
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1))
(eq_of_rel Q ((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q))
(eq_of_rel Q q),
from
natural_square (λw : (Σz : B × B, Q z.1 z.2), eq_of_rel Q w.2)
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1),
krewrite (ap_compose' (class_of Q)) at H2,
krewrite (ap_compose' (λz : B × B, z.1)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr1 at H2, krewrite prod_eq_pr1 at H2,
krewrite (ap_compose' (class_of Q) (λx : (Σz : B × B, Q z.1 z.2), x.1.2)) at H2,
krewrite (ap_compose' (λz : B × B, z.2)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr2 at H2, krewrite prod_eq_pr2 at H2,
apply H2 } },
{ intro qa, induction qa with a a a' r,
{ apply ap (class_of R), apply left_inv },
{ apply eq_pathover, rewrite [ap_id,(ap_compose' (quotient.elim _ _))],
do 2 krewrite elim_eq_of_rel,
have H1 : pathover (λz : A × A, R z.1 z.2)
((left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r)
(prod_eq (left_inv f a) (left_inv f a')) r,
begin apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] end,
have H2 : square
(ap (λx : (Σz : A × A, R z.1 z.2), class_of R x.1.1)
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1))
(ap (λx : (Σz : A × A, R z.1 z.2), class_of R x.1.2)
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1))
(eq_of_rel R ((left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r))
(eq_of_rel R r),
begin
exact
natural_square (λw : (Σz : A × A, R z.1 z.2), eq_of_rel R w.2)
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1)
end,
krewrite (ap_compose' (class_of R)) at H2,
krewrite (ap_compose' (λz : A × A, z.1)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr1 at H2, krewrite prod_eq_pr1 at H2,
krewrite (ap_compose' (class_of R) (λx : (Σz : A × A, R z.1 z.2), x.1.2)) at H2,
krewrite (ap_compose' (λz : A × A, z.2)) at H2,
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
krewrite prod.ap_pr2 at H2, krewrite prod_eq_pr2 at H2,
have H3 :
(k (f⁻¹ (f a)) (f⁻¹ (f a')))⁻¹
((right_inv f (f a))⁻¹ ▸ (right_inv f (f a'))⁻¹ ▸ k a a' r)
= (left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r,
begin
rewrite [adj f a,adj f a',ap_inv',ap_inv'],
rewrite [-(tr_compose _ f (left_inv f a')⁻¹ (k a a' r)),
-(tr_compose _ f (left_inv f a)⁻¹)],
rewrite [-(fn_tr_eq_tr_fn (left_inv f a')⁻¹ (λx, k a x) r),
-(fn_tr_eq_tr_fn (left_inv f a)⁻¹
(λx, k x (f⁻¹ (f a')))),
left_inv (k _ _)]
end,
rewrite H3, apply H2 } }
end
end
section
variables {A : Type} (R : A → A → Type)
{B : Type} (Q : B → B → Type)
(f : A ≃ B) (k : Πa a' : A, R a a' ≃ Q (f a) (f a'))
include f k
/- This could also be proved using ua, but then it wouldn't compute -/
protected definition equiv : quotient R ≃ quotient Q :=
equiv.mk (quotient.functor R Q f k) _
end
end quotient

View file

@ -606,7 +606,7 @@ namespace trunc_two_quotient
begin
note Ht' := ap_ap_e_closure_elim tr (elim P0 P1 P2) (two_quotient.incl1 R Q) t',
note Ht := ap_ap_e_closure_elim tr (elim P0 P1 P2) (two_quotient.incl1 R Q) t,
note Hn := natural_square (ap_compose (elim P0 P1 P2) tr) (two_quotient.incl2 R Q q),
note Hn := natural_square_tr (ap_compose (elim P0 P1 P2) tr) (two_quotient.incl2 R Q q),
note H7 := eq_top_of_square (Ht⁻¹ʰ ⬝h Hn⁻¹ᵛ ⬝h Ht'), clear [Hn, Ht, Ht'],
unfold [ap02,incl2], rewrite [+ap_con,ap_inv,-ap_compose (ap _)],
xrewrite [H7, ↑function.compose, eq_top_of_square (elim_incl2 P0 P1 P2 q)], clear [H7],

View file

@ -151,7 +151,7 @@ namespace chain_complex
begin
refine _ ⬝ !con.assoc⁻¹,
apply whisker_left,
refine transport_eq_Fl _ _ ⬝ _,
refine eq_transport_Fl _ _ ⬝ _,
apply whisker_right,
refine inverse2 !ap_inv ⬝ !inv_inv ⬝ _,
refine ap_compose (fiber_sequence_fun n) pr₁ _ ⬝
@ -219,7 +219,7 @@ namespace chain_complex
{ reflexivity},
{ reflexivity},
refine ap (whisker_left _)
(transport_eq_Fl_idp_left (fiber_sequence_fun n)
(eq_transport_Fl_idp_left (fiber_sequence_fun n)
(respect_pt (fiber_sequence_fun n))) ⬝ _,
apply whisker_left_idp_con_eq_assoc
end
@ -264,26 +264,26 @@ namespace chain_complex
/- Now we are ready to define the long exact sequence of homotopy groups.
First we define its carrier -/
definition loopns : → Type*
definition loop_spaces : → Type*
| 0 := Y
| 1 := X
| 2 := pfiber f
| (k+3) := Ω (loopns k)
| (k+3) := Ω (loop_spaces k)
/- The maps between the homotopy groups -/
definition loopns_fun
: Π(n : ), loopns (n+1) →* loopns n
definition loop_spaces_fun
: Π(n : ), loop_spaces (n+1) →* loop_spaces n
| 0 := proof f qed
| 1 := proof ppoint f qed
| 2 := proof boundary_map qed
| (k+3) := proof ap1 (loopns_fun k) qed
| (k+3) := proof ap1 (loop_spaces_fun k) qed
definition loopns_fun_add3 [unfold_full] (n : ) :
loopns_fun (n + 3) = ap1 (loopns_fun n) :=
definition loop_spaces_fun_add3 [unfold_full] (n : ) :
loop_spaces_fun (n + 3) = ap1 (loop_spaces_fun n) :=
proof idp qed
definition fiber_sequence_pequiv_loopns :
Πn, fiber_sequence_carrier n ≃* loopns n
definition fiber_sequence_pequiv_loop_spaces :
Πn, fiber_sequence_carrier n ≃* loop_spaces n
| 0 := by reflexivity
| 1 := by reflexivity
| 2 := by reflexivity
@ -291,22 +291,22 @@ namespace chain_complex
begin
refine fiber_sequence_carrier_pequiv k ⬝e* _,
apply loop_pequiv_loop,
exact fiber_sequence_pequiv_loopns k
exact fiber_sequence_pequiv_loop_spaces k
end
definition fiber_sequence_pequiv_loopns_add3 (n : )
: fiber_sequence_pequiv_loopns (n + 3) =
ap1 (fiber_sequence_pequiv_loopns n) ∘* fiber_sequence_carrier_pequiv n :=
definition fiber_sequence_pequiv_loop_spaces_add3 (n : )
: fiber_sequence_pequiv_loop_spaces (n + 3) =
ap1 (fiber_sequence_pequiv_loop_spaces n) ∘* fiber_sequence_carrier_pequiv n :=
by reflexivity
definition fiber_sequence_pequiv_loopns_3_phomotopy
: fiber_sequence_pequiv_loopns 3 ~* proof fiber_sequence_carrier_pequiv nat.zero qed :=
definition fiber_sequence_pequiv_loop_spaces_3_phomotopy
: fiber_sequence_pequiv_loop_spaces 3 ~* proof fiber_sequence_carrier_pequiv nat.zero qed :=
begin
refine pwhisker_right _ ap1_pid ⬝* _,
apply pid_pcompose
end
definition pid_or_pinverse : Π(n : ), loopns n ≃* loopns n
definition pid_or_pinverse : Π(n : ), loop_spaces n ≃* loop_spaces n
| 0 := pequiv.rfl
| 1 := pequiv.rfl
| 2 := pequiv.rfl
@ -321,15 +321,15 @@ namespace chain_complex
pid_or_pinverse (n + 4) ~* pinverse ∘* Ω→(pid_or_pinverse (n + 1))
| 0 := begin rewrite [pid_or_pinverse_add4, + to_pmap_pequiv_trans],
replace pid_or_pinverse (0 + 1) with pequiv.refl X,
rewrite [loop_pequiv_loop_rfl, ▸*], refine !pid_pcompose ⬝* _,
refine pwhisker_right _ !loop_pequiv_loop_rfl ⬝* _, refine !pid_pcompose ⬝* _,
exact !pcompose_pid⁻¹* ⬝* pwhisker_left _ !ap1_pid⁻¹* end
| 1 := begin rewrite [pid_or_pinverse_add4, + to_pmap_pequiv_trans],
replace pid_or_pinverse (1 + 1) with pequiv.refl (pfiber f),
rewrite [loop_pequiv_loop_rfl, ▸*], refine !pid_pcompose ⬝* _,
refine pwhisker_right _ !loop_pequiv_loop_rfl ⬝* _, refine !pid_pcompose ⬝* _,
exact !pcompose_pid⁻¹* ⬝* pwhisker_left _ !ap1_pid⁻¹* end
| 2 := begin rewrite [pid_or_pinverse_add4, + to_pmap_pequiv_trans],
replace pid_or_pinverse (2 + 1) with pequiv.refl (Ω Y),
rewrite [loop_pequiv_loop_rfl, ▸*], refine !pid_pcompose ⬝* _,
refine pwhisker_right _ !loop_pequiv_loop_rfl ⬝* _, refine !pid_pcompose ⬝* _,
exact !pcompose_pid⁻¹* ⬝* pwhisker_left _ !ap1_pid⁻¹* end
| (k+3) :=
begin
@ -343,30 +343,30 @@ namespace chain_complex
{ refine !ap1_pinverse⁻¹*}
end
theorem fiber_sequence_phomotopy_loopns : Π(n : ),
fiber_sequence_pequiv_loopns n ∘* fiber_sequence_fun n ~*
(loopns_fun n ∘* pid_or_pinverse (n + 1)) ∘* fiber_sequence_pequiv_loopns (n + 1)
theorem fiber_sequence_phomotopy_loop_spaces : Π(n : ),
fiber_sequence_pequiv_loop_spaces n ∘* fiber_sequence_fun n ~*
(loop_spaces_fun n ∘* pid_or_pinverse (n + 1)) ∘* fiber_sequence_pequiv_loop_spaces (n + 1)
| 0 := proof proof phomotopy.rfl qed ⬝* pwhisker_right _ !pcompose_pid⁻¹* qed
| 1 := by reflexivity
| 2 :=
begin
refine !pid_pcompose ⬝* _,
replace loopns_fun 2 with boundary_map,
refine _ ⬝* pwhisker_left _ fiber_sequence_pequiv_loopns_3_phomotopy⁻¹*,
replace loop_spaces_fun 2 with boundary_map,
refine _ ⬝* pwhisker_left _ fiber_sequence_pequiv_loop_spaces_3_phomotopy⁻¹*,
apply phomotopy_of_pinv_right_phomotopy,
exact !pid_pcompose⁻¹*
end
| (k+3) :=
begin
replace (k + 3 + 1) with (k + 1 + 3),
rewrite [fiber_sequence_pequiv_loopns_add3 k,
fiber_sequence_pequiv_loopns_add3 (k+1)],
rewrite [fiber_sequence_pequiv_loop_spaces_add3 k,
fiber_sequence_pequiv_loop_spaces_add3 (k+1)],
refine !passoc ⬝* _,
refine pwhisker_left _ (fiber_sequence_fun_phomotopy k) ⬝* _,
refine !passoc⁻¹* ⬝* _ ⬝* !passoc,
apply pwhisker_right,
replace (k + 1 + 3) with (k + 4),
xrewrite [loopns_fun_add3, pid_or_pinverse_add4, to_pmap_pequiv_trans],
xrewrite [loop_spaces_fun_add3, pid_or_pinverse_add4, to_pmap_pequiv_trans],
refine _ ⬝* !passoc⁻¹*,
refine _ ⬝* pwhisker_left _ !passoc⁻¹*,
refine _ ⬝* pwhisker_left _ (pwhisker_left _ !ap1_pcompose_pinverse),
@ -374,16 +374,16 @@ namespace chain_complex
apply pwhisker_right,
refine !ap1_pcompose⁻¹* ⬝* _ ⬝* !ap1_pcompose ⬝* pwhisker_right _ !ap1_pcompose,
apply ap1_phomotopy,
exact fiber_sequence_phomotopy_loopns k
exact fiber_sequence_phomotopy_loop_spaces k
end
definition pid_or_pinverse_right : Π(n : ), loopns n →* loopns n
definition pid_or_pinverse_right : Π(n : ), loop_spaces n →* loop_spaces n
| 0 := !pid
| 1 := !pid
| 2 := !pid
| (k+3) := Ω→(pid_or_pinverse_right k) ∘* pinverse
definition pid_or_pinverse_left : Π(n : ), loopns n →* loopns n
definition pid_or_pinverse_left : Π(n : ), loop_spaces n →* loop_spaces n
| 0 := pequiv.rfl
| 1 := pequiv.rfl
| 2 := pequiv.rfl
@ -400,14 +400,14 @@ namespace chain_complex
by reflexivity
theorem pid_or_pinverse_commute_right : Π(n : ),
loopns_fun n ~* pid_or_pinverse_right n ∘* loopns_fun n ∘* pid_or_pinverse (n + 1)
loop_spaces_fun n ~* pid_or_pinverse_right n ∘* loop_spaces_fun n ∘* pid_or_pinverse (n + 1)
| 0 := proof !pcompose_pid⁻¹* ⬝* !pid_pcompose⁻¹* qed
| 1 := proof !pcompose_pid⁻¹* ⬝* !pid_pcompose⁻¹* qed
| 2 := proof !pcompose_pid⁻¹* ⬝* !pid_pcompose⁻¹* qed
| (k+3) :=
begin
replace (k + 3 + 1) with (k + 4),
rewrite [pid_or_pinverse_right_add3, loopns_fun_add3],
rewrite [pid_or_pinverse_right_add3, loop_spaces_fun_add3],
refine _ ⬝* pwhisker_left _ (pwhisker_left _ !pid_or_pinverse_add4_rev⁻¹*),
refine ap1_phomotopy (pid_or_pinverse_commute_right k) ⬝* _,
refine !ap1_pcompose ⬝* _ ⬝* !passoc⁻¹*,
@ -421,7 +421,7 @@ namespace chain_complex
end
theorem pid_or_pinverse_commute_left : Π(n : ),
loopns_fun n ∘* pid_or_pinverse_left (n + 1) ~* pid_or_pinverse n ∘* loopns_fun n
loop_spaces_fun n ∘* pid_or_pinverse_left (n + 1) ~* pid_or_pinverse n ∘* loop_spaces_fun n
| 0 := proof !pcompose_pid ⬝* !pid_pcompose⁻¹* qed
| 1 := proof !pcompose_pid ⬝* !pid_pcompose⁻¹* qed
| 2 := proof !pcompose_pid ⬝* !pid_pcompose⁻¹* qed
@ -431,7 +431,7 @@ namespace chain_complex
replace (k + 4 + 1) with (k + 5),
rewrite [pid_or_pinverse_left_add5, pid_or_pinverse_add4, to_pmap_pequiv_trans],
replace (k + 4) with (k + 1 + 3),
rewrite [loopns_fun_add3],
rewrite [loop_spaces_fun_add3],
refine !passoc⁻¹* ⬝* _ ⬝* !passoc⁻¹*,
refine _ ⬝* pwhisker_left _ !ap1_pcompose_pinverse,
refine _ ⬝* !passoc,
@ -440,23 +440,23 @@ namespace chain_complex
exact ap1_phomotopy (pid_or_pinverse_commute_left (k+1))
end
definition LES_of_loopns' [constructor] : type_chain_complex + :=
definition LES_of_loop_spaces' [constructor] : type_chain_complex + :=
transfer_type_chain_complex
fiber_sequence
(λn, loopns_fun n ∘* pid_or_pinverse (n + 1))
fiber_sequence_pequiv_loopns
fiber_sequence_phomotopy_loopns
(λn, loop_spaces_fun n ∘* pid_or_pinverse (n + 1))
fiber_sequence_pequiv_loop_spaces
fiber_sequence_phomotopy_loop_spaces
definition LES_of_loopns [constructor] : type_chain_complex + :=
definition LES_of_loop_spaces [constructor] : type_chain_complex + :=
type_chain_complex_cancel_aut
LES_of_loopns'
loopns_fun
LES_of_loop_spaces'
loop_spaces_fun
pid_or_pinverse
pid_or_pinverse_right
(λn x, idp)
pid_or_pinverse_commute_right
definition is_exact_LES_of_loopns : is_exact_t LES_of_loopns :=
definition is_exact_LES_of_loop_spaces : is_exact_t LES_of_loop_spaces :=
begin
intro n,
refine is_exact_at_t_cancel_aut n pid_or_pinverse_left _ _ pid_or_pinverse_commute_left _,
@ -470,37 +470,37 @@ namespace chain_complex
PART 3
--------------/
definition loopns2 [reducible] : +3 → Type*
definition loop_spaces2 [reducible] : +3 → Type*
| (n, fin.mk 0 H) := Ω[n] Y
| (n, fin.mk 1 H) := Ω[n] X
| (n, fin.mk k H) := Ω[n] (pfiber f)
definition loopns2_add1 (n : ) : Π(x : fin 3),
loopns2 (n+1, x) = Ω (loopns2 (n, x))
definition loop_spaces2_add1 (n : ) : Π(x : fin 3),
loop_spaces2 (n+1, x) = Ω (loop_spaces2 (n, x))
| (fin.mk 0 H) := by reflexivity
| (fin.mk 1 H) := by reflexivity
| (fin.mk 2 H) := by reflexivity
| (fin.mk (k+3) H) := begin exfalso, apply lt_le_antisymm H, apply le_add_left end
definition loopns_fun2 : Π(n : +3), loopns2 (S n) →* loopns2 n
definition loop_spaces_fun2 : Π(n : +3), loop_spaces2 (S n) →* loop_spaces2 n
| (n, fin.mk 0 H) := proof Ω→[n] f qed
| (n, fin.mk 1 H) := proof Ω→[n] (ppoint f) qed
| (n, fin.mk 2 H) := proof Ω→[n] boundary_map ∘* loopn_succ_in Y n qed
| (n, fin.mk (k+3) H) := begin exfalso, apply lt_le_antisymm H, apply le_add_left end
definition loopns_fun2_add1_0 (n : ) (H : 0 < succ 2)
: loopns_fun2 (n+1, fin.mk 0 H) ~*
cast proof idp qed ap1 (loopns_fun2 (n, fin.mk 0 H)) :=
definition loop_spaces_fun2_add1_0 (n : ) (H : 0 < succ 2)
: loop_spaces_fun2 (n+1, fin.mk 0 H) ~*
cast proof idp qed ap1 (loop_spaces_fun2 (n, fin.mk 0 H)) :=
by reflexivity
definition loopns_fun2_add1_1 (n : ) (H : 1 < succ 2)
: loopns_fun2 (n+1, fin.mk 1 H) ~*
cast proof idp qed ap1 (loopns_fun2 (n, fin.mk 1 H)) :=
definition loop_spaces_fun2_add1_1 (n : ) (H : 1 < succ 2)
: loop_spaces_fun2 (n+1, fin.mk 1 H) ~*
cast proof idp qed ap1 (loop_spaces_fun2 (n, fin.mk 1 H)) :=
by reflexivity
definition loopns_fun2_add1_2 (n : ) (H : 2 < succ 2)
: loopns_fun2 (n+1, fin.mk 2 H) ~*
cast proof idp qed ap1 (loopns_fun2 (n, fin.mk 2 H)) :=
definition loop_spaces_fun2_add1_2 (n : ) (H : 2 < succ 2)
: loop_spaces_fun2 (n+1, fin.mk 2 H) ~*
cast proof idp qed ap1 (loop_spaces_fun2 (n, fin.mk 2 H)) :=
proof !ap1_pcompose⁻¹* qed
definition nat_of_str [unfold 2] [reducible] {n : } : × fin (succ n) → :=
@ -537,39 +537,39 @@ namespace chain_complex
note: in the following theorem the (n+1) case is 3 times the same,
so maybe this can be simplified
-/
definition loopns2_pequiv' : Π(n : ) (x : fin (nat.succ 2)),
loopns (nat_of_str (n, x)) ≃* loopns2 (n, x)
definition loop_spaces2_pequiv' : Π(n : ) (x : fin (nat.succ 2)),
loop_spaces (nat_of_str (n, x)) ≃* loop_spaces2 (n, x)
| 0 (fin.mk 0 H) := by reflexivity
| 0 (fin.mk 1 H) := by reflexivity
| 0 (fin.mk 2 H) := by reflexivity
| (n+1) (fin.mk 0 H) :=
begin
apply loop_pequiv_loop,
rexact loopns2_pequiv' n (fin.mk 0 H)
rexact loop_spaces2_pequiv' n (fin.mk 0 H)
end
| (n+1) (fin.mk 1 H) :=
begin
apply loop_pequiv_loop,
rexact loopns2_pequiv' n (fin.mk 1 H)
rexact loop_spaces2_pequiv' n (fin.mk 1 H)
end
| (n+1) (fin.mk 2 H) :=
begin
apply loop_pequiv_loop,
rexact loopns2_pequiv' n (fin.mk 2 H)
rexact loop_spaces2_pequiv' n (fin.mk 2 H)
end
| n (fin.mk (k+3) H) := begin exfalso, apply lt_le_antisymm H, apply le_add_left end
definition loopns2_pequiv : Π(x : +3),
loopns (nat_of_str x) ≃* loopns2 x
| (n, x) := loopns2_pequiv' n x
definition loop_spaces2_pequiv : Π(x : +3),
loop_spaces (nat_of_str x) ≃* loop_spaces2 x
| (n, x) := loop_spaces2_pequiv' n x
local attribute loop_pequiv_loop [reducible]
/- all cases where n>0 are basically the same -/
definition loopns_fun2_phomotopy (x : +3) :
loopns2_pequiv x ∘* loopns_fun (nat_of_str x) ~*
(loopns_fun2 x ∘* loopns2_pequiv (S x))
∘* pcast (ap (loopns) (nat_of_str_3S x)) :=
definition loop_spaces_fun2_phomotopy (x : +3) :
loop_spaces2_pequiv x ∘* loop_spaces_fun (nat_of_str x) ~*
(loop_spaces_fun2 x ∘* loop_spaces2_pequiv (S x))
∘* pcast (ap (loop_spaces) (nat_of_str_3S x)) :=
begin
cases x with n x, cases x with k H,
do 3 (cases k with k; rotate 1),
@ -579,7 +579,7 @@ namespace chain_complex
{ refine !pid_pcompose ⬝* _ ⬝* !pcompose_pid⁻¹* ⬝* !pcompose_pid⁻¹*,
reflexivity},
{ refine _ ⬝* !pcompose_pid⁻¹*,
refine _ ⬝* pwhisker_right _ !loopns_fun2_add1_0⁻¹*,
refine _ ⬝* pwhisker_right _ !loop_spaces_fun2_add1_0⁻¹*,
refine !ap1_pcompose⁻¹* ⬝* _ ⬝* !ap1_pcompose, apply ap1_phomotopy,
exact IH ⬝* !pcompose_pid}},
{ /-k=1-/
@ -587,7 +587,7 @@ namespace chain_complex
{ refine !pid_pcompose ⬝* _ ⬝* !pcompose_pid⁻¹* ⬝* !pcompose_pid⁻¹*,
reflexivity},
{ refine _ ⬝* !pcompose_pid⁻¹*,
refine _ ⬝* pwhisker_right _ !loopns_fun2_add1_1⁻¹*,
refine _ ⬝* pwhisker_right _ !loop_spaces_fun2_add1_1⁻¹*,
refine !ap1_pcompose⁻¹* ⬝* _ ⬝* !ap1_pcompose, apply ap1_phomotopy,
exact IH ⬝* !pcompose_pid}},
{ /-k=2-/
@ -595,36 +595,36 @@ namespace chain_complex
{ refine !pid_pcompose ⬝* _ ⬝* !pcompose_pid⁻¹*,
refine !pcompose_pid⁻¹* ⬝* pconcat2 _ _,
{ exact (pcompose_pid (chain_complex.boundary_map f))⁻¹*},
{ refine cast (ap (λx, _ ~* x) !loop_pequiv_loop_rfl)⁻¹ _, reflexivity}},
{ refine !loop_pequiv_loop_rfl⁻¹* }},
{ refine _ ⬝* !pcompose_pid⁻¹*,
refine _ ⬝* pwhisker_right _ !loopns_fun2_add1_2⁻¹*,
refine _ ⬝* pwhisker_right _ !loop_spaces_fun2_add1_2⁻¹*,
refine !ap1_pcompose⁻¹* ⬝* _ ⬝* !ap1_pcompose, apply ap1_phomotopy,
exact IH ⬝* !pcompose_pid}},
end
definition LES_of_loopns2 [constructor] : type_chain_complex +3 :=
definition LES_of_loop_spaces2 [constructor] : type_chain_complex +3 :=
transfer_type_chain_complex2
LES_of_loopns
LES_of_loop_spaces
!fin_prod_nat_equiv_nat
nat_of_str_3S
@loopns_fun2
@loopns2_pequiv
@loop_spaces_fun2
@loop_spaces2_pequiv
begin
intro m x,
refine loopns_fun2_phomotopy m x ⬝ _,
apply ap (loopns_fun2 m), apply ap (loopns2_pequiv (S m)),
refine loop_spaces_fun2_phomotopy m x ⬝ _,
apply ap (loop_spaces_fun2 m), apply ap (loop_spaces2_pequiv (S m)),
esimp, exact ap010 cast !ap_compose⁻¹ x
end
definition is_exact_LES_of_loopns2 : is_exact_t LES_of_loopns2 :=
definition is_exact_LES_of_loop_spaces2 : is_exact_t LES_of_loop_spaces2 :=
begin
intro n,
apply is_exact_at_t_transfer2,
apply is_exact_LES_of_loopns
apply is_exact_LES_of_loop_spaces
end
definition LES_of_homotopy_groups' [constructor] : chain_complex +3 :=
trunc_chain_complex LES_of_loopns2
trunc_chain_complex LES_of_loop_spaces2
/--------------
PART 4
@ -635,8 +635,8 @@ namespace chain_complex
| (n, fin.mk 1 H) := π[n] X
| (n, fin.mk k H) := π[n] (pfiber f)
definition homotopy_groups_pequiv_loopns2 [reducible]
: Π(n : +3), ptrunc 0 (loopns2 n) ≃* homotopy_groups n
definition homotopy_groups_pequiv_loop_spaces2 [reducible]
: Π(n : +3), ptrunc 0 (loop_spaces2 n) ≃* homotopy_groups n
| (n, fin.mk 0 H) := by reflexivity
| (n, fin.mk 1 H) := by reflexivity
| (n, fin.mk 2 H) := by reflexivity
@ -649,9 +649,9 @@ namespace chain_complex
proof π→[n] boundary_map ∘* homotopy_group_succ_in Y n qed
| (n, fin.mk (k+3) H) := begin exfalso, apply lt_le_antisymm H, apply le_add_left end
definition homotopy_groups_fun_phomotopy_loopns_fun2 [reducible]
: Π(n : +3), homotopy_groups_pequiv_loopns2 n ∘* ptrunc_functor 0 (loopns_fun2 n) ~*
homotopy_groups_fun n ∘* homotopy_groups_pequiv_loopns2 (S n)
definition homotopy_groups_fun_phomotopy_loop_spaces_fun2 [reducible]
: Π(n : +3), homotopy_groups_pequiv_loop_spaces2 n ∘* ptrunc_functor 0 (loop_spaces_fun2 n) ~*
homotopy_groups_fun n ∘* homotopy_groups_pequiv_loop_spaces2 (S n)
| (n, fin.mk 0 H) := by reflexivity
| (n, fin.mk 1 H) := by reflexivity
| (n, fin.mk 2 H) :=
@ -665,15 +665,15 @@ namespace chain_complex
transfer_chain_complex
LES_of_homotopy_groups'
homotopy_groups_fun
homotopy_groups_pequiv_loopns2
homotopy_groups_fun_phomotopy_loopns_fun2
homotopy_groups_pequiv_loop_spaces2
homotopy_groups_fun_phomotopy_loop_spaces_fun2
definition is_exact_LES_of_homotopy_groups : is_exact LES_of_homotopy_groups :=
begin
intro n,
apply is_exact_at_transfer,
apply is_exact_at_trunc,
apply is_exact_LES_of_loopns2
apply is_exact_LES_of_loop_spaces2
end
variable (n : )
@ -777,14 +777,14 @@ namespace chain_complex
| (n, fin.mk (k+3) H) := begin exfalso, apply lt_le_antisymm H, apply le_add_left end
definition fibration_sequence_pequiv : Π(x : +3),
loopns2 f x ≃* fibration_sequence_car x
loop_spaces2 f x ≃* fibration_sequence_car x
| (n, fin.mk 0 H) := by reflexivity
| (n, fin.mk 1 H) := by reflexivity
| (n, fin.mk 2 H) := loopn_pequiv_loopn n e
| (n, fin.mk (k+3) H) := begin exfalso, apply lt_le_antisymm H, apply le_add_left end
definition fibration_sequence_fun_phomotopy : Π(x : +3),
fibration_sequence_pequiv x ∘* loopns_fun2 f x ~*
fibration_sequence_pequiv x ∘* loop_spaces_fun2 f x ~*
(fibration_sequence_fun x ∘* fibration_sequence_pequiv (S x))
| (n, fin.mk 0 H) := by reflexivity
| (n, fin.mk 1 H) :=
@ -796,7 +796,7 @@ namespace chain_complex
definition type_fibration_sequence [constructor] : type_chain_complex +3 :=
transfer_type_chain_complex
(LES_of_loopns2 f)
(LES_of_loop_spaces2 f)
fibration_sequence_fun
fibration_sequence_pequiv
fibration_sequence_fun_phomotopy
@ -805,7 +805,7 @@ namespace chain_complex
begin
intro n,
apply is_exact_at_t_transfer,
apply is_exact_LES_of_loopns2
apply is_exact_LES_of_loop_spaces2
end
definition fibration_sequence [constructor] : chain_complex +3 :=

View file

@ -18,7 +18,7 @@ chain complexes:
-/
import types.int algebra.group_theory types.fin
import types.int algebra.group_theory types.fin types.unit
open eq pointed int unit is_equiv equiv is_trunc trunc function algebra group sigma.ops
sum prod nat bool fin
@ -61,8 +61,21 @@ notation `-6` := stratified - 5
notation `+6` := stratified + 5
notation `-6` := stratified - 5
namespace succ_str
protected definition add [reducible] {N : succ_str} (n : N) (k : ) : N :=
iterate S k n
infix ` +' `:65 := succ_str.add
definition add_succ {N : succ_str} (n : N) (k : ) : n +' (k + 1) = (S n) +' k :=
by induction k with k p; reflexivity; exact ap S p
end succ_str
namespace chain_complex
export [notation] succ_str
/-
We define "type chain complexes" which are chain complexes without the
"set"-requirement. Exactness is formulated without propositional truncation.
@ -74,7 +87,6 @@ namespace chain_complex
section
variables {N : succ_str} (X : type_chain_complex N) (n : N)
definition tcc_to_car [unfold 2] [coercion] := @type_chain_complex.car
definition tcc_to_fn [unfold 2] : X (S n) →* X n := type_chain_complex.fn X n
definition tcc_is_chain_complex [unfold 2]

View file

@ -15,10 +15,15 @@ namespace cofiber
section
parameters {A B : Type} (f : A → B)
protected definition base [constructor] : cofiber f := inl ⋆
protected definition base : cofiber f := inl ⋆
protected definition cod [constructor] : B → cofiber f := inr
protected definition cod : B → cofiber f := inr
parameter {f}
protected definition glue (a : A) : cofiber.base f = cofiber.cod f (f a) :=
pushout.glue a
parameter (f)
protected definition contr_of_equiv [H : is_equiv f] : is_contr (cofiber f) :=
begin
fapply is_contr.mk, exact base,
@ -30,53 +35,42 @@ namespace cofiber
refine _ ⬝hp (ap (ap inr) !adj⁻¹), refine _ ⬝hp !ap_compose, apply square_Flr_idp_ap },
end
protected definition rec {A : Type} {B : Type} {f : A → B} {P : cofiber f → Type}
(Pinl : P (inl ⋆)) (Pinr : Π (x : B), P (inr x))
(Pglue : Π (x : A), pathover P Pinl (glue x) (Pinr (f x))) :
parameter {f}
protected definition rec {P : cofiber f → Type}
(Pbase : P base) (Pcod : Π (b : B), P (cod b))
(Pglue : Π (a : A), pathover P Pbase (glue a) (Pcod (f a))) :
(Π y, P y) :=
begin
intro y, induction y, induction x, exact Pinl, exact Pinr x, esimp, exact Pglue x,
intro y, induction y, induction x, exact Pbase, exact Pcod x, esimp, exact Pglue x,
end
protected definition rec_on {A : Type} {B : Type} {f : A → B} {P : cofiber f → Type}
(y : cofiber f) (Pinl : P (inl ⋆)) (Pinr : Π (x : B), P (inr x))
(Pglue : Π (x : A), pathover P Pinl (glue x) (Pinr (f x))) : P y :=
begin
induction y, induction x, exact Pinl, exact Pinr x, esimp, exact Pglue x,
end
protected definition rec_on {P : cofiber f → Type} (y : cofiber f)
(Pbase : P base) (Pcod : Π (b : B), P (cod b))
(Pglue : Π (a : A), pathover P Pbase (glue a) (Pcod (f a))) : P y :=
cofiber.rec Pbase Pcod Pglue y
protected definition elim {P : Type} (Pbase : P) (Pcod : B → P)
(Pglue : Π (x : A), Pbase = Pcod (f x)) (y : cofiber f) : P :=
pushout.elim (λu, Pbase) Pcod Pglue y
protected definition elim_on {P : Type} (y : cofiber f) (Pbase : P) (Pcod : B → P)
(Pglue : Π (x : A), Pbase = Pcod (f x)) : P :=
cofiber.elim Pbase Pcod Pglue y
end
end cofiber
attribute cofiber.base cofiber.cod [constructor]
attribute cofiber.rec cofiber.elim [recursor 8] [unfold 8]
attribute cofiber.rec_on cofiber.elim_on [unfold 5]
-- pointed version
definition pcofiber {A B : Type*} (f : A →* B) : Type* := ppushout (pconst A punit) f
definition pcofiber [constructor] {A B : Type*} (f : A →* B) : Type* :=
pointed.MK (cofiber f) !cofiber.base
namespace cofiber
protected definition prec {A B : Type*} {f : A →* B} {P : pcofiber f → Type}
(Pinl : P (inl ⋆)) (Pinr : Π (x : B), P (inr x))
(Pglue : Π (x : A), pathover P Pinl (pglue x) (Pinr (f x))) :
(Π (y : pcofiber f), P y) :=
begin
intro y, induction y, induction x, exact Pinl, exact Pinr x, esimp, exact Pglue x
end
protected definition prec_on {A B : Type*} {f : A →* B} {P : pcofiber f → Type}
(y : pcofiber f) (Pinl : P (inl ⋆)) (Pinr : Π (x : B), P (inr x))
(Pglue : Π (x : A), pathover P Pinl (pglue x) (Pinr (f x))) : P y :=
begin
induction y, induction x, exact Pinl, exact Pinr x, esimp, exact Pglue x
end
protected definition pelim_on {A B C : Type*} {f : A →* B} (y : pcofiber f)
(c : C) (g : B → C) (p : Π x, c = g (f x)) : C :=
begin
fapply pushout.elim_on y, exact (λ x, c), exact g, exact p
end
--TODO more pointed recursors
variables (A : Type*)
definition cofiber_unit : pcofiber (pconst A punit) ≃* psusp A :=

View file

@ -22,7 +22,7 @@ namespace hopf
induction x,
{ reflexivity },
{ apply eq_pathover, induction y,
{ exact natural_square_tr
{ exact natural_square
(λa : S¹, ap (λb : S¹, b * z) (circle_mul_base a))
loop },
{ apply is_prop.elimo, apply is_trunc_square } }

View file

@ -366,3 +366,76 @@ namespace is_conn
end
end is_conn
/-
(bundled) connected types, possibly also truncated or with a point
The notation is n-Type*[k] for k-connected n-truncated pointed types, and you can remove
`n-`, `[k]` or `*` in any combination to remove some conditions
-/
structure conntype (n : ℕ₋₂) : Type :=
(carrier : Type)
(struct : is_conn n carrier)
notation `Type[`:95 n:0 `]`:0 := conntype n
attribute conntype.carrier [coercion]
attribute conntype.struct [instance] [priority 1300]
section
universe variable u
structure pconntype (n : ℕ₋₂) extends conntype.{u} n, pType.{u}
notation `Type*[`:95 n:0 `]`:0 := pconntype n
/-
There are multiple coercions from pconntype to Type. Type class inference doesn't recognize
that all of them are definitionally equal (for performance reasons). One instance is
automatically generated, and we manually add the missing instances.
-/
definition is_conn_pconntype [instance] {n : ℕ₋₂} (X : Type*[n]) : is_conn n X :=
conntype.struct X
structure truncconntype (n k : ℕ₋₂) extends trunctype.{u} n,
conntype.{u} k renaming struct→conn_struct
notation n `-Type[`:95 k:0 `]`:0 := truncconntype n k
definition is_conn_truncconntype [instance] {n k : ℕ₋₂} (X : n-Type[k]) :
is_conn k (truncconntype._trans_of_to_trunctype X) :=
conntype.struct X
definition is_trunc_truncconntype [instance] {n k : ℕ₋₂} (X : n-Type[k]) : is_trunc n X :=
trunctype.struct X
structure ptruncconntype (n k : ℕ₋₂) extends ptrunctype.{u} n,
pconntype.{u} k renaming struct→conn_struct
notation n `-Type*[`:95 k:0 `]`:0 := ptruncconntype n k
attribute ptruncconntype._trans_of_to_pconntype ptruncconntype._trans_of_to_ptrunctype
ptruncconntype._trans_of_to_pconntype_1 ptruncconntype._trans_of_to_ptrunctype_1
ptruncconntype._trans_of_to_pconntype_2 ptruncconntype._trans_of_to_ptrunctype_2
ptruncconntype.to_pconntype ptruncconntype.to_ptrunctype
truncconntype._trans_of_to_conntype truncconntype._trans_of_to_trunctype
truncconntype.to_conntype truncconntype.to_trunctype [unfold 3]
attribute pconntype._trans_of_to_conntype pconntype._trans_of_to_pType
pconntype.to_pType pconntype.to_conntype [unfold 2]
definition is_conn_ptruncconntype [instance] {n k : ℕ₋₂} (X : n-Type*[k]) :
is_conn k (ptruncconntype._trans_of_to_ptrunctype X) :=
conntype.struct X
definition is_trunc_ptruncconntype [instance] {n k : ℕ₋₂} (X : n-Type*[k]) :
is_trunc n (ptruncconntype._trans_of_to_pconntype X) :=
trunctype.struct X
definition ptruncconntype_eq {n k : ℕ₋₂} {X Y : n-Type*[k]} (p : X ≃* Y) : X = Y :=
begin
induction X with X Xt Xp Xc, induction Y with Y Yt Yp Yc,
note q := pType_eq_elim (eq_of_pequiv p),
cases q with r s, esimp at *, induction r,
exact ap0111 (ptruncconntype.mk X) !is_prop.elim (eq_of_pathover_idp s) !is_prop.elim
end
end

View file

@ -6,7 +6,7 @@ Authors: Floris van Doorn
Declaration of mapping cylinders
-/
import hit.quotient
import hit.quotient types.fiber
open quotient eq sum equiv fiber

View file

@ -17,12 +17,12 @@ namespace freudenthal section
/-
This proof is ported from Agda
This is the 95% version of the Freudenthal Suspension Theorem, which means that we don't
prove that loop_susp_unit : A →* Ω(psusp A) is 2n-connected (if A is n-connected),
prove that loop_psusp_unit : A →* Ω(psusp A) is 2n-connected (if A is n-connected),
but instead we only prove that it induces an equivalence on the first 2n homotopy groups.
-/
private definition up (a : A) : north = north :> susp A :=
loop_susp_unit A a
loop_psusp_unit A a
definition code_merid : A → ptrunc (n + n) A → ptrunc (n + n) A :=
begin
@ -162,7 +162,7 @@ namespace freudenthal section
pequiv_of_equiv equiv' decode_north_pt
-- We don't prove this:
-- theorem freudenthal_suspension : is_conn_fun (n+n) (loop_susp_unit A) := sorry
-- theorem freudenthal_suspension : is_conn_fun (n+n) (loop_psusp_unit A) := sorry
end end freudenthal
@ -198,6 +198,23 @@ begin
apply homotopy_group_succ_in_con}
end
definition to_pmap_freudenthal_pequiv {A : Type*} (n k : ) [is_conn n A] (H : k ≤ 2 * n)
: freudenthal_pequiv A H ~* ptrunc_functor k (loop_psusp_unit A) :=
begin
fapply phomotopy.mk,
{ intro x, induction x with a, reflexivity },
{ refine !idp_con ⬝ _, refine _ ⬝ ap02 _ !idp_con⁻¹, refine _ ⬝ !ap_compose, apply ap_compose }
end
definition ptrunc_elim_freudenthal_pequiv {A B : Type*} (n k : ) [is_conn n A] (H : k ≤ 2 * n)
(f : A →* Ω B) [is_trunc (k.+1) (B)] :
ptrunc.elim k (Ω→ (psusp.elim f)) ∘* freudenthal_pequiv A H ~* ptrunc.elim k f :=
begin
refine pwhisker_left _ !to_pmap_freudenthal_pequiv ⬝* _,
refine !ptrunc_elim_ptrunc_functor ⬝* _,
exact ptrunc_elim_phomotopy k !ap1_psusp_elim,
end
namespace susp
definition iterate_psusp_stability_pequiv (A : Type*) {k n : } [is_conn 0 A]
@ -242,4 +259,23 @@ namespace susp
freudenthal_homotopy_group_isomorphism (pointed.MK (iterate_susp (n + 1) A) !north)
(stability_helper1 H)
definition iterated_freudenthal_pequiv (A : Type*) {n k m : } [HA : is_conn n A] (H : k ≤ 2 * n)
: ptrunc k A ≃* ptrunc k (Ω[m] (iterate_psusp m A)) :=
begin
revert A n k HA H, induction m with m IH: intro A n k HA H,
{ reflexivity},
{ have H2 : succ k ≤ 2 * succ n,
from calc
succ k ≤ succ (2 * n) : succ_le_succ H
... ≤ 2 * succ n : self_le_succ,
exact calc
ptrunc k A ≃* ptrunc k (Ω (psusp A)) : freudenthal_pequiv A H
... ≃* Ω (ptrunc (succ k) (psusp A)) : loop_ptrunc_pequiv
... ≃* Ω (ptrunc (succ k) (Ω[m] (iterate_psusp m (psusp A)))) :
loop_pequiv_loop (IH (psusp A) (succ n) (succ k) _ H2)
... ≃* ptrunc k (Ω[succ m] (iterate_psusp m (psusp A))) : loop_ptrunc_pequiv
... ≃* ptrunc k (Ω[succ m] (iterate_psusp (succ m) A)) :
ptrunc_pequiv_ptrunc _ (loopn_pequiv_loopn _ !iterate_psusp_succ_in)}
end
end susp

View file

@ -94,7 +94,7 @@ section
definition transport_codes_merid (a a' : A)
: transport codes (merid a) a' = a * a' :> A :=
by krewrite elim_type_merid
ap10 (elim_type_merid _ _ _ a) a'
definition is_trunc_codes [instance] (x : susp A) : is_trunc 1 (codes x) :=
begin
@ -112,19 +112,14 @@ section
definition transport_codes_merid_one_inv (a : A)
: transport codes (merid 1)⁻¹ a = a :=
begin
rewrite tr_inv,
apply @inv_eq_of_eq A A (transport codes (merid 1)) _ a a,
krewrite elim_type_merid,
change a = 1 * a,
rewrite one_mul
end
ap10 (elim_type_merid_inv _ _ _ 1) a ⬝ begin apply to_inv_eq_of_eq, esimp, refine !one_mul⁻¹ end
proposition encode_decode' (a : A) : encode (decode' a) = a :=
begin
unfold decode', unfold encode, unfold encode₀,
rewrite [con_tr,transport_codes_merid,mul_one,tr_inv],
apply transport_codes_merid_one_inv
esimp [encode, decode', encode₀],
refine !con_tr ⬝ _,
refine (ap (transport _ _) !transport_codes_merid ⬝ !transport_codes_merid_one_inv) ⬝ _,
apply mul_one
end
include coh
@ -172,17 +167,23 @@ section
{ exact (λa, tr (merid a)) },
{ apply pi.arrow_pathover_left, esimp, intro a',
apply pathover_of_tr_eq, krewrite susp.elim_type_merid, esimp,
krewrite [trunc_transport,transport_eq_r], apply inverse,
krewrite [trunc_transport,eq_transport_r], apply inverse,
apply homomorphism }
end
proposition decode_encode {x : susp A} : Πt : P x, decode (encode t) = t :=
definition decode_encode {x : susp A} : Πt : P x, decode (encode t) = t :=
begin
apply trunc.rec, intro p, cases p, apply ap tr, apply con.right_inv
end
/-
We define main_lemma by first defining its inverse, because normally equiv.MK changes
the left_inv-component of an equivalence to adjointify it, but in this case we want the
left_inv-component to be encode_decode'. So we adjointify its inverse, so that only the
right_inv-component is changed.
-/
definition main_lemma : trunc 1 (north = north :> susp A) ≃ A :=
equiv.MK encode decode' encode_decode' decode_encode
(equiv.MK decode' encode decode_encode encode_decode')⁻¹ᵉ
definition main_lemma_point
: ptrunc 1 (Ω(psusp A)) ≃* pointed.MK A 1 :=
@ -191,6 +192,33 @@ section
protected definition delooping : Ω (ptrunc 2 (psusp A)) ≃* pointed.MK A 1 :=
loop_ptrunc_pequiv 1 (psusp A) ⬝e* main_lemma_point
/- characterization of the underlying pointed maps -/
definition to_pmap_main_lemma_point_pinv
: main_lemma_point⁻¹ᵉ* ~* !ptr ∘* loop_psusp_unit (pointed.MK A 1) :=
begin
fapply phomotopy.mk,
{ intro a, reflexivity },
{ reflexivity }
end
definition to_pmap_delooping_pinv :
delooping⁻¹ᵉ* ~* Ω→ !ptr ∘* loop_psusp_unit (pointed.MK A 1) :=
begin
refine !trans_pinv ⬝* _,
refine pwhisker_left _ !to_pmap_main_lemma_point_pinv ⬝* _,
refine !passoc⁻¹* ⬝* _,
refine pwhisker_right _ !ap1_ptr⁻¹*,
end
definition hopf_delooping_elim {B : Type*} (f : pointed.MK A 1 →* Ω B) [H2 : is_trunc 2 B] :
Ω→(ptrunc.elim 2 (psusp.elim f)) ∘* (hopf.delooping A coh)⁻¹ᵉ* ~* f :=
begin
refine pwhisker_left _ !to_pmap_delooping_pinv ⬝* _,
refine !passoc⁻¹* ⬝* _,
refine pwhisker_right _ (!ap1_pcompose⁻¹* ⬝* ap1_phomotopy !ptrunc_elim_ptr) ⬝* _,
apply ap1_psusp_elim
end
end
end hopf

View file

@ -60,7 +60,7 @@ namespace interval
by rewrite [tr_eq_cast_ap_fn,↑interval.elim_type,elim_seg];apply cast_ua_fn
definition is_contr_interval [instance] [priority 900] : is_contr interval :=
is_contr.mk zero (λx, interval.rec_on x idp seg !pathover_eq_r_idp)
is_contr.mk zero (λx, interval.rec_on x idp seg !eq_pathover_r_idp)
definition naive_funext_of_interval : naive_funext :=
λA P f g p, ap (λ(i : interval) (x : A), interval.elim_on i (f x) (g x) (p x)) seg

View file

@ -8,7 +8,7 @@ Declaration of a join as a special case of a pushout
import hit.pushout .sphere cubical.cube
open eq function prod equiv is_trunc bool sigma.ops
open eq function prod equiv is_trunc bool sigma.ops pointed
definition join (A B : Type) : Type := @pushout.pushout (A × B) A B pr1 pr2
@ -65,7 +65,9 @@ namespace join
end
end join
end join open join
definition pjoin [constructor] (A B : Type*) : Type* := pointed.MK (join A B) (inl pt)
attribute join.inl join.inr [constructor]
attribute join.rec [recursor]
@ -231,7 +233,7 @@ namespace join
fapply is_contr.mk, exact inl (center A),
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,
apply concat, apply eq_transport_Fr, esimp, rewrite ap_id,
generalize center_eq a, intro p, cases p, apply idp_con,
end
@ -436,7 +438,7 @@ section join_switch
{b : B} {p₀₁ : Π a, b₀ a = b₁ a} {p₀ : Π a, b₀ a = b} {p₁ : Π a, b₁ a = b}
{x y : A} {q : x = y} {sqx : square (p₀₁ x) idp (p₀ x) (p₁ x)}
{sqy : square (p₀₁ y) idp (p₀ y) (p₁ y)}
(c : cube (natural_square_tr _ _) ids (square_Flr_ap_idp p₀ q) (square_Flr_ap_idp p₁ q)
(c : cube (natural_square _ _) ids (square_Flr_ap_idp p₀ q) (square_Flr_ap_idp p₁ q)
sqx sqy) :
sqx =[q] sqy :=
by cases q; apply pathover_of_eq_tr; apply eq_of_deg12_cube; exact c
@ -446,10 +448,10 @@ section join_switch
(sq : square (ap g (ap f p)) p u v) : u =[p] v :=
by cases p; apply eq_pathover; apply transpose; exact sq
private definition natural_square_tr_beta {A B : Type} {f₁ f₂ : A → B}
private definition natural_square_beta {A B : Type} {f₁ f₂ : A → B}
(p : Π a, f₁ a = f₂ a) {x y : A} (q : x = y) {sq : square (p x) (p y) (ap f₁ q) (ap f₂ q)}
(e : apd p q = eq_pathover sq) :
natural_square_tr p q = sq :=
natural_square p q = sq :=
begin
cases q, esimp at *, apply concat, apply inverse, apply vdeg_square_idp,
apply concat, apply ap vdeg_square, apply ap eq_of_pathover_idp e,
@ -464,7 +466,7 @@ section join_switch
refine pathover_of_triangle_cube _,
esimp, apply cube_transport011,
apply inverse, rotate 1, apply switch_inv_cube,
apply natural_square_tr_beta, apply join.rec_glue,
apply natural_square_beta, apply join.rec_glue,
end
protected definition switch_involutive (x : join (join A B) C) :

View file

@ -80,7 +80,7 @@ namespace hopf
induction x,
{ apply inverse, exact circle_mul_base (y*) },
{ apply eq_pathover, induction y,
{ exact natural_square_tr
{ exact natural_square
(λa : S¹, ap (λb : S¹, b*) (circle_mul_base a)) loop },
{ apply is_prop.elimo } }
end

View file

@ -271,12 +271,12 @@ namespace sphere
end
definition equator [constructor] (n : ) : S* n →* Ω (S* (succ n)) :=
loop_susp_unit (S* n)
loop_psusp_unit (S* n)
definition surf {n : } : Ω[n] (S* n) :=
begin
induction n with n s,
{ exact @base 0},
{ exact south },
{ exact (loopn_succ_in (S* (succ n)) n)⁻¹ᵉ* (apn n (equator n) s), }
end
@ -287,13 +287,13 @@ namespace sphere
| ff := proof north qed
| tt := proof south qed
definition sphere_equiv_bool : S 0 ≃ bool :=
definition sphere_equiv_bool [constructor] : S 0 ≃ bool :=
equiv.MK bool_of_sphere
sphere_of_bool
(λb, match b with | tt := idp | ff := idp end)
(λx, proof susp.rec_on x idp idp (empty.rec _) qed)
definition psphere_pequiv_pbool : S* 0 ≃* pbool :=
definition psphere_pequiv_pbool [constructor] : S* 0 ≃* pbool :=
pequiv_of_equiv sphere_equiv_bool idp
definition sphere_eq_bool : S 0 = bool :=
@ -302,36 +302,26 @@ namespace sphere
definition sphere_eq_pbool : S* 0 = pbool :=
pType_eq sphere_equiv_bool idp
-- TODO1: the commented-out part makes the forward function below "apn _ surf" (the next def also)
-- TODO2: we could make this a pointed equivalence
definition psphere_pmap_equiv (A : Type*) (n : ) : (S* n →* A) ≃ Ω[n] A :=
definition psphere_pmap_pequiv' (A : Type*) (n : ) : ppmap (S* n) A ≃* Ω[n] A :=
begin
-- fapply equiv_change_fun,
-- {
revert A, induction n with n IH: intro A,
{ refine _ ⬝e !pmap_bool_equiv, exact pequiv_ppcompose_right psphere_pequiv_pbool⁻¹ᵉ*},
{ refine susp_adjoint_loop (S* n) A ⬝e !IH ⬝e !loopn_succ_in⁻¹ᵉ* }
-- },
-- { intro f, exact apn n f surf},
-- { revert A, induction n with n IH: intro A f,
-- { exact sorry},
-- { exact sorry}}
revert A, induction n with n IH: intro A,
{ refine _ ⬝e* !pmap_pbool_pequiv, exact pequiv_ppcompose_right psphere_pequiv_pbool⁻¹ᵉ* },
{ refine psusp_adjoint_loop (S* n) A ⬝e* IH (Ω A) ⬝e* !loopn_succ_in⁻¹ᵉ* }
end
-- definition psphere_pmap_equiv' (A : Type*) (n : ) : (S* n →* A) ≃ Ω[n] A :=
-- begin
-- fapply equiv.MK,
-- { intro f, exact apn n f surf },
-- { revert A, induction n with n IH: intro A p,
-- { exact !pmap_bool_equiv⁻¹ᵉ p ∘* psphere_pequiv_pbool },
-- { refine (susp_adjoint_loop (S* n) A)⁻¹ᵉ (IH (Ω A) _),
-- exact loopn_succ_in A n p }},
-- { exact sorry},
-- { exact sorry}
-- end
definition psphere_pmap_pequiv (A : Type*) (n : ) : ppmap (S* n) A ≃* Ω[n] A :=
begin
fapply pequiv_change_fun,
{ exact psphere_pmap_pequiv' A n },
{ exact papn_fun A surf },
{ revert A, induction n with n IH: intro A,
{ reflexivity },
{ intro f, refine ap !loopn_succ_in⁻¹ᵉ* (IH (Ω A) _ ⬝ !apn_pcompose _) ⬝ _,
exact !loopn_succ_in_inv_natural⁻¹* _ }}
end
protected definition elim {n : } {P : Type*} (p : Ω[n] P) : S* n →* P :=
to_inv !psphere_pmap_equiv p
!psphere_pmap_pequiv⁻¹ᵉ* p
-- definition elim_surf {n : } {P : Type*} (p : Ω[n] P) : apn n (sphere.elim p) surf = p :=
-- begin
@ -369,7 +359,7 @@ namespace is_trunc
begin
apply iff.elim_right !is_trunc_iff_is_contr_loop,
intro a,
apply is_trunc_equiv_closed, apply psphere_pmap_equiv,
apply is_trunc_equiv_closed, exact !psphere_pmap_pequiv,
fapply is_contr.mk,
{ exact pmap.mk (λx, a) idp},
{ intro f, fapply pmap_eq,
@ -388,8 +378,9 @@ namespace is_trunc
(a : A) (f : S* n →* pointed.Mk a) (x : S n) : f x = f base :=
begin
let H' := iff.elim_left (is_trunc_iff_is_contr_loop n A) H a,
note H'' := @is_trunc_equiv_closed_rev _ _ _ !psphere_pmap_equiv H',
have p : (f = pmap.mk (λx, f base) (respect_pt f)),
note H'' := @is_trunc_equiv_closed_rev _ _ _ !psphere_pmap_pequiv H',
esimp at H'',
have p : f = pmap.mk (λx, f base) (respect_pt f),
by apply is_prop.elim,
exact ap10 (ap pmap.to_fun p) x
end

View file

@ -103,7 +103,7 @@ namespace susp
apply trunc.rec,
intro a',
apply pathover_of_tr_eq,
rewrite [transport_eq_Fr,idp_con],
rewrite [eq_transport_Fr,idp_con],
revert H, induction n with [n, IH],
{ intro H, apply is_prop.elim },
{ intros H,
@ -119,10 +119,7 @@ namespace susp
}
end
end susp
/- Flattening lemma -/
namespace susp
/- Flattening lemma -/
open prod prod.ops
section
@ -206,9 +203,13 @@ definition psusp [constructor] (X : Type) : Type* :=
pointed.mk' (susp X)
namespace susp
open pointed
open pointed is_trunc
variables {X Y Z : Type*}
definition is_conn_psusp [instance] (n : trunc_index) (A : Type*)
[H : is_conn n A] : is_conn (n .+1) (psusp A) :=
is_conn_susp n A
definition psusp_functor [constructor] (f : X →* Y) : psusp X →* psusp Y :=
begin
fconstructor,
@ -238,15 +239,15 @@ namespace susp
-- adjunction from Coq-HoTT
definition loop_susp_unit [constructor] (X : Type*) : X →* Ω(psusp X) :=
definition loop_psusp_unit [constructor] (X : Type*) : X →* Ω(psusp X) :=
begin
fconstructor,
{ intro x, exact merid x ⬝ (merid pt)⁻¹},
{ apply con.right_inv},
{ intro x, exact merid x ⬝ (merid pt)⁻¹ },
{ apply con.right_inv },
end
definition loop_susp_unit_natural (f : X →* Y)
: loop_susp_unit Y ∘* f ~* ap1 (psusp_functor f) ∘* loop_susp_unit X :=
definition loop_psusp_unit_natural (f : X →* Y)
: loop_psusp_unit Y ∘* f ~* ap1 (psusp_functor f) ∘* loop_psusp_unit X :=
begin
induction X with X x, induction Y with Y y, induction f with f pf, esimp at *, induction pf,
fconstructor,
@ -255,105 +256,140 @@ namespace susp
!idp_con ⬝
(!ap_con ⬝
whisker_left _ !ap_inv) ⬝
(!elim_merid ◾ (inverse2 !elim_merid))
},
(!elim_merid ◾ (inverse2 !elim_merid)) },
{ rewrite [▸*,idp_con (con.right_inv _)],
apply inv_con_eq_of_eq_con,
refine _ ⬝ !con.assoc',
rewrite inverse2_right_inv,
refine _ ⬝ !con.assoc',
rewrite [ap_con_right_inv],
xrewrite [idp_con_idp, -ap_compose (concat idp)]},
xrewrite [idp_con_idp, -ap_compose (concat idp)] },
end
definition loop_susp_counit [constructor] (X : Type*) : psusp (Ω X) →* X :=
definition loop_psusp_counit [constructor] (X : Type*) : psusp (Ω X) →* X :=
begin
fconstructor,
{ intro x, induction x, exact pt, exact pt, exact a},
{ reflexivity},
{ intro x, induction x, exact pt, exact pt, exact a },
{ reflexivity },
end
definition loop_susp_counit_natural (f : X →* Y)
: f ∘* loop_susp_counit X ~* loop_susp_counit Y ∘* (psusp_functor (ap1 f)) :=
definition loop_psusp_counit_natural (f : X →* Y)
: f ∘* loop_psusp_counit X ~* loop_psusp_counit Y ∘* (psusp_functor (ap1 f)) :=
begin
induction X with X x, induction Y with Y y, induction f with f pf, esimp at *, induction pf,
fconstructor,
{ intro x', induction x' with p,
{ reflexivity},
{ reflexivity},
{ reflexivity },
{ reflexivity },
{ esimp, apply eq_pathover, apply hdeg_square,
xrewrite [ap_compose' f, ap_compose' (susp.elim (f x) (f x) (λ (a : f x = f x), a)),▸*],
xrewrite [+elim_merid,▸*,idp_con]}},
{ reflexivity}
xrewrite [+elim_merid,▸*,idp_con] }},
{ reflexivity }
end
definition loop_susp_counit_unit (X : Type*)
: ap1 (loop_susp_counit X) ∘* loop_susp_unit (Ω X) ~* pid (Ω X) :=
definition loop_psusp_counit_unit (X : Type*)
: ap1 (loop_psusp_counit X) ∘* loop_psusp_unit (Ω X) ~* pid (Ω X) :=
begin
induction X with X x, fconstructor,
{ intro p, esimp,
refine !idp_con ⬝
(!ap_con ⬝
whisker_left _ !ap_inv) ⬝
(!elim_merid ◾ inverse2 !elim_merid)},
(!elim_merid ◾ inverse2 !elim_merid) },
{ rewrite [▸*,inverse2_right_inv (elim_merid id idp)],
refine !con.assoc ⬝ _,
xrewrite [ap_con_right_inv (susp.elim x x (λa, a)) (merid idp),idp_con_idp,-ap_compose]}
xrewrite [ap_con_right_inv (susp.elim x x (λa, a)) (merid idp),idp_con_idp,-ap_compose] }
end
definition loop_susp_unit_counit (X : Type*)
: loop_susp_counit (psusp X) ∘* psusp_functor (loop_susp_unit X) ~* pid (psusp X) :=
definition loop_psusp_unit_counit (X : Type*)
: loop_psusp_counit (psusp X) ∘* psusp_functor (loop_psusp_unit X) ~* pid (psusp X) :=
begin
induction X with X x, fconstructor,
{ intro x', induction x',
{ reflexivity},
{ exact merid pt},
{ reflexivity },
{ exact merid pt },
{ apply eq_pathover,
xrewrite [▸*, ap_id, ap_compose' (susp.elim north north (λa, a)), +elim_merid,▸*],
apply square_of_eq, exact !idp_con ⬝ !inv_con_cancel_right⁻¹}},
{ reflexivity}
apply square_of_eq, exact !idp_con ⬝ !inv_con_cancel_right⁻¹ }},
{ reflexivity }
end
definition psusp.elim [constructor] {X Y : Type*} (f : X →* Ω Y) : psusp X →* Y :=
loop_psusp_counit Y ∘* psusp_functor f
definition loop_psusp_intro [constructor] {X Y : Type*} (f : psusp X →* Y) : X →* Ω Y :=
ap1 f ∘* loop_psusp_unit X
definition psusp_adjoint_loop_right_inv {X Y : Type*} (g : X →* Ω Y) :
loop_psusp_intro (psusp.elim g) ~* g :=
begin
refine !pwhisker_right !ap1_pcompose ⬝* _,
refine !passoc ⬝* _,
refine !pwhisker_left !loop_psusp_unit_natural⁻¹* ⬝* _,
refine !passoc⁻¹* ⬝* _,
refine !pwhisker_right !loop_psusp_counit_unit ⬝* _,
apply pid_pcompose
end
definition psusp_adjoint_loop_left_inv {X Y : Type*} (f : psusp X →* Y) :
psusp.elim (loop_psusp_intro f) ~* f :=
begin
refine !pwhisker_left !psusp_functor_compose ⬝* _,
refine !passoc⁻¹* ⬝* _,
refine !pwhisker_right !loop_psusp_counit_natural⁻¹* ⬝* _,
refine !passoc ⬝* _,
refine !pwhisker_left !loop_psusp_unit_counit ⬝* _,
apply pcompose_pid
end
-- TODO: rename to psusp_adjoint_loop (also in above lemmas)
definition susp_adjoint_loop [constructor] (X Y : Type*) : psusp X →* Y ≃ X →* Ω Y :=
definition psusp_adjoint_loop_unpointed [constructor] (X Y : Type*) : psusp X →* Y ≃ X →* Ω Y :=
begin
fapply equiv.MK,
{ intro f, exact ap1 f ∘* loop_susp_unit X},
{ intro g, exact loop_susp_counit Y ∘* psusp_functor g},
{ intro g, apply eq_of_phomotopy, esimp,
refine !pwhisker_right !ap1_pcompose ⬝* _,
refine !passoc ⬝* _,
refine !pwhisker_left !loop_susp_unit_natural⁻¹* ⬝* _,
refine !passoc⁻¹* ⬝* _,
refine !pwhisker_right !loop_susp_counit_unit ⬝* _,
apply pid_pcompose},
{ intro f, apply eq_of_phomotopy, esimp,
refine !pwhisker_left !psusp_functor_compose ⬝* _,
refine !passoc⁻¹* ⬝* _,
refine !pwhisker_right !loop_susp_counit_natural⁻¹* ⬝* _,
refine !passoc ⬝* _,
refine !pwhisker_left !loop_susp_unit_counit ⬝* _,
apply pcompose_pid},
{ exact loop_psusp_intro },
{ exact psusp.elim },
{ intro g, apply eq_of_phomotopy, exact psusp_adjoint_loop_right_inv g },
{ intro f, apply eq_of_phomotopy, exact psusp_adjoint_loop_left_inv f }
end
definition susp_adjoint_loop_nat_right (f : psusp X →* Y) (g : Y →* Z)
: susp_adjoint_loop X Z (g ∘* f) ~* ap1 g ∘* susp_adjoint_loop X Y f :=
definition psusp_adjoint_loop_pconst (X Y : Type*) :
psusp_adjoint_loop_unpointed X Y (pconst (psusp X) Y) ~* pconst X (Ω Y) :=
begin
esimp [susp_adjoint_loop],
refine pwhisker_right _ !ap1_pconst ⬝* _,
apply pconst_pcompose
end
definition psusp_adjoint_loop [constructor] (X Y : Type*) : ppmap (psusp X) Y ≃* ppmap X (Ω Y) :=
begin
apply pequiv_of_equiv (psusp_adjoint_loop_unpointed X Y),
apply eq_of_phomotopy,
apply psusp_adjoint_loop_pconst
end
definition ap1_psusp_elim {A : Type*} {X : Type*} (p : A →* Ω X) :
Ω→(psusp.elim p) ∘* loop_psusp_unit A ~* p :=
psusp_adjoint_loop_right_inv p
definition psusp_adjoint_loop_nat_right (f : psusp X →* Y) (g : Y →* Z)
: psusp_adjoint_loop X Z (g ∘* f) ~* ap1 g ∘* psusp_adjoint_loop X Y f :=
begin
esimp [psusp_adjoint_loop],
refine _ ⬝* !passoc,
apply pwhisker_right,
apply ap1_pcompose
end
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 ∘* psusp_functor g :=
definition psusp_adjoint_loop_nat_left (f : Y →* Ω Z) (g : X →* Y)
: (psusp_adjoint_loop X Z)⁻¹ᵉ (f ∘* g) ~* (psusp_adjoint_loop Y Z)⁻¹ᵉ f ∘* psusp_functor g :=
begin
esimp [susp_adjoint_loop],
esimp [psusp_adjoint_loop],
refine _ ⬝* !passoc⁻¹*,
apply pwhisker_left,
apply psusp_functor_compose
end
/- iterated suspension -/
definition iterate_susp (n : ) (A : Type) : Type := iterate susp n A
definition iterate_psusp (n : ) (A : Type*) : Type* := iterate (λX, psusp X) n A
@ -379,4 +415,30 @@ namespace susp
[H : is_conn 0 A] : is_conn m (iterate_psusp m A) :=
begin induction m with m IH, exact H, exact @is_conn_susp _ _ IH end
definition iterate_psusp_functor (n : ) {A B : Type*} (f : A →* B) :
iterate_psusp n A →* iterate_psusp n B :=
begin
induction n with n g,
{ exact f },
{ exact psusp_functor g }
end
definition iterate_psusp_succ_in (n : ) (A : Type*) :
iterate_psusp (succ n) A ≃* iterate_psusp n (psusp A) :=
begin
induction n with n IH,
{ reflexivity},
{ exact psusp_equiv IH}
end
definition iterate_psusp_adjoint_loopn [constructor] (X Y : Type*) (n : ) :
ppmap (iterate_psusp n X) Y ≃* ppmap X (Ω[n] Y) :=
begin
revert X Y, induction n with n IH: intro X Y,
{ reflexivity },
{ refine !psusp_adjoint_loop ⬝e* !IH ⬝e* _, apply pequiv_ppcompose_left,
symmetry, apply loopn_succ_in }
end
end susp

View file

@ -38,7 +38,7 @@ namespace is_equiv
variables {A B C : Type} (g : B → C) (f : A → B) {f' : A → B}
-- The variant of mk' where f is explicit.
protected abbreviation mk [constructor] := @is_equiv.mk' A B f
protected definition mk [constructor] := @is_equiv.mk' A B f
-- The identity function is an equivalence.
definition is_equiv_id [instance] [constructor] (A : Type) : (is_equiv (id : A → A)) :=
@ -127,41 +127,6 @@ namespace is_equiv
variables {A B C : Type} (f : A → B) {f' : A → B} [Hf : is_equiv f] (g : B → C)
include Hf
--The inverse of an equivalence is, again, an equivalence.
definition is_equiv_inv [instance] [constructor] : is_equiv f⁻¹ :=
adjointify f⁻¹ f (left_inv f) (right_inv f)
-- The 2-out-of-3 properties
definition cancel_right (g : B → C) [Hgf : is_equiv (g ∘ f)] : (is_equiv g) :=
have Hfinv : is_equiv f⁻¹, from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose (g ∘ f) f⁻¹) (λb, ap g (@right_inv _ _ f _ b))
definition cancel_left (g : C → A) [Hgf : is_equiv (f ∘ g)] : (is_equiv g) :=
have Hfinv : is_equiv f⁻¹, from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose f⁻¹ (f ∘ g)) (λa, left_inv f (g a))
definition eq_of_fn_eq_fn' {x y : A} (q : f x = f y) : x = y :=
(left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y
definition ap_eq_of_fn_eq_fn' {x y : A} (q : f x = f y) : ap f (eq_of_fn_eq_fn' f q) = q :=
!ap_con ⬝ whisker_right !ap_con _
⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹)
◾ (inverse (ap_compose f f⁻¹ _))
◾ (adj f _)⁻¹)
⬝ con_ap_con_eq_con_con (right_inv f) _ _
⬝ whisker_right !con.left_inv _
⬝ !idp_con
definition eq_of_fn_eq_fn'_ap {x y : A} (q : x = y) : eq_of_fn_eq_fn' f (ap f q) = q :=
by induction q; apply con.left_inv
definition is_equiv_ap [instance] [constructor] (x y : A) : is_equiv (ap f : x = y → f x = f y) :=
adjointify
(ap f)
(eq_of_fn_eq_fn' f)
(ap_eq_of_fn_eq_fn' f)
(eq_of_fn_eq_fn'_ap f)
-- The function equiv_rect says that given an equivalence f : A → B,
-- and a hypothesis from B, one may always assume that the hypothesis
-- is in the image of e.
@ -192,6 +157,41 @@ namespace is_equiv
!ap_compose ⬝ ap02 f⁻¹ (adj f a)⁻¹)
b
--The inverse of an equivalence is, again, an equivalence.
definition is_equiv_inv [instance] [constructor] [priority 500] : is_equiv f⁻¹ :=
is_equiv.mk f⁻¹ f (left_inv f) (right_inv f) (adj_inv f)
-- The 2-out-of-3 properties
definition cancel_right (g : B → C) [Hgf : is_equiv (g ∘ f)] : (is_equiv g) :=
have Hfinv : is_equiv f⁻¹, from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose (g ∘ f) f⁻¹) (λb, ap g (@right_inv _ _ f _ b))
definition cancel_left (g : C → A) [Hgf : is_equiv (f ∘ g)] : (is_equiv g) :=
have Hfinv : is_equiv f⁻¹, from is_equiv_inv f,
@homotopy_closed _ _ _ _ (is_equiv_compose f⁻¹ (f ∘ g)) (λa, left_inv f (g a))
definition eq_of_fn_eq_fn' [unfold 4] {x y : A} (q : f x = f y) : x = y :=
(left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y
definition ap_eq_of_fn_eq_fn' {x y : A} (q : f x = f y) : ap f (eq_of_fn_eq_fn' f q) = q :=
!ap_con ⬝ whisker_right !ap_con _
⬝ ((!ap_inv ⬝ inverse2 (adj f _)⁻¹)
◾ (inverse (ap_compose f f⁻¹ _))
◾ (adj f _)⁻¹)
⬝ con_ap_con_eq_con_con (right_inv f) _ _
⬝ whisker_right !con.left_inv _
⬝ !idp_con
definition eq_of_fn_eq_fn'_ap {x y : A} (q : x = y) : eq_of_fn_eq_fn' f (ap f q) = q :=
by induction q; apply con.left_inv
definition is_equiv_ap [instance] [constructor] (x y : A) : is_equiv (ap f : x = y → f x = f y) :=
adjointify
(ap f)
(eq_of_fn_eq_fn' f)
(ap_eq_of_fn_eq_fn' f)
(eq_of_fn_eq_fn'_ap f)
end
section
@ -364,10 +364,10 @@ namespace equiv
: equiv_of_eq (refl A) = equiv.refl A :=
idp
definition eq_of_fn_eq_fn (f : A ≃ B) {x y : A} (q : f x = f y) : x = y :=
definition eq_of_fn_eq_fn [unfold 3] (f : A ≃ B) {x y : A} (q : f x = f y) : x = y :=
(left_inv f x)⁻¹ ⬝ ap f⁻¹ q ⬝ left_inv f y
definition eq_of_fn_eq_fn_inv (f : A ≃ B) {x y : B} (q : f⁻¹ x = f⁻¹ y) : x = y :=
definition eq_of_fn_eq_fn_inv [unfold 3] (f : A ≃ B) {x y : B} (q : f⁻¹ x = f⁻¹ y) : x = y :=
(right_inv f x)⁻¹ ⬝ ap f q ⬝ right_inv f y
--we need this theorem for the funext_of_ua proof
@ -399,6 +399,23 @@ namespace equiv
section
variables {A B : Type} (f : A ≃ B) {a : A} {b : B}
definition to_eq_of_eq_inv (p : a = f⁻¹ b) : f a = b :=
ap f p ⬝ right_inv f b
definition to_eq_of_inv_eq (p : f⁻¹ b = a) : b = f a :=
(eq_of_eq_inv p⁻¹)⁻¹
definition to_inv_eq_of_eq (p : b = f a) : f⁻¹ b = a :=
ap f⁻¹ p ⬝ left_inv f a
definition to_eq_inv_of_eq (p : f a = b) : a = f⁻¹ b :=
(inv_eq_of_eq p⁻¹)⁻¹
end
section
variables {A : Type} {B C : A → Type} (f : Π{a}, B a ≃ C a)
{g : A → A} {g' : A → A} (h : Π{a}, B (g' a) → B (g a)) (h' : Π{a}, C (g' a) → C (g a))
@ -426,7 +443,7 @@ namespace is_equiv
definition is_equiv_of_equiv_of_homotopy [constructor] {A B : Type} (f : A ≃ B)
{f' : A → B} (Hty : f ~ f') : is_equiv f' :=
homotopy_closed f Hty
@(homotopy_closed f) f' _ Hty
end is_equiv

View file

@ -576,6 +576,10 @@ namespace eq
transport (P ∘ f) p z = transport P (ap f p) z :=
by induction p; reflexivity
definition tr_ap (P : B → Type) (f : A → B) (p : x = y) (z : P (f x)) :
transport P (ap f p) z = transport (P ∘ f) p z :=
(tr_compose P f p z)⁻¹
definition ap_precompose (f : A → B) (g g' : B → C) (p : g = g') :
ap (λh, h ∘ f) p = transport (λh : B → C, g ∘ f = h ∘ f) p idp :=
by induction p; reflexivity
@ -619,7 +623,7 @@ namespace eq
definition inverse2 [unfold 6] {p q : x = y} (h : p = q) : p⁻¹ = q⁻¹ :=
ap inverse h
infixl ` ◾ `:75 := concat2
infixl ` ◾ `:80 := concat2
postfix [parsing_only] `⁻²`:(max+10) := inverse2 --this notation is abusive, should we use it?
/- Whiskering -/
@ -678,7 +682,7 @@ namespace eq
-- The interchange law for concatenation.
definition con2_con_con2 {p p' p'' : x = y} {q q' q'' : y = z}
(a : p = p') (b : p' = p'') (c : q = q') (d : q' = q'') :
(a ◾ c)(b ◾ d) = (a ⬝ b) ◾ (c ⬝ d) :=
a ◾ c ⬝ b ◾ d = (a ⬝ b) ◾ (c ⬝ d) :=
by induction d; induction c; induction b;induction a; reflexivity
definition con2_eq_rl {A : Type} {x y z : A} {p p' : x = y} {q q' : y = z}
@ -742,7 +746,7 @@ namespace eq
definition ap02_con2 (f : A → B) {x y z : A} {p p' : x = y} {q q' :y = z} (r : p = p')
(s : q = q') :
ap02 f (r ◾ s) = ap_con f p q
⬝ (ap02 f r ap02 f s)
⬝ (ap02 f r ap02 f s)
⬝ (ap_con f p' q')⁻¹ :=
by induction r; induction s; induction q; induction p; reflexivity
@ -750,15 +754,6 @@ namespace eq
apdt f p = transport2 P r (f x) ⬝ apdt f q :=
by induction r; exact !idp_con⁻¹
-- And now for a lemma whose statement is much longer than its proof.
definition apdt02_con {P : A → Type} (f : Π x:A, P x) {x y : A}
{p1 p2 p3 : x = y} (r1 : p1 = p2) (r2 : p2 = p3) :
apdt02 f (r1 ⬝ r2) = apdt02 f r1
⬝ whisker_left (transport2 P r1 (f x)) (apdt02 f r2)
⬝ con.assoc' _ _ _
⬝ (whisker_right (tr2_con r1 r2 (f x))⁻¹ (apdt f p3)) :=
by induction r2; induction r1; induction p1; reflexivity
end eq
/-

View file

@ -26,7 +26,7 @@ namespace pointed
definition pt [reducible] [unfold 2] [H : pointed A] := point A
definition Point [reducible] [unfold 1] (A : Type*) := pType.Point A
abbreviation carrier [unfold 1] (A : Type*) := pType.carrier A
definition carrier [reducible] [unfold 1] (A : Type*) := pType.carrier A
protected definition Mk [constructor] {A : Type} (a : A) := pType.mk A a
protected definition MK [constructor] (A : Type) (a : A) := pType.mk A a
protected definition mk' [constructor] (A : Type) [H : pointed A] : Type* :=
@ -39,7 +39,7 @@ open pointed
section
universe variable u
structure ptrunctype (n : trunc_index) extends trunctype.{u} n, pType.{u}
structure ptrunctype (n : ℕ₋₂) extends trunctype.{u} n, pType.{u}
definition is_trunc_ptrunctype [instance] {n : ℕ₋₂} (X : ptrunctype n)
: is_trunc n (ptrunctype.to_pType X) :=
@ -53,17 +53,18 @@ notation `Set*` := pSet
namespace pointed
protected definition ptrunctype.mk' [constructor] (n : trunc_index)
protected definition ptrunctype.mk' [constructor] (n : ℕ₋₂)
(A : Type) [pointed A] [is_trunc n A] : n-Type* :=
ptrunctype.mk A _ pt
protected definition pSet.mk [constructor] := @ptrunctype.mk (-1.+1)
protected definition pSet.mk' [constructor] := ptrunctype.mk' (-1.+1)
definition ptrunctype_of_trunctype [constructor] {n : trunc_index} (A : n-Type) (a : A) : n-Type* :=
definition ptrunctype_of_trunctype [constructor] {n : ℕ₋₂} (A : n-Type) (a : A)
: n-Type* :=
ptrunctype.mk A _ a
definition ptrunctype_of_pType [constructor] {n : trunc_index} (A : Type*) (H : is_trunc n A)
definition ptrunctype_of_pType [constructor] {n : ℕ₋₂} (A : Type*) (H : is_trunc n A)
: n-Type* :=
ptrunctype.mk A _ pt
@ -83,6 +84,11 @@ namespace pointed
end pointed
/- pointed maps -/
structure ppi (A : Type*) (P : A → Type*) :=
(to_fun : Π a : A, P a)
(resp_pt : to_fun (Point A) = Point (P (Point A)))
-- definition pmap (A B : Type*) := @ppi A (λa, B)
structure pmap (A B : Type*) :=
(to_fun : A → B)
(resp_pt : to_fun (Point A) = Point B)
@ -91,7 +97,12 @@ namespace pointed
abbreviation respect_pt [unfold 3] := @pmap.resp_pt
notation `map₊` := pmap
infix ` →* `:30 := pmap
attribute pmap.to_fun [coercion]
attribute pmap.to_fun ppi.to_fun [coercion]
notation `Π*` binders `, ` r:(scoped P, ppi _ P) := r
-- definition pmap.mk [constructor] {A B : Type*} (f : A → B) (p : f pt = pt) : A →* B :=
-- ppi.mk f p
-- definition pmap.to_fun [coercion] [unfold 3] {A B : Type*} (f : A →* B) : A → B := f
end pointed open pointed
/- pointed homotopies -/
@ -112,8 +123,8 @@ namespace pointed
attribute pequiv._trans_of_to_pmap pequiv._trans_of_to_equiv pequiv.to_pmap pequiv.to_equiv
[unfold 3]
infix ` ≃* `:25 := pequiv
attribute pequiv.to_pmap [coercion]
attribute pequiv.to_is_equiv [instance]
attribute pequiv.to_pmap [coercion]
infix ` ≃* `:25 := pequiv
end pointed

View file

@ -137,13 +137,13 @@ namespace arrow
unfold inv_homotopy_of_homotopy_post,
unfold inv_homotopy_of_homotopy_pre,
rewrite [adj f a,-(ap_compose β f)],
rewrite [eq_of_square (natural_square_tr p (left_inv f a))],
rewrite [eq_of_square (natural_square 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))
apply natural_square_tr (left_inv f') (ap α (left_inv f a))
end
definition ap_bot_inv_commute_of_commute (p : f' ∘ α ~ β ∘ f) (b : B)
@ -163,7 +163,7 @@ namespace arrow
(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')
apply natural_square_tr (right_inv f')
(p (f⁻¹ b) ⬝ ap β (right_inv f b))⁻¹
end

View file

@ -10,8 +10,6 @@ Theorems about path types (identity types)
import types.sigma
open eq sigma sigma.ops equiv is_equiv is_trunc
-- TODO: Rename transport_eq_... and pathover_eq_... to eq_transport_... and eq_pathover_...
namespace eq
/- Path spaces -/
section
@ -126,40 +124,40 @@ namespace eq
- `F` means application of a function to that (varying) endpoint.
-/
definition transport_eq_l (p : a₁ = a₂) (q : a₁ = a₃)
definition eq_transport_l (p : a₁ = a₂) (q : a₁ = a₃)
: transport (λx, x = a₃) p q = p⁻¹ ⬝ q :=
by induction p; induction q; reflexivity
definition transport_eq_r (p : a₂ = a₃) (q : a₁ = a₂)
definition eq_transport_r (p : a₂ = a₃) (q : a₁ = a₂)
: transport (λx, a₁ = x) p q = q ⬝ p :=
by induction p; induction q; reflexivity
definition transport_eq_lr (p : a₁ = a₂) (q : a₁ = a₁)
definition eq_transport_lr (p : a₁ = a₂) (q : a₁ = a₁)
: transport (λx, x = x) p q = p⁻¹ ⬝ q ⬝ p :=
by induction p; rewrite [▸*,idp_con]
definition transport_eq_Fl (p : a₁ = a₂) (q : f a₁ = b)
definition eq_transport_Fl (p : a₁ = a₂) (q : f a₁ = b)
: transport (λx, f x = b) p q = (ap f p)⁻¹ ⬝ q :=
by induction p; induction q; reflexivity
definition transport_eq_Fr (p : a₁ = a₂) (q : b = f a₁)
definition eq_transport_Fr (p : a₁ = a₂) (q : b = f a₁)
: transport (λx, b = f x) p q = q ⬝ (ap f p) :=
by induction p; reflexivity
definition transport_eq_FlFr (p : a₁ = a₂) (q : f a₁ = g a₁)
definition eq_transport_FlFr (p : a₁ = a₂) (q : f a₁ = g a₁)
: transport (λx, f x = g x) p q = (ap f p)⁻¹ ⬝ q ⬝ (ap g p) :=
by induction p; rewrite [▸*,idp_con]
definition transport_eq_FlFr_D {B : A → Type} {f g : Πa, B a}
definition eq_transport_FlFr_D {B : A → Type} {f g : Πa, B a}
(p : a₁ = a₂) (q : f a₁ = g a₁)
: transport (λx, f x = g x) p q = (apdt f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apdt g p) :=
by induction p; rewrite [▸*,idp_con,ap_id]
definition transport_eq_FFlr (p : a₁ = a₂) (q : h (f a₁) = a₁)
definition eq_transport_FFlr (p : a₁ = a₂) (q : h (f a₁) = a₁)
: transport (λx, h (f x) = x) p q = (ap h (ap f p))⁻¹ ⬝ q ⬝ p :=
by induction p; rewrite [▸*,idp_con]
definition transport_eq_lFFr (p : a₁ = a₂) (q : a₁ = h (f a₁))
definition eq_transport_lFFr (p : a₁ = a₂) (q : a₁ = h (f a₁))
: transport (λx, x = h (f x)) p q = p⁻¹ ⬝ q ⬝ (ap h (ap f p)) :=
by induction p; rewrite [▸*,idp_con]
@ -170,44 +168,48 @@ namespace eq
-- we should probably try to do everything just with pathover_eq (defined in cubical.square),
-- the following definitions may be removed in future.
definition pathover_eq_l (p : a₁ = a₂) (q : a₁ = a₃) : q =[p] p⁻¹ ⬝ q := /-(λx, x = a₃)-/
definition eq_pathover_l (p : a₁ = a₂) (q : a₁ = a₃) : q =[p] p⁻¹ ⬝ q := /-(λx, x = a₃)-/
by induction p; induction q; exact idpo
definition pathover_eq_r (p : a₂ = a₃) (q : a₁ = a₂) : q =[p] q ⬝ p := /-(λx, a₁ = x)-/
definition eq_pathover_r (p : a₂ = a₃) (q : a₁ = a₂) : q =[p] q ⬝ p := /-(λx, a₁ = x)-/
by induction p; induction q; exact idpo
definition pathover_eq_lr (p : a₁ = a₂) (q : a₁ = a₁) : q =[p] p⁻¹ ⬝ q ⬝ p := /-(λx, x = x)-/
definition eq_pathover_lr (p : a₁ = a₂) (q : a₁ = a₁) : q =[p] p⁻¹ ⬝ q ⬝ p := /-(λx, x = x)-/
by induction p; rewrite [▸*,idp_con]; exact idpo
definition pathover_eq_Fl (p : a₁ = a₂) (q : f a₁ = b) : q =[p] (ap f p)⁻¹ ⬝ q := /-(λx, f x = b)-/
definition eq_pathover_Fl (p : a₁ = a₂) (q : f a₁ = b) : q =[p] (ap f p)⁻¹ ⬝ q := /-(λx, f x = b)-/
by induction p; induction q; exact idpo
definition pathover_eq_Fr (p : a₁ = a₂) (q : b = f a₁) : q =[p] q ⬝ (ap f p) := /-(λx, b = f x)-/
definition eq_pathover_Fl' (p : a₁ = a₂) (q : f a₂ = b) : (ap f p) ⬝ q =[p] q := /-(λx, f x = b)-/
by induction p; induction q; exact idpo
definition eq_pathover_Fr (p : a₁ = a₂) (q : b = f a₁) : q =[p] q ⬝ (ap f p) := /-(λx, b = f x)-/
by induction p; exact idpo
definition pathover_eq_FlFr (p : a₁ = a₂) (q : f a₁ = g a₁) : q =[p] (ap f p)⁻¹ ⬝ q ⬝ (ap g p) :=
definition eq_pathover_FlFr (p : a₁ = a₂) (q : f a₁ = g a₁) : q =[p] (ap f p)⁻¹ ⬝ q ⬝ (ap g p) :=
/-(λx, f x = g x)-/
by induction p; rewrite [▸*,idp_con]; exact idpo
definition pathover_eq_FlFr_D {B : A → Type} {f g : Πa, B a} (p : a₁ = a₂) (q : f a₁ = g a₁)
definition eq_pathover_FlFr_D {B : A → Type} {f g : Πa, B a} (p : a₁ = a₂) (q : f a₁ = g a₁)
: q =[p] (apdt f p)⁻¹ ⬝ ap (transport B p) q ⬝ (apdt g p) := /-(λx, f x = g x)-/
by induction p; rewrite [▸*,idp_con,ap_id];exact idpo
definition pathover_eq_FFlr (p : a₁ = a₂) (q : h (f a₁) = a₁) : q =[p] (ap h (ap f p))⁻¹ ⬝ q ⬝ p :=
definition eq_pathover_FFlr (p : a₁ = a₂) (q : h (f a₁) = a₁) : q =[p] (ap h (ap f p))⁻¹ ⬝ q ⬝ p :=
/-(λx, h (f x) = x)-/
by induction p; rewrite [▸*,idp_con];exact idpo
definition pathover_eq_lFFr (p : a₁ = a₂) (q : a₁ = h (f a₁)) : q =[p] p⁻¹ ⬝ q ⬝ (ap h (ap f p)) :=
definition eq_pathover_lFFr (p : a₁ = a₂) (q : a₁ = h (f a₁)) : q =[p] p⁻¹ ⬝ q ⬝ (ap h (ap f p)) :=
/-(λx, x = h (f x))-/
by induction p; rewrite [▸*,idp_con];exact idpo
definition pathover_eq_r_idp (p : a₁ = a₂) : idp =[p] p := /-(λx, a₁ = x)-/
definition eq_pathover_r_idp (p : a₁ = a₂) : idp =[p] p := /-(λx, a₁ = x)-/
by induction p; exact idpo
definition pathover_eq_l_idp (p : a₁ = a₂) : idp =[p] p⁻¹ := /-(λx, x = a₁)-/
definition eq_pathover_l_idp (p : a₁ = a₂) : idp =[p] p⁻¹ := /-(λx, x = a₁)-/
by induction p; exact idpo
definition pathover_eq_l_idp' (p : a₁ = a₂) : idp =[p⁻¹] p := /-(λx, x = a₂)-/
definition eq_pathover_l_idp' (p : a₁ = a₂) : idp =[p⁻¹] p := /-(λx, x = a₂)-/
by induction p; exact idpo
-- The Functorial action of paths is [ap].
@ -254,6 +256,10 @@ namespace eq
definition eq_equiv_eq_closed [constructor] (p : a₁ = a₂) (q : a₃ = a₄) : (a₁ = a₃) ≃ (a₂ = a₄) :=
equiv.trans (equiv_eq_closed_left a₃ p) (equiv_eq_closed_right a₂ q)
definition loop_equiv_eq_closed [constructor] {A : Type} {a a' : A} (p : a = a')
: (a = a) ≃ (a' = a') :=
eq_equiv_eq_closed p p
definition is_equiv_whisker_left [constructor] (p : a₁ = a₂) (q r : a₂ = a₃)
: is_equiv (whisker_left p : q = r → p ⬝ q = p ⬝ r) :=
begin
@ -398,36 +404,36 @@ namespace eq
/- Pathover Equivalences -/
definition pathover_eq_equiv_l (p : a₁ = a₂) (q : a₁ = a₃) (r : a₂ = a₃) : q =[p] r ≃ q = p ⬝ r :=
definition eq_pathover_equiv_l (p : a₁ = a₂) (q : a₁ = a₃) (r : a₂ = a₃) : q =[p] r ≃ q = p ⬝ r :=
/-(λx, x = a₃)-/
by induction p; exact !pathover_idp ⬝e !equiv_eq_closed_right !idp_con⁻¹
definition pathover_eq_equiv_r (p : a₂ = a₃) (q : a₁ = a₂) (r : a₁ = a₃) : q =[p] r ≃ q ⬝ p = r :=
definition eq_pathover_equiv_r (p : a₂ = a₃) (q : a₁ = a₂) (r : a₁ = a₃) : q =[p] r ≃ q ⬝ p = r :=
/-(λx, a₁ = x)-/
by induction p; apply pathover_idp
definition pathover_eq_equiv_lr (p : a₁ = a₂) (q : a₁ = a₁) (r : a₂ = a₂)
definition eq_pathover_equiv_lr (p : a₁ = a₂) (q : a₁ = a₁) (r : a₂ = a₂)
: q =[p] r ≃ q ⬝ p = p ⬝ r := /-(λx, x = x)-/
by induction p; exact !pathover_idp ⬝e !equiv_eq_closed_right !idp_con⁻¹
definition pathover_eq_equiv_Fl (p : a₁ = a₂) (q : f a₁ = b) (r : f a₂ = b)
definition eq_pathover_equiv_Fl (p : a₁ = a₂) (q : f a₁ = b) (r : f a₂ = b)
: q =[p] r ≃ q = ap f p ⬝ r := /-(λx, f x = b)-/
by induction p; exact !pathover_idp ⬝e !equiv_eq_closed_right !idp_con⁻¹
definition pathover_eq_equiv_Fr (p : a₁ = a₂) (q : b = f a₁) (r : b = f a₂)
definition eq_pathover_equiv_Fr (p : a₁ = a₂) (q : b = f a₁) (r : b = f a₂)
: q =[p] r ≃ q ⬝ ap f p = r := /-(λx, b = f x)-/
by induction p; apply pathover_idp
definition pathover_eq_equiv_FlFr (p : a₁ = a₂) (q : f a₁ = g a₁) (r : f a₂ = g a₂)
definition eq_pathover_equiv_FlFr (p : a₁ = a₂) (q : f a₁ = g a₁) (r : f a₂ = g a₂)
: q =[p] r ≃ q ⬝ ap g p = ap f p ⬝ r := /-(λx, f x = g x)-/
by induction p; exact !pathover_idp ⬝e !equiv_eq_closed_right !idp_con⁻¹
definition pathover_eq_equiv_FFlr (p : a₁ = a₂) (q : h (f a₁) = a₁) (r : h (f a₂) = a₂)
definition eq_pathover_equiv_FFlr (p : a₁ = a₂) (q : h (f a₁) = a₁) (r : h (f a₂) = a₂)
: q =[p] r ≃ q ⬝ p = ap h (ap f p) ⬝ r :=
/-(λx, h (f x) = x)-/
by induction p; exact !pathover_idp ⬝e !equiv_eq_closed_right !idp_con⁻¹
definition pathover_eq_equiv_lFFr (p : a₁ = a₂) (q : a₁ = h (f a₁)) (r : a₂ = h (f a₂))
definition eq_pathover_equiv_lFFr (p : a₁ = a₂) (q : a₁ = h (f a₁)) (r : a₂ = h (f a₂))
: q =[p] r ≃ q ⬝ ap h (ap f p) = p ⬝ r :=
/-(λx, x = h (f x))-/
by induction p; exact !pathover_idp ⬝e !equiv_eq_closed_right !idp_con⁻¹

View file

@ -7,7 +7,7 @@ Ported from Coq HoTT
Theorems about the types equiv and is_equiv
-/
import .fiber .arrow arity ..prop_trunc cubical.square
import .fiber .arrow arity ..prop_trunc cubical.square .pointed
open eq is_trunc sigma sigma.ops pi fiber function equiv
@ -108,7 +108,7 @@ namespace is_equiv
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))⁻¹ʰ,
eq_bot_of_square (natural_square_tr (λ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]
@ -136,7 +136,7 @@ namespace is_equiv
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)
apply natural_square q (right_inv f b)
end end
abstract begin
intro p, apply eq_of_homotopy, intro a,
@ -149,7 +149,7 @@ namespace is_equiv
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)
apply natural_square p (left_inv f a)
end end
end pre_compose
@ -172,7 +172,7 @@ namespace is_equiv
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)
apply natural_square_tr (left_inv f) (q c)
end end
abstract begin
intro p, apply eq_of_homotopy, intro c,
@ -186,7 +186,7 @@ namespace is_equiv
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)
apply natural_square_tr (right_inv f) (p c)
end end
end post_compose
@ -306,3 +306,13 @@ namespace equiv
end
end equiv
namespace pointed
open equiv is_equiv
definition pequiv_eq {A B : Type*} {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.elimo
end
end pointed

View file

@ -1,14 +1,14 @@
/-
Copyright (c) 2015 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Floris van Doorn
Authors: Floris van Doorn, Mike Shulman
Ported from Coq HoTT
Theorems about fibers
-/
import .sigma .eq .pi cubical.squareover
open equiv sigma sigma.ops eq pi
import .sigma .eq .pi cubical.squareover .pointed .eq
open equiv sigma sigma.ops eq pi pointed
structure fiber {A B : Type} (f : A → B) (b : B) :=
(point : A)
@ -27,7 +27,7 @@ namespace fiber
{intro x, exact abstract begin cases x, apply idp end end},
end
definition fiber_eq_equiv (x y : fiber f b)
definition fiber_eq_equiv [constructor] (x y : fiber f b)
: (x = y) ≃ (Σ(p : point x = point y), point_eq x = ap f p ⬝ point_eq y) :=
begin
apply equiv.trans,
@ -36,7 +36,7 @@ namespace fiber
apply sigma_eq_equiv,
apply sigma_equiv_sigma_right,
intro p,
apply pathover_eq_equiv_Fl,
apply eq_pathover_equiv_Fl,
end
definition fiber_eq {x y : fiber f b} (p : point x = point y)
@ -131,12 +131,83 @@ namespace fiber
... ≃ a₀ = a : sigma_unit_left
-- the pointed fiber of a pointed map, which is the fiber over the basepoint
open pointed
definition pfiber [constructor] {X Y : Type*} (f : X →* Y) : Type* :=
pointed.MK (fiber f pt) (fiber.mk pt !respect_pt)
definition ppoint [constructor] {X Y : Type*} (f : X →* Y) : pfiber f →* X :=
pmap.mk point idp
definition pfiber.sigma_char [constructor] {A B : Type*} (f : A →* B)
: pfiber f ≃* pointed.MK (Σa, f a = pt) ⟨pt, respect_pt f⟩ :=
pequiv_of_equiv (fiber.sigma_char f pt) idp
definition ppoint_sigma_char [constructor] {A B : Type*} (f : A →* B)
: ppoint f ~* pmap.mk pr1 idp ∘* pfiber.sigma_char f :=
!phomotopy.refl
definition pfiber_loop_space {A B : Type*} (f : A →* B) : pfiber (Ω→ f) ≃* Ω (pfiber f) :=
pequiv_of_equiv
(calc pfiber (Ω→ f) ≃ Σ(p : Point A = Point A), ap1 f p = rfl
: (fiber.sigma_char (ap1 f) (Point (Ω B)))
... ≃ Σ(p : Point A = Point A), (respect_pt f) = ap f p ⬝ (respect_pt f)
: (sigma_equiv_sigma_right (λp,
calc (ap1 f p = rfl) ≃ !respect_pt⁻¹ ⬝ (ap f p ⬝ !respect_pt) = rfl
: equiv_eq_closed_left _ (con.assoc _ _ _)
... ≃ ap f p ⬝ (respect_pt f) = (respect_pt f)
: eq_equiv_inv_con_eq_idp
... ≃ (respect_pt f) = ap f p ⬝ (respect_pt f)
: eq_equiv_eq_symm))
... ≃ fiber.mk (Point A) (respect_pt f) = fiber.mk pt (respect_pt f)
: fiber_eq_equiv
... ≃ Ω (pfiber f)
: erfl)
(begin cases f with f p, cases A with A a, cases B with B b, esimp at p, esimp at f,
induction p, reflexivity end)
definition pfiber_equiv_of_phomotopy {A B : Type*} {f g : A →* B} (h : f ~* g)
: pfiber f ≃* pfiber g :=
begin
fapply pequiv_of_equiv,
{ refine (fiber.sigma_char f pt ⬝e _ ⬝e (fiber.sigma_char g pt)⁻¹ᵉ),
apply sigma_equiv_sigma_right, intros a,
apply equiv_eq_closed_left, apply (to_homotopy h) },
{ refine (fiber_eq rfl _),
change (h pt)⁻¹ ⬝ respect_pt f = idp ⬝ respect_pt g,
rewrite idp_con, apply inv_con_eq_of_eq_con, symmetry, exact (to_homotopy_pt h) }
end
definition transport_fiber_equiv [constructor] {A B : Type} (f : A → B) {b1 b2 : B} (p : b1 = b2)
: fiber f b1 ≃ fiber f b2 :=
calc fiber f b1 ≃ Σa, f a = b1 : fiber.sigma_char
... ≃ Σa, f a = b2 : sigma_equiv_sigma_right (λa, equiv_eq_closed_right (f a) p)
... ≃ fiber f b2 : fiber.sigma_char
definition pequiv_postcompose {A B B' : Type*} (f : A →* B) (g : B ≃* B')
: pfiber (g ∘* f) ≃* pfiber f :=
begin
fapply pequiv_of_equiv, esimp,
refine transport_fiber_equiv (g ∘* f) (respect_pt g)⁻¹ ⬝e fiber.equiv_postcompose f g (Point B),
esimp, apply (ap (fiber.mk (Point A))), refine !con.assoc ⬝ _, apply inv_con_eq_of_eq_con,
rewrite [con.assoc, con.right_inv, con_idp, -ap_compose'], apply ap_con_eq_con
end
definition pequiv_precompose {A A' B : Type*} (f : A →* B) (g : A' ≃* A)
: pfiber (f ∘* g) ≃* pfiber f :=
begin
fapply pequiv_of_equiv, esimp,
refine fiber.equiv_precompose f g (Point B),
esimp, apply (eq_of_fn_eq_fn (fiber.sigma_char _ _)), fapply sigma_eq: esimp,
{ apply respect_pt g },
{ apply eq_pathover_Fl' }
end
definition pfiber_equiv_of_square {A B C D : Type*} {f : A →* B} {g : C →* D} (h : A ≃* C)
(k : B ≃* D) (s : k ∘* f ~* g ∘* h) : pfiber f ≃* pfiber g :=
calc pfiber f ≃* pfiber (k ∘* f) : pequiv_postcompose
... ≃* pfiber (g ∘* h) : pfiber_equiv_of_phomotopy s
... ≃* pfiber g : pequiv_precompose
end fiber
open function is_equiv

View file

@ -30,13 +30,13 @@ namespace int
notation `ag` := CommGroup_int
end
definition is_equiv_succ [instance] : is_equiv succ :=
definition is_equiv_succ [constructor] [instance] : is_equiv succ :=
adjointify succ pred (λa, !add_sub_cancel) (λa, !sub_add_cancel)
definition equiv_succ : := equiv.mk succ _
definition equiv_succ [constructor] : := equiv.mk succ _
definition is_equiv_neg [instance] : is_equiv (neg : ) :=
definition is_equiv_neg [constructor] [instance] : is_equiv (neg : ) :=
adjointify neg neg (λx, !neg_neg) (λa, !neg_neg)
definition equiv_neg : := equiv.mk neg _
definition equiv_neg [constructor] : := equiv.mk neg _
definition iterate {A : Type} (f : A ≃ A) (a : ) : A ≃ A :=
rec_nat_on a erfl

View file

@ -168,10 +168,10 @@ namespace lift
fiber (lift_functor f) (up b) ≃ fiber f b :=
begin
fapply equiv.MK: intro v; cases v with a p,
{ cases a with a, exact fiber.mk a (eq_of_fn_eq_fn' up p)},
{ exact fiber.mk (up a) (ap up p)},
{ esimp, apply ap (fiber.mk a), apply eq_of_fn_eq_fn'_ap},
{ cases a with a, esimp, apply ap (fiber.mk (up a)), apply ap_eq_of_fn_eq_fn'}
{ cases a with a, exact fiber.mk a (eq_of_fn_eq_fn' up p) },
{ exact fiber.mk (up a) (ap up p) },
{ apply ap (fiber.mk a), apply eq_of_fn_eq_fn'_ap },
{ cases a with a, esimp, apply ap (fiber.mk (up a)), apply ap_eq_of_fn_eq_fn' }
end

View file

@ -48,7 +48,7 @@ calc
... = succ (x / z) : {!nat.add_sub_cancel}
theorem add_div_self_left {x : } (z : ) (H : x > 0) : (x + z) / x = succ (z / x) :=
!add.comm ▸ !add_div_self H
by rewrite add.comm; exact !add_div_self H
local attribute succ_mul [simp]
@ -110,7 +110,7 @@ by_cases_zero_pos z
... = x % z : nat.add_sub_cancel)
theorem add_mod_self_left [simp] (x z : ) : (x + z) % x = z % x :=
!add.comm ▸ !add_mod_self
by rewrite add.comm; apply add_mod_self
local attribute succ_mul [simp]

View file

@ -90,7 +90,7 @@ calc
protected theorem sub.right_comm (m n k : ) : m - n - k = m - k - n :=
calc
m - n - k = m - (n + k) : !nat.sub_sub
... = m - (k + n) : {!add.comm}
... = m - (k + n) : by rewrite add.comm
... = m - k - n : !nat.sub_sub⁻¹
theorem sub_one (n : ) : n - 1 = pred n :=

View file

@ -384,12 +384,4 @@ namespace pi
: pointed (Πx, P x) :=
pointed.mk (λx, pt)
definition ppi [constructor] {A : Type} (P : A → Type*) : Type* :=
pointed.mk' (Πa, P a)
notation `Π*` binders `, ` r:(scoped P, ppi P) := r
definition ptpi [constructor] {n : ℕ₋₂} {A : Type} (P : A → n-Type*) : n-Type* :=
ptrunctype.mk' n (Πa, P a)
end pi

View file

@ -7,7 +7,7 @@ Ported from Coq HoTT
The basic definitions are in init.pointed
-/
import .equiv .nat.basic
import .nat.basic ..arity ..prop_trunc
open is_trunc eq prod sigma nat equiv option is_equiv bool unit algebra sigma.ops sum
namespace pointed
@ -147,16 +147,28 @@ namespace pointed
/- equivalences and equalities -/
definition pmap.sigma_char [constructor] {A B : Type*} : (A →* B) ≃ Σ(f : A → B), f pt = pt :=
begin
fapply equiv.MK : intros f,
{ exact ⟨f , respect_pt f⟩ },
all_goals cases f with f p,
{ exact pmap.mk f p },
all_goals reflexivity
end
definition pmap_eq (r : Πa, f a = g a) (s : respect_pt f = (r pt) ⬝ respect_pt g) : f = g :=
begin
cases f with f p, cases g with g q,
esimp at *,
fapply apd011 pmap.mk,
{ exact eq_of_homotopy r},
{ apply concato_eq, apply pathover_eq_Fl, apply inv_con_eq_of_eq_con,
{ apply concato_eq, apply eq_pathover_Fl, apply inv_con_eq_of_eq_con,
rewrite [ap_eq_apd10, apd10_eq_of_homotopy, s]}
end
definition pmap_eq_of_homotopy {A B : Type*} {f g : A →* B} [is_set B] (p : f ~ g) : f = g :=
pmap_eq p !is_set.elim
definition pmap_equiv_left (A : Type) (B : Type*) : A₊ →* B ≃ (A → B) :=
begin
fapply equiv.MK,
@ -183,7 +195,8 @@ namespace pointed
reflexivity}
end
definition pmap_bool_equiv (B : Type*) : (pbool →* B) ≃ B :=
-- pmap_pbool_pequiv is the pointed equivalence
definition pmap_pbool_equiv [constructor] (B : Type*) : (pbool →* B) ≃ B :=
begin
fapply equiv.MK,
{ intro f, cases f with f p, exact f tt},
@ -196,6 +209,8 @@ namespace pointed
{ esimp, exact !con.left_inv⁻¹}},
end
/- some specific pointed maps -/
-- The constant pointed map between any two types
definition pconst [constructor] (A B : Type*) : A →* B :=
pmap.mk (λ a, Point B) idp
@ -204,8 +219,6 @@ namespace pointed
definition ppmap [constructor] (A B : Type*) : Type* :=
pType.mk (A →* B) (pconst A B)
/- instances of pointed maps -/
definition pcast [constructor] {A B : Type*} (p : A = B) : A →* B :=
pmap.mk (cast (ap pType.carrier p)) (by induction p; reflexivity)
@ -229,6 +242,65 @@ namespace pointed
prefix `Ω→`:(max+5) := ap1
notation `Ω→[`:95 n:0 `]`:0 := apn n
definition ptransport [constructor] {A : Type} (B : A → Type*) {a a' : A} (p : a = a')
: B a →* B a' :=
pmap.mk (transport B p) (apdt (λa, Point (B a)) p)
definition pmap_of_eq_pt [constructor] {A : Type} {a a' : A} (p : a = a') :
pointed.MK A a →* pointed.MK A a' :=
pmap.mk id p
definition pbool_pmap [constructor] {A : Type*} (a : A) : pbool →* A :=
pmap.mk (bool.rec pt a) idp
-- properties of pointed maps
definition apn_zero [unfold_full] (f : A →* B) : Ω→[0] f = f := idp
definition apn_succ [unfold_full] (n : ) (f : A →* B) : Ω→[n + 1] f = Ω→ (Ω→[n] f) := idp
theorem ap1_con (f : A →* B) (p q : Ω A) : ap1 f (p ⬝ q) = ap1 f p ⬝ ap1 f q :=
begin
rewrite [▸*,ap_con, +con.assoc, con_inv_cancel_left], repeat apply whisker_left
end
theorem ap1_inv (f : A →* B) (p : Ω A) : ap1 f p⁻¹ = (ap1 f p)⁻¹ :=
begin
rewrite [▸*,ap_inv, +con_inv, inv_inv, +con.assoc], repeat apply whisker_left
end
definition pinverse_con [constructor] {X : Type*} (p q : Ω X)
: pinverse (p ⬝ q) = pinverse q ⬝ pinverse p :=
!con_inv
definition pinverse_inv [constructor] {X : Type*} (p : Ω X)
: pinverse p⁻¹ = (pinverse p)⁻¹ :=
idp
theorem apn_con (n : ) (f : A →* B) (p q : Ω[n+1] A)
: apn (n+1) f (p ⬝ q) = apn (n+1) f p ⬝ apn (n+1) f q :=
by rewrite [+apn_succ, ap1_con]
theorem apn_inv (n : ) (f : A →* B) (p : Ω[n+1] A) : apn (n+1) f p⁻¹ = (apn (n+1) f p)⁻¹ :=
by rewrite [+apn_succ, ap1_inv]
definition is_equiv_ap1 (f : A →* B) [is_equiv f] : is_equiv (ap1 f) :=
begin
induction B with B b, induction f with f pf, esimp at *, cases pf, esimp,
apply is_equiv.homotopy_closed (ap f),
intro p, exact !idp_con⁻¹
end
definition is_equiv_apn (n : ) (f : A →* B) [H : is_equiv f]
: is_equiv (apn n f) :=
begin
induction n with n IH,
{ exact H},
{ exact is_equiv_ap1 (apn n f)}
end
definition is_equiv_pcast [instance] {A B : Type*} (p : A = B) : is_equiv (pcast p) :=
!is_equiv_cast
/- categorical properties of pointed homotopies -/
protected definition phomotopy.refl [constructor] [refl] (f : A →* B) : f ~* f :=
@ -259,25 +331,95 @@ namespace pointed
infix ` ⬝* `:75 := phomotopy.trans
postfix `⁻¹*`:(max+1) := phomotopy.symm
/- properties about the given pointed maps -/
/- equalities and equivalences relating pointed homotopies -/
definition is_equiv_ap1 (f : A →* B) [is_equiv f] : is_equiv (ap1 f) :=
definition phomotopy.sigma_char [constructor] {A B : Type*} (f g : A →* B)
: (f ~* g) ≃ Σ(p : f ~ g), p pt ⬝ respect_pt g = respect_pt f :=
begin
induction B with B b, induction f with f pf, esimp at *, cases pf, esimp,
apply is_equiv.homotopy_closed (ap f),
intro p, exact !idp_con⁻¹
fapply equiv.MK : intros h,
{ exact ⟨h , to_homotopy_pt h⟩ },
all_goals cases h with h p,
{ exact phomotopy.mk h p },
all_goals reflexivity
end
definition is_equiv_apn (n : ) (f : A →* B) [H : is_equiv f]
: is_equiv (apn n f) :=
definition is_trunc_pmap [instance] (n : ℕ₋₂) (A B : Type*) [is_trunc n B] :
is_trunc n (A →* B) :=
is_trunc_equiv_closed_rev _ !pmap.sigma_char
definition is_trunc_ppmap [instance] (n : ℕ₋₂) {A B : Type*} [is_trunc n B] :
is_trunc n (ppmap A B) :=
!is_trunc_pmap
definition phomotopy_of_eq [constructor] {A B : Type*} {f g : A →* B} (p : f = g) : f ~* g :=
phomotopy.mk (ap010 pmap.to_fun p) begin induction p, apply idp_con end
definition pconcat_eq [constructor] {A B : Type*} {f g h : A →* B} (p : f ~* g) (q : g = h)
: f ~* h :=
p ⬝* phomotopy_of_eq q
definition eq_pconcat [constructor] {A B : Type*} {f g h : A →* B} (p : f = g) (q : g ~* h)
: f ~* h :=
phomotopy_of_eq p ⬝* q
infix ` ⬝*p `:75 := pconcat_eq
infix ` ⬝p* `:75 := eq_pconcat
definition eq_of_phomotopy (p : f ~* g) : f = g :=
begin
fapply pmap_eq,
{ intro a, exact p a},
{ exact !to_homotopy_pt⁻¹}
end
definition pmap_eq_equiv {A B : Type*} (f g : A →* B) : (f = g) ≃ (f ~* g) :=
calc (f = g) ≃ pmap.sigma_char f = pmap.sigma_char g
: eq_equiv_fn_eq pmap.sigma_char f g
... ≃ Σ(p : pmap.to_fun f = pmap.to_fun g),
pathover (λh, h pt = pt) (respect_pt f) p (respect_pt g)
: sigma_eq_equiv _ _
... ≃ Σ(p : pmap.to_fun f = pmap.to_fun g), respect_pt f = ap (λh, h pt) p ⬝ respect_pt g
: sigma_equiv_sigma_right (λp, eq_pathover_equiv_Fl p (respect_pt f)
(respect_pt g))
... ≃ Σ(p : pmap.to_fun f = pmap.to_fun g), respect_pt f = ap10 p pt ⬝ respect_pt g
: sigma_equiv_sigma_right
(λp, equiv_eq_closed_right _ (whisker_right (ap_eq_apd10 p _) _))
... ≃ Σ(p : pmap.to_fun f ~ pmap.to_fun g), respect_pt f = p pt ⬝ respect_pt g
: sigma_equiv_sigma_left' eq_equiv_homotopy
... ≃ Σ(p : pmap.to_fun f ~ pmap.to_fun g), p pt ⬝ respect_pt g = respect_pt f
: sigma_equiv_sigma_right (λp, eq_equiv_eq_symm _ _)
... ≃ (f ~* g) : phomotopy.sigma_char f g
/-
Pointed maps respecting pointed homotopies.
In general we need function extensionality for pap,
but for particular F we can do it without function extensionality.
This is preferred, because such pointed homotopies compute
-/
definition pap (F : (A →* B) → (C →* D)) {f g : A →* B} (p : f ~* g) : F f ~* F g :=
phomotopy.mk (ap010 (λf, pmap.to_fun (F f)) (eq_of_phomotopy p))
begin cases eq_of_phomotopy p, apply idp_con end
definition ap1_phomotopy {f g : A →* B} (p : f ~* g)
: ap1 f ~* ap1 g :=
begin
induction p with p q, induction f with f pf, induction g with g pg, induction B with B b,
esimp at *, induction q, induction pg,
fapply phomotopy.mk,
{ intro l, esimp, refine _ ⬝ !idp_con⁻¹, refine !con.assoc ⬝ _, apply inv_con_eq_of_eq_con,
apply ap_con_eq_con_ap},
{ induction A with A a, unfold [ap_con_eq_con_ap], generalize p a, generalize g a, intro b q,
induction q, reflexivity}
end
definition apn_phomotopy {f g : A →* B} (n : ) (p : f ~* g) : apn n f ~* apn n g :=
begin
induction n with n IH,
{ exact H},
{ exact is_equiv_ap1 (apn n f)}
{ exact p},
{ exact ap1_phomotopy IH}
end
definition is_equiv_pcast [instance] {A B : Type*} (p : A = B) : is_equiv (pcast p) :=
!is_equiv_cast
/- pointed homotopies between the given pointed maps -/
definition ap1_pid [constructor] {A : Type*} : ap1 (pid A) ~* pid (Ω A) :=
begin
@ -311,106 +453,17 @@ namespace pointed
{ induction B with B b, induction f with f pf, esimp at *, induction pf, reflexivity},
end
theorem ap1_con (f : A →* B) (p q : Ω A) : ap1 f (p ⬝ q) = ap1 f p ⬝ ap1 f q :=
begin
rewrite [▸*,ap_con, +con.assoc, con_inv_cancel_left], repeat apply whisker_left
end
definition ap1_pconst (A B : Type*) : Ω→(pconst A B) ~* pconst (Ω A) (Ω B) :=
phomotopy.mk (λp, idp_con _ ⬝ ap_constant p pt) rfl
theorem ap1_inv (f : A →* B) (p : Ω A) : ap1 f p⁻¹ = (ap1 f p)⁻¹ :=
begin
rewrite [▸*,ap_inv, +con_inv, inv_inv, +con.assoc], repeat apply whisker_left
end
definition ptransport_change_eq [constructor] {A : Type} (B : A → Type*) {a a' : A} {p q : a = a'}
(r : p = q) : ptransport B p ~* ptransport B q :=
phomotopy.mk (λb, ap (λp, transport B p b) r) begin induction r, apply idp_con end
definition pinverse_con [constructor] {X : Type*} (p q : Ω X)
: pinverse (p ⬝ q) = pinverse q ⬝ pinverse p :=
!con_inv
definition pinverse_inv [constructor] {X : Type*} (p : Ω X)
: pinverse p⁻¹ = (pinverse p)⁻¹ :=
idp
definition pcast_idp [constructor] {A : Type*} : pcast (idpath A) ~* pid A :=
by reflexivity
definition apn_zero [unfold_full] (f : A →* B) : Ω→[0] f = f := idp
definition apn_succ [unfold_full] (n : ) (f : A →* B) : Ω→[n + 1] f = Ω→ (Ω→[n] f) := idp
/- more on pointed homotopies -/
definition phomotopy_of_eq [constructor] {A B : Type*} {f g : A →* B} (p : f = g) : f ~* g :=
phomotopy.mk (ap010 pmap.to_fun p) begin induction p, apply idp_con end
definition pconcat_eq [constructor] {A B : Type*} {f g h : A →* B} (p : f ~* g) (q : g = h)
: f ~* h :=
p ⬝* phomotopy_of_eq q
definition eq_pconcat [constructor] {A B : Type*} {f g h : A →* B} (p : f = g) (q : g ~* h)
: f ~* h :=
phomotopy_of_eq p ⬝* q
infix ` ⬝*p `:75 := pconcat_eq
infix ` ⬝p* `:75 := eq_pconcat
definition pwhisker_left [constructor] (h : B →* C) (p : f ~* g) : h ∘* f ~* h ∘* g :=
phomotopy.mk (λa, ap h (p a))
abstract begin
induction A, induction B, induction C,
induction f with f pf, induction g with g pg, induction h with h ph,
induction p with p p', esimp at *, induction ph, induction pg, induction p', reflexivity
end end
definition pwhisker_right [constructor] (h : C →* A) (p : f ~* g) : f ∘* h ~* g ∘* h :=
phomotopy.mk (λa, p (h a))
abstract begin
induction A, induction B, induction C,
induction f with f pf, induction g with g pg, induction h with h ph,
induction p with p p', esimp at *, induction ph, induction pg, induction p', esimp,
exact !idp_con⁻¹
end end
definition pconcat2 [constructor] {A B C : Type*} {h i : B →* C} {f g : A →* B}
(q : h ~* i) (p : f ~* g) : h ∘* f ~* i ∘* g :=
pwhisker_left _ p ⬝* pwhisker_right _ q
definition eq_of_phomotopy (p : f ~* g) : f = g :=
begin
fapply pmap_eq,
{ intro a, exact p a},
{ exact !to_homotopy_pt⁻¹}
end
/-
In general we need function extensionality for pap,
but for particular F we can do it without function extensionality.
-/
definition pap (F : (A →* B) → (C →* D)) {f g : A →* B} (p : f ~* g) : F f ~* F g :=
phomotopy.mk (ap010 F (eq_of_phomotopy p)) begin cases eq_of_phomotopy p, apply idp_con end
definition ap1_phomotopy {f g : A →* B} (p : f ~* g)
: ap1 f ~* ap1 g :=
begin
induction p with p q, induction f with f pf, induction g with g pg, induction B with B b,
esimp at *, induction q, induction pg,
fapply phomotopy.mk,
{ intro l, esimp, refine _ ⬝ !idp_con⁻¹, refine !con.assoc ⬝ _, apply inv_con_eq_of_eq_con,
apply ap_con_eq_con_ap},
{ unfold [ap_con_eq_con_ap], generalize p (Point A), generalize g (Point A), intro b q,
induction q, reflexivity}
end
definition apn_phomotopy {f g : A →* B} (n : ) (p : f ~* g) : apn n f ~* apn n g :=
begin
induction n with n IH,
{ exact p},
{ exact ap1_phomotopy IH}
end
definition apn_pcompose (n : ) (g : B →* C) (f : A →* B) : apn n (g ∘* f) ~* apn n g ∘* apn n f :=
begin
induction n with n IH,
{ reflexivity},
{ refine ap1_phomotopy IH ⬝* _, apply ap1_pcompose}
end
definition pnatural_square {A B : Type} (X : B → Type*) {f g : A → B}
(h : Πa, X (f a) →* X (g a)) {a a' : A} (p : a = a') :
h a' ∘* ptransport X (ap f p) ~* ptransport X (ap g p) ∘* h a :=
by induction p; exact !pcompose_pid ⬝* !pid_pcompose⁻¹*
definition apn_pid [constructor] {A : Type*} (n : ) : apn n (pid A) ~* pid (Ω[n] A) :=
begin
@ -419,12 +472,24 @@ namespace pointed
{ exact ap1_phomotopy IH ⬝* ap1_pid}
end
theorem apn_con (n : ) (f : A →* B) (p q : Ω[n+1] A)
: apn (n+1) f (p ⬝ q) = apn (n+1) f p ⬝ apn (n+1) f q :=
by rewrite [+apn_succ, ap1_con]
definition apn_pconst (A B : Type*) (n : ) :
apn n (pconst A B) ~* pconst (Ω[n] A) (Ω[n] B) :=
begin
induction n with n IH,
{ reflexivity },
{ exact ap1_phomotopy IH ⬝* !ap1_pconst }
end
theorem apn_inv (n : ) (f : A →* B) (p : Ω[n+1] A) : apn (n+1) f p⁻¹ = (apn (n+1) f p)⁻¹ :=
by rewrite [+apn_succ, ap1_inv]
definition apn_pcompose (n : ) (g : B →* C) (f : A →* B) :
apn n (g ∘* f) ~* apn n g ∘* apn n f :=
begin
induction n with n IH,
{ reflexivity},
{ refine ap1_phomotopy IH ⬝* _, apply ap1_pcompose}
end
definition pcast_idp [constructor] {A : Type*} : pcast (idpath A) ~* pid A :=
by reflexivity
definition pinverse_pinverse (A : Type*) : pinverse ∘* pinverse ~* pid (Ω A) :=
begin
@ -449,12 +514,15 @@ namespace pointed
{ reflexivity}
end
definition pmap_of_eq_pt [constructor] {A : Type} {a a' : A} (p : a = a') :
pointed.MK A a →* pointed.MK A a' :=
pmap.mk id p
definition pcast_commute [constructor] {A : Type} {B C : A → Type*} (f : Πa, B a →* C a)
{a₁ a₂ : A} (p : a₁ = a₂) : pcast (ap C p) ∘* f a₁ ~* f a₂ ∘* pcast (ap B p) :=
phomotopy.mk
begin induction p, reflexivity end
begin induction p, esimp, refine !idp_con ⬝ !idp_con ⬝ !ap_id⁻¹ end
/- pointed equivalences -/
/- constructors / projections + variants -/
definition pequiv_of_pmap [constructor] (f : A →* B) (H : is_equiv f) : A ≃* B :=
pequiv.mk f _ (respect_pt f)
@ -496,15 +564,14 @@ namespace pointed
abstract [irreducible] begin
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)),
note H3 := eq_top_of_square (natural_square (to_homotopy fg) (respect_pt f)),
rewrite [▸* at *, H, H3, H2, ap_id, - +con.assoc, ap_compose' f g, con_inv,
- ap_inv, - +ap_con g],
apply whisker_right, apply ap02 g,
rewrite [ap_con, - + con.assoc, +ap_inv, +inv_con_cancel_right, con.left_inv],
end end
definition pua {A B : Type*} (f : A ≃* B) : A = B :=
pType_eq (equiv_of_pequiv f) !respect_pt
/- categorical properties of pointed equivalences -/
protected definition pequiv.refl [refl] [constructor] (A : Type*) : A ≃* A :=
pequiv_of_pmap !pid !is_equiv_id
@ -512,15 +579,25 @@ namespace pointed
protected definition pequiv.rfl [constructor] : A ≃* A :=
pequiv.refl A
protected definition pequiv.symm [symm] (f : A ≃* B) : B ≃* A :=
protected definition pequiv.symm [symm] [constructor] (f : A ≃* B) : B ≃* A :=
pequiv_of_pmap (to_pinv f) !is_equiv_inv
protected definition pequiv.trans [trans] (f : A ≃* B) (g : B ≃* C) : A ≃* C :=
protected definition pequiv.trans [trans] [constructor] (f : A ≃* B) (g : B ≃* C) : A ≃* C :=
pequiv_of_pmap (g ∘* f) !is_equiv_compose
definition pequiv_compose {A B C : Type*} (g : B ≃* C) (f : A ≃* B) : A ≃* C :=
pequiv_of_pmap (g ∘* f) (is_equiv_compose g f)
infixr ` ∘*ᵉ `:60 := pequiv_compose
postfix `⁻¹ᵉ*`:(max + 1) := pequiv.symm
infix ` ⬝e* `:75 := pequiv.trans
/- more on pointed equivalences -/
definition pequiv_ap [constructor] {A : Type} (B : A → Type*) {a a' : A} (p : a = a')
: B a ≃* B a' :=
pequiv_of_pmap (ptransport B p) !is_equiv_tr
definition to_pmap_pequiv_trans {A B C : Type*} (f : A ≃* B) (g : B ≃* C)
: pequiv.to_pmap (f ⬝e* g) = g ∘* f :=
!to_pmap_pequiv_of_pmap
@ -536,6 +613,9 @@ namespace pointed
(g : Πb, P (f⁻¹ b) b) (a : A) : P a (f a) :=
left_inv f a ▸ g (f a)
definition pua {A B : Type*} (f : A ≃* B) : A = B :=
pType_eq (equiv_of_pequiv f) !respect_pt
definition pequiv_of_eq [constructor] {A B : Type*} (p : A = B) : A ≃* B :=
pequiv_of_pmap (pcast p) !is_equiv_tr
@ -551,16 +631,40 @@ namespace pointed
definition peap {A B : Type*} (F : Type* → Type*) (p : A ≃* B) : F A ≃* F B :=
pequiv_of_pmap (pcast (ap F (eq_of_pequiv p))) begin cases eq_of_pequiv p, apply is_equiv_id end
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.elimo
end
infix ` ⬝e*p `:75 := peconcat_eq
infix ` ⬝pe* `:75 := eq_peconcat
local attribute pequiv.symm [constructor]
definition pequiv_of_eq_commute [constructor] {A : Type} {B C : A → Type*} (f : Πa, B a →* C a)
{a₁ a₂ : A} (p : a₁ = a₂) : pequiv_of_eq (ap C p) ∘* f a₁ ~* f a₂ ∘* pequiv_of_eq (ap B p) :=
pcast_commute f p
/-
the theorem pequiv_eq, which gives a condition for two pointed equivalences are equal
is in types.equiv to avoid circular imports
-/
/- computation rules of pointed homotopies, possibly combined with pointed equivalences -/
definition pwhisker_left [constructor] (h : B →* C) (p : f ~* g) : h ∘* f ~* h ∘* g :=
phomotopy.mk (λa, ap h (p a))
abstract begin
induction A, induction B, induction C,
induction f with f pf, induction g with g pg, induction h with h ph,
induction p with p p', esimp at *, induction ph, induction pg, induction p', reflexivity
end end
definition pwhisker_right [constructor] (h : C →* A) (p : f ~* g) : f ∘* h ~* g ∘* h :=
phomotopy.mk (λa, p (h a))
abstract begin
induction A, induction B, induction C,
induction f with f pf, induction g with g pg, induction h with h ph,
induction p with p p', esimp at *, induction ph, induction pg, induction p', esimp,
exact !idp_con⁻¹
end end
definition pconcat2 [constructor] {A B C : Type*} {h i : B →* C} {f g : A →* B}
(q : h ~* i) (p : f ~* g) : h ∘* f ~* i ∘* g :=
pwhisker_left _ p ⬝* pwhisker_right _ q
definition pleft_inv (f : A ≃* B) : f⁻¹ᵉ* ∘* f ~* pid A :=
phomotopy.mk (left_inv f)
abstract begin
@ -575,7 +679,7 @@ namespace pointed
note q := natural_square (right_inv f) p,
rewrite [ap_id at q],
apply eq_bot_of_square,
exact transpose q
exact q
end end
definition pcancel_left (f : B ≃* C) {g h : A →* B} (p : f ∘* g ~* f ∘* h) : g ~* h :=
@ -646,6 +750,73 @@ namespace pointed
(p : h ~* f⁻¹ᵉ* ∘* g) : f ∘* h ~* g :=
(phomotopy_of_pinv_left_phomotopy p⁻¹*)⁻¹*
definition pcompose2 {A B C : Type*} {g g' : B →* C} {f f' : A →* B} (p : f ~* f') (q : g ~* g') :
g ∘* f ~* g' ∘* f' :=
pwhisker_right f q ⬝* pwhisker_left g' p
infixr ` ◾* `:80 := pcompose2
definition phomotopy_pinv_of_phomotopy_pid {A B : Type*} {f : A →* B} {g : B ≃* A}
(p : g ∘* f ~* pid A) : f ~* g⁻¹ᵉ* :=
phomotopy_pinv_left_of_phomotopy p ⬝* !pcompose_pid
definition phomotopy_pinv_of_phomotopy_pid' {A B : Type*} {f : A →* B} {g : B ≃* A}
(p : f ∘* g ~* pid B) : f ~* g⁻¹ᵉ* :=
phomotopy_pinv_right_of_phomotopy p ⬝* !pid_pcompose
definition pinv_phomotopy_of_pid_phomotopy {A B : Type*} {f : A →* B} {g : B ≃* A}
(p : pid A ~* g ∘* f) : g⁻¹ᵉ* ~* f :=
(phomotopy_pinv_of_phomotopy_pid p⁻¹*)⁻¹*
definition pinv_phomotopy_of_pid_phomotopy' {A B : Type*} {f : A →* B} {g : B ≃* A}
(p : pid B ~* f ∘* g) : g⁻¹ᵉ* ~* f :=
(phomotopy_pinv_of_phomotopy_pid' p⁻¹*)⁻¹*
definition pinv_pinv {A B : Type*} (f : A ≃* B) : (f⁻¹ᵉ*)⁻¹ᵉ* ~* f :=
(phomotopy_pinv_of_phomotopy_pid (pleft_inv f))⁻¹*
definition pinv2 {A B : Type*} {f f' : A ≃* B} (p : f ~* f') : f⁻¹ᵉ* ~* f'⁻¹ᵉ* :=
phomotopy_pinv_of_phomotopy_pid (pinv_right_phomotopy_of_phomotopy (!pid_pcompose ⬝* p)⁻¹*)
postfix [parsing_only] `⁻²*`:(max+10) := pinv2
definition trans_pinv {A B C : Type*} (f : A ≃* B) (g : B ≃* C) :
(f ⬝e* g)⁻¹ᵉ* ~* f⁻¹ᵉ* ∘* g⁻¹ᵉ* :=
begin
refine (phomotopy_pinv_of_phomotopy_pid _)⁻¹*,
refine !passoc ⬝* _,
refine pwhisker_left _ (!passoc⁻¹* ⬝* pwhisker_right _ !pright_inv ⬝* !pid_pcompose) ⬝* _,
apply pright_inv
end
definition pinv_trans_pinv_left {A B C : Type*} (f : B ≃* A) (g : B ≃* C) :
(f⁻¹ᵉ* ⬝e* g)⁻¹ᵉ* ~* f ∘* g⁻¹ᵉ* :=
!trans_pinv ⬝* pwhisker_right _ !pinv_pinv
definition pinv_trans_pinv_right {A B C : Type*} (f : A ≃* B) (g : C ≃* B) :
(f ⬝e* g⁻¹ᵉ*)⁻¹ᵉ* ~* f⁻¹ᵉ* ∘* g :=
!trans_pinv ⬝* pwhisker_left _ !pinv_pinv
definition pinv_trans_pinv_pinv {A B C : Type*} (f : B ≃* A) (g : C ≃* B) :
(f⁻¹ᵉ* ⬝e* g⁻¹ᵉ*)⁻¹ᵉ* ~* f ∘* g :=
!trans_pinv ⬝* !pinv_pinv ◾* !pinv_pinv
definition pinv_pcompose_cancel_left {A B C : Type*} (g : B ≃* C) (f : A →* B) :
g⁻¹ᵉ* ∘* (g ∘* f) ~* f :=
!passoc⁻¹* ⬝* pwhisker_right f !pleft_inv ⬝* !pid_pcompose
definition pcompose_pinv_cancel_left {A B C : Type*} (g : C ≃* B) (f : A →* B) :
g ∘* (g⁻¹ᵉ* ∘* f) ~* f :=
!passoc⁻¹* ⬝* pwhisker_right f !pright_inv ⬝* !pid_pcompose
definition pinv_pcompose_cancel_right {A B C : Type*} (g : B →* C) (f : B ≃* A) :
(g ∘* f⁻¹ᵉ*) ∘* f ~* g :=
!passoc ⬝* pwhisker_left g !pleft_inv ⬝* !pcompose_pid
definition pcompose_pinv_cancel_right {A B C : Type*} (g : B →* C) (f : A ≃* B) :
(g ∘* f) ∘* f⁻¹ᵉ* ~* g :=
!passoc ⬝* pwhisker_left g !pright_inv ⬝* !pcompose_pid
/- pointed equivalences between particular pointed types -/
definition loopn_pequiv_loopn [constructor] (n : ) (f : A ≃* B) : Ω[n] A ≃* Ω[n] B :=
@ -690,14 +861,13 @@ namespace pointed
loopn_pequiv_loopn_con 0 f p q
definition loopn_pequiv_loopn_rfl (n : ) (A : Type*) :
loopn_pequiv_loopn n (pequiv.refl A) = pequiv.refl (Ω[n] A) :=
loopn_pequiv_loopn n (pequiv.refl A) ~* pequiv.refl (Ω[n] A) :=
begin
apply pequiv_eq, apply eq_of_phomotopy,
exact !to_pmap_loopn_pequiv_loopn ⬝* apn_pid n,
end
definition loop_pequiv_loop_rfl (A : Type*) :
loop_pequiv_loop (pequiv.refl A) = pequiv.refl (Ω A) :=
loop_pequiv_loop (pequiv.refl A) ~* pequiv.refl (Ω A) :=
loopn_pequiv_loopn_rfl 1 A
definition pmap_functor [constructor] {A A' B B' : Type*} (f : A' →* A) (g : B →* B') :
@ -796,6 +966,20 @@ namespace pointed
{ exact !ap1_pcompose⁻¹* ⬝* ap1_phomotopy IH ⬝* !ap1_pcompose}
end
definition loopn_succ_in_natural {A B : Type*} (n : ) (f : A →* B) :
loopn_succ_in B n ∘* Ω→[n+1] f ~* Ω→[n] (Ω→ f) ∘* loopn_succ_in A n :=
!apn_succ_phomotopy_in
definition loopn_succ_in_inv_natural {A B : Type*} (n : ) (f : A →* B) :
Ω→[n + 1] f ∘* (loopn_succ_in A n)⁻¹ᵉ* ~* (loopn_succ_in B n)⁻¹ᵉ* ∘* Ω→[n] (Ω→ f):=
begin
apply pinv_right_phomotopy_of_phomotopy,
refine _ ⬝* !passoc⁻¹*,
apply phomotopy_pinv_left_of_phomotopy,
apply apn_succ_phomotopy_in
end
/- properties of ppmap, the pointed type of pointed maps -/
definition ppcompose_left [constructor] (g : B →* C) : ppmap A B →* ppmap A C :=
pmap.mk (pcompose g) (eq_of_phomotopy (phomotopy.mk (λa, respect_pt g) (idp_con _)⁻¹))
@ -805,12 +989,14 @@ namespace pointed
fapply is_equiv.adjointify,
{ exact (ppcompose_left (pequiv_of_pmap g H)⁻¹ᵉ*) },
all_goals (intros f; esimp; apply eq_of_phomotopy),
{ exact calc g ∘* ((pequiv_of_pmap g H)⁻¹ᵉ* ∘* f) ~* (g ∘* (pequiv_of_pmap g H)⁻¹ᵉ*) ∘* f : passoc
... ~* pid _ ∘* f : pwhisker_right f (pright_inv (pequiv_of_pmap g H))
... ~* f : pid_pcompose f },
{ exact calc (pequiv_of_pmap g H)⁻¹ᵉ* ∘* (g ∘* f) ~* ((pequiv_of_pmap g H)⁻¹ᵉ* ∘* g) ∘* f : passoc
... ~* pid _ ∘* f : pwhisker_right f (pleft_inv (pequiv_of_pmap g H))
... ~* f : pid_pcompose f }
{ exact calc g ∘* ((pequiv_of_pmap g H)⁻¹ᵉ* ∘* f)
~* (g ∘* (pequiv_of_pmap g H)⁻¹ᵉ*) ∘* f : passoc
... ~* pid _ ∘* f : pwhisker_right f (pright_inv (pequiv_of_pmap g H))
... ~* f : pid_pcompose f },
{ exact calc (pequiv_of_pmap g H)⁻¹ᵉ* ∘* (g ∘* f)
~* ((pequiv_of_pmap g H)⁻¹ᵉ* ∘* g) ∘* f : passoc
... ~* pid _ ∘* f : pwhisker_right f (pleft_inv (pequiv_of_pmap g H))
... ~* f : pid_pcompose f }
end
definition pequiv_ppcompose_left [constructor] (g : B ≃* C) : ppmap A B ≃* ppmap A C :=
@ -840,5 +1026,36 @@ namespace pointed
refine pwhisker_left g !pleft_inv ⬝* !pcompose_pid, },
end
definition loop_pmap_commute (A B : Type*) : Ω(ppmap A B) ≃* (ppmap A (Ω B)) :=
pequiv_of_equiv
(calc Ω(ppmap A B) ≃ (pconst A B ~* pconst A B) : pmap_eq_equiv _ _
... ≃ Σ(p : pconst A B ~ pconst A B), p pt ⬝ rfl = rfl : phomotopy.sigma_char
... ≃ (A →* Ω B) : pmap.sigma_char)
(by reflexivity)
definition papply [constructor] {A : Type*} (B : Type*) (a : A) : ppmap A B →* B :=
pmap.mk (λ(f : A →* B), f a) idp
definition papply_pcompose [constructor] {A : Type*} (B : Type*) (a : A) : ppmap A B →* B :=
pmap.mk (λ(f : A →* B), f a) idp
definition pmap_pbool_pequiv [constructor] (B : Type*) : ppmap pbool B ≃* B :=
begin
fapply pequiv.MK,
{ exact papply B tt },
{ exact pbool_pmap },
{ intro f, fapply pmap_eq,
{ intro b, cases b, exact !respect_pt⁻¹, reflexivity },
{ exact !con.left_inv⁻¹ }},
{ intro b, reflexivity },
end
definition papn_pt [constructor] (n : ) (A B : Type*) : ppmap A B →* ppmap (Ω[n] A) (Ω[n] B) :=
pmap.mk (λf, apn n f) (eq_of_phomotopy !apn_pconst)
definition papn_fun [constructor] {n : } {A : Type*} (B : Type*) (p : Ω[n] A) :
ppmap A B →* Ω[n] B :=
papply _ p ∘* papn_pt n A B
end pointed

View file

@ -264,9 +264,12 @@ namespace sigma
sigma_equiv_sigma equiv.rfl Hg
definition sigma_equiv_sigma_left [constructor] (Hf : A ≃ A') :
(Σa, B a) ≃ (Σa', B (to_inv Hf a')) :=
(Σa, B a) ≃ (Σa', B (to_inv Hf a')) :=
sigma_equiv_sigma Hf (λ a, equiv_ap B !right_inv⁻¹)
definition sigma_equiv_sigma_left' [constructor] (Hf : A' ≃ A) : (Σa, B (Hf a)) ≃ (Σa', B a') :=
sigma_equiv_sigma Hf (λa, erfl)
definition ap_sigma_functor_eq_dpair (p : a = a') (q : b =[p] b') :
ap (sigma_functor f g) (sigma_eq p q) = sigma_eq (ap f p) (pathover.rec_on q idpo) :=
by induction q; reflexivity

View file

@ -322,17 +322,17 @@ namespace is_trunc
definition tua_refl {n : ℕ₋₂} (A : n-Type) : tua (@erfl A) = idp :=
begin
refine ap (trunctype_eq_equiv n A A)⁻¹ᶠ (ua_refl A) ⬝ _,
esimp, refine ap (eq_of_fn_eq_fn _) _ ⬝ !eq_of_fn_eq_fn'_idp ,
esimp, apply ap (dpair_eq_dpair idp), apply is_prop.elim
refine ap (eq_of_fn_eq_fn _) _ ⬝ !eq_of_fn_eq_fn'_idp ,
apply ap (dpair_eq_dpair idp), apply is_prop.elim
end
definition tua_trans {n : ℕ₋₂} {A B C : n-Type} (f : A ≃ B) (g : B ≃ C)
: tua (f ⬝e g) = tua f ⬝ tua g :=
begin
refine ap (trunctype_eq_equiv n A C)⁻¹ᶠ (ua_trans f g) ⬝ _,
esimp, refine ap (eq_of_fn_eq_fn _) _ ⬝ !eq_of_fn_eq_fn'_con,
refine ap (eq_of_fn_eq_fn _) _ ⬝ !eq_of_fn_eq_fn'_con,
refine _ ⬝ !dpair_eq_dpair_con,
apply ap (dpair_eq_dpair _), apply is_prop.elim
apply ap (dpair_eq_dpair _), esimp, apply is_prop.elim
end
definition tua_symm {n : ℕ₋₂} {A B : n-Type} (f : A ≃ B) : tua f⁻¹ᵉ = (tua f)⁻¹ :=
@ -372,7 +372,7 @@ namespace is_trunc
to_fun (equiv.symm !heq_pi) H2,
have H4 : imp (refl a) ⬝ p = imp (refl a), from
calc
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : transport_eq_r
imp (refl a) ⬝ p = transport (λx, a = x) p (imp (refl a)) : eq_transport_r
... = imp (transport (λx, R a x) p (refl a)) : H3
... = imp (refl a) : is_prop.elim,
cancel_left (imp (refl a)) H4)
@ -674,19 +674,31 @@ namespace trunc
exact tr (fiber.mk (tr a) (ap tr p))}
end
/- truncation of pointed types and its functorial action -/
/- truncation of pointed types -/
definition ptrunc [constructor] (n : ℕ₋₂) (X : Type*) : n-Type* :=
ptrunctype.mk (trunc n X) _ (tr pt)
/- pointed maps involving ptrunc -/
definition ptrunc_functor [constructor] {X Y : Type*} (n : ℕ₋₂) (f : X →* Y)
: ptrunc n X →* ptrunc n Y :=
pmap.mk (trunc_functor n f) (ap tr (respect_pt f))
definition ptr [constructor] (n : ℕ₋₂) (A : Type*) : A →* ptrunc n A :=
pmap.mk tr idp
definition puntrunc [constructor] (n : ℕ₋₂) (A : Type*) [is_trunc n A] : ptrunc n A →* A :=
pmap.mk untrunc_of_is_trunc idp
definition ptrunc.elim [constructor] (n : ℕ₋₂) {X Y : Type*} [is_trunc n Y] (f : X →* Y) :
ptrunc n X →* Y :=
pmap.mk (trunc.elim f) (respect_pt f)
/- pointed equivalences involving ptrunc -/
definition ptrunc_pequiv_ptrunc [constructor] (n : ℕ₋₂) {X Y : Type*} (H : X ≃* Y)
: ptrunc n X ≃* ptrunc n Y :=
pequiv_of_equiv (trunc_equiv_trunc n H) (ap tr (respect_pt H))
definition ptrunc_pequiv [constructor] (n : ℕ₋₂) (X : Type*) (H : is_trunc n X)
definition ptrunc_pequiv [constructor] (n : ℕ₋₂) (X : Type*) [H : is_trunc n X]
: ptrunc n X ≃* X :=
pequiv_of_equiv (trunc_equiv n X) idp
@ -705,7 +717,7 @@ namespace trunc
ptrunc_ptrunc_pequiv_left B H
definition ptrunc_ptrunc_pequiv_ptrunc_ptrunc [constructor] (n m : ℕ₋₂) (A : Type*)
: ptrunc n (ptrunc m A) ≃ ptrunc m (ptrunc n A) :=
: ptrunc n (ptrunc m A) ≃* ptrunc m (ptrunc n A) :=
pequiv_of_equiv (trunc_trunc_equiv_trunc_trunc n m A) idp
definition loop_ptrunc_pequiv [constructor] (n : ℕ₋₂) (A : Type*) :
@ -717,7 +729,6 @@ namespace trunc
tconcat (loop_ptrunc_pequiv n A p) (loop_ptrunc_pequiv n A q) :=
encode_con p q
-- rename
definition loopn_ptrunc_pequiv (n : ℕ₋₂) (k : ) (A : Type*) :
Ω[k] (ptrunc (n+k) A) ≃* ptrunc n (Ω[k] A) :=
begin
@ -747,6 +758,7 @@ namespace trunc
equiv.inv_preserve_binary (loopn_ptrunc_pequiv n (succ k) A) concat tconcat
(@loopn_ptrunc_pequiv_con n k A) p q
/- pointed homotopies involving ptrunc -/
definition ptrunc_functor_pcompose [constructor] {X Y Z : Type*} (n : ℕ₋₂) (g : Y →* Z)
(f : X →* Y) : ptrunc_functor n (g ∘* f) ~* ptrunc_functor n g ∘* ptrunc_functor n f :=
begin
@ -789,6 +801,59 @@ namespace trunc
{ induction p, reflexivity}
end
definition ptrunc_elim_ptr [constructor] (n : ℕ₋₂) {X Y : Type*} [is_trunc n Y] (f : X →* Y) :
ptrunc.elim n f ∘* ptr n X ~* f :=
begin
fapply phomotopy.mk,
{ reflexivity },
{ reflexivity }
end
definition ptrunc_elim_phomotopy (n : ℕ₋₂) {X Y : Type*} [is_trunc n Y] {f g : X →* Y}
(H : f ~* g) : ptrunc.elim n f ~* ptrunc.elim n g :=
begin
fapply phomotopy.mk,
{ intro x, induction x with x, exact H x },
{ exact to_homotopy_pt H }
end
definition ap1_ptrunc_functor (n : ℕ₋₂) {A B : Type*} (f : A →* B) :
Ω→ (ptrunc_functor (n.+1) f) ∘* (loop_ptrunc_pequiv n A)⁻¹ᵉ* ~*
(loop_ptrunc_pequiv n B)⁻¹ᵉ* ∘* ptrunc_functor n (Ω→ f) :=
begin
fapply phomotopy.mk,
{ intro p, induction p with p,
refine (!ap_inv⁻¹ ◾ !ap_compose⁻¹ ◾ idp) ⬝ _ ⬝ !ap_con⁻¹,
apply whisker_right, refine _ ⬝ !ap_con⁻¹, exact whisker_left _ !ap_compose },
{ induction B with B b, induction f with f p, esimp at f, esimp at p, induction p, reflexivity }
end
definition ap1_ptrunc_elim (n : ℕ₋₂) {A B : Type*} (f : A →* B) [is_trunc (n.+1) B] :
Ω→ (ptrunc.elim (n.+1) f) ∘* (loop_ptrunc_pequiv n A)⁻¹ᵉ* ~*
ptrunc.elim n (Ω→ f) :=
begin
fapply phomotopy.mk,
{ intro p, induction p with p, exact idp ◾ !ap_compose⁻¹ ◾ idp },
{ reflexivity }
end
definition ap1_ptr (n : ℕ₋₂) (A : Type*) :
Ω→ (ptr (n.+1) A) ~* (loop_ptrunc_pequiv n A)⁻¹ᵉ* ∘* ptr n (Ω A) :=
begin
fapply phomotopy.mk,
{ intro p, apply idp_con },
{ reflexivity }
end
definition ptrunc_elim_ptrunc_functor (n : ℕ₋₂) {A B C : Type*} (g : B →* C) (f : A →* B)
[is_trunc n C] :
ptrunc.elim n g ∘* ptrunc_functor n f ~* ptrunc.elim n (g ∘* f) :=
begin
fapply phomotopy.mk,
{ intro x, induction x with a, reflexivity },
{ esimp, exact !idp_con ⬝ whisker_right !ap_compose _ },
end
end trunc open trunc
/- The truncated encode-decode method -/

View file

@ -47,5 +47,3 @@ namespace unit
end
end unit
open unit is_trunc