Documentation

Mathlib.RingTheory.Nakayama

Nakayama's lemma #

This file contains some alternative statements of Nakayama's Lemma as found in Stacks: Nakayama's Lemma.

Main statements #

Note that a version of Statement (1) in Stacks: Nakayama's Lemma can be found in RingTheory.Finiteness under the name Submodule.exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul

References #

Tags #

Nakayama, Jacobson

theorem Submodule.eq_smul_of_le_smul_of_le_jacobson {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {J : Ideal R} {N : Submodule R M} (hN : Submodule.FG N) (hIN : N I N) (hIjac : I Ideal.jacobson J) :
N = J N

Nakayama's Lemma - A slightly more general version of (2) in Stacks 00DV. See also eq_bot_of_le_smul_of_le_jacobson_bot for the special case when J = ⊥.

theorem Submodule.eq_bot_of_le_smul_of_le_jacobson_bot {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) (hN : Submodule.FG N) (hIN : N I N) (hIjac : I Ideal.jacobson ) :
N =

Nakayama's Lemma - Statement (2) in Stacks 00DV. See also eq_smul_of_le_smul_of_le_jacobson for a generalisation to the jacobson of any ideal

theorem Submodule.sup_eq_sup_smul_of_le_smul_of_le_jacobson {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {J : Ideal R} {N : Submodule R M} {N' : Submodule R M} (hN' : Submodule.FG N') (hIJ : I Ideal.jacobson J) (hNN : N' N I N') :
N N' = N J N'
theorem Submodule.sup_smul_eq_sup_smul_of_le_smul_of_le_jacobson {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {J : Ideal R} {N : Submodule R M} {N' : Submodule R M} (hN' : Submodule.FG N') (hIJ : I Ideal.jacobson J) (hNN : N' N I N') :
N I N' = N J N'

Nakayama's Lemma - A slightly more general version of (4) in Stacks 00DV. See also smul_le_of_le_smul_of_le_jacobson_bot for the special case when J = ⊥.

theorem Submodule.le_of_le_smul_of_le_jacobson_bot {R : Type u_3} {M : Type u_4} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {N : Submodule R M} {N' : Submodule R M} (hN' : Submodule.FG N') (hIJ : I Ideal.jacobson ) (hNN : N' N I N') :
N' N
theorem Submodule.smul_le_of_le_smul_of_le_jacobson_bot {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {I : Ideal R} {N : Submodule R M} {N' : Submodule R M} (hN' : Submodule.FG N') (hIJ : I Ideal.jacobson ) (hNN : N' N I N') :
I N' N

Nakayama's Lemma - Statement (4) in Stacks 00DV. See also sup_smul_eq_sup_smul_of_le_smul_of_le_jacobson for a generalisation to the jacobson of any ideal