Add the missing 'p'.
This commit is contained in:
parent
a8592e0184
commit
5cda45ae1d
2 changed files with 20 additions and 21 deletions
|
@ -18,12 +18,12 @@ namespace homology
|
|||
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*} (f : Y →* Z) (g : X →* Y),
|
||||
(Hpid : Π(n : ℤ) {X : Type*} (x : HH n X), Hh n (pid X) x = x)
|
||||
(Hpcompose : Π(n : ℤ) {X Y Z : Type*} (f : Y →* Z) (g : X →* Y),
|
||||
Hh n (f ∘* g) ~ Hh n f ∘ Hh n g)
|
||||
(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)
|
||||
(Hpsusp : Π(n : ℤ) (X : Type*), HH (succ n) (psusp X) ≃g HH n X)
|
||||
(Hpsusp_natural : Π(n : ℤ) {X Y : Type*} (f : X →* Y),
|
||||
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)))
|
||||
(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)))
|
||||
|
@ -37,20 +37,20 @@ namespace homology
|
|||
|
||||
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) :=
|
||||
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)
|
||||
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 Hpsusp theory n (pType.mk A b)
|
||||
|
||||
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,
|
||||
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))
|
||||
: by exact ap (Hh theory n (pmap.mk f p)) (equiv.to_right_inv (equiv_of_isomorphism (Hsusp theory n A)) x)⁻¹
|
||||
... = Hsusp theory n B (Hh theory (succ n) (pmap.mk (susp.functor f) !refl) ((Hsusp theory n A)⁻¹ x))
|
||||
: by exact (Hsusp_natural theory n (pmap.mk f p) ((Hsusp theory n A)⁻¹ᵍ x))⁻¹
|
||||
... = Hh theory n (pmap.mk f q) (Hsusp theory n A ((Hsusp theory n A)⁻¹ x))
|
||||
: by exact Hsusp_natural theory n (pmap.mk f q) ((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 (Hpsusp 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 (Hpsusp_natural theory n (pmap.mk f p) ((Hpsusp theory n A)⁻¹ᵍ x))⁻¹
|
||||
... = Hh theory n (pmap.mk f q) (Hpsusp theory n A ((Hpsusp 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
|
||||
: 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,
|
||||
calc Hh theory n f x
|
||||
|
@ -67,15 +67,15 @@ namespace homology
|
|||
{ exact Hh theory n e⁻¹ᵉ* },
|
||||
{ intro x, exact calc
|
||||
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
|
||||
... = x : by exact Hid theory n x
|
||||
... = x : by exact Hpid theory n x
|
||||
},
|
||||
{ intro x, exact calc
|
||||
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
|
||||
... = x : by exact Hid theory n x
|
||||
... = x : by exact Hpid theory n x
|
||||
}
|
||||
end
|
||||
|
||||
|
|
|
@ -5,7 +5,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Author: Kuen-Bang Hou (Favonia)
|
||||
-/
|
||||
|
||||
import core
|
||||
import .homology
|
||||
|
||||
open eq pointed group algebra circle sphere nat equiv susp
|
||||
|
@ -18,7 +17,7 @@ section
|
|||
|
||||
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
|
||||
intros n m, revert n, induction m with m,
|
||||
{ exact λ n, isomorphism_ap (λ n, HH theory n (plift (psphere 0))) (sub_zero n)⁻¹ },
|
||||
|
@ -27,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 (succ (pred n)) (psusp (plift (psphere m)))
|
||||
: 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 (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))
|
||||
|
|
Loading…
Add table
Reference in a new issue