Documentation

Mathlib.RingTheory.Finiteness.Defs

Finiteness conditions in commutative algebra #

In this file we define a notion of finiteness that is common in commutative algebra.

Main declarations #

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
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_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} :
    N.FG ∃ (n : ) (s : Fin nM), Submodule.span R (Set.range s) = N
    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} :
    N.FG ∃ (G : Type w) (_ : Finite G) (g : GM), Submodule.span A (Set.range g) = N
    def Ideal.FG {R : Type u_1} [Semiring R] (I : Ideal R) :

    An ideal of R is finitely generated if it is the span of a finite subset of R.

    This is defeq to Submodule.FG, but unfolds more nicely.

    Equations
    Instances For
      class Module.Finite (R : Type u_1) (M : Type u_4) [Semiring R] [AddCommMonoid M] [Module R M] :

      A module over a semiring is Finite if it is finitely generated as a module.

      • out : .FG

        A module over a semiring is Finite if it is finitely generated as a module.

      Instances
        theorem Module.finite_def {R : Type u_6} {M : Type u_7} [Semiring R] [AddCommMonoid M] [Module R M] :
        theorem Module.Finite.exists_fin {R : Type u_1} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] [Module.Finite R M] :
        ∃ (n : ) (s : Fin nM), Submodule.span R (Set.range s) =

        See also Module.Finite.exists_fin'.

        def RingHom.Finite {A : Type u_1} {B : Type u_2} [CommRing A] [CommRing B] (f : A →+* B) :

        A ring morphism A →+* B is Finite if B is finitely generated as A-module.

        Equations
        Instances For
          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
          Instances For