Merge branch 'master' of github.com:fpvandoorn/Spectral

This commit is contained in:
Yuri Sulyma 2017-06-08 14:04:58 -06:00
parent a3146d0d2a
commit 3fd6e8e852
10 changed files with 205 additions and 73 deletions

View file

@ -6,9 +6,9 @@ Authors: Floris van Doorn, Egbert Rijke
Constructions with groups Constructions with groups
-/ -/
import .quotient_group .free_commutative_group import .quotient_group .free_commutative_group .product_group
open eq algebra is_trunc set_quotient relation sigma prod sum list trunc function equiv sigma.ops open eq is_equiv algebra is_trunc set_quotient relation sigma prod sum list trunc function equiv sigma.ops
namespace group namespace group
@ -73,10 +73,10 @@ namespace group
definition dirsum_elim [constructor] (f : Πi, Y i →g A') : dirsum →g A' := definition dirsum_elim [constructor] (f : Πi, Y i →g A') : dirsum →g A' :=
gqg_elim _ (free_ab_group_elim (λv, f v.1 v.2)) (dirsum_elim_resp_quotient f) gqg_elim _ (free_ab_group_elim (λv, f v.1 v.2)) (dirsum_elim_resp_quotient f)
definition dirsum_elim_compute (f : Πi, Y i →g A') (i : I) : definition dirsum_elim_compute (f : Πi, Y i →g A') (i : I) (y : Y i) :
dirsum_elim f ∘g dirsum_incl i ~ f i := dirsum_elim f (dirsum_incl i y) = f i y :=
begin begin
intro g, apply one_mul apply one_mul
end end
definition dirsum_elim_unique (f : Πi, Y i →g A') (k : dirsum →g A') definition dirsum_elim_unique (f : Πi, Y i →g A') (k : dirsum →g A')
@ -89,6 +89,35 @@ namespace group
end end
definition binary_dirsum (G H : AbGroup) : dirsum (bool.rec G H) ≃g G ×ag H :=
let branch := bool.rec G H in
let to_hom := (dirsum_elim (bool.rec (product_inl G H) (product_inr G H))
: dirsum (bool.rec G H) →g G ×ag H) in
let from_hom := (Group_sum_elim (dirsum (bool.rec G H))
(dirsum_incl branch bool.ff) (dirsum_incl branch bool.tt)
: G ×g H →g dirsum branch) in
begin
fapply isomorphism.mk,
{ exact dirsum_elim (bool.rec (product_inl G H) (product_inr G H)) },
fapply adjointify,
{ exact from_hom },
{ intro gh, induction gh with g h,
exact prod_eq (mul_one (1 * g) ⬝ one_mul g) (ap (λ o, o * h) (mul_one 1) ⬝ one_mul h) },
{ refine dirsum.rec _ _ _,
{ intro b x,
refine ap from_hom (dirsum_elim_compute (bool.rec (product_inl G H) (product_inr G H)) b x) ⬝ _,
induction b,
{ exact ap (λ y, dirsum_incl branch bool.ff x * y) (to_respect_one (dirsum_incl branch bool.tt)) ⬝ mul_one _ },
{ exact ap (λ y, y * dirsum_incl branch bool.tt x) (to_respect_one (dirsum_incl branch bool.ff)) ⬝ one_mul _ }
},
{ refine ap from_hom (to_respect_one to_hom) ⬝ to_respect_one from_hom },
{ intro g h gβ hβ,
refine ap from_hom (to_respect_mul to_hom _ _) ⬝ to_respect_mul from_hom _ _ ⬝ _,
exact ap011 mul gβ hβ
}
}
end
variables {I J : Set} {Y Y' Y'' : I → AbGroup} variables {I J : Set} {Y Y' Y'' : I → AbGroup}
definition dirsum_functor [constructor] (f : Πi, Y i →g Y' i) : dirsum Y →g dirsum Y' := definition dirsum_functor [constructor] (f : Πi, Y i →g Y' i) : dirsum Y →g dirsum Y' :=

View file

@ -8,7 +8,7 @@ Constructions with groups
import algebra.group_theory hit.set_quotient types.list types.sum .subgroup .quotient_group import algebra.group_theory hit.set_quotient types.list types.sum .subgroup .quotient_group
open eq algebra is_trunc set_quotient relation sigma sigma.ops prod prod.ops sum list trunc function open eq algebra is_trunc set_quotient relation sigma prod prod.ops sum list trunc function
equiv equiv
namespace group namespace group
@ -61,6 +61,19 @@ namespace group
infix ` ×g `:60 := group.product infix ` ×g `:60 := group.product
infix ` ×ag `:60 := group.ab_product infix ` ×ag `:60 := group.ab_product
definition product_inl [constructor] (G H : Group) : G →g G ×g H :=
homomorphism.mk (λx, (x, one)) (λx y, prod_eq !refl !one_mul⁻¹)
definition product_inr [constructor] (G H : Group) : H →g G ×g H :=
homomorphism.mk (λx, (one, x)) (λx y, prod_eq !one_mul⁻¹ !refl)
definition Group_sum_elim [constructor] {G H : Group} (I : AbGroup) (φ : G →g I) (ψ : H →g I) : G ×g H →g I :=
homomorphism.mk (λx, φ x.1 * ψ x.2) abstract (λx y, calc
φ (x.1 * y.1) * ψ (x.2 * y.2) = (φ x.1 * φ y.1) * (ψ x.2 * ψ y.2)
: by exact ap011 mul (to_respect_mul φ x.1 y.1) (to_respect_mul ψ x.2 y.2)
... = (φ x.1 * ψ x.2) * (φ y.1 * ψ y.2)
: by exact interchange I (φ x.1) (φ y.1) (ψ x.2) (ψ y.2)) end
definition product_functor [constructor] {G G' H H' : Group} (φ : G →g H) (ψ : G' →g H') : definition product_functor [constructor] {G G' H H' : Group} (φ : G →g H) (ψ : G' →g H') :
G ×g G' →g H ×g H' := G ×g G' →g H ×g H' :=
homomorphism.mk (λx, (φ x.1, ψ x.2)) (λx y, prod_eq !to_respect_mul !to_respect_mul) homomorphism.mk (λx, (φ x.1, ψ x.2)) (λx y, prod_eq !to_respect_mul !to_respect_mul)
@ -73,4 +86,7 @@ namespace group
infix ` ×≃g `:60 := group.product_isomorphism infix ` ×≃g `:60 := group.product_isomorphism
definition product_group_mul_eq {G H : Group} (g h : G ×g H) : g * h = product_mul g h :=
idp
end group end group

View file

@ -227,10 +227,10 @@ namespace group
unfold qg_map, esimp, exact to_respect_mul f g h } unfold qg_map, esimp, exact to_respect_mul f g h }
end end
definition quotient_group_compute (f : G →g G') (H : Π⦃g⦄, N g → f g = 1) : definition quotient_group_compute (f : G →g G') (H : Π⦃g⦄, N g → f g = 1) (g : G) :
quotient_group_elim f H ∘g qg_map N ~ f := quotient_group_elim f H (qg_map N g) = f g :=
begin begin
intro g, reflexivity reflexivity
end end
definition gelim_unique (f : G →g G') (H : Π⦃g⦄, N g → f g = 1) (k : quotient_group N →g G') definition gelim_unique (f : G →g G') (H : Π⦃g⦄, N g → f g = 1) (k : quotient_group N →g G')
@ -247,7 +247,7 @@ namespace group
end end
definition qg_universal_property (f : G →g G') (H : Π⦃g⦄, N g → f g = 1) : definition qg_universal_property (f : G →g G') (H : Π⦃g⦄, N g → f g = 1) :
is_contr (Σ(g : quotient_group N →g G'), g ∘g qg_map N ~ f) := is_contr (Σ(g : quotient_group N →g G'), g ∘ qg_map N ~ f) :=
begin begin
fapply is_contr.mk, fapply is_contr.mk,
-- give center of contraction -- give center of contraction
@ -442,7 +442,7 @@ definition kernel_quotient_extension {A B : AbGroup} (f : A →g B) : quotient_a
end end
definition kernel_quotient_extension_triangle {A B : AbGroup} (f : A →g B) : definition kernel_quotient_extension_triangle {A B : AbGroup} (f : A →g B) :
kernel_quotient_extension f ∘g ab_qg_map (kernel_subgroup f) ~ f := kernel_quotient_extension f ∘ ab_qg_map (kernel_subgroup f) ~ f :=
begin begin
intro a, intro a,
apply quotient_group_compute apply quotient_group_compute

View file

@ -18,14 +18,21 @@ namespace group
definition seq_colim_incl [constructor] (i : ) : A i →g seq_colim := definition seq_colim_incl [constructor] (i : ) : A i →g seq_colim :=
qg_map _ ∘g dirsum_incl A i qg_map _ ∘g dirsum_incl A i
definition seq_colim_quotient (h : Πi, A i →g A') (k : Πi a, h i a = h (i + 1) (f i a)) definition seq_colim_quotient (h : Πi, A i →g A') (k : Πi a, h i a = h (succ i) (f i a))
(v : seq_colim_carrier) (r : ∥seq_colim_rel v∥) : dirsum_elim h v = 1 := (v : seq_colim_carrier) (r : ∥seq_colim_rel v∥) : dirsum_elim h v = 1 :=
begin begin
induction r with r, induction r, exact sorry induction r with r, induction r,
refine !to_respect_mul ⬝ _,
refine ap (λγ, group_fun (dirsum_elim h) (group_fun (dirsum_incl A i) a) * group_fun (dirsum_elim h) γ) (!to_respect_inv)⁻¹ ⬝ _,
refine ap (λγ, γ * group_fun (dirsum_elim h) (group_fun (dirsum_incl A (succ i)) (f i a)⁻¹)) !dirsum_elim_compute ⬝ _,
refine ap (λγ, (h i a) * γ) !dirsum_elim_compute ⬝ _,
refine ap (λγ, γ * group_fun (h (succ i)) (f i a)⁻¹) !k ⬝ _,
refine ap (λγ, group_fun (h (succ i)) (f i a) * γ) (!to_respect_inv) ⬝ _,
exact !mul.right_inv
end end
definition seq_colim_elim [constructor] (h : Πi, A i →g A') definition seq_colim_elim [constructor] (h : Πi, A i →g A')
(k : Πi a, h i a = h (i + 1) (f i a)) : seq_colim →g A' := (k : Πi a, h i a = h (succ i) (f i a)) : seq_colim →g A' :=
gqg_elim _ (dirsum_elim h) (seq_colim_quotient h k) gqg_elim _ (dirsum_elim h) (seq_colim_quotient h k)
end end

View file

@ -251,7 +251,7 @@ namespace seq_colim
definition pshift_equiv_pinclusion {A : → Type*} (f : Πn, A n →* A (succ n)) (n : ) : definition pshift_equiv_pinclusion {A : → Type*} (f : Πn, A n →* A (succ n)) (n : ) :
psquare (pinclusion f n) (pinclusion (λn, f (n+1)) n) (f n) (pshift_equiv f) := psquare (pinclusion f n) (pinclusion (λn, f (n+1)) n) (f n) (pshift_equiv f) :=
phomotopy.mk homotopy.rfl begin phomotopy.mk homotopy.rfl begin
refine !idp_con ⬝ _, esimp, refine !idp_con ⬝ _, esimp,
induction n with n IH, induction n with n IH,
{ esimp[inclusion_pt], esimp[shift_diag], exact !idp_con⁻¹ }, { esimp[inclusion_pt], esimp[shift_diag], exact !idp_con⁻¹ },
@ -259,7 +259,7 @@ namespace seq_colim
rewrite ap_con, rewrite ap_con, rewrite ap_con, rewrite ap_con,
refine _ ⬝ whisker_right _ !con.assoc, refine _ ⬝ whisker_right _ !con.assoc,
refine _ ⬝ (con.assoc (_ ⬝ _) _ _)⁻¹, refine _ ⬝ (con.assoc (_ ⬝ _) _ _)⁻¹,
xrewrite [-IH], xrewrite [-IH],
esimp[shift_up], rewrite [elim_glue, ap_inv, -ap_compose'], esimp, esimp[shift_up], rewrite [elim_glue, ap_inv, -ap_compose'], esimp,
rewrite [-+con.assoc], apply whisker_right, rewrite [-+con.assoc], apply whisker_right,
rewrite con.assoc, apply !eq_inv_con_of_con_eq, rewrite con.assoc, apply !eq_inv_con_of_con_eq,
@ -285,7 +285,6 @@ namespace seq_colim
!elim_glue !elim_glue
omit p omit p
definition is_equiv_seq_colim_functor [constructor] [H : Πn, is_equiv (@g n)] definition is_equiv_seq_colim_functor [constructor] [H : Πn, is_equiv (@g n)]
: is_equiv (seq_colim_functor @g p) := : is_equiv (seq_colim_functor @g p) :=
adjointify _ (seq_colim_functor (λn, (@g _)⁻¹) (λn a, inv_commute' g f f' p a)) adjointify _ (seq_colim_functor (λn, (@g _)⁻¹) (λn a, inv_commute' g f f' p a))
@ -323,29 +322,60 @@ namespace seq_colim
: Π(x : seq_colim f), P x := : Π(x : seq_colim f), P x :=
by induction v with Pincl Pglue; exact seq_colim.rec f Pincl Pglue by induction v with Pincl Pglue; exact seq_colim.rec f Pincl Pglue
definition pseq_colim_pequiv [constructor] {A A' : → Type*} {f : Π{n}, A n →* A (n+1)} definition pseq_colim_pequiv' [constructor] {A A' : → Type*} {f : Πn, A n →* A (n+1)}
{f' : Π{n}, A' n →* A' (n+1)} (g : Π{n}, A n ≃* A' n) {f' : Πn, A' n →* A' (n+1)} (g : Πn, A n ≃* A' n)
(p : Π⦃n⦄, g ∘* f ~ f' ∘* g) : pseq_colim @f ≃* pseq_colim @f' := (p : Π⦃n⦄, g (n+1) ∘* f n ~ f' n ∘* g n) : pseq_colim @f ≃* pseq_colim @f' :=
pequiv_of_equiv (seq_colim_equiv @g p) (ap (ι _) (respect_pt g)) pequiv_of_equiv (seq_colim_equiv g p) (ap (ι _) (respect_pt (g _)))
definition pseq_colim_pequiv [constructor] {A A' : → Type*} {f : Πn, A n →* A (n+1)}
{f' : Πn, A' n →* A' (n+1)} (g : Πn, A n ≃* A' n)
(p : Πn, g (n+1) ∘* f n ~* f' n ∘* g n) : pseq_colim @f ≃* pseq_colim @f' :=
pseq_colim_pequiv' g (λn, @p n)
definition seq_colim_equiv_constant [constructor] {A : → Type*} {f f' : Π⦃n⦄, A n → A (n+1)} definition seq_colim_equiv_constant [constructor] {A : → Type*} {f f' : Π⦃n⦄, A n → A (n+1)}
(p : Π⦃n⦄ (a : A n), f a = f' a) : seq_colim f ≃ seq_colim f' := (p : Π⦃n⦄ (a : A n), f a = f' a) : seq_colim f ≃ seq_colim f' :=
seq_colim_equiv (λn, erfl) p seq_colim_equiv (λn, erfl) p
definition pseq_colim_equiv_constant [constructor] {A : → Type*} {f f' : Π{n}, A n →* A (n+1)} definition pseq_colim_equiv_constant' [constructor] {A : → Type*} {f f' : Πn, A n →* A (n+1)}
(p : Π⦃n⦄, f ~ f') : pseq_colim @f ≃* pseq_colim @f' := (p : Π⦃n⦄, f n ~ f' n) : pseq_colim @f ≃* pseq_colim @f' :=
pseq_colim_pequiv (λn, pequiv.rfl) p pseq_colim_pequiv' (λn, pequiv.rfl) p
definition pseq_colim_pequiv_pinclusion {A A' : → Type*} {f : Π(n), A n →* A (n+1)} definition pseq_colim_equiv_constant [constructor] {A : → Type*} {f f' : Πn, A n →* A (n+1)}
{f' : Π(n), A' n →* A' (n+1)} (g : Π(n), A n ≃* A' n) (p : Πn, f n ~* f' n) : pseq_colim @f ≃* pseq_colim @f' :=
(p : Π⦃n⦄, g (n+1) ∘* f n ~ f' n ∘* g n) (n : ) : pseq_colim_pequiv (λn, pequiv.rfl) (λn, !pid_pcompose ⬝* p n ⬝* !pcompose_pid⁻¹*)
definition pseq_colim_pequiv_pinclusion {A A' : → Type*} {f : Πn, A n →* A (n+1)}
{f' : Πn, A' n →* A' (n+1)} (g : Πn, A n ≃* A' n)
(p : Π⦃n⦄, g (n+1) ∘* f n ~* f' n ∘* g n) (n : ) :
psquare (pinclusion f n) (pinclusion f' n) (g n) (pseq_colim_pequiv g p) := psquare (pinclusion f n) (pinclusion f' n) (g n) (pseq_colim_pequiv g p) :=
sorry phomotopy.mk homotopy.rfl begin
esimp, refine !idp_con ⬝ _,
induction n with n IH,
{ esimp[inclusion_pt], exact !idp_con⁻¹ },
{ esimp[inclusion_pt], rewrite [+ap_con, -+ap_inv, +con.assoc, +seq_colim_functor_glue],
xrewrite[-IH],
rewrite[-+ap_compose', -+con.assoc],
apply whisker_right, esimp,
rewrite[(eq_con_inv_of_con_eq (!to_homotopy_pt))],
rewrite[ap_con], esimp,
rewrite[-+con.assoc, ap_con, -ap_compose', +ap_inv],
rewrite[-+con.assoc],
refine _ ⬝ whisker_right _ (whisker_right _ (whisker_right _ (whisker_right _ !con.left_inv⁻¹))),
rewrite[idp_con, +con.assoc], apply whisker_left,
rewrite[ap_con, -ap_compose', con_inv, +con.assoc], apply whisker_left,
refine eq_inv_con_of_con_eq _,
symmetry, exact eq_of_square !natural_square
}
end
definition seq_colim_equiv_constant_pinclusion {A : → Type*} {f f' : Π⦃n⦄, A n →* A (n+1)} definition seq_colim_equiv_constant_pinclusion {A : → Type*} {f f' : Πn, A n →* A (n+1)}
(p : Π⦃n⦄ (a : A n), f a = f' a) (n : ) : (p : Πn, f n ~* f' n) (n : ) :
pseq_colim_equiv_constant p ∘* pinclusion f n ~* pinclusion f' n := pseq_colim_equiv_constant p ∘* pinclusion f n ~* pinclusion f' n :=
sorry begin
transitivity pinclusion f' n ∘* !pid,
refine phomotopy_of_psquare !pseq_colim_pequiv_pinclusion,
exact !pcompose_pid
end
definition is_equiv_seq_colim_rec (P : seq_colim f → Type) : definition is_equiv_seq_colim_rec (P : seq_colim f → Type) :
is_equiv (seq_colim_rec_unc : is_equiv (seq_colim_rec_unc :
@ -369,7 +399,7 @@ namespace seq_colim
equiv.mk _ !is_equiv_seq_colim_rec equiv.mk _ !is_equiv_seq_colim_rec
end functor end functor
definition pseq_colim.elim [constructor] {A : → Type*} {B : Type*} {f : Π{n}, A n →* A (n+1)} definition pseq_colim.elim' [constructor] {A : → Type*} {B : Type*} {f : Π{n}, A n →* A (n+1)}
(g : Πn, A n →* B) (p : Πn, g (n+1) ∘* f ~ g n) : pseq_colim @f →* B := (g : Πn, A n →* B) (p : Πn, g (n+1) ∘* f ~ g n) : pseq_colim @f →* B :=
begin begin
fapply pmap.mk, fapply pmap.mk,
@ -379,6 +409,10 @@ namespace seq_colim
{ esimp, apply respect_pt } { esimp, apply respect_pt }
end end
definition pseq_colim.elim [constructor] {A : → Type*} {B : Type*} {f : Π{n}, A n →* A (n+1)}
(g : Πn, A n →* B) (p : Πn, g (n+1) ∘* f ~* g n) : pseq_colim @f →* B :=
pseq_colim.elim' g p
definition prep0 [constructor] {A : → Type*} (f : pseq_diagram A) (k : ) : A 0 →* A k := definition prep0 [constructor] {A : → Type*} (f : pseq_diagram A) (k : ) : A 0 →* A k :=
pmap.mk (rep0 (λn x, f x) k) pmap.mk (rep0 (λn x, f x) k)
begin induction k with k p, reflexivity, exact ap (@f k) p ⬝ !respect_pt end begin induction k with k p, reflexivity, exact ap (@f k) p ⬝ !respect_pt end

View file

@ -1,3 +1,11 @@
/-
Copyright (c) 2017 Yuri Sulyma, Favonia
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Yuri Sulyma, Favonia
Reduced homology theories
-/
import ..homotopy.spectrum ..homotopy.EM ..algebra.arrow_group ..algebra.direct_sum ..homotopy.fwedge ..choice ..homotopy.pushout ..move_to_lib import ..homotopy.spectrum ..homotopy.EM ..algebra.arrow_group ..algebra.direct_sum ..homotopy.fwedge ..choice ..homotopy.pushout ..move_to_lib
open eq spectrum int pointed group algebra sphere nat equiv susp is_trunc open eq spectrum int pointed group algebra sphere nat equiv susp is_trunc
@ -10,36 +18,39 @@ namespace homology
structure homology_theory.{u} : Type.{u+1} := structure homology_theory.{u} : Type.{u+1} :=
(HH : → pType.{u} → AbGroup.{u}) (HH : → pType.{u} → AbGroup.{u})
(Hh : Π(n : ) {X Y : Type*} (f : X →* Y), HH n X →g HH n Y) (Hh : Π(n : ) {X Y : Type*} (f : X →* Y), HH n X →g HH n Y)
(Hid : Π(n : ) {X : Type*} (x : HH n X), Hh n (pid X) x = x) (Hpid : Π(n : ) {X : Type*} (x : HH n X), Hh n (pid X) x = x)
(Hcompose : Π(n : ) {X Y Z : Type*} (f : Y →* Z) (g : X →* Y), (Hpcompose : Π(n : ) {X Y Z : Type*} (f : Y →* Z) (g : X →* Y),
Hh n (f ∘* g) ~ Hh n f ∘ Hh n g) Hh n (f ∘* g) ~ Hh n f ∘ Hh n g)
(Hsusp : Π(n : ) (X : Type*), HH (succ n) (psusp X) ≃g HH n X) (Hpsusp : Π(n : ) (X : Type*), HH (succ n) (psusp X) ≃g HH n X)
(Hsusp_natural : Π(n : ) {X Y : Type*} (f : X →* Y), (Hpsusp_natural : Π(n : ) {X Y : Type*} (f : X →* Y),
Hsusp n Y ∘ Hh (succ n) (psusp_functor f) ~ Hh n f ∘ Hsusp n X) Hpsusp n Y ∘ Hh (succ n) (psusp_functor f) ~ Hh n f ∘ Hpsusp n X)
(Hexact : Π(n : ) {X Y : Type*} (f : X →* Y), is_exact_g (Hh n f) (Hh n (pcod f))) (Hexact : Π(n : ) {X Y : Type*} (f : X →* Y), is_exact_g (Hh n f) (Hh n (pcod f)))
(Hadditive : Π(n : ) {I : Set.{u}} (X : I → Type*), is_equiv (Hadditive : Π(n : ) {I : Set.{u}} (X : I → Type*), is_equiv
(dirsum_elim (λi, Hh n (pinl i)) : dirsum (λi, HH n (X i)) → HH n ( X))) (dirsum_elim (λi, Hh n (pinl i)) : dirsum (λi, HH n (X i)) → HH n ( X)))
structure ordinary_homology_theory.{u} extends homology_theory.{u} : Type.{u+1} :=
(Hdimension : Π(n : ), n ≠ 0 → is_contr (HH n (plift (psphere 0))))
section section
parameter (theory : homology_theory) parameter (theory : homology_theory)
open homology_theory open homology_theory
theorem HH_base_indep (n : ) {A : Type} (a b : A) theorem HH_base_indep (n : ) {A : Type} (a b : A)
: HH theory n (pType.mk A a) ≃g HH theory n (pType.mk A b) := : HH theory n (pType.mk A a) ≃g HH theory n (pType.mk A b) :=
calc HH theory n (pType.mk A a) ≃g HH theory (int.succ n) (psusp A) : by exact (Hsusp theory n (pType.mk A a)) ⁻¹ᵍ calc HH theory n (pType.mk A a) ≃g HH theory (int.succ n) (psusp A) : by exact (Hpsusp theory n (pType.mk A a)) ⁻¹ᵍ
... ≃g HH theory n (pType.mk A b) : by exact Hsusp theory n (pType.mk A b) ... ≃g HH theory n (pType.mk A b) : by exact Hpsusp theory n (pType.mk A b)
theorem Hh_homotopy' (n : ) {A B : Type*} (f : A → B) (p q : f pt = pt) theorem Hh_homotopy' (n : ) {A B : Type*} (f : A → B) (p q : f pt = pt)
: Hh theory n (pmap.mk f p) ~ Hh theory n (pmap.mk f q) := λ x, : Hh theory n (pmap.mk f p) ~ Hh theory n (pmap.mk f q) := λ x,
calc Hh theory n (pmap.mk f p) x calc Hh theory n (pmap.mk f p) x
= Hh theory n (pmap.mk f p) (Hsusp theory n A ((Hsusp theory n A)⁻¹ᵍ x)) = Hh theory n (pmap.mk f p) (Hpsusp theory n A ((Hpsusp theory n A)⁻¹ᵍ x))
: by exact ap (Hh theory n (pmap.mk f p)) (equiv.to_right_inv (equiv_of_isomorphism (Hsusp theory n A)) x)⁻¹ : by exact ap (Hh theory n (pmap.mk f p)) (equiv.to_right_inv (equiv_of_isomorphism (Hpsusp theory n A)) x)⁻¹
... = Hsusp theory n B (Hh theory (succ n) (pmap.mk (susp.functor f) !refl) ((Hsusp theory n A)⁻¹ x)) ... = Hpsusp theory n B (Hh theory (succ n) (pmap.mk (susp.functor f) !refl) ((Hpsusp theory n A)⁻¹ x))
: by exact (Hsusp_natural theory n (pmap.mk f p) ((Hsusp theory n A)⁻¹ᵍ x))⁻¹ : by exact (Hpsusp_natural theory n (pmap.mk f p) ((Hpsusp theory n A)⁻¹ᵍ x))⁻¹
... = Hh theory n (pmap.mk f q) (Hsusp theory n A ((Hsusp theory n A)⁻¹ x)) ... = Hh theory n (pmap.mk f q) (Hpsusp theory n A ((Hpsusp theory n A)⁻¹ x))
: by exact Hsusp_natural theory n (pmap.mk f q) ((Hsusp theory n A)⁻¹ x) : by exact Hpsusp_natural theory n (pmap.mk f q) ((Hpsusp theory n A)⁻¹ x)
... = Hh theory n (pmap.mk f q) x ... = Hh theory n (pmap.mk f q) x
: by exact ap (Hh theory n (pmap.mk f q)) (equiv.to_right_inv (equiv_of_isomorphism (Hsusp theory n A)) x) : by exact ap (Hh theory n (pmap.mk f q)) (equiv.to_right_inv (equiv_of_isomorphism (Hpsusp theory n A)) x)
theorem Hh_homotopy (n : ) {A B : Type*} (f g : A →* B) (h : f ~ g) : Hh theory n f ~ Hh theory n g := λ x, theorem Hh_homotopy (n : ) {A B : Type*} (f g : A →* B) (h : f ~ g) : Hh theory n f ~ Hh theory n g := λ x,
calc Hh theory n f x calc Hh theory n f x
@ -56,15 +67,15 @@ namespace homology
{ exact Hh theory n e⁻¹ᵉ* }, { exact Hh theory n e⁻¹ᵉ* },
{ intro x, exact calc { intro x, exact calc
Hh theory n e (Hh theory n e⁻¹ᵉ* x) Hh theory n e (Hh theory n e⁻¹ᵉ* x)
= Hh theory n (e ∘* e⁻¹ᵉ*) x : by exact (Hcompose theory n e e⁻¹ᵉ* x)⁻¹ = Hh theory n (e ∘* e⁻¹ᵉ*) x : by exact (Hpcompose theory n e e⁻¹ᵉ* x)⁻¹
... = Hh theory n !pid x : by exact Hh_homotopy n (e ∘* e⁻¹ᵉ*) !pid (to_right_inv e) x ... = Hh theory n !pid x : by exact Hh_homotopy n (e ∘* e⁻¹ᵉ*) !pid (to_right_inv e) x
... = x : by exact Hid theory n x ... = x : by exact Hpid theory n x
}, },
{ intro x, exact calc { intro x, exact calc
Hh theory n e⁻¹ᵉ* (Hh theory n e x) Hh theory n e⁻¹ᵉ* (Hh theory n e x)
= Hh theory n (e⁻¹ᵉ* ∘* e) x : by exact (Hcompose theory n e⁻¹ᵉ* e x)⁻¹ = Hh theory n (e⁻¹ᵉ* ∘* e) x : by exact (Hpcompose theory n e⁻¹ᵉ* e x)⁻¹
... = Hh theory n !pid x : by exact Hh_homotopy n (e⁻¹ᵉ* ∘* e) !pid (to_left_inv e) x ... = Hh theory n !pid x : by exact Hh_homotopy n (e⁻¹ᵉ* ∘* e) !pid (to_left_inv e) x
... = x : by exact Hid theory n x ... = x : by exact Hpid theory n x
} }
end end

View file

@ -1,6 +1,10 @@
-- Author: Kuen-Bang Hou (Favonia) /-
Copyright (c) 2017 Kuen-Bang Hou (Favonia).
Released under Apache 2.0 license as described in the file LICENSE.
Author: Kuen-Bang Hou (Favonia)
-/
import core
import .homology import .homology
open eq pointed group algebra circle sphere nat equiv susp open eq pointed group algebra circle sphere nat equiv susp
@ -13,7 +17,7 @@ section
open homology_theory open homology_theory
theorem Hsphere : Π(n : )(m : ), HH theory n (plift (psphere m)) ≃g HH theory (n - m) (plift (psphere 0)) := theorem Hpsphere : Π(n : )(m : ), HH theory n (plift (psphere m)) ≃g HH theory (n - m) (plift (psphere 0)) :=
begin begin
intros n m, revert n, induction m with m, intros n m, revert n, induction m with m,
{ exact λ n, isomorphism_ap (λ n, HH theory n (plift (psphere 0))) (sub_zero n)⁻¹ }, { exact λ n, isomorphism_ap (λ n, HH theory n (plift (psphere 0))) (sub_zero n)⁻¹ },
@ -22,7 +26,7 @@ section
≃g HH theory n (psusp (plift (psphere m))) : by exact HH_isomorphism theory n (plift_psusp (psphere m)) ≃g HH theory n (psusp (plift (psphere m))) : by exact HH_isomorphism theory n (plift_psusp (psphere m))
... ≃g HH theory (succ (pred n)) (psusp (plift (psphere m))) ... ≃g HH theory (succ (pred n)) (psusp (plift (psphere m)))
: by exact isomorphism_ap (λ n, HH theory n (psusp (plift (psphere m)))) (succ_pred n)⁻¹ : by exact isomorphism_ap (λ n, HH theory n (psusp (plift (psphere m)))) (succ_pred n)⁻¹
... ≃g HH theory (pred n) (plift (psphere m)) : by exact Hsusp theory (pred n) (plift (psphere m)) ... ≃g HH theory (pred n) (plift (psphere m)) : by exact Hpsusp theory (pred n) (plift (psphere m))
... ≃g HH theory (pred n - m) (plift (psphere 0)) : by exact v_0 (pred n) ... ≃g HH theory (pred n - m) (plift (psphere 0)) : by exact v_0 (pred n)
... ≃g HH theory (n - succ m) (plift (psphere 0)) ... ≃g HH theory (n - succ m) (plift (psphere 0))
: by exact isomorphism_ap (λ n, HH theory n (plift (psphere 0))) (sub_sub n 1 m ⬝ ap (λ m, n - m) (add.comm 1 m)) : by exact isomorphism_ap (λ n, HH theory n (plift (psphere 0))) (sub_sub n 1 m ⬝ ap (λ m, n - m) (add.comm 1 m))

View file

@ -189,7 +189,7 @@ structure cohomology_theory.{u} : Type.{u+1} :=
(Hadditive : Π(n : ) {I : Type.{u}} (X : I → Type*), has_choice 0 I → (Hadditive : Π(n : ) {I : Type.{u}} (X : I → Type*), has_choice 0 I →
is_equiv (Group_pi_intro (λi, Hh n (pinl i)) : HH n ( X) → Πᵍ i, HH n (X i))) is_equiv (Group_pi_intro (λi, Hh n (pinl i)) : HH n ( X) → Πᵍ i, HH n (X i)))
structure ordinary_theory.{u} extends cohomology_theory.{u} : Type.{u+1} := structure ordinary_cohomology_theory.{u} extends cohomology_theory.{u} : Type.{u+1} :=
(Hdimension : Π(n : ), n ≠ 0 → is_contr (HH n (plift pbool))) (Hdimension : Π(n : ), n ≠ 0 → is_contr (HH n (plift pbool)))
attribute cohomology_theory.HH [coercion] attribute cohomology_theory.HH [coercion]
@ -270,7 +270,7 @@ cohomology_theory.mk
-- print has_choice_lift -- print has_choice_lift
-- print equiv_lift -- print equiv_lift
-- print has_choice_equiv_closed -- print has_choice_equiv_closed
definition ordinary_theory_EM [constructor] (G : AbGroup) : ordinary_theory := definition ordinary_cohomology_theory_EM [constructor] (G : AbGroup) : ordinary_cohomology_theory :=
⦃ordinary_theory, cohomology_theory_spectrum (EM_spectrum G), Hdimension := EM_dimension G ⦄ ⦃ordinary_cohomology_theory, cohomology_theory_spectrum (EM_spectrum G), Hdimension := EM_dimension G ⦄
end cohomology end cohomology

View file

@ -373,13 +373,17 @@ namespace spectrum
definition spectrify_type_term {N : succ_str} (X : gen_prespectrum N) (n : N) (k : ) : Type* := definition spectrify_type_term {N : succ_str} (X : gen_prespectrum N) (n : N) (k : ) : Type* :=
Ω[k] (X (n +' k)) Ω[k] (X (n +' k))
definition spectrify_type_fun' {N : succ_str} (X : gen_prespectrum N) (k : ) (n : N) : definition spectrify_type_fun' {N : succ_str} (X : gen_prespectrum N) (n : N) (k : ) :
Ω[k] (X n) →* Ω[k+1] (X (S n)) := Ω[k] (X n) →* Ω[k+1] (X (S n)) :=
!loopn_succ_in⁻¹ᵉ* ∘* Ω→[k] (glue X n) !loopn_succ_in⁻¹ᵉ* ∘* Ω→[k] (glue X n)
definition spectrify_type_fun {N : succ_str} (X : gen_prespectrum N) (n : N) (k : ) : definition spectrify_type_fun {N : succ_str} (X : gen_prespectrum N) (n : N) (k : ) :
spectrify_type_term X n k →* spectrify_type_term X n (k+1) := spectrify_type_term X n k →* spectrify_type_term X n (k+1) :=
spectrify_type_fun' X k (n +' k) spectrify_type_fun' X (n +' k) k
definition spectrify_type_fun_zero {N : succ_str} (X : gen_prespectrum N) (n : N) :
spectrify_type_fun X n 0 ~* glue X n :=
!pid_pcompose
definition spectrify_type {N : succ_str} (X : gen_prespectrum N) (n : N) : Type* := definition spectrify_type {N : succ_str} (X : gen_prespectrum N) (n : N) : Type* :=
pseq_colim (spectrify_type_fun X n) pseq_colim (spectrify_type_fun X n)
@ -393,20 +397,28 @@ namespace spectrum
... ≡ Y n ... ≡ Y n
-/ -/
definition spectrify_type_fun'_succ {N : succ_str} (X : gen_prespectrum N) (n : N) (k : ) :
spectrify_type_fun' X n (succ k) ~* Ω→ (spectrify_type_fun' X n k) :=
begin
refine _ ⬝* !ap1_pcompose⁻¹*,
apply !pwhisker_right,
refine !to_pinv_pequiv_MK2
end
definition spectrify_pequiv {N : succ_str} (X : gen_prespectrum N) (n : N) : definition spectrify_pequiv {N : succ_str} (X : gen_prespectrum N) (n : N) :
spectrify_type X n ≃* Ω (spectrify_type X (S n)) := spectrify_type X n ≃* Ω (spectrify_type X (S n)) :=
begin begin
refine !pshift_equiv ⬝e* _, refine !pshift_equiv ⬝e* _,
transitivity pseq_colim (λk, spectrify_type_fun' X (succ k) (S n +' k)), transitivity pseq_colim (λk, spectrify_type_fun' X (S n +' k) (succ k)),
fapply pseq_colim_pequiv, fapply pseq_colim_pequiv,
{ intro n, apply loopn_pequiv_loopn, apply pequiv_ap X, apply succ_str.add_succ }, { intro n, apply loopn_pequiv_loopn, apply pequiv_ap X, apply succ_str.add_succ },
{ intro k, apply to_homotopy, { exact abstract begin intro k,
refine !passoc⁻¹* ⬝* _, refine pwhisker_right _ (loopn_succ_in_inv_natural (succ k) _) ⬝* _, refine !passoc⁻¹* ⬝* _, refine pwhisker_right _ (loopn_succ_in_inv_natural (succ k) _) ⬝* _,
refine !passoc ⬝* _ ⬝* !passoc⁻¹*, apply pwhisker_left, refine !passoc ⬝* _ ⬝* !passoc⁻¹*, apply pwhisker_left,
refine !apn_pcompose⁻¹* ⬝* _ ⬝* !apn_pcompose, apply apn_phomotopy, refine !apn_pcompose⁻¹* ⬝* _ ⬝* !apn_pcompose, apply apn_phomotopy,
exact !glue_ptransport⁻¹* }, exact !glue_ptransport⁻¹* end end },
refine _ ⬝e* !pseq_colim_loop⁻¹ᵉ*, refine _ ⬝e* !pseq_colim_loop⁻¹ᵉ*,
refine pseq_colim_equiv_constant (λn, !ap1_pcompose⁻¹*), exact pseq_colim_equiv_constant (λn, !spectrify_type_fun'_succ),
end end
definition spectrify [constructor] {N : succ_str} (X : gen_prespectrum N) : gen_spectrum N := definition spectrify [constructor] {N : succ_str} (X : gen_prespectrum N) : gen_spectrum N :=
@ -419,21 +431,31 @@ namespace spectrum
-- note: the forward map is (currently) not definitionally equal to gluen. Is that a problem? -- note: the forward map is (currently) not definitionally equal to gluen. Is that a problem?
definition equiv_gluen {N : succ_str} (X : gen_spectrum N) (n : N) (k : ) definition equiv_gluen {N : succ_str} (X : gen_spectrum N) (n : N) (k : )
: X n ≃* Ω[k] (X (n +' k)) := : X n ≃* Ω[k] (X (n +' k)) :=
by induction k with k f; reflexivity; exact f ⬝e* loopn_pequiv_loopn k (equiv_glue X (n +' k)) by induction k with k f; reflexivity; exact f ⬝e* (loopn_pequiv_loopn k (equiv_glue X (n +' k))
⬝e* !loopn_succ_in⁻¹ᵉ* ⬝e* !loopn_succ_in⁻¹ᵉ*)
definition equiv_gluen_inv_succ {N : succ_str} (X : gen_spectrum N) (n : N) (k : ) :
(equiv_gluen X n (k+1))⁻¹ᵉ* ~*
(equiv_gluen X n k)⁻¹ᵉ* ∘* Ω→[k] (equiv_glue X (n +' k))⁻¹ᵉ* ∘* !loopn_succ_in :=
begin
refine !trans_pinv ⬝* pwhisker_left _ _, refine !trans_pinv ⬝* _, refine !to_pinv_pequiv_MK2 ◾* !pinv_pinv
end
definition spectrify_map {N : succ_str} {X : gen_prespectrum N} : X →ₛ spectrify X := definition spectrify_map {N : succ_str} {X : gen_prespectrum N} : X →ₛ spectrify X :=
begin begin
fapply smap.mk, fapply smap.mk,
{ intro n, exact pinclusion _ 0 }, { intro n, exact pinclusion _ 0 },
{ intro n, apply phomotopy_of_psquare, refine !pid_pcompose⁻¹* ⬝ph* _, { intro n, apply phomotopy_of_psquare,
refine !pid_pcompose⁻¹* ⬝ph* _, refine !pid_pcompose⁻¹* ⬝ph* _,
--pshift_equiv_pinclusion (spectrify_type_fun X n) 0 refine !passoc ⬝* pwhisker_left _ (pshift_equiv_pinclusion (spectrify_type_fun X n) 0) ⬝* _,
refine _ ⬝v* _, refine !passoc⁻¹* ⬝* _,
rotate 1, exact pshift_equiv_pinclusion (spectrify_type_fun X n) 0, refine _ ◾* (spectrify_type_fun_zero X n ⬝* !pid_pcompose⁻¹*),
-- refine !passoc⁻¹* ⬝* pwhisker_left _ _ ⬝* _, refine !passoc ⬝* pwhisker_left _ !pseq_colim_pequiv_pinclusion ⬝* _,
exact sorry refine pwhisker_left _ (pwhisker_left _ (ap1_pid) ⬝* !pcompose_pid) ⬝* _,
} refine !passoc ⬝* pwhisker_left _ !seq_colim_equiv_constant_pinclusion ⬝* _,
apply pinv_left_phomotopy_of_phomotopy,
exact !pseq_colim_loop_pinclusion⁻¹*
}
end end
definition spectrify.elim {N : succ_str} {X : gen_prespectrum N} {Y : gen_spectrum N} definition spectrify.elim {N : succ_str} {X : gen_prespectrum N} {Y : gen_spectrum N}
@ -442,7 +464,9 @@ namespace spectrum
fapply smap.mk, fapply smap.mk,
{ intro n, fapply pseq_colim.elim, { intro n, fapply pseq_colim.elim,
{ intro k, refine !equiv_gluen⁻¹ᵉ* ∘* apn k (f (n +' k)) }, { intro k, refine !equiv_gluen⁻¹ᵉ* ∘* apn k (f (n +' k)) },
{ intro k, apply to_homotopy, exact sorry }}, { intro k, refine !passoc ⬝* pwhisker_right _ !equiv_gluen_inv_succ ⬝* _,
refine !passoc ⬝* _, apply pwhisker_left,
refine !passoc ⬝* _, exact sorry }},
{ intro n, exact sorry } { intro n, exact sorry }
end end

View file

@ -180,6 +180,13 @@ namespace group
definition isomorphism_ap {A : Type} (F : A → Group) {a b : A} (p : a = b) : F a ≃g F b := definition isomorphism_ap {A : Type} (F : A → Group) {a b : A} (p : a = b) : F a ≃g F b :=
isomorphism_of_eq (ap F p) isomorphism_of_eq (ap F p)
definition interchange (G : AbGroup) (a b c d : G) : (a * b) * (c * d) = (a * c) * (b * d) :=
calc (a * b) * (c * d) = a * (b * (c * d)) : by exact mul.assoc a b (c * d)
... = a * ((b * c) * d) : by exact ap (λ bcd, a * bcd) (mul.assoc b c d)⁻¹
... = a * ((c * b) * d) : by exact ap (λ bc, a * (bc * d)) (mul.comm b c)
... = a * (c * (b * d)) : by exact ap (λ bcd, a * bcd) (mul.assoc c b d)
... = (a * c) * (b * d) : by exact (mul.assoc a c (b * d))⁻¹
end group open group end group open group
namespace function namespace function