Prove basic properties of spectral sequences

Also separate exact_couple and spectral_sequence in separate files
This commit is contained in:
Floris van Doorn 2018-10-01 17:44:49 -04:00
parent 4481935a83
commit 179575794a
9 changed files with 1316 additions and 1213 deletions

View file

@ -3,7 +3,7 @@
Formalization project of the CMU HoTT group to formalize the Serre spectral sequence.
*Update July 16*: The construction of the Serre spectral sequence has been completed. The result is `serre_convergence` in `cohomology.serre`.
The main algebra part is in `algebra.spectral_sequence`.
The main algebra part is in `algebra.exact_couple`.
This repository also contains the contents of the MRC group on formalizing homology in Lean.

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,288 @@
/-
Copyright (c) 2016 Egbert Rijke. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Egbert Rijke, Steve Awodey
Exact couple, derived couples, and so on
-/
/-
import algebra.group_theory hit.set_quotient types.sigma types.list types.sum .quotient_group .subgroup .ses
open eq algebra is_trunc set_quotient relation sigma sigma.ops prod prod.ops sum list trunc function group trunc
equiv is_equiv
-- This definition needs to be moved to exactness.hlean. However we had trouble doing so. Please help.
definition iso_ker_im_of_exact {A B C : AbGroup} (f : A →g B) (g : B →g C) (E : is_exact f g) : ab_Kernel g ≃g ab_image f :=
begin
fapply ab_subgroup_iso,
intro a,
induction E,
exact ker_in_im a,
intro a b, induction b with q, induction q with b p, induction p,
induction E,
exact im_in_ker b,
end
definition is_differential {B : AbGroup} (d : B →g B) := Π(b:B), d (d b) = 1
definition image_subgroup_of_diff {B : AbGroup} (d : B →g B) (H : is_differential d) : subgroup_rel (ab_kernel d) :=
subgroup_rel_of_subgroup (image_subgroup d) (kernel_subgroup d)
begin
intro g p,
induction p with f, induction f with h p,
rewrite [p⁻¹],
esimp,
exact H h
end
definition diff_im_in_ker {B : AbGroup} (d : B →g B) (H : is_differential d) : Π(b : B), image_subgroup d b → kernel_subgroup d b :=
begin
intro b p,
induction p with q, induction q with b' p, induction p, exact H b'
end
definition homology {B : AbGroup} (d : B →g B) (H : is_differential d) : AbGroup :=
@quotient_ab_group (ab_kernel d) (image_subgroup_of_diff d H)
definition homology_ugly {B : AbGroup} (d : B →g B) (H : is_differential d) : AbGroup :=
@quotient_ab_group (ab_kernel d) (image_subgroup (ab_subgroup_of_subgroup_incl (diff_im_in_ker d H)))
definition homology_iso_ugly {B : AbGroup} (d : B →g B) (H : is_differential d) : (homology d H) ≃g (homology_ugly d H) :=
begin
fapply @iso_of_ab_qg_group (ab_kernel d),
intro a,
intro p, induction p with f, induction f with b p,
fapply tr, fapply fiber.mk, fapply sigma.mk, exact d b, fapply tr, fapply fiber.mk, exact b, reflexivity,
induction a with c q, fapply subtype_eq, refine p ⬝ _, reflexivity,
intro b p, induction p with f, induction f with c p, induction p,
induction c with a q, induction q with f, induction f with a' p, induction p,
fapply tr, fapply fiber.mk, exact a', reflexivity
end
definition SES_iso_C {A B C C' : AbGroup} (ses : SES A B C) (k : C ≃g C') : SES A B C' :=
begin
fapply SES.mk,
exact SES.f ses,
exact k ∘g SES.g ses,
exact SES.Hf ses,
fapply @is_surjective_compose _ _ _ k (SES.g ses),
exact is_surjective_of_is_equiv k,
exact SES.Hg ses,
fapply is_exact.mk,
intro a,
esimp,
note h := SES.ex ses,
note h2 := is_exact.im_in_ker h a,
refine ap k h2 ⬝ _ ,
exact to_respect_one k,
intro b,
intro k3,
note h := SES.ex ses,
note h3 := is_exact.ker_in_im h b,
fapply is_exact.ker_in_im h,
refine _ ⬝ ap k⁻¹ᵍ k3 ⬝ _ ,
esimp,
exact (to_left_inv (equiv_of_isomorphism k) ((SES.g ses) b))⁻¹,
exact to_respect_one k⁻¹ᵍ
end
definition SES_of_differential_ugly {B : AbGroup} (d : B →g B) (H : is_differential d) : SES (ab_image d) (ab_kernel d) (homology_ugly d H) :=
begin
exact SES_of_inclusion (ab_subgroup_of_subgroup_incl (diff_im_in_ker d H)) (is_embedding_ab_subgroup_of_subgroup_incl (diff_im_in_ker d H)),
end
definition SES_of_differential {B : AbGroup} (d : B →g B) (H : is_differential d) : SES (ab_image d) (ab_kernel d) (homology d H) :=
begin
fapply SES_iso_C,
fapply SES_of_inclusion (ab_subgroup_of_subgroup_incl (diff_im_in_ker d H)) (is_embedding_ab_subgroup_of_subgroup_incl (diff_im_in_ker d H)),
exact (homology_iso_ugly d H)⁻¹ᵍ
end
structure exact_couple (A B : AbGroup) : Type :=
( i : A →g A) (j : A →g B) (k : B →g A)
( exact_ij : is_exact i j)
( exact_jk : is_exact j k)
( exact_ki : is_exact k i)
definition differential {A B : AbGroup} (EC : exact_couple A B) : B →g B :=
(exact_couple.j EC) ∘g (exact_couple.k EC)
definition differential_is_differential {A B : AbGroup} (EC : exact_couple A B) : is_differential (differential EC) :=
begin
induction EC,
induction exact_jk,
intro b,
exact (ap (group_fun j) (im_in_ker (group_fun k b))) ⬝ (respect_one j)
end
section derived_couple
/-
A - i -> A
k ^ |
| v j
B ====== B
-/
parameters {A B : AbGroup} (EC : exact_couple A B)
local abbreviation i := exact_couple.i EC
local abbreviation j := exact_couple.j EC
local abbreviation k := exact_couple.k EC
local abbreviation d := differential EC
local abbreviation H := differential_is_differential EC
-- local abbreviation u := exact_couple.i (SES_of_differential d H)
definition derived_couple_A : AbGroup :=
ab_subgroup (image_subgroup i)
definition derived_couple_B : AbGroup :=
homology (differential EC) (differential_is_differential EC)
definition derived_couple_i : derived_couple_A →g derived_couple_A :=
(image_lift (exact_couple.i EC)) ∘g (image_incl (exact_couple.i EC))
definition SES_of_exact_couple_at_i : SES (ab_kernel i) A (ab_image i) :=
begin
fapply SES_iso_C,
fapply SES_of_subgroup (kernel_subgroup i),
fapply ab_group_first_iso_thm i,
end
definition kj_zero (a : A) : k (j a) = 1 :=
is_exact.im_in_ker (exact_couple.exact_jk EC) a
definition j_factor : A →g (ab_kernel d) :=
begin
fapply ab_hom_lift j,
intro a,
unfold kernel_subgroup,
exact calc
d (j a) = j (k (j a)) : rfl
... = j 1 : by exact ap j (kj_zero a)
... = 1 : to_respect_one,
end
definition subgroup_iso_exact_at_A : ab_kernel i ≃g ab_image k :=
begin
fapply ab_subgroup_iso,
intro a,
induction EC,
induction exact_ki,
exact ker_in_im a,
intro a b, induction b with f, induction f with b p, induction p,
induction EC,
induction exact_ki,
exact im_in_ker b,
end
definition subgroup_iso_exact_at_A_triangle : ab_kernel_incl i ~ ab_image_incl k ∘g subgroup_iso_exact_at_A :=
begin
fapply ab_subgroup_iso_triangle,
intro a b, induction b with f, induction f with b p, induction p,
induction EC, induction exact_ki, exact im_in_ker b,
end
definition subgroup_homom_ker_to_im : ab_kernel i →g ab_image d :=
(image_homomorphism k j) ∘g subgroup_iso_exact_at_A
open eq
definition left_square_derived_ses_aux : j_factor ∘g ab_image_incl k ~ (SES.f (SES_of_differential d H)) ∘g (image_homomorphism k j) :=
begin
intro x,
induction x with a p, induction p with f, induction f with b p, induction p,
fapply subtype_eq,
reflexivity,
end
definition left_square_derived_ses : j_factor ∘g (ab_kernel_incl i) ~ (SES.f (SES_of_differential d H)) ∘g subgroup_homom_ker_to_im :=
begin
intro x,
exact (ap j_factor (subgroup_iso_exact_at_A_triangle x)) ⬝ (left_square_derived_ses_aux (subgroup_iso_exact_at_A x)),
end
definition derived_couple_j_unique :
is_contr (Σ hC, group_fun (hC ∘g SES.g SES_of_exact_couple_at_i) ~ group_fun
(SES.g (SES_of_differential d H) ∘g j_factor)) :=
quotient_extend_unique_SES (SES_of_exact_couple_at_i) (SES_of_differential d H) (subgroup_homom_ker_to_im) (j_factor) (left_square_derived_ses)
definition derived_couple_j : derived_couple_A →g derived_couple_B :=
begin
exact pr1 (center' (derived_couple_j_unique)),
end
definition derived_couple_j_htpy : group_fun (derived_couple_j ∘g SES.g SES_of_exact_couple_at_i) ~ group_fun
(SES.g (SES_of_differential d H) ∘g j_factor) :=
begin
exact pr2 (center' (derived_couple_j_unique)),
end
definition SES_im_i_trivial : SES trivial_ab_group derived_couple_A derived_couple_A :=
begin
fapply SES_of_isomorphism_right,
fapply isomorphism.refl,
end
definition subgroup_iso_exact_kerj_imi : ab_kernel j ≃g ab_image i :=
begin
fapply iso_ker_im_of_exact,
induction EC,
exact exact_ij,
end
definition k_restrict_aux : ab_kernel d →g ab_kernel j :=
begin
fapply ab_hom_lift_kernel,
exact k ∘g ab_kernel_incl d,
intro p, induction p with b p, exact p,
end
definition k_restrict : ab_kernel d →g derived_couple_A :=
subgroup_iso_exact_kerj_imi ∘g k_restrict_aux
definition k_restrict_square_left : k_restrict ∘g (SES.f (SES_of_differential d H)) ~ λ x, 1 :=
begin
intro x,
induction x with b' p,
induction p with q,
induction q with b p,
induction p,
fapply subtype_eq,
induction EC,
induction exact_jk,
fapply im_in_ker,
end
definition derived_couple_k_unique : is_contr
(Σ hC, group_fun (hC ∘g SES.g (SES_of_differential d H)) ~ group_fun
(SES.g SES_im_i_trivial ∘g k_restrict))
:=
quotient_extend_unique_SES (SES_of_differential d H) (SES_im_i_trivial) (trivial_homomorphism (ab_image d) trivial_ab_group) (k_restrict) (k_restrict_square_left)
definition derived_couple_k : derived_couple_B →g derived_couple_A :=
begin
exact pr1 (center' (derived_couple_k_unique)),
end
definition derived_couple_k_htpy : group_fun (derived_couple_k ∘g SES.g (SES_of_differential d H)) ~ group_fun
(SES.g (SES_im_i_trivial) ∘g k_restrict) :=
begin
exact pr2 (center' (derived_couple_k_unique)),
end
definition derived_couple_exact_ij : is_exact_ag derived_couple_i derived_couple_j :=
begin
fapply is_exact.mk,
intro a,
induction a with a' t,
induction t with q, induction q with a p, induction p,
repeat exact sorry,
end
end derived_couple
-/

View file

@ -473,6 +473,63 @@ end
end
/- we say that an left module D is built from the sequence E if
D is a "twisted sum" of the E, and E has only finitely many nontrivial values -/
open nat
structure is_built_from.{u v w} {R : Ring} (D : LeftModule.{u v} R)
(E : → LeftModule.{u w} R) : Type.{max u (v+1) w} :=
(part : → LeftModule.{u v} R)
(ses : Πn, short_exact_mod (E n) (part n) (part (n+1)))
(e0 : part 0 ≃lm D)
(n₀ : )
(HD' : Π(s : ), n₀ ≤ s → is_contr (part s))
open is_built_from
universe variables u v w
variables {R : Ring.{u}} {D D' : LeftModule.{u v} R} {E E' : → LeftModule.{u w} R}
definition is_built_from_shift (H : is_built_from D E) :
is_built_from (part H 1) (λn, E (n+1)) :=
is_built_from.mk (λn, part H (n+1)) (λn, ses H (n+1)) isomorphism.rfl (pred (n₀ H))
(λs Hle, HD' H _ (le_succ_of_pred_le Hle))
definition isomorphism_of_is_contr_part (H : is_built_from D E) (n : ) (HE : is_contr (E n)) :
part H n ≃lm part H (n+1) :=
isomorphism_of_is_contr_left (ses H n) HE
definition is_contr_submodules (H : is_built_from D E) (HD : is_contr D) (n : ) :
is_contr (part H n) :=
begin
induction n with n IH,
{ exact is_trunc_equiv_closed_rev -2 (equiv_of_isomorphism (e0 H)) _ },
{ exact is_contr_right_of_short_exact_mod (ses H n) IH }
end
definition is_contr_quotients (H : is_built_from D E) (HD : is_contr D) (n : ) :
is_contr (E n) :=
begin
apply is_contr_left_of_short_exact_mod (ses H n),
exact is_contr_submodules H HD n
end
definition is_contr_total (H : is_built_from D E) (HE : Πn, is_contr (E n)) : is_contr D :=
have is_contr (part H 0), from
nat.rec_down (λn, is_contr (part H n)) _ (HD' H _ !le.refl)
(λn H2, is_contr_middle_of_short_exact_mod (ses H n) (HE n) H2),
is_contr_equiv_closed (equiv_of_isomorphism (e0 H)) _
definition is_built_from_isomorphism (e : D ≃lm D') (f : Πn, E n ≃lm E' n)
(H : is_built_from D E) : is_built_from D' E' :=
⦃is_built_from, H, e0 := e0 H ⬝lm e,
ses := λn, short_exact_mod_isomorphism (f n)⁻¹ˡᵐ isomorphism.rfl isomorphism.rfl (ses H n)⦄
definition is_built_from_isomorphism_left (e : D ≃lm D') (H : is_built_from D E) :
is_built_from D' E :=
⦃is_built_from, H, e0 := e0 H ⬝lm e⦄
definition isomorphism_zero_of_is_built_from (H : is_built_from D E) (p : n₀ H = 1) : E 0 ≃lm D :=
isomorphism_of_is_contr_right (ses H 0) (HD' H 1 (le_of_eq p)) ⬝lm e0 H
section int
open int
definition left_module_int_of_ab_group [constructor] (A : Type) [add_ab_group A] : left_module r A :=

File diff suppressed because it is too large Load diff

View file

@ -160,11 +160,11 @@ begin
apply h (of_mem_property_of Sm)
end
definition is_prop_submodule (S : property M) [is_submodule M S] [H : is_prop M] : is_prop (submodule S) :=
begin apply @is_trunc_sigma, exact H end
local attribute is_prop_submodule [instance]
definition is_contr_submodule [instance] (S : property M) [is_submodule M S] [is_contr M] : is_contr (submodule S) :=
is_contr_of_inhabited_prop 0 _
definition is_contr_submodule (S : property M) [is_submodule M S] (H : is_contr M) :
is_contr (submodule S) :=
have is_prop M, from _,
have is_prop (submodule S), from @is_trunc_sigma _ _ _ this _,
is_contr_of_inhabited_prop 0 this
definition submodule_isomorphism [constructor] (S : property M) [is_submodule M S] (h : Πg, g ∈ S) :
submodule S ≃lm M :=
@ -240,11 +240,12 @@ lm_homomorphism_of_group_homomorphism
definition is_prop_quotient_module (S : property M) [is_submodule M S] [H : is_prop M] : is_prop (quotient_module S) :=
begin apply @set_quotient.is_trunc_set_quotient, exact H end
local attribute is_prop_quotient_module [instance]
definition is_contr_quotient_module [instance] (S : property M) [is_submodule M S] [is_contr M] :
is_contr (quotient_module S) :=
is_contr_of_inhabited_prop 0 _
definition is_contr_quotient_module [instance] (S : property M) [is_submodule M S]
(H : is_contr M) : is_contr (quotient_module S) :=
have is_prop M, from _,
have is_prop (quotient_module S), from @set_quotient.is_trunc_set_quotient _ _ _ this,
is_contr_of_inhabited_prop 0 this
definition rel_of_is_contr_quotient_module (S : property M) [is_submodule M S]
(H : is_contr (quotient_module S)) (m : M) : S m :=
@ -331,11 +332,11 @@ end
-- reflexivity
-- end
definition is_contr_image_module [instance] (φ : M₁ →lm M₂) [is_contr M₂] :
definition is_contr_image_module [instance] (φ : M₁ →lm M₂) (H : is_contr M₂) :
is_contr (image_module φ) :=
!is_contr_submodule
is_contr_submodule _ _
definition is_contr_image_module_of_is_contr_dom (φ : M₁ →lm M₂) [is_contrM₁ : is_contr M₁] :
definition is_contr_image_module_of_is_contr_dom (φ : M₁ →lm M₂) (H : is_contr M₁) :
is_contr (image_module φ) :=
is_contr.mk 0
begin
@ -343,7 +344,7 @@ is_contr.mk 0
apply @total_image.rec,
exact this,
intro m,
have h : is_contr (LeftModule.carrier M₁), from is_contrM₁,
have h : is_contr (LeftModule.carrier M₁), from H,
induction (eq_of_is_contr 0 m), apply subtype_eq,
exact (to_respect_zero φ)⁻¹
end
@ -369,19 +370,21 @@ is_submodule.mk
begin apply @is_subgroup.mul_mem, apply is_subgroup_kernel, exact hg₁, exact hg₂ end)
(has_scalar_kernel φ)
definition kernel_full (φ : M₁ →lm M₂) [is_contr M₂] (m : M₁) : m ∈ lm_kernel φ :=
definition kernel_full (φ : M₁ →lm M₂) (H : is_contr M₂) (m : M₁) : m ∈ lm_kernel φ :=
!is_prop.elim
definition kernel_module [reducible] (φ : M₁ →lm M₂) : LeftModule R := submodule (lm_kernel φ)
definition is_contr_kernel_module [instance] (φ : M₁ →lm M₂) [is_contr M₁] :
definition is_contr_kernel_module [instance] (φ : M₁ →lm M₂) (H : is_contr M₁) :
is_contr (kernel_module φ) :=
!is_contr_submodule
is_contr_submodule _ _
definition kernel_module_isomorphism [constructor] (φ : M₁ →lm M₂) [is_contr M₂] : kernel_module φ ≃lm M₁ :=
submodule_isomorphism _ (kernel_full φ)
definition kernel_module_isomorphism [constructor] (φ : M₁ →lm M₂) (H : is_contr M₂) :
kernel_module φ ≃lm M₁ :=
submodule_isomorphism _ (kernel_full φ _)
definition homology_quotient_property (ψ : M₂ →lm M₃) (φ : M₁ →lm M₂) : property (kernel_module ψ) :=
definition homology_quotient_property (ψ : M₂ →lm M₃) (φ : M₁ →lm M₂) :
property (kernel_module ψ) :=
property_submodule (lm_kernel ψ) (image (homomorphism_fn φ))
definition is_submodule_homology_property [instance] (ψ : M₂ →lm M₃) (φ : M₁ →lm M₂) :
@ -417,14 +420,14 @@ quotient_elim (θ ∘lm submodule_incl _)
exact ap θ p⁻¹ ⬝ H v -- m'
end
definition is_contr_homology [instance] (ψ : M₂ →lm M₃) (φ : M₁ →lm M₂) [is_contr M₂] :
definition is_contr_homology [instance] (ψ : M₂ →lm M₃) (φ : M₁ →lm M₂) (H : is_contr M₂) :
is_contr (homology ψ φ) :=
begin apply @is_contr_quotient_module end
is_contr_quotient_module _ (is_contr_kernel_module _ _)
definition homology_isomorphism [constructor] (ψ : M₂ →lm M₃) (φ : M₁ →lm M₂)
[is_contr M₁] [is_contr M₃] : homology ψ φ ≃lm M₂ :=
(H₁ : is_contr M₁) (H₃ : is_contr M₃) : homology ψ φ ≃lm M₂ :=
(quotient_module_isomorphism (homology_quotient_property ψ φ)
(eq_zero_of_mem_property_submodule_trivial (image_trivial _))) ⬝lm (kernel_module_isomorphism ψ)
(eq_zero_of_mem_property_submodule_trivial (image_trivial _))) ⬝lm (kernel_module_isomorphism ψ _)
definition ker_in_im_of_is_contr_homology (ψ : M₂ →lm M₃) {φ : M₁ →lm M₂}
(H₁ : is_contr (homology ψ φ)) {m : M₂} (p : ψ m = 0) : image φ m :=

View file

@ -5,7 +5,7 @@ Author: Floris van Doorn-/
import .serre
open eq spectrum EM EM.ops int pointed cohomology left_module algebra group fiber is_equiv equiv
prod is_trunc function
prod is_trunc function exact_couple
namespace temp
@ -22,9 +22,9 @@ namespace temp
λx, pt
definition fserre :
(λp q, uoH^p[K ag 2, H^q[circle₊]]) ⟹ᵍ (λn, H^-n[unit₊]) :=
(λp q, uoH^p[K ag 2, H^q[circle₊]]) ⟹ᵍ (λn, H^n[unit₊]) :=
proof
converges_to_g_isomorphism
convergent_exact_couple_g_isomorphism
(serre_convergence_map_of_is_conn pt f (EM_spectrum ag) 0
(is_strunc_EM_spectrum ag) (is_conn_EM ag 2))
begin

View file

@ -126,6 +126,8 @@ pfiber_postnikov_map_pred' A n k _ idp ⬝e*
pequiv_ap (EM_spectrum (πₛ[n] A)) (add.comm n k)
qed
open exact_couple
section atiyah_hirzebruch
parameters {X : Type*} (Y : X → spectrum) (s₀ : ) (H : Πx, is_strunc s₀ (Y x))
include H
@ -249,14 +251,14 @@ convergent_exact_couple_g_isomorphism
intro n, reflexivity
end
definition unreduced_atiyah_hirzebruch_spectral_sequence {X : Type} (Y : X → spectrum) (s₀ : )
(H : Πx, is_strunc s₀ (Y x)) :
convergent_spectral_sequence_g (λp q, uopH^p[(x : X), πₛ[-q] (Y x)]) (λn, upH^n[(x : X), Y x]) :=
begin
apply convergent_spectral_sequence_of_exact_couple (unreduced_atiyah_hirzebruch_convergence Y s₀ H),
{ intro n, exact add.comm (s₀ - -n) (-s₀) ⬝ !neg_add_cancel_left ⬝ !neg_neg },
{ reflexivity }
end
definition unreduced_atiyah_hirzebruch_spectral_sequence {X : Type} (Y : X → spectrum) (s₀ : )
(H : Πx, is_strunc s₀ (Y x)) :
convergent_spectral_sequence_g (λp q, uopH^p[(x : X), πₛ[-q] (Y x)]) (λn, upH^n[(x : X), Y x]) :=
begin
apply convergent_spectral_sequence_of_exact_couple (unreduced_atiyah_hirzebruch_convergence Y s₀ H),
{ intro n, exact add.comm (s₀ - -n) (-s₀) ⬝ !neg_add_cancel_left ⬝ !neg_neg },
{ reflexivity }
end
end unreduced_atiyah_hirzebruch

View file

@ -28,7 +28,7 @@ namespace pushout
open three_by_three_span
variable (E : three_by_three_span)
check (pushout.functor (f₂₁ E) (f₀₁ E) (f₄₁ E) (s₁₁ E) (s₃₁ E))
-- check (pushout.functor (f₂₁ E) (f₀₁ E) (f₄₁ E) (s₁₁ E) (s₃₁ E))
definition pushout2hv (E : three_by_three_span) : Type :=
pushout (pushout.functor (f₂₁ E) (f₀₁ E) (f₄₁ E) (s₁₁ E) (s₃₁ E))
(pushout.functor (f₂₃ E) (f₀₃ E) (f₄₃ E) (s₁₃ E) (s₃₃ E))