Documentation

Mathlib.RingTheory.Ideal.GoingUp

Ideals over/under ideals in integral extensions #

This file proves some going-up results for integral algebras.

Implementation notes #

The proofs of the comap_ne_bot and comap_lt_comap families use an approach specific for their situation: we construct an element in I.comap f from the coefficients of a minimal polynomial. Once mathlib has more material on the localization at a prime ideal, the results can be proven using more general going-up/going-down theory.

theorem Ideal.coeff_zero_mem_comap_of_root_mem_of_eval_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I : Ideal S} {r : S} (hr : r I) {p : Polynomial R} (hp : Polynomial.eval₂ f r p I) :
p.coeff 0 comap f I
theorem Ideal.coeff_zero_mem_comap_of_root_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I : Ideal S} {r : S} (hr : r I) {p : Polynomial R} (hp : Polynomial.eval₂ f r p = 0) :
p.coeff 0 comap f I
theorem Ideal.exists_coeff_ne_zero_mem_comap_of_non_zero_divisor_root_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I : Ideal S} {r : S} (r_non_zero_divisor : ∀ {x : S}, x * r = 0x = 0) (hr : r I) {p : Polynomial R} :
p 0Polynomial.eval₂ f r p = 0∃ (i : ), p.coeff i 0 p.coeff i comap f I

Let P be an ideal in R[x]. The map R[x]/P → (R / (P ∩ R))[x] / (P / (P ∩ R)) is injective.

The identity in this lemma asserts that the "obvious" square

    R    → (R / (P ∩ R))
    ↓          ↓
R[x] / P → (R / (P ∩ R))[x] / (P / (P ∩ R))

commutes. It is used, for instance, in the proof of quotient_mk_comp_C_is_integral_of_jacobson, in the file Mathlib.RingTheory.Jacobson.Polynomial.

theorem Ideal.exists_nonzero_mem_of_ne_bot {R : Type u_1} [CommRing R] {P : Ideal (Polynomial R)} (Pb : P ) (hP : ∀ (x : R), Polynomial.C x Px = 0) :

This technical lemma asserts the existence of a polynomial p in an ideal P ⊂ R[x] that is non-zero in the quotient R / (P ∩ R) [x]. The assumptions are equivalent to P ≠ 0 and P ∩ R = (0).

theorem Ideal.exists_coeff_ne_zero_mem_comap_of_root_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I : Ideal S} [IsDomain S] {r : S} (r_ne_zero : r 0) (hr : r I) {p : Polynomial R} :
p 0Polynomial.eval₂ f r p = 0∃ (i : ), p.coeff i 0 p.coeff i comap f I
theorem Ideal.exists_coeff_mem_comap_sdiff_comap_of_root_mem_sdiff {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I J : Ideal S} [I.IsPrime] (hIJ : I J) {r : S} (hr : r J \ I) {p : Polynomial R} (p_ne_zero : Polynomial.map (Quotient.mk (comap f I)) p 0) (hpI : Polynomial.eval₂ f r p I) :
∃ (i : ), p.coeff i (comap f J) \ (comap f I)
theorem Ideal.comap_lt_comap_of_root_mem_sdiff {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I J : Ideal S} [I.IsPrime] (hIJ : I J) {r : S} (hr : r J \ I) {p : Polynomial R} (p_ne_zero : Polynomial.map (Quotient.mk (comap f I)) p 0) (hp : Polynomial.eval₂ f r p I) :
comap f I < comap f J
theorem Ideal.mem_of_one_mem {S : Type u_2} [CommRing S] {I : Ideal S} (h : 1 I) (x : S) :
x I
theorem Ideal.comap_lt_comap_of_integral_mem_sdiff {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {I J : Ideal S} [Algebra R S] [hI : I.IsPrime] (hIJ : I J) {x : S} (mem : x J \ I) (integral : IsIntegral R x) :
theorem Ideal.comap_ne_bot_of_root_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {f : R →+* S} {I : Ideal S} [IsDomain S] {r : S} (r_ne_zero : r 0) (hr : r I) {p : Polynomial R} (p_ne_zero : p 0) (hp : Polynomial.eval₂ f r p = 0) :
theorem Ideal.isMaximal_of_isIntegral_of_isMaximal_comap' {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] (f : R →+* S) (hf : f.IsIntegral) (I : Ideal S) [I.IsPrime] (hI : (comap f I).IsMaximal) :
theorem Ideal.comap_ne_bot_of_algebraic_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {I : Ideal S} [Algebra R S] [IsDomain S] {x : S} (x_ne_zero : x 0) (x_mem : x I) (hx : IsAlgebraic R x) :
theorem Ideal.comap_ne_bot_of_integral_mem {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {I : Ideal S} [Algebra R S] [Nontrivial R] [IsDomain S] {x : S} (x_ne_zero : x 0) (x_mem : x I) (hx : IsIntegral R x) :
theorem Ideal.eq_bot_of_comap_eq_bot {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] {I : Ideal S} [Algebra R S] [Nontrivial R] [IsDomain S] [Algebra.IsIntegral R S] (hI : comap (algebraMap R S) I = ) :
I =
theorem Ideal.isMaximal_comap_of_isIntegral_of_isMaximal' {R : Type u_3} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) (hf : f.IsIntegral) (I : Ideal S) [I.IsMaximal] :
theorem Ideal.IsIntegralClosure.comap_lt_comap {R : Type u_1} [CommRing R] (S : Type u_2) [CommRing S] [Algebra R S] {A : Type u_3} [CommRing A] [Algebra R A] [Algebra A S] [IsScalarTower R A S] [IsIntegralClosure A R S] {I J : Ideal A} [I.IsPrime] (I_lt_J : I < J) :
theorem Ideal.IsIntegralClosure.isMaximal_of_isMaximal_comap {R : Type u_1} [CommRing R] (S : Type u_2) [CommRing S] [Algebra R S] {A : Type u_3} [CommRing A] [Algebra R A] [Algebra A S] [IsScalarTower R A S] [IsIntegralClosure A R S] (I : Ideal A) [I.IsPrime] (hI : (comap (algebraMap R A) I).IsMaximal) :
theorem Ideal.IsIntegralClosure.comap_ne_bot {R : Type u_1} [CommRing R] (S : Type u_2) [CommRing S] [Algebra R S] {A : Type u_3} [CommRing A] [Algebra R A] [Algebra A S] [IsScalarTower R A S] [IsIntegralClosure A R S] [IsDomain A] [Nontrivial R] {I : Ideal A} (I_ne_bot : I ) :
theorem Ideal.IsIntegralClosure.eq_bot_of_comap_eq_bot {R : Type u_1} [CommRing R] (S : Type u_2) [CommRing S] [Algebra R S] {A : Type u_3} [CommRing A] [Algebra R A] [Algebra A S] [IsScalarTower R A S] [IsIntegralClosure A R S] [IsDomain A] [Nontrivial R] {I : Ideal A} :
comap (algebraMap R A) I = I =
theorem Ideal.IntegralClosure.comap_lt_comap {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] {I J : Ideal (integralClosure R S)} [I.IsPrime] (I_lt_J : I < J) :
theorem Ideal.IntegralClosure.comap_ne_bot {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] [IsDomain S] [Nontrivial R] {I : Ideal (integralClosure R S)} (I_ne_bot : I ) :
theorem Ideal.exists_ideal_over_prime_of_isIntegral_of_isDomain {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] [IsDomain S] [Algebra.IsIntegral R S] (P : Ideal R) [P.IsPrime] (hP : RingHom.ker (algebraMap R S) P) :
∃ (Q : Ideal S), Q.IsPrime comap (algebraMap R S) Q = P

comap (algebraMap R S) is a surjection from the prime spec of R to prime spec of S. hP : (algebraMap R S).ker ≤ P is a slight generalization of the extension being injective

theorem Ideal.exists_ideal_over_prime_of_isIntegral_of_isPrime {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] [Algebra.IsIntegral R S] (P : Ideal R) [P.IsPrime] (I : Ideal S) [I.IsPrime] (hIP : comap (algebraMap R S) I P) :
QI, Q.IsPrime comap (algebraMap R S) Q = P

More general going-up theorem than exists_ideal_over_prime_of_isIntegral_of_isDomain. TODO: Version of going-up theorem with arbitrary length chains (by induction on this)? Not sure how best to write an ascending chain in Lean

theorem Ideal.exists_ideal_comap_le_prime {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] (P : Ideal R) [P.IsPrime] (I : Ideal S) (hI : comap (algebraMap R S) I P) :
QI, Q.IsPrime comap (algebraMap R S) Q P
theorem Ideal.exists_ideal_over_prime_of_isIntegral {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] [Algebra.IsIntegral R S] (P : Ideal R) [P.IsPrime] (I : Ideal S) (hIP : comap (algebraMap R S) I P) :
QI, Q.IsPrime comap (algebraMap R S) Q = P
theorem Ideal.exists_ideal_over_maximal_of_isIntegral {R : Type u_1} [CommRing R] {S : Type u_2} [CommRing S] [Algebra R S] [Algebra.IsIntegral R S] (P : Ideal R) [P_max : P.IsMaximal] (hP : RingHom.ker (algebraMap R S) P) :
∃ (Q : Ideal S), Q.IsMaximal comap (algebraMap R S) Q = P

comap (algebraMap R S) is a surjection from the max spec of S to max spec of R. hP : (algebraMap R S).ker ≤ P is a slight generalization of the extension being injective

theorem Ideal.map_eq_top_iff_of_ker_le {R : Type u_3} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) {I : Ideal R} (hf₁ : RingHom.ker f I) (hf₂ : f.IsIntegral) :
map f I = I =
theorem Ideal.map_eq_top_iff {R : Type u_3} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) {I : Ideal R} (hf₁ : Function.Injective f) (hf₂ : f.IsIntegral) :
map f I = I =
instance Ideal.IsMaximal.under (A : Type u_1) [CommRing A] {B : Type u_2} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] (P : Ideal B) [P.IsMaximal] :

If B is an integral A-algebra, P is a maximal ideal of B, then the pull back of P is also a maximal ideal of A.

theorem Ideal.IsMaximal.of_liesOver_isMaximal {A : Type u_1} [CommRing A] {B : Type u_2} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] (P : Ideal B) (p : Ideal A) [P.LiesOver p] [hpm : p.IsMaximal] [P.IsPrime] :
theorem Ideal.IsMaximal.of_isMaximal_liesOver {A : Type u_1} [CommRing A] {B : Type u_2} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] (P : Ideal B) (p : Ideal A) [P.LiesOver p] [P.IsMaximal] :
instance Ideal.Quotient.algebra_isIntegral_of_liesOver {A : Type u_1} [CommRing A] {B : Type u_2} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] (P : Ideal B) (p : Ideal A) [P.LiesOver p] :

B ⧸ P is an integral A ⧸ p-algebra if B is a integral A-algebra.

instance Ideal.primesOver.isMaximal {A : Type u_1} [CommRing A] {p : Ideal A} [p.IsMaximal] {B : Type u_2} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B] (Q : (p.primesOver B)) :
(↑Q).IsMaximal