This file contains a proof that the radical of any homogeneous ideal is a homogeneous ideal
Main statements #
Ideal.IsHomogeneous.isPrime_iff: for anyI : Ideal A, ifIis homogeneous, thenIis prime if and only ifIis homogeneously prime, i.e.I ≠ ⊤and ifx, yare homogeneous elements such thatx * y ∈ I, then at least one ofx,yis inI.Ideal.IsPrime.homogeneousCore: for anyI : Ideal A, ifIis prime, thenI.homogeneous_core 𝒜(i.e. the largest homogeneous ideal contained inI) is also prime.Ideal.IsHomogeneous.radical: for anyI : Ideal A, ifIis homogeneous, then the radical ofIis homogeneous as well.HomogeneousIdeal.radical: for anyI : HomogeneousIdeal 𝒜,I.radicalis the radical ofIas aHomogeneousIdeal 𝒜.
Implementation details #
Throughout this file, the indexing type ι of grading is assumed to be a
LinearOrderedCancelAddCommMonoid. This might be stronger than necessary but cancelling
property is strictly necessary; for a counterexample of how Ideal.IsHomogeneous.isPrime_iff
fails for a non-cancellative set see Counterexamples/HomogeneousPrimeNotPrime.lean.
Tags #
homogeneous, radical
theorem
Ideal.IsHomogeneous.isPrime_of_homogeneous_mem_or_mem
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
{I : Ideal A}
(hI : IsHomogeneous 𝒜 I)
(I_ne_top : I ≠ ⊤)
(homogeneous_mem_or_mem :
∀ {x y : A}, SetLike.IsHomogeneousElem 𝒜 x → SetLike.IsHomogeneousElem 𝒜 y → x * y ∈ I → x ∈ I ∨ y ∈ I)
:
I.IsPrime
theorem
Ideal.IsHomogeneous.isPrime_iff
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
{I : Ideal A}
(h : IsHomogeneous 𝒜 I)
:
theorem
Ideal.IsPrime.homogeneousCore
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
{I : Ideal A}
(h : I.IsPrime)
:
theorem
Ideal.IsHomogeneous.radical_eq
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
{I : Ideal A}
(hI : IsHomogeneous 𝒜 I)
:
theorem
Ideal.IsHomogeneous.radical
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
{I : Ideal A}
(h : IsHomogeneous 𝒜 I)
:
IsHomogeneous 𝒜 I.radical
def
HomogeneousIdeal.radical
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
(I : HomogeneousIdeal 𝒜)
:
The radical of a homogeneous ideal, as another homogeneous ideal.
Instances For
@[simp]
theorem
HomogeneousIdeal.coe_radical
{ι : Type u_1}
{σ : Type u_2}
{A : Type u_3}
[CommRing A]
[AddCommMonoid ι]
[LinearOrder ι]
[IsOrderedCancelAddMonoid ι]
[SetLike σ A]
[AddSubmonoidClass σ A]
{𝒜 : ι → σ}
[GradedRing 𝒜]
(I : HomogeneousIdeal 𝒜)
: