Finiteness conditions in commutative algebra #
In this file we define a notion of finiteness that is common in commutative algebra.
Main declarations #
Submodule.FG
,Ideal.FG
These express that some object is finitely generated as submodule over some base ring.Module.Finite
,RingHom.Finite
,AlgHom.Finite
all of these express that some object is finitely generated as module over some base ring.
def
Submodule.FG
{R : Type u_1}
{M : Type u_2}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(N : Submodule R M)
:
A submodule of M
is finitely generated if it is the span of a finite subset of M
.
Equations
- N.FG = ∃ (S : Finset M), Submodule.span R ↑S = N
Instances For
theorem
Submodule.fg_def
{R : Type u_1}
{M : Type u_2}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{N : Submodule R M}
:
N.FG ↔ ∃ (S : Set M), S.Finite ∧ Submodule.span R S = N
theorem
Submodule.fg_iff_addSubmonoid_fg
{M : Type u_2}
[AddCommMonoid M]
(P : Submodule ℕ M)
:
P.FG ↔ P.FG
theorem
Submodule.fg_iff_add_subgroup_fg
{G : Type u_3}
[AddCommGroup G]
(P : Submodule ℤ G)
:
P.FG ↔ P.toAddSubgroup.FG
theorem
Submodule.fg_iff_exists_fin_generating_family
{R : Type u_1}
{M : Type u_2}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{N : Submodule R M}
:
theorem
Submodule.fg_iff_exists_finite_generating_family
{A : Type u}
[Semiring A]
{M : Type v}
[AddCommMonoid M]
[Module A M]
{N : Submodule A M}
:
theorem
Module.finite_def
{R : Type u_6}
{M : Type u_7}
[Semiring R]
[AddCommMonoid M]
[Module R M]
:
Module.Finite R M ↔ ⊤.FG
theorem
Module.Finite.iff_addMonoid_fg
{M : Type u_6}
[AddCommMonoid M]
:
Module.Finite ℕ M ↔ AddMonoid.FG M
theorem
Module.Finite.iff_addGroup_fg
{G : Type u_6}
[AddCommGroup G]
:
Module.Finite ℤ G ↔ AddGroup.FG G
theorem
Module.Finite.exists_fin
{R : Type u_1}
{M : Type u_4}
[Semiring R]
[AddCommMonoid M]
[Module R M]
[Module.Finite R M]
:
See also Module.Finite.exists_fin'
.
def
AlgHom.Finite
{R : Type u_1}
{A : Type u_2}
{B : Type u_3}
[CommRing R]
[CommRing A]
[CommRing B]
[Algebra R A]
[Algebra R B]
(f : A →ₐ[R] B)
:
An algebra morphism A →ₐ[R] B
is finite if it is finite as ring morphism.
In other words, if B
is finitely generated as A
-module.
Equations
- f.Finite = f.Finite