Spectral/homology/homology.hlean
2017-06-06 14:29:41 -06:00

35 lines
1.6 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import ..homotopy.spectrum ..homotopy.EM ..algebra.arrow_group ..algebra.direct_sum ..homotopy.fwedge ..choice ..homotopy.pushout ..move_to_lib
open eq spectrum int trunc pointed EM group algebra circle sphere nat EM.ops equiv susp is_trunc
function fwedge cofiber bool lift sigma is_equiv choice pushout algebra unit pi smash
namespace homology
/- homology theory -/
structure homology_theory.{u} : Type.{u+1} :=
(HH : → pType.{u} → AbGroup.{u})
(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)
(Hcompose : Π(n : ) {X Y Z : Type*} (g : Y →* Z) (f : X →* Y) (x : HH n X),
Hh n (g ∘* f) x = Hh n g (Hh n f x))
(Hsusp : Π(n : ) (X : Type*), HH (succ n) (psusp X) ≃g HH n X)
(Hsusp_natural : Π(n : ) {X Y : Type*} (f : X →* Y),
Hsusp n Y ∘ Hh (succ n) (psusp_functor f) ~ Hh n f ∘ Hsusp n X)
(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 (
dirsum_elim (λi, Hh n (pinl i)) : dirsum (λi, HH n (X i)) → HH n ( X))
)
section
parameter (theory : homology_theory)
open homology_theory
definition HH_base_indep (n : ) {A : Type} (a b : A)
: 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)) ⁻¹ᵍ
... ≃g HH theory n (pType.mk A b) : by exact Hsusp theory n (pType.mk A b)
end
end homology