mathlib3 documentation

ring_theory.finite_type

Finiteness conditions in commutative algebra #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

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

Main declarations #

@[class]
structure algebra.finite_type (R : Type u_1) (A : Type u_2) [comm_semiring R] [semiring A] [algebra R A] :
Prop

An algebra over a commutative semiring is of finite_type if it is finitely generated over the base ring as algebra.

Instances of this typeclass
@[protected, instance]
def module.finite.finite_type {R : Type u_1} (A : Type u_2) [comm_semiring R] [semiring A] [algebra R A] [hRA : module.finite R A] :
@[protected]
theorem algebra.finite_type.of_surjective {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] [algebra R B] (hRA : algebra.finite_type R A) (f : A →ₐ[R] B) (hf : function.surjective f) :
theorem algebra.finite_type.equiv {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] [algebra R B] (hRA : algebra.finite_type R A) (e : A ≃ₐ[R] B) :
theorem algebra.finite_type.trans {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] [algebra R B] [algebra A B] [is_scalar_tower R A B] (hRA : algebra.finite_type R A) (hAB : algebra.finite_type A B) :

An algebra is finitely generated if and only if it is a quotient of a polynomial ring whose variables are indexed by a finset.

An algebra is finitely generated if and only if it is a quotient of a polynomial ring whose variables are indexed by a fintype.

An algebra is finitely generated if and only if it is a quotient of a polynomial ring in n variables.

@[protected, instance]
def algebra.finite_type.prod {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] [algebra R B] [hA : algebra.finite_type R A] [hB : algebra.finite_type R B] :
theorem subalgebra.fg_iff_finite_type {R : Type u_1} {A : Type u_2} [comm_semiring R] [semiring A] [algebra R A] (S : subalgebra R A) :
def ring_hom.finite_type {A : Type u_1} {B : Type u_2} [comm_ring A] [comm_ring B] (f : A →+* B) :
Prop

A ring morphism A →+* B is of finite_type if B is finitely generated as A-algebra.

Equations
theorem ring_hom.finite.finite_type {A : Type u_1} {B : Type u_2} [comm_ring A] [comm_ring B] {f : A →+* B} (hf : f.finite) :
theorem ring_hom.finite_type.comp_surjective {A : Type u_1} {B : Type u_2} {C : Type u_3} [comm_ring A] [comm_ring B] [comm_ring C] {f : A →+* B} {g : B →+* C} (hf : f.finite_type) (hg : function.surjective g) :
theorem ring_hom.finite_type.of_surjective {A : Type u_1} {B : Type u_2} [comm_ring A] [comm_ring B] (f : A →+* B) (hf : function.surjective f) :
theorem ring_hom.finite_type.comp {A : Type u_1} {B : Type u_2} {C : Type u_3} [comm_ring A] [comm_ring B] [comm_ring C] {g : B →+* C} {f : A →+* B} (hg : g.finite_type) (hf : f.finite_type) :
theorem ring_hom.finite_type.of_finite {A : Type u_1} {B : Type u_2} [comm_ring A] [comm_ring B] {f : A →+* B} (hf : f.finite) :
theorem ring_hom.finite.to_finite_type {A : Type u_1} {B : Type u_2} [comm_ring A] [comm_ring B] {f : A →+* B} (hf : f.finite) :

Alias of ring_hom.finite_type.of_finite.

theorem ring_hom.finite_type.of_comp_finite_type {A : Type u_1} {B : Type u_2} {C : Type u_3} [comm_ring A] [comm_ring B] [comm_ring C] {f : A →+* B} {g : B →+* C} (h : (g.comp f).finite_type) :
def alg_hom.finite_type {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [comm_ring B] [algebra R A] [algebra R B] (f : A →ₐ[R] B) :
Prop

An algebra morphism A →ₐ[R] B is of finite_type if it is of finite type as ring morphism. In other words, if B is finitely generated as A-algebra.

Equations
theorem alg_hom.finite.finite_type {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [comm_ring B] [algebra R A] [algebra R B] {f : A →ₐ[R] B} (hf : f.finite) :
theorem alg_hom.finite_type.id (R : Type u_1) (A : Type u_2) [comm_ring R] [comm_ring A] [algebra R A] :
theorem alg_hom.finite_type.comp {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [comm_ring R] [comm_ring A] [comm_ring B] [comm_ring C] [algebra R A] [algebra R B] [algebra R C] {g : B →ₐ[R] C} {f : A →ₐ[R] B} (hg : g.finite_type) (hf : f.finite_type) :
theorem alg_hom.finite_type.comp_surjective {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [comm_ring R] [comm_ring A] [comm_ring B] [comm_ring C] [algebra R A] [algebra R B] [algebra R C] {f : A →ₐ[R] B} {g : B →ₐ[R] C} (hf : f.finite_type) (hg : function.surjective g) :
theorem alg_hom.finite_type.of_surjective {R : Type u_1} {A : Type u_2} {B : Type u_3} [comm_ring R] [comm_ring A] [comm_ring B] [algebra R A] [algebra R B] (f : A →ₐ[R] B) (hf : function.surjective f) :
theorem alg_hom.finite_type.of_comp_finite_type {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [comm_ring R] [comm_ring A] [comm_ring B] [comm_ring C] [algebra R A] [algebra R B] [algebra R C] {f : A →ₐ[R] B} {g : B →ₐ[R] C} (h : (g.comp f).finite_type) :

An element of add_monoid_algebra R M is in the subalgebra generated by its support.

If a set S generates, as algebra, add_monoid_algebra R M, then the set of supports of elements of S generates add_monoid_algebra R M.

If a set S generates, as algebra, add_monoid_algebra R M, then the image of the union of the supports of elements of S generates add_monoid_algebra R M.

If add_monoid_algebra R M is of finite type, there there is a G : finset M such that its image generates, as algera, add_monoid_algebra R M.

The image of an element m : M in add_monoid_algebra R M belongs the submodule generated by S : set M if and only if m ∈ S.

If the image of an element m : M in add_monoid_algebra R M belongs the submodule generated by the closure of some S : set M then m ∈ closure S.

If a set S generates an additive monoid M, then the image of M generates, as algebra, add_monoid_algebra R M.

@[protected, instance]

If an additive monoid M is finitely generated then add_monoid_algebra R M is of finite type.

An additive monoid M is finitely generated if and only if add_monoid_algebra R M is of finite type.

If add_monoid_algebra R M is of finite type then M is finitely generated.

An additive group G is finitely generated if and only if add_monoid_algebra R G is of finite type.

An element of monoid_algebra R M is in the subalgebra generated by its support.

theorem monoid_algebra.support_gen_of_gen {R : Type u_1} {M : Type u_2} [comm_semiring R] [monoid M] {S : set (monoid_algebra R M)} (hS : algebra.adjoin R S = ) :

If a set S generates, as algebra, monoid_algebra R M, then the set of supports of elements of S generates monoid_algebra R M.

theorem monoid_algebra.support_gen_of_gen' {R : Type u_1} {M : Type u_2} [comm_semiring R] [monoid M] {S : set (monoid_algebra R M)} (hS : algebra.adjoin R S = ) :

If a set S generates, as algebra, monoid_algebra R M, then the image of the union of the supports of elements of S generates monoid_algebra R M.

If monoid_algebra R M is of finite type, there there is a G : finset M such that its image generates, as algera, monoid_algebra R M.

theorem monoid_algebra.of_mem_span_of_iff {R : Type u_1} {M : Type u_2} [comm_ring R] [comm_monoid M] [nontrivial R] {m : M} {S : set M} :

The image of an element m : M in monoid_algebra R M belongs the submodule generated by S : set M if and only if m ∈ S.

If the image of an element m : M in monoid_algebra R M belongs the submodule generated by the closure of some S : set M then m ∈ closure S.

If a set S generates a monoid M, then the image of M generates, as algebra, monoid_algebra R M.

@[protected, instance]

If a monoid M is finitely generated then monoid_algebra R M is of finite type.

A monoid M is finitely generated if and only if monoid_algebra R M is of finite type.

If monoid_algebra R M is of finite type then M is finitely generated.

A group G is finitely generated if and only if add_monoid_algebra R G is of finite type.

noncomputable def module_polynomial_of_endo {R : Type u_1} [comm_ring R] {M : Type u_2} [add_comm_group M] [module R M] (f : M →ₗ[R] M) :

The structure of a module M over a ring R as a module over R[X] when given a choice of how X acts by choosing a linear map f : M →ₗ[R] M

Equations
theorem module_polynomial_of_endo_smul_def {R : Type u_1} [comm_ring R] {M : Type u_2} [add_comm_group M] [module R M] (f : M →ₗ[R] M) (n : polynomial R) (a : M) :

A theorem/proof by Vasconcelos, given a finite module M over a commutative ring, any surjective endomorphism of M is also injective. Based on, https://math.stackexchange.com/a/239419/31917, https://www.ams.org/journals/tran/1969-138-00/S0002-9947-1969-0238839-5/. This is similar to is_noetherian.injective_of_surjective_endomorphism but only applies in the commutative case, but does not use a Noetherian hypothesis.