Documentation

Mathlib.RingTheory.Noetherian

Noetherian rings and modules #

The following are equivalent for a module M over a ring R:

  1. Every increasing chain of submodules M₁ ⊆ M₂ ⊆ M₃ ⊆ ⋯ eventually stabilises.
  2. Every submodule is finitely generated.

A module satisfying these equivalent conditions is said to be a Noetherian R-module. A ring is a Noetherian ring if it is Noetherian as a module over itself.

(Note that we do not assume yet that our rings are commutative, so perhaps this should be called "left Noetherian". To avoid cumbersome names once we specialize to the commutative case, we don't make this explicit in the declaration names.)

Main definitions #

Let R be a ring and let M and P be R-modules. Let N be an R-submodule of M.

Main statements #

Note that the Hilbert basis theorem, that if a commutative ring R is Noetherian then so is R[X], is proved in RingTheory.Polynomial.

References #

Tags #

Noetherian, noetherian, Noetherian ring, Noetherian module, noetherian ring, noetherian module

class IsNoetherian (R : Type u_1) (M : Type u_2) [Semiring R] [AddCommMonoid M] [Module R M] :

IsNoetherian R M is the proposition that M is a Noetherian R-module, implemented as the predicate that all R-submodules of M are finitely generated.

  • noetherian : ∀ (s : Submodule R M), Submodule.FG s

    IsNoetherian R M is the proposition that M is a Noetherian R-module, implemented as the predicate that all R-submodules of M are finitely generated.

Instances
    theorem isNoetherian_def {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] :
    IsNoetherian R M ∀ (s : Submodule R M), Submodule.FG s

    An R-module is Noetherian iff all its submodules are finitely-generated.

    theorem isNoetherian_submodule {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] {N : Submodule R M} :
    IsNoetherian R N sN, Submodule.FG s
    theorem isNoetherian_submodule_left {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] {N : Submodule R M} :
    IsNoetherian R N ∀ (s : Submodule R M), Submodule.FG (N s)
    theorem isNoetherian_submodule_right {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] {N : Submodule R M} :
    IsNoetherian R N ∀ (s : Submodule R M), Submodule.FG (s N)
    instance isNoetherian_submodule' {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] [IsNoetherian R M] (N : Submodule R M) :
    Equations
    • =
    theorem isNoetherian_of_le {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] {s : Submodule R M} {t : Submodule R M} [ht : IsNoetherian R t] (h : s t) :
    theorem isNoetherian_of_surjective {R : Type u_1} (M : Type u_2) {P : Type u_3} [Semiring R] [AddCommMonoid M] [AddCommMonoid P] [Module R M] [Module R P] (f : M →ₗ[R] P) (hf : LinearMap.range f = ) [IsNoetherian R M] :
    theorem isNoetherian_of_linearEquiv {R : Type u_1} {M : Type u_2} {P : Type u_3} [Semiring R] [AddCommMonoid M] [AddCommMonoid P] [Module R M] [Module R P] (f : M ≃ₗ[R] P) [IsNoetherian R M] :
    theorem isNoetherian_top_iff {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] :
    theorem isNoetherian_of_injective {R : Type u_1} {M : Type u_2} {P : Type u_3} [Semiring R] [AddCommMonoid M] [AddCommMonoid P] [Module R M] [Module R P] [IsNoetherian R P] (f : M →ₗ[R] P) (hf : Function.Injective f) :
    theorem fg_of_injective {R : Type u_1} {M : Type u_2} {P : Type u_3} [Semiring R] [AddCommMonoid M] [AddCommMonoid P] [Module R M] [Module R P] [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) (hf : Function.Injective f) :
    instance Module.IsNoetherian.finite (R : Type u_1) (M : Type u_2) [Semiring R] [AddCommMonoid M] [Module R M] [IsNoetherian R M] :
    Equations
    • =
    theorem Module.Finite.of_injective {R : Type u_1} {M : Type u_2} {N : Type u_3} [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [IsNoetherian R N] (f : M →ₗ[R] N) (hf : Function.Injective f) :
    theorem isNoetherian_of_ker_bot {R : Type u_1} {M : Type u_2} {P : Type u_3} [Ring R] [AddCommGroup M] [AddCommGroup P] [Module R M] [Module R P] [IsNoetherian R P] (f : M →ₗ[R] P) (hf : LinearMap.ker f = ) :
    theorem fg_of_ker_bot {R : Type u_1} {M : Type u_2} {P : Type u_3} [Ring R] [AddCommGroup M] [AddCommGroup P] [Module R M] [Module R P] [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) (hf : LinearMap.ker f = ) :
    instance isNoetherian_prod {R : Type u_1} {M : Type u_2} {P : Type u_3} [Ring R] [AddCommGroup M] [AddCommGroup P] [Module R M] [Module R P] [IsNoetherian R M] [IsNoetherian R P] :
    Equations
    • =
    instance isNoetherian_pi {R : Type u_4} {ι : Type u_5} {M : ιType u_6} [Ring R] [(i : ι) → AddCommGroup (M i)] [(i : ι) → Module R (M i)] [Finite ι] [∀ (i : ι), IsNoetherian R (M i)] :
    IsNoetherian R ((i : ι) → M i)
    Equations
    • =
    instance isNoetherian_pi' {R : Type u_4} {ι : Type u_5} {M : Type u_6} [Ring R] [AddCommGroup M] [Module R M] [Finite ι] [IsNoetherian R M] :
    IsNoetherian R (ιM)

    A version of isNoetherian_pi for non-dependent functions. We need this instance because sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to prove that ι → ℝ is finite dimensional over ).

    Equations
    • =
    instance isNoetherian_linearMap_pi (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsNoetherian R M] {ι : Type u_4} [Finite ι] :
    IsNoetherian R ((ιR) →ₗ[R] M)
    Equations
    • =
    instance isNoetherian_linearMap (R : Type u_1) (M : Type u_2) (N : Type u_3) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [IsNoetherian R M] [Module.Finite R N] :
    Equations
    • =
    theorem isNoetherian_iff_wellFounded {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] :
    IsNoetherian R M WellFounded fun (x x_1 : Submodule R M) => x > x_1
    theorem isNoetherian_iff_fg_wellFounded {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] :
    IsNoetherian R M WellFounded fun (x x_1 : { N : Submodule R M // Submodule.FG N }) => x > x_1
    theorem wellFounded_submodule_gt (R : Type u_4) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] [IsNoetherian R M] :
    WellFounded fun (x x_1 : Submodule R M) => x > x_1
    theorem set_has_maximal_iff_noetherian {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] :
    (∀ (a : Set (Submodule R M)), Set.Nonempty a∃ M' ∈ a, Ia, ¬M' < I) IsNoetherian R M

    A module is Noetherian iff every nonempty set of submodules has a maximal submodule among them.

    theorem monotone_stabilizes_iff_noetherian {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] :
    (∀ (f : →o Submodule R M), ∃ (n : ), ∀ (m : ), n mf n = f m) IsNoetherian R M

    A module is Noetherian iff every increasing chain of submodules stabilizes.

    theorem eventuallyConst_of_isNoetherian {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] [IsNoetherian R M] (f : →o Submodule R M) :
    Filter.EventuallyConst (f) Filter.atTop
    theorem IsNoetherian.induction {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] [IsNoetherian R M] {P : Submodule R MProp} (hgt : ∀ (I : Submodule R M), (J > I, P J)P I) (I : Submodule R M) :
    P I

    If ∀ I > J, P I implies P J, then P holds for all submodules.

    theorem Submodule.finite_ne_bot_of_independent {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] {ι : Type u_4} {N : ιSubmodule R M} (h : CompleteLattice.Independent N) :
    Set.Finite {i : ι | N i }
    theorem LinearIndependent.finite_of_isNoetherian {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] [Nontrivial R] {ι : Type u_4} {v : ιM} (hv : LinearIndependent R v) :

    A linearly-independent family of vectors in a module over a non-trivial ring must be finite if the module is Noetherian.

    theorem LinearIndependent.set_finite_of_isNoetherian {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] [Nontrivial R] {s : Set M} (hi : LinearIndependent (ι := { x : M // x s }) R Subtype.val) :
    @[deprecated LinearIndependent.set_finite_of_isNoetherian]
    theorem finite_of_linearIndependent {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] [Nontrivial R] {s : Set M} (hi : LinearIndependent (ι := { x : M // x s }) R Subtype.val) :

    Alias of LinearIndependent.set_finite_of_isNoetherian.

    theorem isNoetherian_of_range_eq_ker {R : Type u_1} {M : Type u_2} {P : Type u_3} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] [IsNoetherian R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) (hf : Function.Injective f) (hg : Function.Surjective g) (h : LinearMap.range f = LinearMap.ker g) :

    If the first and final modules in a short exact sequence are Noetherian, then the middle module is also Noetherian.

    theorem LinearMap.eventually_disjoint_ker_pow_range_pow {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] (f : M →ₗ[R] M) :
    ∀ᶠ (n : ) in Filter.atTop, Disjoint (LinearMap.ker (f ^ n)) (LinearMap.range (f ^ n))

    For an endomorphism of a Noetherian module, any sufficiently large iterate has disjoint kernel and range.

    theorem LinearMap.eventually_iSup_ker_pow_eq {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] (f : M →ₗ[R] M) :
    ∀ᶠ (n : ) in Filter.atTop, ⨆ (m : ), LinearMap.ker (f ^ m) = LinearMap.ker (f ^ n)

    Any surjective endomorphism of a Noetherian module is injective.

    Any surjective endomorphism of a Noetherian module is bijective.

    theorem IsNoetherian.disjoint_partialSups_eventually_bot {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherian R M] (f : Submodule R M) (h : ∀ (n : ), Disjoint ((partialSups f) n) (f (n + 1))) :
    ∃ (n : ), ∀ (m : ), n mf m =

    A sequence f of submodules of a noetherian module, with f (n+1) disjoint from the supremum of f 0, ..., f n, is eventually zero.

    noncomputable def IsNoetherian.equivPUnitOfProdInjective {R : Type u_1} {M : Type u_2} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] [IsNoetherian R M] (f : M × N →ₗ[R] M) (i : Function.Injective f) :

    If M ⊕ N embeds into M, for M noetherian over R, then N is trivial.

    Equations
    Instances For
      @[reducible]
      def IsNoetherianRing (R : Type u_1) [Semiring R] :

      A (semi)ring is Noetherian if it is Noetherian as a module over itself, i.e. all its ideals are finitely generated.

      Equations
      Instances For

        A ring is Noetherian if and only if all its ideals are finitely-generated.

        instance isNoetherian_of_finite (R : Type u_1) (M : Type u_2) [Finite M] [Semiring R] [AddCommMonoid M] [Module R M] :
        Equations
        • =
        instance isNoetherian_of_subsingleton (R : Type u_1) (M : Type u_2) [Subsingleton R] [Semiring R] [AddCommMonoid M] [Module R M] :

        Modules over the trivial ring are Noetherian.

        Equations
        • =
        theorem isNoetherian_of_submodule_of_noetherian (R : Type u_1) (M : Type u_2) [Semiring R] [AddCommMonoid M] [Module R M] (N : Submodule R M) (h : IsNoetherian R M) :
        instance Submodule.Quotient.isNoetherian {R : Type u_1} [Ring R] {M : Type u_2} [AddCommGroup M] [Module R M] (N : Submodule R M) [h : IsNoetherian R M] :
        Equations
        • =
        theorem isNoetherian_of_tower (R : Type u_1) {S : Type u_2} {M : Type u_3} [Semiring R] [Semiring S] [AddCommMonoid M] [SMul R S] [Module S M] [Module R M] [IsScalarTower R S M] (h : IsNoetherian R M) :

        If M / S / R is a scalar tower, and M / R is Noetherian, then M / S is also noetherian.

        theorem isNoetherian_of_fg_of_noetherian {R : Type u_1} {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] (N : Submodule R M) [I : IsNoetherianRing R] (hN : Submodule.FG N) :
        Equations
        • =
        theorem isNoetherian_span_of_finite (R : Type u_1) {M : Type u_2} [Ring R] [AddCommGroup M] [Module R M] [IsNoetherianRing R] {A : Set M} (hA : Set.Finite A) :

        In a module over a Noetherian ring, the submodule generated by finitely many vectors is Noetherian.

        theorem isNoetherianRing_of_surjective (R : Type u_1) [Ring R] (S : Type u_2) [Ring S] (f : R →+* S) (hf : Function.Surjective f) [H : IsNoetherianRing R] :
        instance isNoetherianRing_range {R : Type u_1} [Ring R] {S : Type u_2} [Ring S] (f : R →+* S) [IsNoetherianRing R] :
        Equations
        • =
        theorem isNoetherianRing_of_ringEquiv (R : Type u_1) [Ring R] {S : Type u_2} [Ring S] (f : R ≃+* S) [IsNoetherianRing R] :