@@ -5,6 +5,7 @@ Author: Anne Baanen
5
5
-/
6
6
7
7
import ring_theory.algebraic
8
+ import ring_theory.localization
8
9
9
10
/-!
10
11
# Ideals over/under ideals
@@ -143,6 +144,15 @@ lemma is_maximal_of_is_integral_of_is_maximal_comap
143
144
λ J I_lt_J, let ⟨I_le_J, x, hxJ, hxI⟩ := lt_iff_le_and_exists.mp I_lt_J
144
145
in comap_eq_top_iff.mp (hI.2 _ (comap_lt_comap_of_integral_mem_sdiff I_le_J ⟨hxJ, hxI⟩ (hRS x))) ⟩
145
146
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
+
146
156
lemma integral_closure.comap_ne_bot [nontrivial R] {I : ideal (integral_closure R S)}
147
157
(I_ne_bot : I ≠ ⊥) : I.comap (algebra_map R (integral_closure R S)) ≠ ⊥ :=
148
158
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
163
173
(hI : is_maximal (I.comap (algebra_map R (integral_closure R S)))) : is_maximal I :=
164
174
is_maximal_of_is_integral_of_is_maximal_comap (λ x, integral_closure.is_integral x) I hI
165
175
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
+
166
226
end integral_domain
167
227
168
228
end ideal
0 commit comments