Documentation

Mathlib.RingTheory.Ideal.MinimalPrime.Localization

Minimal primes and localization #

We provide various results concerning the minimal primes above an ideal that needs the theory of localizations

Main results #

theorem Ideal.iUnion_minimalPrimes {R : Type u_1} [CommSemiring R] {I : Ideal R} :
pI.minimalPrimes, p = {x : R | yI.radical, x * y I.radical}
theorem Ideal.exists_mul_mem_of_mem_minimalPrimes {R : Type u_1} [CommSemiring R] {I p : Ideal R} (hp : p I.minimalPrimes) {x : R} (hx : x p) :
yI, x * y I

minimal primes are contained in zero divisors.

theorem Ideal.exists_comap_eq_of_mem_minimalPrimes_of_injective {R : Type u_1} {S : Type u_2} [CommSemiring R] [CommSemiring S] {f : R →+* S} (hf : Function.Injective f) (p : Ideal R) (H : p minimalPrimes R) :
∃ (p' : Ideal S), p'.IsPrime comap f p' = p
theorem Ideal.exists_comap_eq_of_mem_minimalPrimes {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] {I : Ideal S} (f : R →+* S) (p : Ideal R) (H : p (comap f I).minimalPrimes) :
∃ (p' : Ideal S), p'.IsPrime I p' comap f p' = p
theorem Ideal.exists_minimalPrimes_comap_eq {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] {I : Ideal S} (f : R →+* S) (p : Ideal R) (H : p (comap f I).minimalPrimes) :
p'I.minimalPrimes, comap f p' = p
theorem Ideal.minimal_primes_comap_of_surjective {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] {f : R →+* S} (hf : Function.Surjective f) {I J : Ideal S} (h : J I.minimalPrimes) :
theorem IsLocalization.AtPrime.prime_unique_of_minimal {R : Type u_1} [CommSemiring R] {I : Ideal R} [hI : I.IsPrime] (hMin : I minimalPrimes R) {S : Type u_2} [CommSemiring S] [Algebra R S] [IsLocalization.AtPrime S I] {J K : Ideal S} [J.IsPrime] [K.IsPrime] :
J = K