Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit a18be37

Browse files
committed
feat(ring_theory/ideal/over): Going up theorem for integral extensions (#4064)
The main statement is `exists_ideal_over_prime_of_is_integral` which shows that for an integral extension, every prime ideal of the original ring lies under some prime ideal of the extension ring. `is_field_of_is_integral_of_is_field` is a brute force proof that if `R → S` is an integral extension, and `S` is a field, then `R` is also a field (using the somewhat new `is_field` proposition). `is_maximal_comap_of_is_integral_of_is_maximal` Gives essentially the same statement in terms of maximal ideals. `disjoint_compl` has also been replaced with `disjoint_compl_left` and `disjoint_compl_right` variants. Co-authored-by: Devon Tuma <dtumad@gmail.com>
1 parent 6babb55 commit a18be37

File tree

11 files changed

+212
-9
lines changed

11 files changed

+212
-9
lines changed

src/data/polynomial/eval.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -179,6 +179,10 @@ ring_hom.of (eval₂ f x)
179179

180180
lemma eval₂_pow (n : ℕ) : (p ^ n).eval₂ f x = p.eval₂ f x ^ n := (eval₂_ring_hom _ _).map_pow _ _
181181

182+
lemma eval₂_eq_sum_range :
183+
p.eval₂ f x = ∑ i in finset.range (p.nat_degree + 1), f (p.coeff i) * x^i :=
184+
trans (congr_arg _ p.as_sum) (trans (eval₂_finset_sum f _ _ x) (congr_arg _ (by simp)))
185+
182186
end eval₂
183187

184188
section eval

src/data/set/lattice.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1050,7 +1050,9 @@ ht.sup_right hu
10501050
theorem disjoint_diff {a b : set α} : disjoint a (b \ a) :=
10511051
disjoint_iff.2 (inter_diff_self _ _)
10521052

1053-
theorem disjoint_compl (s : set α) : disjoint s sᶜ := assume a ⟨h₁, h₂⟩, h₂ h₁
1053+
theorem disjoint_compl_left (s : set α) : disjoint sᶜ s := assume a ⟨h₁, h₂⟩, h₁ h₂
1054+
1055+
theorem disjoint_compl_right (s : set α) : disjoint s sᶜ := assume a ⟨h₁, h₂⟩, h₂ h₁
10541056

10551057
theorem disjoint_singleton_left {a : α} {s : set α} : disjoint {a} s ↔ a ∉ s :=
10561058
by simp [set.disjoint_iff, subset_def]; exact iff.rfl

src/linear_algebra/basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -2486,8 +2486,8 @@ lemma disjoint_std_basis_std_basis (I J : set ι) (h : disjoint I J) :
24862486
disjoint (⨆i∈I, range (std_basis R φ i)) (⨆i∈J, range (std_basis R φ i)) :=
24872487
begin
24882488
refine disjoint.mono
2489-
(supr_range_std_basis_le_infi_ker_proj _ _ _ _ $ set.disjoint_compl I)
2490-
(supr_range_std_basis_le_infi_ker_proj _ _ _ _ $ set.disjoint_compl J) _,
2489+
(supr_range_std_basis_le_infi_ker_proj _ _ _ _ $ set.disjoint_compl_right I)
2490+
(supr_range_std_basis_le_infi_ker_proj _ _ _ _ $ set.disjoint_compl_right J) _,
24912491
simp only [disjoint, submodule.le_def', mem_infi, mem_inf, mem_ker, mem_bot, proj_apply,
24922492
funext_iff],
24932493
rintros b ⟨hI, hJ⟩ i,

src/linear_algebra/finsupp.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -70,8 +70,8 @@ lemma disjoint_lsingle_lsingle (s t : set α) (hs : disjoint s t) :
7070
disjoint (⨆a∈s, (lsingle a : M →ₗ[R] (α →₀ M)).range) (⨆a∈t, (lsingle a).range) :=
7171
begin
7272
refine disjoint.mono
73-
(lsingle_range_le_ker_lapply _ _ $ disjoint_compl s)
74-
(lsingle_range_le_ker_lapply _ _ $ disjoint_compl t)
73+
(lsingle_range_le_ker_lapply _ _ $ disjoint_compl_right s)
74+
(lsingle_range_le_ker_lapply _ _ $ disjoint_compl_right t)
7575
(le_trans (le_infi $ assume i, _) infi_ker_lapply_le_bot),
7676
classical,
7777
by_cases his : i ∈ s,

src/linear_algebra/matrix.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -588,7 +588,7 @@ begin
588588
simp only [comap_infi, (ker_comp _ _).symm, proj_diagonal, ker_smul'],
589589
have : univ ⊆ {i : m | w i = 0} ∪ {i : m | w i = 0}ᶜ, { rw set.union_compl_self },
590590
exact (supr_range_std_basis_eq_infi_ker_proj K (λi:m, K)
591-
(disjoint_compl {i | w i = 0}) this (finite.of_fintype _)).symm
591+
(disjoint_compl_right {i | w i = 0}) this (finite.of_fintype _)).symm
592592
end
593593

594594
lemma range_diagonal [decidable_eq m] (w : m → K) :
@@ -604,7 +604,7 @@ lemma rank_diagonal [decidable_eq m] [decidable_eq K] (w : m → K) :
604604
rank (diagonal w).to_lin = fintype.card { i // w i ≠ 0 } :=
605605
begin
606606
have hu : univ ⊆ {i : m | w i = 0}ᶜ ∪ {i : m | w i = 0}, { rw set.compl_union_self },
607-
have hd : disjoint {i : m | w i ≠ 0} {i : m | w i = 0} := (disjoint_compl {i | w i = 0}).symm,
607+
have hd : disjoint {i : m | w i ≠ 0} {i : m | w i = 0} := (disjoint_compl_right {i | w i = 0}).symm,
608608
have h₁ := supr_range_std_basis_eq_infi_ker_proj K (λi:m, K) hd hu (finite.of_fintype _),
609609
have h₂ := @infi_ker_proj_equiv K _ _ (λi:m, K) _ _ _ _ (by simp; apply_instance) hd hu,
610610
rw [rank, range_diagonal, h₁, ←@dim_fun' K],

src/measure_theory/measure_space.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -723,7 +723,7 @@ ext $ λ t' ht', restrict_union_apply (h.mono inf_le_right inf_le_right) hs ht h
723723

724724
@[simp] lemma restrict_add_restrict_compl {s : set α} (hs : is_measurable s) :
725725
μ.restrict s + μ.restrict sᶜ = μ :=
726-
by rw [← restrict_union (disjoint_compl _) hs hs.compl, union_compl_self, restrict_univ]
726+
by rw [← restrict_union (disjoint_compl_right _) hs hs.compl, union_compl_self, restrict_univ]
727727

728728
@[simp] lemma restrict_compl_add_restrict {s : set α} (hs : is_measurable s) :
729729
μ.restrict sᶜ + μ.restrict s = μ :=

src/measure_theory/set_integral.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -331,7 +331,7 @@ lemma integral_univ : ∫ x in univ, f x ∂μ = ∫ x, f x ∂μ := by rw [meas
331331

332332
lemma integral_add_compl (hs : is_measurable s) (hfi : integrable f μ) :
333333
∫ x in s, f x ∂μ + ∫ x in sᶜ, f x ∂μ = ∫ x, f x ∂μ :=
334-
by rw [← integral_union (disjoint_compl s) hs hs.compl hfi.integrable_on hfi.integrable_on,
334+
by rw [← integral_union (disjoint_compl_right s) hs hs.compl hfi.integrable_on hfi.integrable_on,
335335
union_compl_self, integral_univ]
336336

337337
/-- For a measurable function `f` and a measurable set `s`, the integral of `indicator s f`

src/ring_theory/ideal/operations.lean

Lines changed: 25 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -909,6 +909,31 @@ begin
909909
refine ⟨hJ' ▸ map_mono hJ.left, hJ' ▸ map_is_prime_of_surjective hf (le_trans h hJ.left)⟩ },
910910
end
911911

912+
section quotient_algebra
913+
914+
/-- The ring hom `R/f⁻¹(I) →+* S/I` induced by a ring hom `f : R →+* S` -/
915+
def quotient_map {I : ideal R} (J : ideal S) (f : R →+* S) (hIJ : I ≤ J.comap f) :
916+
I.quotient →+* J.quotient :=
917+
(quotient.lift I ((quotient.mk J).comp f) (λ _ ha,
918+
by simpa [function.comp_app, ring_hom.coe_comp, quotient.eq_zero_iff_mem] using hIJ ha))
919+
920+
variables {I : ideal R} {J: ideal S} [algebra R S]
921+
922+
@[priority 100]
923+
instance quotient_algebra : algebra (J.comap (algebra_map R S)).quotient J.quotient :=
924+
(quotient_map J (algebra_map R S) (le_of_eq rfl)).to_algebra
925+
926+
lemma algebra_map_quotient_injective :
927+
function.injective (algebra_map (J.comap (algebra_map R S)).quotient J.quotient) :=
928+
begin
929+
rintros ⟨a⟩ ⟨b⟩ hab,
930+
replace hab := quotient.eq.mp hab,
931+
rw ← ring_hom.map_sub at hab,
932+
exact quotient.eq.mpr hab
933+
end
934+
935+
end quotient_algebra
936+
912937
end ideal
913938

914939
namespace submodule

src/ring_theory/ideal/over.lean

Lines changed: 60 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,7 @@ Author: Anne Baanen
55
-/
66

77
import ring_theory.algebraic
8+
import ring_theory.localization
89

910
/-!
1011
# Ideals over/under ideals
@@ -143,6 +144,15 @@ lemma is_maximal_of_is_integral_of_is_maximal_comap
143144
λ J I_lt_J, let ⟨I_le_J, x, hxJ, hxI⟩ := lt_iff_le_and_exists.mp I_lt_J
144145
in comap_eq_top_iff.mp (hI.2 _ (comap_lt_comap_of_integral_mem_sdiff I_le_J ⟨hxJ, hxI⟩ (hRS x))) ⟩
145146

147+
lemma is_maximal_comap_of_is_integral_of_is_maximal (hRS_integral : ∀ (x : S), is_integral R x)
148+
(I : ideal S) [hI : I.is_maximal] : is_maximal (I.comap (algebra_map R S)) :=
149+
begin
150+
refine quotient.maximal_of_is_field _ _,
151+
haveI : is_prime (I.comap (algebra_map R S)) := comap_is_prime _ _,
152+
exact is_field_of_is_integral_of_is_field (is_integral_quotient_of_is_integral hRS_integral)
153+
algebra_map_quotient_injective (by rwa ← quotient.maximal_ideal_iff_is_field_quotient),
154+
end
155+
146156
lemma integral_closure.comap_ne_bot [nontrivial R] {I : ideal (integral_closure R S)}
147157
(I_ne_bot : I ≠ ⊥) : I.comap (algebra_map R (integral_closure R S)) ≠ ⊥ :=
148158
let ⟨x, x_mem, x_ne_zero⟩ := I.ne_bot_iff.mp I_ne_bot in
@@ -163,6 +173,56 @@ lemma integral_closure.is_maximal_of_is_maximal_comap
163173
(hI : is_maximal (I.comap (algebra_map R (integral_closure R S)))) : is_maximal I :=
164174
is_maximal_of_is_integral_of_is_maximal_comap (λ x, integral_closure.is_integral x) I hI
165175

176+
/-- `comap (algebra_map R S)` is a surjection from the prime spec of `R` to prime spec of `S`.
177+
`hP : (algebra_map R S).ker ≤ P` is a slight generalization of the extension being injective -/
178+
lemma exists_ideal_over_prime_of_is_integral' (H : ∀ x : S, is_integral R x)
179+
(P : ideal R) [is_prime P] (hP : (algebra_map R S).ker ≤ P) :
180+
∃ (Q : ideal S), is_prime Q ∧ P = Q.comap (algebra_map R S) :=
181+
begin
182+
have hP0 : (0 : S) ∉ algebra.algebra_map_submonoid S P.prime_compl,
183+
{ rintro ⟨x, ⟨hx, x0⟩⟩,
184+
exact absurd (hP x0) hx },
185+
let Rₚ := localization P.prime_compl,
186+
let f := localization.of P.prime_compl,
187+
let Sₚ := localization (algebra.algebra_map_submonoid S P.prime_compl),
188+
let g := localization.of (algebra.algebra_map_submonoid S P.prime_compl),
189+
letI : integral_domain (localization (algebra.algebra_map_submonoid S P.prime_compl)) :=
190+
localization_map.integral_domain_localization (le_non_zero_divisors_of_domain hP0),
191+
obtain ⟨Qₚ : ideal Sₚ, Qₚ_maximal⟩ := @exists_maximal Sₚ _ (by apply_instance),
192+
haveI Qₚ_max : is_maximal (comap _ Qₚ) := @is_maximal_comap_of_is_integral_of_is_maximal Rₚ _ Sₚ _
193+
(localization_algebra P.prime_compl f g)
194+
(is_integral_localization f g H) _ Qₚ_maximal,
195+
refine ⟨comap g.to_map Qₚ, ⟨comap_is_prime g.to_map Qₚ, _⟩⟩,
196+
convert localization.at_prime.comap_maximal_ideal.symm,
197+
rw [comap_comap, ← local_ring.eq_maximal_ideal Qₚ_max, ← f.map_comp _],
198+
refl
199+
end
200+
201+
/-- More general going-up theorem than `exists_ideal_over_prime_of_is_integral'`.
202+
TODO: Version of going-up theorem with arbitrary length chains (by induction on this)?
203+
Not sure how best to write an ascending chain in Lean -/
204+
theorem exists_ideal_over_prime_of_is_integral (H : ∀ x : S, is_integral R x)
205+
(P : ideal R) [is_prime P] (I : ideal S) [is_prime I] (hIP : I.comap (algebra_map R S) ≤ P) :
206+
∃ Q ≥ I, is_prime Q ∧ P = Q.comap (algebra_map R S) :=
207+
begin
208+
obtain ⟨Q' : ideal I.quotient, ⟨Q'_prime, hQ'⟩⟩ := @exists_ideal_over_prime_of_is_integral'
209+
(I.comap (algebra_map R S)).quotient _ I.quotient _
210+
ideal.quotient_algebra
211+
(is_integral_quotient_of_is_integral H)
212+
(map (quotient.mk (I.comap (algebra_map R S))) P)
213+
(map_is_prime_of_surjective quotient.mk_surjective (by simp [hIP]))
214+
(le_trans
215+
(le_of_eq ((ring_hom.injective_iff_ker_eq_bot _).1 algebra_map_quotient_injective))
216+
bot_le),
217+
haveI := Q'_prime,
218+
refine ⟨Q'.comap _, le_trans (le_of_eq mk_ker.symm) (ker_le_comap _), ⟨comap_is_prime _ Q', _⟩⟩,
219+
rw comap_comap,
220+
refine trans _ (trans (congr_arg (comap (quotient.mk (comap (algebra_map R S) I))) hQ') _),
221+
{ refine trans ((sup_eq_left.2 _).symm) (comap_map_of_surjective _ quotient.mk_surjective _).symm,
222+
simpa [← ring_hom.ker_eq_comap_bot] using hIP},
223+
{ simpa [comap_comap] },
224+
end
225+
166226
end integral_domain
167227

168228
end ideal

src/ring_theory/integral_closure.lean

Lines changed: 51 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -371,6 +371,57 @@ begin
371371
exact hp',
372372
end
373373

374+
lemma is_integral_quotient_of_is_integral {I : ideal A} (hRS : ∀ (x : A), is_integral R x) :
375+
∀ (x : I.quotient), is_integral (I.comap (algebra_map R A)).quotient x :=
376+
begin
377+
rintros ⟨x⟩,
378+
obtain ⟨p, ⟨p_monic, hpx⟩⟩ := hRS x,
379+
refine ⟨p.map (ideal.quotient.mk _), ⟨monic_map _ p_monic, _⟩⟩,
380+
simpa only [aeval_def, hom_eval₂, eval₂_map] using congr_arg (ideal.quotient.mk I) hpx
381+
end
382+
383+
/-- If the integral extension `R → S` is injective, and `S` is a field, then `R` is also a field -/
384+
lemma is_field_of_is_integral_of_is_field {R S : Type*} [integral_domain R] [integral_domain S]
385+
[algebra R S] (H : ∀ x : S, is_integral R x) (hRS : function.injective (algebra_map R S))
386+
(hS : is_field S) : is_field R :=
387+
begin
388+
refine ⟨⟨0, 1, zero_ne_one⟩, mul_comm, λ a ha, _⟩,
389+
-- Let `a_inv` be the inverse of `algebra_map R S a`,
390+
-- then we need to show that `a_inv` is of the form `algebra_map R S b`.
391+
obtain ⟨a_inv, ha_inv⟩ := hS.mul_inv_cancel (λ h, ha (hRS (trans h (ring_hom.map_zero _).symm))),
392+
393+
-- Let `p : polynomial R` be monic with root `a_inv`,
394+
-- and `q` be `p` with coefficients reversed (so `q(a) = q'(a) * a + 1`).
395+
-- We claim that `q(a) = 0`, so `-q'(a)` is the inverse of `a`.
396+
obtain ⟨p, p_monic, hp⟩ := H a_inv,
397+
use -∑ (i : ℕ) in finset.range p.nat_degree, (p.coeff i) * a ^ (p.nat_degree - i - 1),
398+
399+
-- `q(a) = 0`, because multiplying everything with `a_inv^n` gives `p(a_inv) = 0`.
400+
-- TODO: this could be a lemma for `polynomial.reverse`.
401+
have hq : ∑ (i : ℕ) in finset.range (p.nat_degree + 1), (p.coeff i) * a ^ (p.nat_degree - i) = 0,
402+
{ apply (algebra_map R S).injective_iff.mp hRS,
403+
have a_inv_ne_zero : a_inv ≠ 0 := right_ne_zero_of_mul (mt ha_inv.symm.trans one_ne_zero),
404+
refine (mul_eq_zero.mp _).resolve_right (pow_ne_zero p.nat_degree a_inv_ne_zero),
405+
rw [aeval_def, eval₂_eq_sum_range] at hp,
406+
rw [ring_hom.map_sum, finset.sum_mul],
407+
refine (finset.sum_congr rfl (λ i hi, _)).trans hp,
408+
rw [ring_hom.map_mul, mul_assoc],
409+
congr,
410+
have : a_inv ^ p.nat_degree = a_inv ^ (p.nat_degree - i) * a_inv ^ i,
411+
{ rw [← pow_add a_inv, nat.sub_add_cancel (nat.le_of_lt_succ (finset.mem_range.mp hi))] },
412+
rw [ring_hom.map_pow, this, ← mul_assoc, ← mul_pow, ha_inv, one_pow, one_mul] },
413+
414+
-- Since `q(a) = 0` and `q(a) = q'(a) * a + 1`, we have `a * -q'(a) = 1`.
415+
-- TODO: we could use a lemma for `polynomial.div_X` here.
416+
rw [finset.sum_range_succ, p_monic.coeff_nat_degree, one_mul, nat.sub_self, pow_zero,
417+
add_eq_zero_iff_eq_neg, eq_comm] at hq,
418+
rw [mul_comm, ← neg_mul_eq_neg_mul, finset.sum_mul],
419+
convert hq using 2,
420+
refine finset.sum_congr rfl (λ i hi, _),
421+
have : 1 ≤ p.nat_degree - i := nat.le_sub_left_of_add_le (finset.mem_range.mp hi),
422+
rw [mul_assoc, ← pow_succ', nat.sub_add_cancel this]
423+
end
424+
374425
end algebra
375426

376427
theorem integral_closure_idem {R : Type*} {A : Type*} [comm_ring R] [comm_ring A] [algebra R A] :

0 commit comments

Comments
 (0)