Skip to content

Commit b479ce6

Browse files
committed
chore(Mathlib/RingTheory/LocalProperties): split LocalProperties.lean (#16879)
Split the section on local properties in general and the properties that satisfy local properties in `LocalProperties.lean` into separate files. See the discussions [here](#16558 (comment)). Co-authored-by: Hu Yongle <2065545849@qq.com> Co-authored-by: Yongle Hu <2065545849@qq.com>
1 parent 60ff730 commit b479ce6

File tree

16 files changed

+842
-791
lines changed

16 files changed

+842
-791
lines changed

Mathlib.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3920,7 +3920,9 @@ import Mathlib.RingTheory.Kaehler.CotangentComplex
39203920
import Mathlib.RingTheory.Kaehler.Polynomial
39213921
import Mathlib.RingTheory.LaurentSeries
39223922
import Mathlib.RingTheory.LittleWedderburn
3923-
import Mathlib.RingTheory.LocalProperties
3923+
import Mathlib.RingTheory.LocalProperties.Basic
3924+
import Mathlib.RingTheory.LocalProperties.IntegrallyClosed
3925+
import Mathlib.RingTheory.LocalProperties.Reduced
39243926
import Mathlib.RingTheory.LocalRing.Basic
39253927
import Mathlib.RingTheory.LocalRing.Defs
39263928
import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic

Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -95,11 +95,11 @@ theorem spec_of_surjective {R S : CommRingCat} (f : R ⟶ S) (h : Function.Surje
9595
surj_on_stalks x := by
9696
haveI : (RingHom.toMorphismProperty (fun f ↦ Function.Surjective f)).RespectsIso := by
9797
rw [← RingHom.toMorphismProperty_respectsIso_iff]
98-
exact surjective_respectsIso
98+
exact RingHom.surjective_respectsIso
9999
apply (MorphismProperty.arrow_mk_iso_iff
100100
(RingHom.toMorphismProperty (fun f ↦ Function.Surjective f))
101101
(Scheme.arrowStalkMapSpecIso f x)).mpr
102-
exact surjective_localRingHom_of_surjective f h x.asIdeal
102+
exact RingHom.surjective_localRingHom_of_surjective f h x.asIdeal
103103

104104
/-- For any ideal `I` in a commutative ring `R`, the quotient map `specObj R ⟶ specObj (R ⧸ I)`
105105
is a closed immersion. -/

Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Andrew Yang
55
-/
66
import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap
7-
import Mathlib.RingTheory.LocalProperties
7+
import Mathlib.RingTheory.RingHom.Surjective
88

99
/-!
1010
@@ -52,7 +52,7 @@ lemma isPreimmersion_eq_inf :
5252
/-- Being surjective on stalks is local at the target. -/
5353
instance isSurjectiveOnStalks_isLocalAtTarget : IsLocalAtTarget
5454
(stalkwise (Function.Surjective ·)) :=
55-
stalkwiseIsLocalAtTarget_of_respectsIso surjective_respectsIso
55+
stalkwiseIsLocalAtTarget_of_respectsIso RingHom.surjective_respectsIso
5656

5757
namespace IsPreimmersion
5858

Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Andrew Yang
55
-/
66
import Mathlib.AlgebraicGeometry.Morphisms.Basic
7-
import Mathlib.RingTheory.LocalProperties
7+
import Mathlib.RingTheory.LocalProperties.Basic
88

99
/-!
1010

Mathlib/AlgebraicGeometry/Properties.lean

Lines changed: 1 addition & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Andrew Yang
55
-/
66
import Mathlib.AlgebraicGeometry.AffineScheme
7-
import Mathlib.RingTheory.Nilpotent.Lemmas
8-
import Mathlib.Topology.Sheaves.SheafCondition.Sites
9-
import Mathlib.Algebra.Category.Ring.Constructions
10-
import Mathlib.RingTheory.LocalProperties
7+
import Mathlib.RingTheory.LocalProperties.Reduced
118

129
/-!
1310
# Basic properties of schemes

Mathlib/RingTheory/Ideal/Norm.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ import Mathlib.LinearAlgebra.FreeModule.Determinant
1212
import Mathlib.LinearAlgebra.FreeModule.IdealQuotient
1313
import Mathlib.RingTheory.DedekindDomain.PID
1414
import Mathlib.RingTheory.Ideal.Basis
15-
import Mathlib.RingTheory.LocalProperties
15+
import Mathlib.RingTheory.LocalProperties.Basic
1616
import Mathlib.RingTheory.Localization.NormTrace
1717

1818
/-!

Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -3,11 +3,11 @@ Copyright (c) 2023 Andrew Yang, Patrick Lutz. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Andrew Yang
55
-/
6-
import Mathlib.RingTheory.IntegralClosure.IntegrallyClosed
7-
import Mathlib.RingTheory.LocalProperties
8-
import Mathlib.RingTheory.Localization.NormTrace
9-
import Mathlib.RingTheory.Localization.LocalizationLocalization
106
import Mathlib.RingTheory.DedekindDomain.IntegralClosure
7+
import Mathlib.RingTheory.RingHom.Finite
8+
import Mathlib.RingTheory.Localization.LocalizationLocalization
9+
import Mathlib.RingTheory.Localization.NormTrace
10+
1111
/-!
1212
# Restriction of various maps between fields to integrally closed subrings.
1313

Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean

Lines changed: 1 addition & 39 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Anne Baanen
55
-/
66
import Mathlib.RingTheory.Localization.Integral
7-
import Mathlib.RingTheory.Localization.LocalizationLocalization
87

98
/-!
109
# Integrally closed rings
@@ -278,44 +277,7 @@ end integralClosure
278277

279278
/-- Any field is integral closed. -/
280279
/- Although `infer_instance` can find this if you import Mathlib, in this file they have not been
281-
proven yet. However, the next theorem is a fundamental property of `IsIntegrallyClosed`,
280+
proven yet. However, it is used to prove a fundamental property of `IsIntegrallyClosed`,
282281
and it is not desirable to involve more content from other files. -/
283282
instance Field.instIsIntegrallyClosed (K : Type*) [Field K] : IsIntegrallyClosed K :=
284283
(isIntegrallyClosed_iff K).mpr fun {x} _ ↦ ⟨x, rfl⟩
285-
286-
open Localization Ideal IsLocalization in
287-
/-- An integral domain `R` is integral closed if `Rₘ` is integral closed
288-
for any maximal ideal `m` of `R`. -/
289-
theorem IsIntegrallyClosed.of_localization_maximal {R : Type*} [CommRing R] [IsDomain R]
290-
(h : ∀ p : Ideal R, p ≠ ⊥ → [p.IsMaximal] → IsIntegrallyClosed (Localization.AtPrime p)) :
291-
IsIntegrallyClosed R := by
292-
by_cases hf : IsField R
293-
· exact hf.toField.instIsIntegrallyClosed
294-
apply (isIntegrallyClosed_iff (FractionRing R)).mpr
295-
rintro ⟨x⟩ hx
296-
let I : Ideal R := span {x.2.1} / span {x.1}
297-
have h1 : 1 ∈ I := by
298-
apply I.eq_top_iff_one.mp
299-
by_contra hn
300-
rcases I.exists_le_maximal hn with ⟨p, hpm, hpi⟩
301-
have hic := h p (Ring.ne_bot_of_isMaximal_of_not_isField hpm hf)
302-
have hxp : IsIntegral (Localization.AtPrime p) (mk x.1 x.2) := hx.tower_top
303-
/- `x.1 / x.2.1 ∈ Rₚ` since it is integral over `Rₚ` and `Rₚ` is integrally closed.
304-
More precisely, `x.1 / x.2.1 = y.1 / y.2.1` where `y.1, y.2.1 ∈ R` and `y.2.1 ∉ p`. -/
305-
rcases (isIntegrallyClosed_iff (FractionRing R)).mp hic hxp with ⟨⟨y⟩, hy⟩
306-
/- `y.2.1 ∈ I` since for all `a ∈ Ideal.span {x.1}`, say `a = b * x.1`,
307-
we have `y.2 * a = b * x.1 * y.2 = b * y.1 * x.2.1 ∈ Ideal.span {x.2.1}`. -/
308-
have hyi : y.2.1 ∈ I := by
309-
intro a ha
310-
rcases mem_span_singleton'.mp ha with ⟨b, hb⟩
311-
apply mem_span_singleton'.mpr ⟨b * y.1, _⟩
312-
rw [← hb, ← mul_assoc, mul_comm y.2.1 b, mul_assoc, mul_assoc]
313-
exact congrArg (HMul.hMul b) <| (mul_comm y.1 x.2.1).trans <|
314-
NoZeroSMulDivisors.algebraMap_injective R (Localization R⁰) <| mk'_eq_iff_eq.mp <|
315-
(mk'_eq_algebraMap_mk'_of_submonoid_le _ _ p.primeCompl_le_nonZeroDivisors y.1 y.2).trans
316-
<| show algebraMap (Localization.AtPrime p) _ (mk' _ y.1 y.2) = mk' _ x.1 x.2
317-
by simpa only [← mk_eq_mk', ← hy] using by rfl
318-
-- `y.2.1 ∈ I` implies `y.2.1 ∈ p` since `I ⊆ p`, which contradicts to the choice of `y`.
319-
exact y.2.2 (hpi hyi)
320-
rcases mem_span_singleton'.mp (h1 x.1 (mem_span_singleton_self x.1)) with ⟨y, hy⟩
321-
exact ⟨y, (eq_mk'_of_mul_eq (hy.trans (one_mul x.1))).trans (mk_eq_mk'_apply x.1 x.2).symm⟩

0 commit comments

Comments
 (0)