Artinian rings and modules #
A module satisfying these equivalent conditions is said to be an Artinian R-module
if every decreasing chain of submodules is eventually constant, or equivalently,
if the relation <
on submodules is well founded.
A ring is said to be left (or right) Artinian if it is Artinian as a left (or right) module over itself, or simply Artinian if it is both left and right Artinian.
Main definitions #
Let R
be a ring and let M
and P
be R
-modules. Let N
be an R
-submodule of M
.
IsArtinian R M
is the proposition thatM
is an ArtinianR
-module. It is a class, implemented as the predicate that the<
relation on submodules is well founded.IsArtinianRing R
is the proposition thatR
is a left Artinian ring.
Main results #
IsArtinianRing.primeSpectrum_finite
,IsArtinianRing.isMaximal_of_isPrime
: there are only finitely prime ideals in a commutative artinian ring, and each of them is maximal.IsArtinianRing.equivPi
: a reduced commutative artinian ringR
is isomorphic to a finite product of fields (and therefore is a semisimple ring and a decomposition monoid; moreoverR[X]
is also a decomposition monoid).
References #
Tags #
Artinian, artinian, Artinian ring, Artinian module, artinian ring, artinian module
IsArtinian R M
is the proposition that M
is an Artinian R
-module,
implemented as the well-foundedness of submodule inclusion.
Equations
- IsArtinian R M = WellFoundedLT (Submodule R M)
Instances For
A module is Artinian iff every nonempty set of submodules has a minimal submodule among them.
If ∀ I > J, P I
implies P J
, then P
holds for all submodules.
Any injective endomorphism of an Artinian module is surjective.
Any injective endomorphism of an Artinian module is bijective.
A sequence f
of submodules of an artinian module,
with the supremum f (n+1)
and the infimum of f 0
, ..., f n
being ⊤,
is eventually ⊤.
A version of isArtinian_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 ℝ
).
For any endomorphism of an Artinian module, any sufficiently high iterate has codisjoint kernel and range.
This is the Fitting decomposition of the module M
with respect to the endomorphism f
.
See also LinearMap.isCompl_iSup_ker_pow_iInf_range_pow
for an alternative spelling.
This is the Fitting decomposition of the module M
with respect to the endomorphism f
.
See also LinearMap.eventually_isCompl_ker_pow_range_pow
for an alternative spelling.
If M / S / R
is a scalar tower, and M / R
is Artinian, then M / S
is also Artinian.
A ring is Artinian if it is Artinian as a module over itself.
Strictly speaking, this should be called IsLeftArtinianRing
but we omit the Left
for
convenience in the commutative case. For a right Artinian ring, use IsArtinian Rᵐᵒᵖ R
.
For equivalent definitions, see Mathlib.RingTheory.Artinian.Ring
.
Equations
- IsArtinianRing R = IsArtinian R R
Instances For
In a module over an artinian ring, the submodule generated by finitely many vectors is artinian.
A temporary field instance on the quotients by maximal ideals.
Equations
Instances For
The quotient of a commutative artinian ring by its nilradical is isomorphic to a finite product of fields, namely the quotients by the maximal ideals.
Equations
Instances For
A reduced commutative artinian ring is isomorphic to a finite product of fields, namely the quotients by the maximal ideals.
Equations
- IsArtinianRing.equivPi R = (RingEquiv.quotientBot R).symm.trans ((Ideal.quotEquivOfEq ⋯).trans (IsArtinianRing.quotNilradicalEquivPi R))