Spectral/homology/torus.hlean
Floris van Doorn 3367c20f9d make pointed suspension and spheres the default
There is one proof in realprojective which I couldn't quite fix, so for now I left a sorry
2017-07-20 18:03:13 +01:00

40 lines
2.2 KiB
Text
Raw Permalink 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.

/-
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 .basic .sphere ..homotopy.susp_product
open eq pointed group algebra circle sphere nat equiv susp
function sphere homology int lift prod smash
namespace homology
section
parameter (theory : ordinary_homology_theory)
open ordinary_homology_theory
theorem Hptorus : Π(n : )(m : ), HH theory n (plift (sphere m ×* sphere m)) ≃g
HH theory n (plift (sphere m)) ×g (HH theory n (plift (sphere m)) ×g HH theory n (plift (sphere (m + m)))) := λ n m,
calc HH theory n (plift (sphere m ×* sphere m))
≃g HH theory (n + 1) (plift (⅀ (sphere m ×* sphere m))) : by exact (Hplift_susp theory n (sphere m ×* sphere m))⁻¹ᵍ
... ≃g HH theory (n + 1) (plift (⅀ (sphere m) (⅀ (sphere m) ⅀ (sphere m ∧ sphere m))))
: by exact Hplift_isomorphism theory (n + 1) (susp_product (sphere m) (sphere m))
... ≃g HH theory (n + 1) (plift (⅀ (sphere m))) ×g HH theory (n + 1) (plift (⅀ (sphere m) (⅀ (sphere m ∧ sphere m))))
: by exact Hplift_wedge theory (n + 1) (⅀ (sphere m)) (⅀ (sphere m) (⅀ (sphere m ∧ sphere m)))
... ≃g HH theory n (plift (sphere m)) ×g (HH theory n (plift (sphere m)) ×g HH theory n (plift (sphere (m + m))))
: by exact product_isomorphism (Hplift_susp theory n (sphere m))
(calc
HH theory (n + 1) (plift (⅀ (sphere m) (⅀ (sphere m ∧ sphere m))))
≃g HH theory (n + 1) (plift (⅀ (sphere m))) ×g HH theory (n + 1) (plift (⅀ (sphere m ∧ sphere m)))
: by exact Hplift_wedge theory (n + 1) (⅀ (sphere m)) (⅀ (sphere m ∧ sphere m))
... ≃g HH theory n (plift (sphere m)) ×g HH theory n (plift (sphere (m + m)))
: by exact product_isomorphism (Hplift_susp theory n (sphere m))
(Hplift_susp theory n (sphere m ∧ sphere m) ⬝g Hplift_isomorphism theory n (sphere_smash_sphere m m)))
end
end homology