2017-03-30 22:33:33 +00:00
|
|
|
|
/- Graded (left-) R-modules for a ring R. -/
|
2016-03-24 18:24:47 +00:00
|
|
|
|
|
2017-03-30 19:02:24 +00:00
|
|
|
|
-- Author: Floris van Doorn
|
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
import .left_module .direct_sum .submodule --..heq
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
open algebra eq left_module pointed function equiv is_equiv is_trunc prod group sigma
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
|
|
|
|
namespace left_module
|
|
|
|
|
|
2017-03-31 22:21:02 +00:00
|
|
|
|
definition graded [reducible] (str : Type) (I : Type) : Type := I → str
|
|
|
|
|
definition graded_module [reducible] (R : Ring) : Type → Type := graded (LeftModule R)
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
variables {R : Ring} {I : Set} {M M₁ M₂ M₃ : graded_module R I}
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
/-
|
|
|
|
|
morphisms between graded modules.
|
|
|
|
|
The definition is unconventional in two ways:
|
|
|
|
|
(1) The degree is determined by an endofunction instead of a element of I (and in this case we
|
|
|
|
|
don't need to assume that I is a group). The "standard" degree i corresponds to the endofunction
|
|
|
|
|
which is addition with i on the right. However, this is more flexible. For example, the
|
|
|
|
|
composition of two graded module homomorphisms φ₂ and φ₁ with degrees i₂ and i₁ has type
|
|
|
|
|
M₁ i → M₂ ((i + i₁) + i₂).
|
|
|
|
|
However, a homomorphism with degree i₁ + i₂ must have type
|
|
|
|
|
M₁ i → M₂ (i + (i₁ + i₂)),
|
|
|
|
|
which means that we need to insert a transport. With endofunctions this is not a problem:
|
|
|
|
|
λi, (i + i₁) + i₂
|
|
|
|
|
is a perfectly fine degree of a map
|
|
|
|
|
(2) Since we cannot eliminate all possible transports, we don't define a homomorphism as function
|
|
|
|
|
M₁ i →lm M₂ (i + deg f) or M₁ i →lm M₂ (deg f i)
|
|
|
|
|
but as a function taking a path as argument. Specifically, for every path
|
|
|
|
|
deg f i = j
|
|
|
|
|
we get a function M₁ i → M₂ j.
|
2017-04-24 17:33:48 +00:00
|
|
|
|
(3) Note: we do assume that I is a set. This is not strictly necessary, but it simplifies things
|
2017-03-30 22:27:09 +00:00
|
|
|
|
-/
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
definition graded_hom_of_deg (d : I ≃ I) (M₁ M₂ : graded_module R I) : Type :=
|
|
|
|
|
Π⦃i j : I⦄ (p : d i = j), M₁ i →lm M₂ j
|
|
|
|
|
|
|
|
|
|
definition gmd_constant [constructor] (d : I ≃ I) (M₁ M₂ : graded_module R I) : graded_hom_of_deg d M₁ M₂ :=
|
|
|
|
|
λi j p, lm_constant (M₁ i) (M₂ j)
|
|
|
|
|
|
|
|
|
|
definition gmd0 [constructor] {d : I ≃ I} {M₁ M₂ : graded_module R I} : graded_hom_of_deg d M₁ M₂ :=
|
|
|
|
|
gmd_constant d M₁ M₂
|
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
structure graded_hom (M₁ M₂ : graded_module R I) : Type :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
mk' :: (d : I ≃ I)
|
|
|
|
|
(fn' : graded_hom_of_deg d M₁ M₂)
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
notation M₁ ` →gm ` M₂ := graded_hom M₁ M₂
|
|
|
|
|
|
|
|
|
|
abbreviation deg [unfold 5] := @graded_hom.d
|
2017-04-24 17:33:48 +00:00
|
|
|
|
postfix ` ↘`:max := graded_hom.fn' -- there is probably a better character for this? Maybe ↷?
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
definition graded_hom_fn [reducible] [unfold 5] [coercion] (f : M₁ →gm M₂) (i : I) : M₁ i →lm M₂ (deg f i) :=
|
|
|
|
|
f ↘ idp
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_hom_fn_out [reducible] [unfold 5] (f : M₁ →gm M₂) (i : I) : M₁ ((deg f)⁻¹ i) →lm M₂ i :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
f ↘ (to_right_inv (deg f) i)
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
infix ` ← `:101 := graded_hom_fn_out -- todo: change notation
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
definition graded_hom.mk [constructor] {M₁ M₂ : graded_module R I} (d : I ≃ I)
|
2017-03-30 22:27:09 +00:00
|
|
|
|
(fn : Πi, M₁ i →lm M₂ (d i)) : M₁ →gm M₂ :=
|
|
|
|
|
graded_hom.mk' d (λi j p, homomorphism_of_eq (ap M₂ p) ∘lm fn i)
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_hom.mk_out [constructor] {M₁ M₂ : graded_module R I} (d : I ≃ I)
|
2017-04-21 02:58:19 +00:00
|
|
|
|
(fn : Πi, M₁ (d⁻¹ i) →lm M₂ i) : M₁ →gm M₂ :=
|
|
|
|
|
graded_hom.mk' d (λi j p, fn j ∘lm homomorphism_of_eq (ap M₁ (eq_inv_of_eq p)))
|
|
|
|
|
|
|
|
|
|
definition graded_hom.mk_out_in [constructor] {M₁ M₂ : graded_module R I} (d₁ : I ≃ I) (d₂ : I ≃ I)
|
|
|
|
|
(fn : Πi, M₁ (d₁ i) →lm M₂ (d₂ i)) : M₁ →gm M₂ :=
|
|
|
|
|
graded_hom.mk' (d₁⁻¹ᵉ ⬝e d₂) (λi j p, homomorphism_of_eq (ap M₂ p) ∘lm fn (d₁⁻¹ᵉ i) ∘lm
|
|
|
|
|
homomorphism_of_eq (ap M₁ (to_right_inv d₁ i)⁻¹))
|
|
|
|
|
|
|
|
|
|
definition graded_hom_eq_transport (f : M₁ →gm M₂) {i j : I} (p : deg f i = j) (m : M₁ i) :
|
|
|
|
|
f ↘ p m = transport M₂ p (f i m) :=
|
|
|
|
|
by induction p; reflexivity
|
|
|
|
|
|
|
|
|
|
definition graded_hom_eq_zero {f : M₁ →gm M₂} {i j k : I} {q : deg f i = j} {p : deg f i = k}
|
|
|
|
|
(m : M₁ i) (r : f ↘ q m = 0) : f ↘ p m = 0 :=
|
|
|
|
|
have f ↘ p m = transport M₂ (q⁻¹ ⬝ p) (f ↘ q m), begin induction p, induction q, reflexivity end,
|
|
|
|
|
this ⬝ ap (transport M₂ (q⁻¹ ⬝ p)) r ⬝ tr_eq_of_pathover (apd (λi, 0) (q⁻¹ ⬝ p))
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
variables {f' : M₂ →gm M₃} {f g h : M₁ →gm M₂}
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
|
|
|
|
definition graded_hom_compose [constructor] (f' : M₂ →gm M₃) (f : M₁ →gm M₂) : M₁ →gm M₃ :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
graded_hom.mk (deg f ⬝e deg f') (λi, f' (deg f i) ∘lm f i)
|
|
|
|
|
|
|
|
|
|
infixr ` ∘gm `:75 := graded_hom_compose
|
|
|
|
|
|
|
|
|
|
definition graded_hom_compose_fn (f' : M₂ →gm M₃) (f : M₁ →gm M₂) (i : I) (m : M₁ i) :
|
|
|
|
|
(f' ∘gm f) i m = f' (deg f i) (f i m) :=
|
|
|
|
|
proof idp qed
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
variable (M)
|
|
|
|
|
definition graded_hom_id [constructor] [refl] : M →gm M :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
graded_hom.mk erfl (λi, lmid)
|
2017-03-30 22:27:09 +00:00
|
|
|
|
variable {M}
|
|
|
|
|
abbreviation gmid [constructor] := graded_hom_id M
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
definition gm_constant [constructor] (M₁ M₂ : graded_module R I) (d : I ≃ I) : M₁ →gm M₂ :=
|
|
|
|
|
graded_hom.mk' d (gmd_constant d M₁ M₂)
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
structure graded_iso (M₁ M₂ : graded_module R I) : Type :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
mk' :: (to_hom : M₁ →gm M₂)
|
|
|
|
|
(is_equiv_to_hom : Π⦃i j⦄ (p : deg to_hom i = j), is_equiv (to_hom ↘ p))
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
infix ` ≃gm `:25 := graded_iso
|
|
|
|
|
attribute graded_iso.to_hom [coercion]
|
|
|
|
|
attribute graded_iso._trans_of_to_hom [unfold 5]
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
definition is_equiv_graded_iso [instance] [priority 1010] (φ : M₁ ≃gm M₂) (i : I) :
|
|
|
|
|
is_equiv (φ i) :=
|
|
|
|
|
graded_iso.is_equiv_to_hom φ idp
|
|
|
|
|
|
|
|
|
|
definition isomorphism_of_graded_iso' [constructor] (φ : M₁ ≃gm M₂) {i j : I} (p : deg φ i = j) :
|
|
|
|
|
M₁ i ≃lm M₂ j :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
isomorphism.mk (φ ↘ p) !graded_iso.is_equiv_to_hom
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
|
|
|
|
definition isomorphism_of_graded_iso [constructor] (φ : M₁ ≃gm M₂) (i : I) :
|
|
|
|
|
M₁ i ≃lm M₂ (deg φ i) :=
|
|
|
|
|
isomorphism.mk (φ i) _
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition isomorphism_of_graded_iso_out [constructor] (φ : M₁ ≃gm M₂) (i : I) :
|
2017-04-21 02:58:19 +00:00
|
|
|
|
M₁ ((deg φ)⁻¹ i) ≃lm M₂ i :=
|
|
|
|
|
isomorphism_of_graded_iso' φ !to_right_inv
|
|
|
|
|
|
|
|
|
|
protected definition graded_iso.mk [constructor] (d : I ≃ I) (φ : Πi, M₁ i ≃lm M₂ (d i)) :
|
2017-03-30 22:27:09 +00:00
|
|
|
|
M₁ ≃gm M₂ :=
|
|
|
|
|
begin
|
2017-04-21 02:58:19 +00:00
|
|
|
|
apply graded_iso.mk' (graded_hom.mk d φ),
|
|
|
|
|
intro i j p, induction p,
|
2017-03-30 22:27:09 +00:00
|
|
|
|
exact to_is_equiv (equiv_of_isomorphism (φ i)),
|
|
|
|
|
end
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
protected definition graded_iso.mk_out [constructor] (d : I ≃ I) (φ : Πi, M₁ (d⁻¹ i) ≃lm M₂ i) :
|
2017-04-21 02:58:19 +00:00
|
|
|
|
M₁ ≃gm M₂ :=
|
|
|
|
|
begin
|
2017-04-24 17:33:48 +00:00
|
|
|
|
apply graded_iso.mk' (graded_hom.mk_out d φ),
|
2017-04-21 02:58:19 +00:00
|
|
|
|
intro i j p, esimp,
|
|
|
|
|
exact @is_equiv_compose _ _ _ _ _ !is_equiv_cast _,
|
|
|
|
|
end
|
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
definition graded_iso_of_eq [constructor] {M₁ M₂ : graded_module R I} (p : M₁ ~ M₂)
|
|
|
|
|
: M₁ ≃gm M₂ :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
graded_iso.mk erfl (λi, isomorphism_of_eq (p i))
|
|
|
|
|
|
|
|
|
|
-- definition to_gminv [constructor] (φ : M₁ ≃gm M₂) : M₂ →gm M₁ :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
-- graded_hom.mk_out (deg φ)⁻¹ᵉ
|
2017-04-21 02:58:19 +00:00
|
|
|
|
-- abstract begin
|
|
|
|
|
-- intro i, apply isomorphism.to_hom, symmetry,
|
|
|
|
|
-- apply isomorphism_of_graded_iso φ
|
|
|
|
|
-- end end
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
|
|
|
|
variable (M)
|
|
|
|
|
definition graded_iso.refl [refl] [constructor] : M ≃gm M :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
graded_iso.mk equiv.rfl (λi, isomorphism.rfl)
|
2017-03-30 22:27:09 +00:00
|
|
|
|
variable {M}
|
|
|
|
|
|
|
|
|
|
definition graded_iso.rfl [refl] [constructor] : M ≃gm M := graded_iso.refl M
|
|
|
|
|
|
|
|
|
|
definition graded_iso.symm [symm] [constructor] (φ : M₁ ≃gm M₂) : M₂ ≃gm M₁ :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
graded_iso.mk_out (deg φ)⁻¹ᵉ (λi, (isomorphism_of_graded_iso φ i)⁻¹ˡᵐ)
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
|
|
|
|
definition graded_iso.trans [trans] [constructor] (φ : M₁ ≃gm M₂) (ψ : M₂ ≃gm M₃) : M₁ ≃gm M₃ :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
graded_iso.mk (deg φ ⬝e deg ψ)
|
2017-03-30 22:27:09 +00:00
|
|
|
|
(λi, isomorphism_of_graded_iso φ i ⬝lm isomorphism_of_graded_iso ψ (deg φ i))
|
|
|
|
|
|
|
|
|
|
definition graded_iso.eq_trans [trans] [constructor]
|
2017-04-24 17:33:48 +00:00
|
|
|
|
{M₁ M₂ M₃ : graded_module R I} (φ : M₁ ~ M₂) (ψ : M₂ ≃gm M₃) : M₁ ≃gm M₃ :=
|
2017-03-30 22:27:09 +00:00
|
|
|
|
proof graded_iso.trans (graded_iso_of_eq φ) ψ qed
|
|
|
|
|
|
|
|
|
|
definition graded_iso.trans_eq [trans] [constructor]
|
2017-04-24 17:33:48 +00:00
|
|
|
|
{M₁ M₂ M₃ : graded_module R I} (φ : M₁ ≃gm M₂) (ψ : M₂ ~ M₃) : M₁ ≃gm M₃ :=
|
2017-03-30 22:27:09 +00:00
|
|
|
|
graded_iso.trans φ (graded_iso_of_eq ψ)
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
postfix `⁻¹ᵉᵍᵐ`:(max + 1) := graded_iso.symm
|
|
|
|
|
infixl ` ⬝egm `:75 := graded_iso.trans
|
|
|
|
|
infixl ` ⬝egmp `:75 := graded_iso.trans_eq
|
|
|
|
|
infixl ` ⬝epgm `:75 := graded_iso.eq_trans
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_hom_of_eq [constructor] {M₁ M₂ : graded_module R I} (p : M₁ ~ M₂) : M₁ →gm M₂ :=
|
|
|
|
|
proof graded_iso_of_eq p qed
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
definition foo {I : Set} (P : I → Type) {i j : I} (M : P i) (N : P j) := unit
|
|
|
|
|
notation M ` ==[`:50 P:0 `] `:0 N:50 := foo P M N
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_homotopy (f g : M₁ →gm M₂) : Type :=
|
|
|
|
|
Π⦃i j k⦄ (p : deg f i = j) (q : deg g i = k) (m : M₁ i), f ↘ p m ==[λi, M₂ i] g ↘ q m
|
|
|
|
|
-- mk' :: (hd : deg f ~ deg g)
|
|
|
|
|
-- (hfn : Π⦃i j : I⦄ (pf : deg f i = j) (pg : deg g i = j), f ↘ pf ~ g ↘ pg)
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
infix ` ~gm `:50 := graded_homotopy
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
-- definition graded_homotopy.mk2 (hd : deg f ~ deg g) (hfn : Πi m, f i m =[hd i] g i m) : f ~gm g :=
|
|
|
|
|
-- graded_homotopy.mk' hd
|
|
|
|
|
-- begin
|
|
|
|
|
-- intro i j pf pg m, induction (is_set.elim (hd i ⬝ pg) pf), induction pg, esimp,
|
|
|
|
|
-- exact graded_hom_eq_transport f (hd i) m ⬝ tr_eq_of_pathover (hfn i m),
|
|
|
|
|
-- end
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_homotopy.mk (h : Πi m, f i m ==[λi, M₂ i] g i m) : f ~gm g :=
|
|
|
|
|
begin
|
2017-04-25 22:26:59 +00:00
|
|
|
|
intros i j k p q m, induction q, induction p, constructor --exact h i m
|
2017-04-24 17:33:48 +00:00
|
|
|
|
end
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
-- definition graded_hom_compose_out {d₁ d₂ : I ≃ I} (f₂ : Πi, M₂ i →lm M₃ (d₂ i))
|
|
|
|
|
-- (f₁ : Πi, M₁ (d₁⁻¹ i) →lm M₂ i) : graded_hom.mk d₂ f₂ ∘gm graded_hom.mk_out d₁ f₁ ~gm
|
|
|
|
|
-- graded_hom.mk_out_in d₁⁻¹ᵉ d₂ _ :=
|
|
|
|
|
-- _
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_hom_out_in_compose_out {d₁ d₂ d₃ : I ≃ I} (f₂ : Πi, M₂ (d₂ i) →lm M₃ (d₃ i))
|
|
|
|
|
(f₁ : Πi, M₁ (d₁⁻¹ i) →lm M₂ i) : graded_hom.mk_out_in d₂ d₃ f₂ ∘gm graded_hom.mk_out d₁ f₁ ~gm
|
|
|
|
|
graded_hom.mk_out_in (d₂ ⬝e d₁⁻¹ᵉ) d₃ (λi, f₂ i ∘lm (f₁ (d₂ i))) :=
|
|
|
|
|
begin
|
|
|
|
|
apply graded_homotopy.mk, intro i m, exact sorry
|
|
|
|
|
end
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_hom_out_in_rfl {d₁ d₂ : I ≃ I} (f : Πi, M₁ i →lm M₂ (d₂ i))
|
|
|
|
|
(p : Πi, d₁ i = i) :
|
|
|
|
|
graded_hom.mk_out_in d₁ d₂ (λi, sorry) ~gm graded_hom.mk d₂ f :=
|
|
|
|
|
begin
|
|
|
|
|
apply graded_homotopy.mk, intro i m, exact sorry
|
|
|
|
|
end
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_homotopy.trans (h₁ : f ~gm g) (h₂ : g ~gm h) : f ~gm h :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
begin
|
2017-04-24 17:33:48 +00:00
|
|
|
|
exact sorry
|
2017-04-21 02:58:19 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
-- postfix `⁻¹ᵍᵐ`:(max + 1) := graded_iso.symm
|
|
|
|
|
infixl ` ⬝gm `:75 := graded_homotopy.trans
|
|
|
|
|
-- infixl ` ⬝gmp `:75 := graded_iso.trans_eq
|
|
|
|
|
-- infixl ` ⬝pgm `:75 := graded_iso.eq_trans
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- definition graded_homotopy_of_deg (d : I ≃ I) (f g : graded_hom_of_deg d M₁ M₂) : Type :=
|
|
|
|
|
-- Π⦃i j : I⦄ (p : d i = j), f p ~ g p
|
|
|
|
|
|
|
|
|
|
-- notation f ` ~[`:50 d:0 `] `:0 g:50 := graded_homotopy_of_deg d f g
|
|
|
|
|
|
|
|
|
|
-- variables {d : I ≃ I} {f₁ f₂ : graded_hom_of_deg d M₁ M₂}
|
|
|
|
|
|
|
|
|
|
-- definition graded_homotopy_of_deg.mk [constructor] (h : Πi, f₁ (idpath (d i)) ~ f₂ (idpath (d i))) :
|
|
|
|
|
-- f₁ ~[d] f₂ :=
|
|
|
|
|
-- begin
|
|
|
|
|
-- intro i j p, induction p, exact h i
|
|
|
|
|
-- end
|
|
|
|
|
|
|
|
|
|
-- definition graded_homotopy.mk_out [constructor] {M₁ M₂ : graded_module R I} (d : I ≃ I)
|
2017-04-21 02:58:19 +00:00
|
|
|
|
-- (fn : Πi, M₁ (d⁻¹ i) →lm M₂ i) : M₁ →gm M₂ :=
|
|
|
|
|
-- graded_hom.mk' d (λi j p, fn j ∘lm homomorphism_of_eq (ap M₁ (eq_inv_of_eq p)))
|
|
|
|
|
-- definition is_gconstant (f : M₁ →gm M₂) : Type :=
|
|
|
|
|
-- f↘ ~[deg f] gmd0
|
|
|
|
|
|
|
|
|
|
definition compose_constant (f' : M₂ →gm M₃) (f : M₁ →gm M₂) : Type :=
|
|
|
|
|
Π⦃i j k : I⦄ (p : deg f i = j) (q : deg f' j = k) (m : M₁ i), f' ↘ q (f ↘ p m) = 0
|
|
|
|
|
|
|
|
|
|
definition compose_constant.mk (h : Πi m, f' (deg f i) (f i m) = 0) : compose_constant f' f :=
|
|
|
|
|
by intros; induction p; induction q; exact h i m
|
|
|
|
|
|
|
|
|
|
definition compose_constant.elim (h : compose_constant f' f) (i : I) (m : M₁ i) : f' (deg f i) (f i m) = 0 :=
|
|
|
|
|
h idp idp m
|
|
|
|
|
|
|
|
|
|
definition is_gconstant (f : M₁ →gm M₂) : Type :=
|
|
|
|
|
Π⦃i j : I⦄ (p : deg f i = j) (m : M₁ i), f ↘ p m = 0
|
|
|
|
|
|
|
|
|
|
definition is_gconstant.mk (h : Πi m, f i m = 0) : is_gconstant f :=
|
|
|
|
|
by intros; induction p; exact h i m
|
|
|
|
|
|
|
|
|
|
definition is_gconstant.elim (h : is_gconstant f) (i : I) (m : M₁ i) : f i m = 0 :=
|
|
|
|
|
h idp m
|
|
|
|
|
|
2017-03-31 22:21:02 +00:00
|
|
|
|
/- direct sum of graded R-modules -/
|
|
|
|
|
|
|
|
|
|
variables {J : Set} (N : graded_module R J)
|
|
|
|
|
definition dirsum' : AddAbGroup :=
|
|
|
|
|
group.dirsum (λj, AddAbGroup_of_LeftModule (N j))
|
|
|
|
|
variable {N}
|
2017-04-13 18:54:48 +00:00
|
|
|
|
definition dirsum_smul [constructor] (r : R) : dirsum' N →a dirsum' N :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
dirsum_functor (λi, smul_homomorphism (N i) r)
|
|
|
|
|
|
|
|
|
|
definition dirsum_smul_right_distrib (r s : R) (n : dirsum' N) :
|
|
|
|
|
dirsum_smul (r + s) n = dirsum_smul r n + dirsum_smul s n :=
|
|
|
|
|
begin
|
|
|
|
|
refine dirsum_functor_homotopy _ n ⬝ !dirsum_functor_add⁻¹,
|
|
|
|
|
intro i ni, exact to_smul_right_distrib r s ni
|
|
|
|
|
end
|
|
|
|
|
|
2017-04-13 18:54:48 +00:00
|
|
|
|
definition dirsum_mul_smul' (r s : R) (n : dirsum' N) :
|
|
|
|
|
dirsum_smul (r * s) n = (dirsum_smul r ∘a dirsum_smul s) n :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
begin
|
2017-04-13 18:54:48 +00:00
|
|
|
|
refine dirsum_functor_homotopy _ n ⬝ (dirsum_functor_compose _ _ n)⁻¹ᵖ,
|
2017-03-31 22:21:02 +00:00
|
|
|
|
intro i ni, exact to_mul_smul r s ni
|
|
|
|
|
end
|
|
|
|
|
|
2017-04-13 18:54:48 +00:00
|
|
|
|
definition dirsum_mul_smul (r s : R) (n : dirsum' N) :
|
|
|
|
|
dirsum_smul (r * s) n = dirsum_smul r (dirsum_smul s n) :=
|
|
|
|
|
proof dirsum_mul_smul' r s n qed
|
|
|
|
|
|
2017-03-31 22:21:02 +00:00
|
|
|
|
definition dirsum_one_smul (n : dirsum' N) : dirsum_smul 1 n = n :=
|
|
|
|
|
begin
|
|
|
|
|
refine dirsum_functor_homotopy _ n ⬝ !dirsum_functor_gid,
|
|
|
|
|
intro i ni, exact to_one_smul ni
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition dirsum : LeftModule R :=
|
2017-04-14 00:39:04 +00:00
|
|
|
|
LeftModule_of_AddAbGroup (dirsum' N) (λr n, dirsum_smul r n)
|
|
|
|
|
(λr, homomorphism.addstruct (dirsum_smul r))
|
2017-03-31 22:21:02 +00:00
|
|
|
|
dirsum_smul_right_distrib
|
|
|
|
|
dirsum_mul_smul
|
|
|
|
|
dirsum_one_smul
|
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
/- graded variants of left-module constructions -/
|
|
|
|
|
|
|
|
|
|
definition graded_submodule [constructor] (S : Πi, submodule_rel (M i)) : graded_module R I :=
|
|
|
|
|
λi, submodule (S i)
|
|
|
|
|
|
|
|
|
|
definition graded_submodule_incl [constructor] (S : Πi, submodule_rel (M i)) :
|
|
|
|
|
graded_submodule S →gm M :=
|
|
|
|
|
graded_hom.mk erfl (λi, submodule_incl (S i))
|
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
definition graded_hom_lift [constructor] {S : Πi, submodule_rel (M₂ i)}
|
|
|
|
|
(φ : M₁ →gm M₂)
|
2017-04-21 02:58:19 +00:00
|
|
|
|
(h : Π(i : I) (m : M₁ i), S (deg φ i) (φ i m)) : M₁ →gm graded_submodule S :=
|
|
|
|
|
graded_hom.mk (deg φ) (λi, hom_lift (φ i) (h i))
|
|
|
|
|
|
|
|
|
|
definition graded_image (f : M₁ →gm M₂) : graded_module R I :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
λi, image_module (f ← i)
|
2017-04-20 18:42:29 +00:00
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
definition graded_image_lift [constructor] (f : M₁ →gm M₂) : M₁ →gm graded_image f :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
graded_hom.mk_out (deg f) (λi, image_lift (f ← i))
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
definition graded_image_elim [constructor] {f : M₁ →gm M₂} (g : M₁ →gm M₃)
|
|
|
|
|
(h : Π⦃i m⦄, f i m = 0 → g i m = 0) :
|
|
|
|
|
graded_image f →gm M₃ :=
|
|
|
|
|
begin
|
|
|
|
|
apply graded_hom.mk_out_in (deg f) (deg g),
|
|
|
|
|
intro i,
|
|
|
|
|
apply image_elim (g ↘ (ap (deg g) (to_left_inv (deg f) i))),
|
|
|
|
|
intro m p,
|
|
|
|
|
refine graded_hom_eq_zero m (h _),
|
|
|
|
|
exact graded_hom_eq_zero m p
|
|
|
|
|
end
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_image' (f : M₁ →gm M₂) : graded_module R I :=
|
|
|
|
|
λi, image_module (f i)
|
|
|
|
|
|
|
|
|
|
definition graded_image'_lift [constructor] (f : M₁ →gm M₂) : M₁ →gm graded_image' f :=
|
|
|
|
|
graded_hom.mk erfl (λi, image_lift (f i))
|
|
|
|
|
|
|
|
|
|
definition graded_image'_elim [constructor] {f : M₁ →gm M₂} (g : M₁ →gm M₃)
|
|
|
|
|
(h : Π⦃i m⦄, f i m = 0 → g i m = 0) :
|
|
|
|
|
graded_image' f →gm M₃ :=
|
|
|
|
|
begin
|
|
|
|
|
apply graded_hom.mk (deg g),
|
|
|
|
|
intro i,
|
|
|
|
|
apply image_elim (g i),
|
|
|
|
|
intro m p, exact h p
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem graded_image'_elim_compute {f : M₁ →gm M₂} {g : M₁ →gm M₃}
|
|
|
|
|
(h : Π⦃i m⦄, f i m = 0 → g i m = 0) :
|
|
|
|
|
graded_image'_elim g h ∘gm graded_image'_lift f ~gm g :=
|
|
|
|
|
begin
|
|
|
|
|
apply graded_homotopy.mk,
|
2017-04-25 22:26:59 +00:00
|
|
|
|
intro i m, exact sorry --reflexivity
|
2017-04-24 17:33:48 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem graded_image_elim_compute {f : M₁ →gm M₂} {g : M₁ →gm M₃}
|
|
|
|
|
(h : Π⦃i m⦄, f i m = 0 → g i m = 0) :
|
|
|
|
|
graded_image_elim g h ∘gm graded_image_lift f ~gm g :=
|
|
|
|
|
begin
|
|
|
|
|
refine _ ⬝gm graded_image'_elim_compute h,
|
|
|
|
|
esimp, exact sorry
|
|
|
|
|
-- refine graded_hom_out_in_compose_out _ _ ⬝gm _, exact sorry
|
|
|
|
|
-- -- apply graded_homotopy.mk,
|
|
|
|
|
-- -- intro i m,
|
|
|
|
|
end
|
|
|
|
|
variables {α β : I ≃ I}
|
|
|
|
|
definition gen_image (f : M₁ →gm M₂) (p : Πi, deg f (α i) = β i) : graded_module R I :=
|
|
|
|
|
λi, image_module (f ↘ (p i))
|
|
|
|
|
|
|
|
|
|
definition gen_image_lift [constructor] (f : M₁ →gm M₂) (p : Πi, deg f (α i) = β i) : M₁ →gm gen_image f p :=
|
|
|
|
|
graded_hom.mk_out α⁻¹ᵉ (λi, image_lift (f ↘ (p i)))
|
|
|
|
|
|
|
|
|
|
definition gen_image_elim [constructor] {f : M₁ →gm M₂} (p : Πi, deg f (α i) = β i) (g : M₁ →gm M₃)
|
|
|
|
|
(h : Π⦃i m⦄, f i m = 0 → g i m = 0) :
|
|
|
|
|
gen_image f p →gm M₃ :=
|
|
|
|
|
begin
|
|
|
|
|
apply graded_hom.mk_out_in α⁻¹ᵉ (deg g),
|
|
|
|
|
intro i,
|
|
|
|
|
apply image_elim (g ↘ (ap (deg g) (to_right_inv α i))),
|
|
|
|
|
intro m p,
|
|
|
|
|
refine graded_hom_eq_zero m (h _),
|
|
|
|
|
exact graded_hom_eq_zero m p
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem gen_image_elim_compute {f : M₁ →gm M₂} {p : deg f ∘ α ~ β} {g : M₁ →gm M₃}
|
|
|
|
|
(h : Π⦃i m⦄, f i m = 0 → g i m = 0) :
|
|
|
|
|
gen_image_elim p g h ∘gm gen_image_lift f p ~gm g :=
|
|
|
|
|
begin
|
|
|
|
|
-- induction β with β βe, esimp at *, induction p using homotopy.rec_on_idp,
|
|
|
|
|
assert q : β ⬝e (deg f)⁻¹ᵉ = α,
|
|
|
|
|
{ apply equiv_eq, intro i, apply inv_eq_of_eq, exact (p i)⁻¹ },
|
|
|
|
|
induction q,
|
|
|
|
|
-- unfold [gen_image_elim, gen_image_lift],
|
|
|
|
|
|
|
|
|
|
-- induction (is_prop.elim (λi, to_right_inv (deg f) (β i)) p),
|
|
|
|
|
-- apply graded_homotopy.mk,
|
|
|
|
|
-- intro i m, reflexivity
|
2017-04-25 22:26:59 +00:00
|
|
|
|
exact sorry
|
2017-04-24 17:33:48 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition graded_kernel (f : M₁ →gm M₂) : graded_module R I :=
|
|
|
|
|
λi, kernel_module (f i)
|
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
definition graded_quotient (S : Πi, submodule_rel (M i)) : graded_module R I :=
|
|
|
|
|
λi, quotient_module (S i)
|
|
|
|
|
|
|
|
|
|
definition graded_quotient_map [constructor] (S : Πi, submodule_rel (M i)) :
|
|
|
|
|
M →gm graded_quotient S :=
|
|
|
|
|
graded_hom.mk erfl (λi, quotient_map (S i))
|
|
|
|
|
|
|
|
|
|
definition graded_homology (g : M₂ →gm M₃) (f : M₁ →gm M₂) : graded_module R I :=
|
|
|
|
|
λi, homology (g i) (f ↘ (to_right_inv (deg f) i))
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition graded_homology_intro [constructor] (g : M₂ →gm M₃) (f : M₁ →gm M₂) :
|
|
|
|
|
graded_kernel g →gm graded_homology g f :=
|
|
|
|
|
graded_quotient_map _
|
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
definition graded_homology_elim {g : M₂ →gm M₃} {f : M₁ →gm M₂} (h : M₂ →gm M)
|
|
|
|
|
(H : compose_constant h f) : graded_homology g f →gm M :=
|
|
|
|
|
graded_hom.mk (deg h) (λi, homology_elim (h i) (H _ _))
|
2017-04-20 18:42:29 +00:00
|
|
|
|
|
|
|
|
|
|
2017-03-30 22:27:09 +00:00
|
|
|
|
/- exact couples -/
|
|
|
|
|
|
|
|
|
|
definition is_exact_gmod (f : M₁ →gm M₂) (f' : M₂ →gm M₃) : Type :=
|
2017-04-21 02:58:19 +00:00
|
|
|
|
Π⦃i j k⦄ (p : deg f i = j) (q : deg f' j = k), is_exact_mod (f ↘ p) (f' ↘ q)
|
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
definition is_exact_gmod.mk {f : M₁ →gm M₂} {f' : M₂ →gm M₃}
|
|
|
|
|
(h₁ : Π⦃i⦄ (m : M₁ i), f' (deg f i) (f i m) = 0)
|
|
|
|
|
(h₂ : Π⦃i⦄ (m : M₂ (deg f i)), f' (deg f i) m = 0 → image (f i) m) : is_exact_gmod f f' :=
|
|
|
|
|
begin intro i j k p q; induction p; induction q; split, apply h₁, apply h₂ end
|
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
definition gmod_im_in_ker (h : is_exact_gmod f f') : compose_constant f' f :=
|
|
|
|
|
λi j k p q, is_exact.im_in_ker (h p q)
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
-- structure exact_couple (M₁ M₂ : graded_module R I) : Type :=
|
|
|
|
|
-- (i : M₁ →gm M₁) (j : M₁ →gm M₂) (k : M₂ →gm M₁)
|
|
|
|
|
-- (exact_ij : is_exact_gmod i j)
|
|
|
|
|
-- (exact_jk : is_exact_gmod j k)
|
|
|
|
|
-- (exact_ki : is_exact_gmod k i)
|
2017-03-30 22:27:09 +00:00
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
end left_module
|
|
|
|
|
|
|
|
|
|
namespace left_module
|
2017-04-25 22:26:59 +00:00
|
|
|
|
|
|
|
|
|
structure exact_couple (R : Ring) (I : Set) : Type :=
|
|
|
|
|
(D E : graded_module R I)
|
|
|
|
|
(i : D →gm D) (j : D →gm E) (k : E →gm D)
|
|
|
|
|
(ij : is_exact_gmod i j)
|
|
|
|
|
(jk : is_exact_gmod j k)
|
|
|
|
|
(ki : is_exact_gmod k i)
|
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
namespace derived_couple
|
|
|
|
|
section
|
2017-04-25 22:26:59 +00:00
|
|
|
|
open exact_couple
|
|
|
|
|
|
|
|
|
|
parameters {R : Ring} {I : Set} (X : exact_couple R I)
|
|
|
|
|
local abbreviation D := D X
|
|
|
|
|
local abbreviation E := E X
|
|
|
|
|
local abbreviation i := i X
|
|
|
|
|
local abbreviation j := j X
|
|
|
|
|
local abbreviation k := k X
|
|
|
|
|
local abbreviation ij := ij X
|
|
|
|
|
local abbreviation jk := jk X
|
|
|
|
|
local abbreviation ki := ki X
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-04-21 02:58:19 +00:00
|
|
|
|
definition d : E →gm E := j ∘gm k
|
|
|
|
|
definition D' : graded_module R I := graded_image i
|
|
|
|
|
definition E' : graded_module R I := graded_homology d d
|
2017-04-24 17:33:48 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
definition is_contr_E' {x : I} (H : is_contr (E x)) : is_contr (E' x) :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition is_contr_D' {x : I} (H : is_contr (D x)) : is_contr (D' x) :=
|
|
|
|
|
sorry
|
2017-04-24 17:33:48 +00:00
|
|
|
|
|
|
|
|
|
definition i' : D' →gm D' :=
|
|
|
|
|
graded_image_lift i ∘gm graded_submodule_incl _
|
|
|
|
|
-- degree i + 0
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
theorem j_lemma1 ⦃x : I⦄ (m : D x) : d ((deg j) x) (j x m) = 0 :=
|
|
|
|
|
begin
|
|
|
|
|
rewrite [graded_hom_compose_fn,↑d,graded_hom_compose_fn],
|
2017-04-25 22:26:59 +00:00
|
|
|
|
refine ap (graded_hom_fn j (deg k (deg j x))) _ ⬝
|
|
|
|
|
!to_respect_zero,
|
|
|
|
|
exact compose_constant.elim (gmod_im_in_ker (jk)) x m
|
2017-04-21 02:58:19 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem j_lemma2 : Π⦃x : I⦄ ⦃m : D x⦄ (p : i x m = 0),
|
|
|
|
|
(graded_quotient_map _ ∘gm graded_hom_lift j j_lemma1) x m = 0 :> E' _ :=
|
|
|
|
|
begin
|
2017-04-25 22:26:59 +00:00
|
|
|
|
have Π⦃x y : I⦄ (q : deg k x = y) (r : deg d x = deg j y)
|
|
|
|
|
(s : ap (deg j) q = r) ⦃m : D y⦄ (p : i y m = 0), image (d ↘ r) (j y m),
|
2017-04-21 02:58:19 +00:00
|
|
|
|
begin
|
2017-04-21 22:00:27 +00:00
|
|
|
|
intros, induction s, induction q,
|
2017-04-25 22:26:59 +00:00
|
|
|
|
note m_in_im_k := is_exact.ker_in_im (ki idp _) _ p,
|
2017-04-21 02:58:19 +00:00
|
|
|
|
induction m_in_im_k with e q,
|
|
|
|
|
induction q,
|
|
|
|
|
apply image.mk e idp
|
|
|
|
|
end,
|
|
|
|
|
have Π⦃x : I⦄ ⦃m : D x⦄ (p : i x m = 0), image (d ← (deg j x)) (j x m),
|
|
|
|
|
begin
|
2017-04-21 22:00:27 +00:00
|
|
|
|
intros,
|
|
|
|
|
refine this _ _ _ p,
|
|
|
|
|
exact to_right_inv (deg k) _ ⬝ to_left_inv (deg j) x,
|
2017-04-24 17:33:48 +00:00
|
|
|
|
apply is_set.elim
|
|
|
|
|
-- rewrite [ap_con, -adj],
|
2017-04-21 02:58:19 +00:00
|
|
|
|
end,
|
|
|
|
|
intros,
|
|
|
|
|
rewrite [graded_hom_compose_fn],
|
|
|
|
|
exact quotient_map_eq_zero _ (this p)
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition j' : D' →gm E' :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
graded_image_elim (graded_homology_intro d d ∘gm graded_hom_lift j j_lemma1) j_lemma2
|
2017-04-25 22:26:59 +00:00
|
|
|
|
-- degree deg j - deg i
|
|
|
|
|
|
|
|
|
|
theorem k_lemma1 ⦃x : I⦄ (m : E x) : image (i ← (deg k x)) (k x m) :=
|
|
|
|
|
begin
|
|
|
|
|
exact sorry
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem k_lemma2 : compose_constant (graded_hom_lift k k_lemma1 : E →gm D') d :=
|
|
|
|
|
begin
|
|
|
|
|
-- apply compose_constant.mk, intro x m,
|
|
|
|
|
-- rewrite [graded_hom_compose_fn],
|
|
|
|
|
-- refine ap (graded_hom_fn (graded_image_lift i) (deg k (deg d x))) _ ⬝ !to_respect_zero,
|
|
|
|
|
-- exact compose_constant.elim (gmod_im_in_ker jk) (deg k x) (k x m)
|
|
|
|
|
exact sorry
|
|
|
|
|
end
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
|
|
|
|
definition k' : E' →gm D' :=
|
2017-04-25 22:26:59 +00:00
|
|
|
|
graded_homology_elim (graded_hom_lift k k_lemma1) k_lemma2
|
|
|
|
|
|
|
|
|
|
definition deg_i' : deg i' ~ deg i := by reflexivity
|
|
|
|
|
definition deg_j' : deg j' ~ deg j ∘ (deg i)⁻¹ := by reflexivity
|
|
|
|
|
definition deg_k' : deg k' ~ deg k := by reflexivity
|
|
|
|
|
|
|
|
|
|
theorem i'j' : is_exact_gmod i' j' :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
begin
|
|
|
|
|
apply is_exact_gmod.mk,
|
2017-04-25 22:26:59 +00:00
|
|
|
|
{ intro x, refine total_image.rec _, intro m, exact sorry
|
|
|
|
|
-- exact calc
|
|
|
|
|
-- j' (deg i' x) (i' x ⟨(i ← x) m, image.mk m idp⟩)
|
|
|
|
|
-- = j' (deg i' x) (graded_image_lift i x ((i ← x) m)) : idp
|
|
|
|
|
-- ... = graded_homology_intro d d (deg j ((deg i)⁻¹ᵉ (deg i x)))
|
|
|
|
|
-- (graded_hom_lift j j_lemma1 ((deg i)⁻¹ᵉ (deg i x))
|
|
|
|
|
-- (i ↘ (!to_right_inv ⬝ !to_left_inv⁻¹) m)) : _
|
|
|
|
|
-- ... = graded_homology_intro d d (deg j ((deg i)⁻¹ᵉ (deg i x)))
|
|
|
|
|
-- (graded_hom_lift j j_lemma1 ((deg i)⁻¹ᵉ (deg i x))
|
|
|
|
|
-- (i ↘ (!to_right_inv ⬝ !to_left_inv⁻¹) m)) : _
|
|
|
|
|
-- ... = 0 : _
|
2017-04-24 17:33:48 +00:00
|
|
|
|
},
|
|
|
|
|
{ exact sorry }
|
|
|
|
|
end
|
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
theorem j'k' : is_exact_gmod j' k' :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
begin
|
|
|
|
|
apply is_exact_gmod.mk,
|
2017-04-25 22:26:59 +00:00
|
|
|
|
{ exact sorry },
|
2017-04-24 17:33:48 +00:00
|
|
|
|
{ exact sorry }
|
|
|
|
|
end
|
2017-04-21 02:58:19 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
theorem k'i' : is_exact_gmod k' i' :=
|
2017-04-24 17:33:48 +00:00
|
|
|
|
begin
|
|
|
|
|
apply is_exact_gmod.mk,
|
2017-04-25 22:26:59 +00:00
|
|
|
|
{ intro x m, exact sorry },
|
2017-04-24 17:33:48 +00:00
|
|
|
|
{ exact sorry }
|
2017-04-21 02:58:19 +00:00
|
|
|
|
end
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
2017-04-24 17:33:48 +00:00
|
|
|
|
end
|
2017-04-25 22:26:59 +00:00
|
|
|
|
end derived_couple
|
|
|
|
|
|
|
|
|
|
section
|
|
|
|
|
open derived_couple exact_couple
|
2017-04-24 17:33:48 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
definition derived_couple [constructor] {R : Ring} {I : Set}
|
|
|
|
|
(X : exact_couple R I) : exact_couple R I :=
|
|
|
|
|
⦃exact_couple, D := D' X, E := E' X, i := i' X, j := j' X, k := k' X,
|
|
|
|
|
ij := i'j' X, jk := j'k' X, ki := k'i' X⦄
|
2017-04-24 17:33:48 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
parameters {R : Ring} {I : Set} (X : exact_couple R I) (B : I → ℕ)
|
|
|
|
|
(Dub : Π⦃x y⦄ ⦃s : ℕ⦄, B x ≤ s → deg (j X) (iterate (deg (i X)) s x) = y → is_contr (D X y))
|
|
|
|
|
(Dlb : Π⦃x y⦄ ⦃s : ℕ⦄, B x ≤ s → deg (j X) (iterate (deg (i X))⁻¹ s x) = y → is_contr (D X y))
|
|
|
|
|
(Eub : Π⦃x y⦄ ⦃s : ℕ⦄, B x ≤ s → deg (j X) (iterate (deg (i X)) s x) = y → is_contr (E X y))
|
|
|
|
|
(Elb : Π⦃x y⦄ ⦃s : ℕ⦄, B x ≤ s → deg (j X) (iterate (deg (i X))⁻¹ s x) = y → is_contr (E X y))
|
|
|
|
|
-- also need a single deg j and/or deg k here
|
2017-04-24 17:33:48 +00:00
|
|
|
|
|
2017-04-25 22:26:59 +00:00
|
|
|
|
-- we start counting pages at 0, not at 2.
|
|
|
|
|
definition page (r : ℕ) : exact_couple R I :=
|
|
|
|
|
iterate derived_couple r X
|
|
|
|
|
|
|
|
|
|
definition is_contr_E (r : ℕ) (x : I) (h : is_contr (E X x)) :
|
|
|
|
|
is_contr (E (page r) x) :=
|
|
|
|
|
by induction r with r IH; exact h; exact is_contr_E' (page r) IH
|
|
|
|
|
|
|
|
|
|
definition is_contr_D (r : ℕ) (x : I) (h : is_contr (D X x)) :
|
|
|
|
|
is_contr (D (page r) x) :=
|
|
|
|
|
by induction r with r IH; exact h; exact is_contr_D' (page r) IH
|
|
|
|
|
|
|
|
|
|
definition deg_i (r : ℕ) : deg (i (page r)) ~ deg (i X) :=
|
|
|
|
|
begin
|
|
|
|
|
induction r with r IH,
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ exact IH }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition deg_k (r : ℕ) : deg (k (page r)) ~ deg (k X) :=
|
|
|
|
|
begin
|
|
|
|
|
induction r with r IH,
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ exact IH }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition deg_j (r : ℕ) :
|
|
|
|
|
deg (j (page r)) ~ deg (j X) ∘ iterate (deg (i X))⁻¹ r :=
|
|
|
|
|
begin
|
|
|
|
|
induction r with r IH,
|
|
|
|
|
{ reflexivity },
|
|
|
|
|
{ refine hwhisker_left (deg (j (page r)))
|
|
|
|
|
(to_inv_homotopy_inv (deg_i r)) ⬝hty _,
|
|
|
|
|
refine hwhisker_right _ IH ⬝hty _,
|
|
|
|
|
apply hwhisker_left, symmetry, apply iterate_succ }
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition deg_d (r : ℕ) :
|
|
|
|
|
deg (d (page r)) ~ deg (j X) ∘ iterate (deg (i X))⁻¹ r ∘ deg (k X) :=
|
|
|
|
|
compose2 (deg_j r) (deg_k r)
|
|
|
|
|
|
|
|
|
|
definition Eub' (x : I) (r : ℕ) (h : B (deg (k X) x) ≤ r) :
|
|
|
|
|
is_contr (E (page r) (deg (d (page r)) x)) :=
|
|
|
|
|
is_contr_E _ _ (Elb h (deg_d r x)⁻¹)
|
|
|
|
|
|
|
|
|
|
definition Estable {x : I} {r : ℕ} (H : B (deg (k X) x) ≤ r) :
|
|
|
|
|
E (page (r + 1)) x ≃ E (page r) x :=
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
definition inf_page : graded_module R I :=
|
|
|
|
|
λx, E (page (B (deg (k X) x))) x
|
|
|
|
|
|
|
|
|
|
end
|
2017-03-30 19:02:24 +00:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end left_module
|