Documentation

Mathlib.RingTheory.Filtration

I-filtrations of modules #

This file contains the definitions and basic results around (stable) I-filtrations of modules.

Main results #

structure Ideal.Filtration {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_3) [AddCommGroup M] [Module R M] :
Type u_3

An I-filtration on the module M is a sequence of decreasing submodules N i such that I • (N i) ≤ N (i + 1). Note that we do not require the filtration to start from .

theorem Ideal.Filtration.ext_iff {R : Type u_1} {inst✝ : CommRing R} {I : Ideal R} {M : Type u_3} {inst✝¹ : AddCommGroup M} {inst✝² : Module R M} {x y : I.Filtration M} :
x = y x.N = y.N
theorem Ideal.Filtration.ext {R : Type u_1} {inst✝ : CommRing R} {I : Ideal R} {M : Type u_3} {inst✝¹ : AddCommGroup M} {inst✝² : Module R M} {x y : I.Filtration M} (N : x.N = y.N) :
x = y
theorem Ideal.Filtration.pow_smul_le {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) (i j : ) :
I ^ i F.N j F.N (i + j)
theorem Ideal.Filtration.pow_smul_le_pow_smul {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) (i j k : ) :
I ^ (i + k) F.N j I ^ k F.N (i + j)
theorem Ideal.Filtration.antitone {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) :
def Ideal.trivialFiltration {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) :

The trivial I-filtration of N.

Equations
@[simp]
theorem Ideal.trivialFiltration_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) (x✝ : ) :
(I.trivialFiltration N).N x✝ = N
instance Ideal.Filtration.instMax {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :

The sup of two I.Filtrations is an I.Filtration.

Equations
instance Ideal.Filtration.instSupSet {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :

The sSup of a family of I.Filtrations is an I.Filtration.

Equations
instance Ideal.Filtration.instMin {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :

The inf of two I.Filtrations is an I.Filtration.

Equations
instance Ideal.Filtration.instInfSet {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :

The sInf of a family of I.Filtrations is an I.Filtration.

Equations
instance Ideal.Filtration.instTop {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :
Equations
instance Ideal.Filtration.instBot {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :
Equations
@[simp]
theorem Ideal.Filtration.sup_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F F' : I.Filtration M) :
(FF').N = F.NF'.N
@[simp]
theorem Ideal.Filtration.sSup_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (S : Set (I.Filtration M)) :
(sSup S).N = sSup (N '' S)
@[simp]
theorem Ideal.Filtration.inf_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F F' : I.Filtration M) :
(FF').N = F.NF'.N
@[simp]
theorem Ideal.Filtration.sInf_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (S : Set (I.Filtration M)) :
(sInf S).N = sInf (N '' S)
@[simp]
theorem Ideal.Filtration.top_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :
@[simp]
theorem Ideal.Filtration.bot_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :
@[simp]
theorem Ideal.Filtration.iSup_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {ι : Sort u_3} (f : ιI.Filtration M) :
(iSup f).N = ⨆ (i : ι), (f i).N
@[simp]
theorem Ideal.Filtration.iInf_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {ι : Sort u_3} (f : ιI.Filtration M) :
(iInf f).N = ⨅ (i : ι), (f i).N
instance Ideal.Filtration.instInhabited {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} :
Equations
def Ideal.Filtration.Stable {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) :

An I filtration is stable if I • F.N n = F.N (n+1) for large enough n.

Equations
def Ideal.stableFiltration {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) :

The trivial stable I-filtration of N.

Equations
@[simp]
theorem Ideal.stableFiltration_N {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) (i : ) :
(I.stableFiltration N).N i = I ^ i N
theorem Ideal.stableFiltration_stable {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) :
theorem Ideal.Filtration.Stable.exists_pow_smul_eq {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F : I.Filtration M} (h : F.Stable) :
∃ (n₀ : ), ∀ (k : ), F.N (n₀ + k) = I ^ k F.N n₀
theorem Ideal.Filtration.Stable.exists_pow_smul_eq_of_ge {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F : I.Filtration M} (h : F.Stable) :
∃ (n₀ : ), nn₀, F.N n = I ^ (n - n₀) F.N n₀
theorem Ideal.Filtration.stable_iff_exists_pow_smul_eq_of_ge {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F : I.Filtration M} :
F.Stable ∃ (n₀ : ), nn₀, F.N n = I ^ (n - n₀) F.N n₀
theorem Ideal.Filtration.Stable.exists_forall_le {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F F' : I.Filtration M} (h : F.Stable) (e : F.N 0 F'.N 0) :
∃ (n₀ : ), ∀ (n : ), F.N (n + n₀) F'.N n
theorem Ideal.Filtration.Stable.bounded_difference {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F F' : I.Filtration M} (h : F.Stable) (h' : F'.Stable) (e : F.N 0 = F'.N 0) :
∃ (n₀ : ), ∀ (n : ), F.N (n + n₀) F'.N n F'.N (n + n₀) F.N n
noncomputable def Ideal.Filtration.submodule {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) :

The R[IX]-submodule of M[X] associated with an I-filtration.

Equations
@[simp]
theorem Ideal.Filtration.mem_submodule {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) (f : PolynomialModule R M) :
f F.submodule ∀ (i : ), f i F.N i
theorem Ideal.Filtration.inf_submodule {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F F' : I.Filtration M) :
(FF').submodule = F.submoduleF'.submodule
theorem Ideal.Filtration.submodule_span_single {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) :
Submodule.span (↥(reesAlgebra I)) (⋃ (i : ), (PolynomialModule.single R i) '' (F.N i)) = F.submodule
theorem Ideal.Filtration.submodule_eq_span_le_iff_stable_ge {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) (n₀ : ) :
F.submodule = Submodule.span (↥(reesAlgebra I)) (⋃ (i : ), ⋃ (_ : i n₀), (PolynomialModule.single R i) '' (F.N i)) nn₀, I F.N n = F.N (n + 1)
theorem Ideal.Filtration.submodule_fg_iff_stable {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} (F : I.Filtration M) (hF' : ∀ (i : ), (F.N i).FG) :

If the components of a filtration are finitely generated, then the filtration is stable iff its associated submodule of is finitely generated.

theorem Ideal.Filtration.Stable.of_le {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F : I.Filtration M} [IsNoetherianRing R] [Module.Finite R M] (hF : F.Stable) {F' : I.Filtration M} (hf : F' F) :
theorem Ideal.Filtration.Stable.inter_right {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F : I.Filtration M} (F' : I.Filtration M) [IsNoetherianRing R] [Module.Finite R M] (hF : F.Stable) :
(FF').Stable
theorem Ideal.Filtration.Stable.inter_left {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {F : I.Filtration M} (F' : I.Filtration M) [IsNoetherianRing R] [Module.Finite R M] (hF : F.Stable) :
(F'F).Stable
theorem Ideal.exists_pow_inf_eq_pow_smul {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) [IsNoetherianRing R] [Module.Finite R M] (N : Submodule R M) :
∃ (k : ), nk, I ^ n N = I ^ (n - k) (I ^ k N)

Artin-Rees lemma

theorem Ideal.mem_iInf_smul_pow_eq_bot_iff {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) [IsNoetherianRing R] [Module.Finite R M] (x : M) :
x ⨅ (i : ), I ^ i ∃ (r : I), r x = x
theorem Ideal.iInf_pow_smul_eq_bot_of_le_jacobson {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) [IsNoetherianRing R] [Module.Finite R M] (h : I .jacobson) :
⨅ (i : ), I ^ i =
theorem Ideal.iInf_pow_smul_eq_bot_of_isLocalRing {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) [IsNoetherianRing R] [IsLocalRing R] [Module.Finite R M] (h : I ) :
⨅ (i : ), I ^ i =
@[deprecated Ideal.iInf_pow_smul_eq_bot_of_isLocalRing (since := "2024-11-12")]
theorem Ideal.iInf_pow_smul_eq_bot_of_localRing {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) [IsNoetherianRing R] [IsLocalRing R] [Module.Finite R M] (h : I ) :
⨅ (i : ), I ^ i =

Alias of Ideal.iInf_pow_smul_eq_bot_of_isLocalRing.

theorem Ideal.iInf_pow_eq_bot_of_isLocalRing {R : Type u_1} [CommRing R] (I : Ideal R) [IsNoetherianRing R] [IsLocalRing R] (h : I ) :
⨅ (i : ), I ^ i =

Krull's intersection theorem for noetherian local rings.

@[deprecated Ideal.iInf_pow_eq_bot_of_isLocalRing (since := "2024-11-12")]
theorem Ideal.iInf_pow_eq_bot_of_localRing {R : Type u_1} [CommRing R] (I : Ideal R) [IsNoetherianRing R] [IsLocalRing R] (h : I ) :
⨅ (i : ), I ^ i =

Alias of Ideal.iInf_pow_eq_bot_of_isLocalRing.


Krull's intersection theorem for noetherian local rings.

Also see Ideal.isIdempotentElem_iff_eq_bot_or_top for integral domains.

@[deprecated Ideal.isIdempotentElem_iff_eq_bot_or_top_of_isLocalRing (since := "2024-11-12")]

Alias of Ideal.isIdempotentElem_iff_eq_bot_or_top_of_isLocalRing.


Also see Ideal.isIdempotentElem_iff_eq_bot_or_top for integral domains.

theorem Ideal.iInf_pow_smul_eq_bot_of_noZeroSMulDivisors {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) [IsNoetherianRing R] [NoZeroSMulDivisors R M] [Module.Finite R M] (h : I ) :
⨅ (i : ), I ^ i =
theorem Ideal.iInf_pow_eq_bot_of_isDomain {R : Type u_1} [CommRing R] (I : Ideal R) [IsNoetherianRing R] [IsDomain R] (h : I ) :
⨅ (i : ), I ^ i =

Krull's intersection theorem for noetherian domains.