Documentation

Mathlib.RingTheory.FilteredAlgebra.Basic

The filtration on abelian groups and rings #

In this file, we define the concept of filtration for abelian groups, rings, and modules.

Main definitions #

class IsFiltration {ι : Type u_1} {A : Type u_2} {σ : Type u_3} [Preorder ι] [SetLike σ A] (F : ισ) (F_lt : outParam (ισ)) :

For a family of subsets σ of A, an increasing series of F in σ is a filtration if there is another series F_lt in σ equal to the supremum of F with smaller index.

In the intended applications, σ is a complete lattice, and F_lt is uniquely-determined as F_lt j = ⨆ i < j, F i. Thus F_lt is an implementation detail which allows us defer depending on a complete lattice structure on σ. It also provides the ancillary benefit of giving us better definition control. This is convenient e.g., when the index is .

  • mono : Monotone F
  • is_le {i j : ι} : i < jF i F_lt j
  • is_sup (B : σ) (j : ι) : (∀ i < j, F i B)F_lt j B
Instances
    theorem IsFiltration.F_lt_le_F {ι : Type u_1} {A : Type u_2} {σ : Type u_3} [Preorder ι] [SetLike σ A] (F : ισ) (F_lt : outParam (ισ)) (i : ι) [IsFiltration F F_lt] :
    F_lt i F i
    theorem IsFiltration.mk_int {A : Type u_2} {σ : Type u_3} [SetLike σ A] (F : σ) (mono : Monotone F) :
    IsFiltration F fun (n : ) => F (n - 1)

    A convenience constructor for IsFiltration when the index is the integers.

    class IsRingFiltration {ι : Type u_1} {R : Type u_2} {σ : Type u_3} [OrderedAddCommMonoid ι] [Semiring R] [SetLike σ R] (F : ισ) (F_lt : outParam (ισ)) extends IsFiltration F F_lt, SetLike.GradedMonoid F :

    For a family of subsets σ of semiring R, an increasing series F in σ is a ring filtration if IsFiltration F F_lt and the pointwise multiplication of F i and F j is in F (i + j).

    Instances
      theorem IsRingFiltration.mk_int {R : Type u_2} {σ : Type u_3} [Semiring R] [SetLike σ R] (F : σ) (mono : Monotone F) [SetLike.GradedMonoid F] :
      IsRingFiltration F fun (n : ) => F (n - 1)

      A convenience constructor for IsRingFiltration when the index is the integers.

      class IsModuleFiltration {ι : Type u_1} {ιM : Type u_2} {R : Type u_3} {M : Type u_4} {σ : Type u_5} {σM : Type u_6} [OrderedAddCommMonoid ι] [OrderedAddCommMonoid ιM] [Semiring R] [SetLike σ R] [AddCommMonoid M] [Module R M] [VAdd ι ιM] [SetLike σM M] (F : ισ) (F_lt : outParam (ισ)) [IsRingFiltration F F_lt] (F' : ιMσM) (F'_lt : outParam (ιMσM)) extends IsFiltration F' F'_lt, SetLike.GradedSMul F F' :

      For F satisfying IsRingFiltration F F_lt in a semiring R and σM a family of subsets of a R module M, an increasing series FM in σM is a module filtration if IsFiltration F F_lt and the pointwise scalar multiplication of F i and FM j is in F (i +ᵥ j).

      The index set ιM for the module can be more general, however usually we take ιM = ι.

      Instances
        theorem IsModuleFiltration.mk_int {R : Type u_3} {M : Type u_4} {σ : Type u_5} {σM : Type u_6} [Semiring R] [SetLike σ R] [AddCommMonoid M] [Module R M] [SetLike σM M] (F : σ) (mono : Monotone F) [SetLike.GradedMonoid F] (F' : σM) (mono' : Monotone F') [SetLike.GradedSMul F F'] :
        IsModuleFiltration F (fun (n : ) => F (n - 1)) F' fun (n : ) => F' (n - 1)

        A convenience constructor for IsModuleFiltration when the index is the integers.