proved one direction

This commit is contained in:
poincare-duality 2023-06-15 20:45:29 -07:00
parent 165d598066
commit 37d8991879

View file

@ -16,6 +16,7 @@ import Mathlib.Order.ConditionallyCompleteLattice.Basic
import Mathlib.Algebra.Ring.Pi import Mathlib.Algebra.Ring.Pi
import Mathlib.RingTheory.Finiteness import Mathlib.RingTheory.Finiteness
import Mathlib.Util.PiNotation import Mathlib.Util.PiNotation
import CommAlg.krull
open PiNotation open PiNotation
@ -23,10 +24,10 @@ namespace Ideal
variable (R : Type _) [CommRing R] (P : PrimeSpectrum R) variable (R : Type _) [CommRing R] (P : PrimeSpectrum R)
noncomputable def height : ℕ∞ := Set.chainHeight {J : PrimeSpectrum R | J < P} -- noncomputable def height : ℕ∞ := Set.chainHeight {J : PrimeSpectrum R | J < P}
noncomputable def krullDim (R : Type) [CommRing R] : -- noncomputable def krullDim (R : Type) [CommRing R] :
WithBot ℕ∞ := ⨆ (I : PrimeSpectrum R), height R I -- WithBot ℕ∞ := ⨆ (I : PrimeSpectrum R), height R I
--variable {R} --variable {R}
@ -66,6 +67,8 @@ lemma ring_Noetherian_iff_spec_Noetherian : IsNoetherianRing R
↔ TopologicalSpace.NoetherianSpace (PrimeSpectrum R) := by ↔ TopologicalSpace.NoetherianSpace (PrimeSpectrum R) := by
constructor constructor
intro RisNoetherian intro RisNoetherian
sorry
sorry
-- how do I apply an instance to prove one direction? -- how do I apply an instance to prove one direction?
@ -99,9 +102,9 @@ lemma containment_radical_power_containment :
-- Stacks Lemma 10.52.6: I is a maximal ideal and IM = 0. Then length of M is -- Stacks Lemma 10.52.6: I is a maximal ideal and IM = 0. Then length of M is
-- the same as the dimension as a vector space over R/I, -- the same as the dimension as a vector space over R/I,
lemma length_eq_dim_if_maximal_annihilates {I : Ideal R} [Ideal.IsMaximal I] -- lemma length_eq_dim_if_maximal_annihilates {I : Ideal R} [Ideal.IsMaximal I]
: I • ( : Submodule R M) = 0 -- : I • ( : Submodule R M) = 0
→ Module.length R M = Module.rank RI M(I • ( : Submodule R M)) := by sorry -- → Module.length R M = Module.rank RI M(I • ( : Submodule R M)) := by sorry
-- Does lean know that M/IM is a R/I module? -- Does lean know that M/IM is a R/I module?
-- Use 10.52.5 -- Use 10.52.5
@ -125,30 +128,34 @@ lemma Artinian_has_finite_max_ideal
let m' : ↪ MaximalSpectrum R := Infinite.natEmbedding (MaximalSpectrum R) let m' : ↪ MaximalSpectrum R := Infinite.natEmbedding (MaximalSpectrum R)
have m'inj := m'.injective have m'inj := m'.injective
let m'' : → Ideal R := fun n : ↦ ⨅ k ∈ range n, (m' k).asIdeal let m'' : → Ideal R := fun n : ↦ ⨅ k ∈ range n, (m' k).asIdeal
let f : → Ideal R := fun n : ↦ (m' n).asIdeal -- let f : → MaximalSpectrum R := fun n : ↦ m' n
let F : Fin n → Ideal R := fun k ↦ (m' k).asIdeal -- let F : (n : ) → Fin n → MaximalSpectrum R := fun n k ↦ m' k
have comaximal : ∀ i j : , i ≠ j → (m' i).asIdeal ⊔ (m' j).asIdeal = have DCC : ∃ n : , ∀ k : , n ≤ k → m'' n = m'' k := by
( : Ideal R) := by apply IsArtinian.monotone_stabilizes {
intro i j distinct toFun := m''
apply Ideal.IsMaximal.coprime_of_ne monotone' := sorry
exact (m' i).IsMaximal }
exact (m' j).IsMaximal cases' DCC with n DCCn
have : (m' i) ≠ (m' j) := by specialize DCCn (n+1)
exact Function.Injective.ne m'inj distinct specialize DCCn (Nat.le_succ n)
intro h have containment1 : m'' n < (m' (n + 1)).asIdeal := by sorry
apply this have : ∀ (j : ), (j ≠ n + 1) → ∃ x, x ∈ (m' j).asIdeal ∧ x ∉ (m' (n+1)).asIdeal := by
exact MaximalSpectrum.ext _ _ h intro j jnotn
have ∀ n : , (R ⨅ (i : Fin n), (F n) i) ≃+* ((i : Fin n) → R (F n) i) := by have notcontain : ¬ (m' j).asIdeal ≤ (m' (n+1)).asIdeal := by
-- apply Ideal.IsMaximal (m' j).asIdeal
sorry
sorry sorry
-- (let F : Fin n → Ideal R := fun k : Fin n ↦ (m' k).asIdeal) sorry
-- let g := Ideal.quotientInfRingEquivPiQuotient f comaximal -- have distinct : (m' j).asIdeal ≠ (m' (n+1)).asIdeal := by
-- intro h
-- apply Function.Injective.ne m'inj jnotn
-- exact MaximalSpectrum.ext _ _ h
-- simp
-- unfold
-- Stacks Lemma 10.53.4: R Artinian => Jacobson ideal of R is nilpotent -- Stacks Lemma 10.53.4: R Artinian => Jacobson ideal of R is nilpotent
lemma Jacobson_of_Artinian_is_nilpotent -- This is in mathlib: IsArtinianRing.isNilpotent_jacobson_bot
[IsArtinianRing R] : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) := by
have J := Ideal.jacobson (⊥ : Ideal R)
-- Stacks Lemma 10.53.5: If R has finitely many maximal ideals and -- Stacks Lemma 10.53.5: If R has finitely many maximal ideals and
-- locally nilpotent Jacobson radical, then R is the product of its localizations at -- locally nilpotent Jacobson radical, then R is the product of its localizations at
@ -187,21 +194,92 @@ lemma primes_of_Artinian_are_maximal
-- Lemma: Krull dimension of a ring is the supremum of height of maximal ideals -- Lemma: Krull dimension of a ring is the supremum of height of maximal ideals
-- Stacks Lemma 10.60.5: R is Artinian iff R is Noetherian of dimension 0 -- Stacks Lemma 10.60.5: R is Artinian iff R is Noetherian of dimension 0
lemma dim_zero_Noetherian_iff_Artinian (R : Type _) [CommRing R] : lemma dim_le_zero_Noetherian_iff_Artinian (R : Type _) [CommRing R] :
IsNoetherianRing R ∧ Ideal.krullDim R = 0 ↔ IsArtinianRing R := by IsNoetherianRing R ∧ Ideal.krullDim R 0 ↔ IsArtinianRing R := by
constructor constructor
sorry sorry
intro RisArtinian intro RisArtinian
constructor constructor
apply finite_length_is_Noetherian apply finite_length_is_Noetherian
rwa [IsArtinian_iff_finite_length] at RisArtinian rwa [IsArtinian_iff_finite_length] at RisArtinian
sorry -- can use Grant's lemma dim_eq_zero_iff rw [Ideal.dim_le_zero_iff]
intro I
apply primes_of_Artinian_are_maximal
-- Trash bin
-- lemma Artinian_has_finite_max_ideal
-- [IsArtinianRing R] : Finite (MaximalSpectrum R) := by
-- by_contra infinite
-- simp only [not_finite_iff_infinite] at infinite
-- let m' : ↪ MaximalSpectrum R := Infinite.natEmbedding (MaximalSpectrum R)
-- have m'inj := m'.injective
-- let m'' : → Ideal R := fun n : ↦ ⨅ k ∈ range n, (m' k).asIdeal
-- let f : → Ideal R := fun n : ↦ (m' n).asIdeal
-- have DCC : ∃ n : , ∀ k : , n ≤ k → m'' n = m'' k := by
-- apply IsArtinian.monotone_stabilizes {
-- toFun := m''
-- monotone' := sorry
-- }
-- cases' DCC with n DCCn
-- specialize DCCn (n+1)
-- specialize DCCn (Nat.le_succ n)
-- let F : Fin (n + 1) → MaximalSpectrum R := fun k ↦ m' k
-- have comaximal : ∀ (i j : Fin (n + 1)), (i ≠ j) → (F i).asIdeal ⊔ (F j).asIdeal =
-- ( : Ideal R) := by
-- intro i j distinct
-- apply Ideal.IsMaximal.coprime_of_ne
-- exact (F i).IsMaximal
-- exact (F j).IsMaximal
-- have : (F i) ≠ (F j) := by
-- apply Function.Injective.ne m'inj
-- contrapose! distinct
-- exact Fin.ext distinct
-- intro h
-- apply this
-- exact MaximalSpectrum.ext _ _ h
-- let CRT1 : (R ⨅ (i : Fin (n + 1)), ((F i).asIdeal))
-- ≃+* ((i : Fin (n + 1)) → R (F i).asIdeal) :=
-- Ideal.quotientInfRingEquivPiQuotient
-- (fun i ↦ (F i).asIdeal) comaximal
-- let CRT2 : (R ⨅ (i : Fin (n + 1)), ((F i).asIdeal))
-- ≃+* ((i : Fin (n + 1)) → R (F i).asIdeal) :=
-- Ideal.quotientInfRingEquivPiQuotient
-- (fun i ↦ (F i).asIdeal) comaximal
-- have comaximal : ∀ (n : ) (i j : Fin n), (i ≠ j) → ((F n) i).asIdeal ⊔ ((F n) j).asIdeal =
-- ( : Ideal R) := by
-- intro n i j distinct
-- apply Ideal.IsMaximal.coprime_of_ne
-- exact (F n i).IsMaximal
-- exact (F n j).IsMaximal
-- have : (F n i) ≠ (F n j) := by
-- apply Function.Injective.ne m'inj
-- contrapose! distinct
-- exact Fin.ext distinct
-- intro h
-- apply this
-- exact MaximalSpectrum.ext _ _ h
-- let CRT : (n : ) → (R ⨅ (i : Fin n), ((F n) i).asIdeal)
-- ≃+* ((i : Fin n) → R ((F n) i).asIdeal) :=
-- fun n ↦ Ideal.quotientInfRingEquivPiQuotient
-- (fun i ↦ (F n i).asIdeal) (comaximal n)
-- have DCC : ∃ n : , ∀ k : , n ≤ k → m'' n = m'' k := by
-- apply IsArtinian.monotone_stabilizes {
-- toFun := m''
-- monotone' := sorry
-- }
-- cases' DCC with n DCCn
-- specialize DCCn (n+1)
-- specialize DCCn (Nat.le_succ n)
-- let CRT1 := CRT n
-- let CRT2 := CRT (n + 1)