Merge pull request #99 from GTBarkley/main

unify with main
This commit is contained in:
ah1112 2023-06-16 17:48:07 -04:00 committed by GitHub
commit aac88adc02
No known key found for this signature in database
GPG key ID: 4AEE18F83AFDEB23
4 changed files with 157 additions and 85 deletions

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 Mathlib.RingTheory.Ideal.MinimalPrime
import CommAlg.krull import CommAlg.krull
open PiNotation open PiNotation
@ -43,6 +44,8 @@ class IsLocallyNilpotent {R : Type _} [CommRing R] (I : Ideal R) : Prop :=
#check Ideal.IsLocallyNilpotent #check Ideal.IsLocallyNilpotent
end Ideal end Ideal
def RingJacobson (R) [Ring R] := Ideal.jacobson (⊥ : Ideal R)
-- Repeats the definition of the length of a module by Monalisa -- Repeats the definition of the length of a module by Monalisa
variable (R : Type _) [CommRing R] (I J : Ideal R) variable (R : Type _) [CommRing R] (I J : Ideal R)
variable (M : Type _) [AddCommMonoid M] [Module R M] variable (M : Type _) [AddCommMonoid M] [Module R M]
@ -169,15 +172,15 @@ abbrev Prod_of_localization :=
def foo : Prod_of_localization R →+* R where def foo : Prod_of_localization R →+* R where
toFun := sorry toFun := sorry
-- invFun := sorry -- invFun := sorry
left_inv := sorry --left_inv := sorry
right_inv := sorry --right_inv := sorry
map_mul' := sorry map_mul' := sorry
map_add' := sorry map_add' := sorry
def product_of_localization_at_maximal_ideal [Finite (MaximalSpectrum R)] def product_of_localization_at_maximal_ideal [Finite (MaximalSpectrum R)]
(h : Ideal.IsLocallyNilpotent (Ideal.jacobson (⊥ : Ideal R))) : (h : Ideal.IsLocallyNilpotent (RingJacobson R)) :
Prod_of_localization R ≃+* R := by sorry R ≃+* Prod_of_localization R := by sorry
-- Stacks Lemma 10.53.6: R is Artinian iff R has finite length as an R-mod -- Stacks Lemma 10.53.6: R is Artinian iff R has finite length as an R-mod
lemma IsArtinian_iff_finite_length : lemma IsArtinian_iff_finite_length :
@ -193,18 +196,61 @@ 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 -- Lemma: X is an irreducible component of Spec(R) ↔ X = V(I) for I a minimal prime
lemma dim_le_zero_Noetherian_iff_Artinian (R : Type _) [CommRing R] : lemma irred_comp_minmimal_prime (X) :
IsNoetherianRing R ∧ Ideal.krullDim R ≤ 0 ↔ IsArtinianRing R := by X ∈ irreducibleComponents (PrimeSpectrum R)
constructor ↔ ∃ (P : minimalPrimes R), X = PrimeSpectrum.zeroLocus P := by
rintro ⟨RisNoetherian, dimzero⟩
rw [ring_Noetherian_iff_spec_Noetherian] at RisNoetherian
let Z := irreducibleComponents (PrimeSpectrum R)
have Zfinite : Set.Finite Z := by
-- apply TopologicalSpace.NoetherianSpace.finite_irreducibleComponents ?_
sorry sorry
-- Lemma: localization of Noetherian ring is Noetherian
-- lemma localization_of_Noetherian_at_prime [IsNoetherianRing R]
-- (atprime: Ideal.IsPrime I) :
-- IsNoetherianRing (Localization.AtPrime I) := by sorry
-- Stacks Lemma 10.60.5: R is Artinian iff R is Noetherian of dimension 0
lemma Artinian_if_dim_le_zero_Noetherian (R : Type _) [CommRing R] :
IsNoetherianRing R ∧ Ideal.krullDim R ≤ 0 → IsArtinianRing R := by
rintro ⟨RisNoetherian, dimzero⟩
rw [ring_Noetherian_iff_spec_Noetherian] at RisNoetherian
have := fun X => (irred_comp_minmimal_prime R X).mp
choose F hf using this
let Z := irreducibleComponents (PrimeSpectrum R)
-- have Zfinite : Set.Finite Z := by
-- apply TopologicalSpace.NoetherianSpace.finite_irreducibleComponents ?_
-- sorry
--let P := fun
rw [← ring_Noetherian_iff_spec_Noetherian] at RisNoetherian
have PrimeIsMaximal : ∀ X : Z, Ideal.IsMaximal (F X X.2).1 := by
intro X
have prime : Ideal.IsPrime (F X X.2).1 := (F X X.2).2.1.1
rw [Ideal.dim_le_zero_iff] at dimzero
exact dimzero ⟨_, prime⟩
have JacLocallyNil : Ideal.IsLocallyNilpotent (RingJacobson R) := by sorry
let Loc := fun X : Z ↦ Localization.AtPrime (F X.1 X.2).1
have LocNoetherian : ∀ X, IsNoetherianRing (Loc X) := by
intro X
sorry sorry
-- apply IsLocalization.isNoetherianRing (F X.1 X.2).1 (Loc X) RisNoetherian
have Locdimzero : ∀ X, Ideal.krullDim (Loc X) ≤ 0 := by sorry
have powerannihilates : ∀ X, ∃ n : ,
((F X.1 X.2).1) ^ n • (: Submodule R (Loc X)) = 0 := by sorry
have LocFinitelength : ∀ X, ∃ n : , Module.length R (Loc X) ≤ n := by
intro X
have idealfg : Ideal.FG (F X.1 X.2).1 := by
rw [isNoetherianRing_iff_ideal_fg] at RisNoetherian
specialize RisNoetherian (F X.1 X.2).1
exact RisNoetherian
have modulefg : Module.Finite R (Loc X) := by sorry -- not sure if this is true
specialize PrimeIsMaximal X
specialize powerannihilates X
apply power_zero_finite_length R (F X.1 X.2).1 (Loc X) idealfg powerannihilates
have RingFinitelength : ∃ n : , Module.length R R ≤ n := by sorry
rw [IsArtinian_iff_finite_length]
exact RingFinitelength
lemma dim_le_zero_Noetherian_if_Artinian (R : Type _) [CommRing R] :
IsArtinianRing R → IsNoetherianRing R ∧ Ideal.krullDim R ≤ 0 := by
intro RisArtinian intro RisArtinian
constructor constructor
apply finite_length_is_Noetherian apply finite_length_is_Noetherian
@ -213,7 +259,6 @@ lemma dim_le_zero_Noetherian_iff_Artinian (R : Type _) [CommRing R] :
intro I intro I
apply primes_of_Artinian_are_maximal apply primes_of_Artinian_are_maximal
-- Use TopologicalSpace.NoetherianSpace.exists_finset_irreducible :

View file

@ -49,10 +49,12 @@ lemma height_le_of_le {I J : PrimeSpectrum R} (I_le_J : I ≤ J) : height I ≤
show J' < J show J' < J
exact lt_of_lt_of_le hJ' I_le_J exact lt_of_lt_of_le hJ' I_le_J
lemma krullDim_le_iff (R : Type _) [CommRing R] (n : ) : @[simp]
lemma krullDim_le_iff {R : Type _} [CommRing R] {n : } :
krullDim R ≤ n ↔ ∀ I : PrimeSpectrum R, (height I : WithBot ℕ∞) ≤ ↑n := iSup_le_iff (α := WithBot ℕ∞) krullDim R ≤ n ↔ ∀ I : PrimeSpectrum R, (height I : WithBot ℕ∞) ≤ ↑n := iSup_le_iff (α := WithBot ℕ∞)
lemma krullDim_le_iff' (R : Type _) [CommRing R] (n : ℕ∞) : @[simp]
lemma krullDim_le_iff' {R : Type _} [CommRing R] {n : ℕ∞} :
krullDim R ≤ n ↔ ∀ I : PrimeSpectrum R, (height I : WithBot ℕ∞) ≤ ↑n := iSup_le_iff (α := WithBot ℕ∞) krullDim R ≤ n ↔ ∀ I : PrimeSpectrum R, (height I : WithBot ℕ∞) ≤ ↑n := iSup_le_iff (α := WithBot ℕ∞)
@[simp] @[simp]
@ -61,11 +63,10 @@ lemma height_le_krullDim (I : PrimeSpectrum R) : height I ≤ krullDim R :=
/-- In a domain, the height of a prime ideal is Bot (0 in this case) iff it's the Bot ideal. -/ /-- In a domain, the height of a prime ideal is Bot (0 in this case) iff it's the Bot ideal. -/
@[simp] @[simp]
lemma height_bot_iff_bot {D: Type _} [CommRing D] [IsDomain D] {P : PrimeSpectrum D} : height P = ⊥ ↔ P = ⊥ := by lemma height_zero_iff_bot {D: Type _} [CommRing D] [IsDomain D] {P : PrimeSpectrum D} : height P = 0 ↔ P = ⊥ := by
constructor constructor
· intro h · intro h
unfold height at h unfold height at h
rw [bot_eq_zero] at h
simp only [Set.chainHeight_eq_zero_iff] at h simp only [Set.chainHeight_eq_zero_iff] at h
apply eq_bot_of_minimal apply eq_bot_of_minimal
intro I intro I
@ -85,13 +86,10 @@ lemma height_bot_iff_bot {D: Type _} [CommRing D] [IsDomain D] {P : PrimeSpectru
have := not_lt_of_lt JneP have := not_lt_of_lt JneP
contradiction contradiction
@[simp]
lemma height_bot_eq {D: Type _} [CommRing D] [IsDomain D] : height (⊥ : PrimeSpectrum D) = ⊥ := by
rw [height_bot_iff_bot]
/-- The Krull dimension of a ring being ≥ n is equivalent to there being an /-- The Krull dimension of a ring being ≥ n is equivalent to there being an
ideal of height ≥ n. -/ ideal of height ≥ n. -/
lemma le_krullDim_iff (R : Type _) [CommRing R] (n : ) : @[simp]
lemma le_krullDim_iff {R : Type _} [CommRing R] {n : } :
n ≤ krullDim R ↔ ∃ I : PrimeSpectrum R, n ≤ (height I : WithBot ℕ∞) := by n ≤ krullDim R ↔ ∃ I : PrimeSpectrum R, n ≤ (height I : WithBot ℕ∞) := by
constructor constructor
· unfold krullDim · unfold krullDim
@ -131,9 +129,19 @@ lemma le_krullDim_iff (R : Type _) [CommRing R] (n : ) :
#check ENat.recTopCoe #check ENat.recTopCoe
/- terrible place for this lemma. Also this probably exists somewhere /- terrible place for these two lemmas. Also this probably exists somewhere
Also this is a terrible proof Also this is a terrible proof
-/ -/
lemma eq_top_iff' (n : ℕ∞) : n = ↔ ∀ m : , m ≤ n := by
refine' ⟨fun a b => _, fun h => _⟩
. rw [a]; exact le_top
. induction' n using ENat.recTopCoe with n
. rfl
. exfalso
apply not_lt_of_ge (h (n + 1))
norm_cast
norm_num
lemma eq_top_iff (n : WithBot ℕ∞) : n = ↔ ∀ m : , m ≤ n := by lemma eq_top_iff (n : WithBot ℕ∞) : n = ↔ ∀ m : , m ≤ n := by
aesop aesop
induction' n using WithBot.recBotCoe with n induction' n using WithBot.recBotCoe with n
@ -151,47 +159,30 @@ lemma eq_top_iff (n : WithBot ℕ∞) : n = ↔ ∀ m : , m ≤ n := by
lemma krullDim_eq_top_iff (R : Type _) [CommRing R] : lemma krullDim_eq_top_iff (R : Type _) [CommRing R] :
krullDim R = ↔ ∀ (n : ), ∃ I : PrimeSpectrum R, n ≤ height I := by krullDim R = ↔ ∀ (n : ), ∃ I : PrimeSpectrum R, n ≤ height I := by
simp [eq_top_iff, le_krullDim_iff] simp_rw [eq_top_iff, le_krullDim_iff]
change (∀ (m : ), ∃ I, ((m : ℕ∞) : WithBot ℕ∞) ≤ height I) ↔ _ change (∀ (m : ), ∃ I, ((m : ℕ∞) : WithBot ℕ∞) ≤ height I) ↔ _
simp [WithBot.coe_le_coe] simp [WithBot.coe_le_coe]
/-- The Krull dimension of a local ring is the height of its maximal ideal. -/ /-- The Krull dimension of a local ring is the height of its maximal ideal. -/
lemma krullDim_eq_height [LocalRing R] : krullDim R = height (closedPoint R) := by lemma krullDim_eq_height [LocalRing R] : krullDim R = height (closedPoint R) := by
apply le_antisymm apply le_antisymm
. rw [krullDim_le_iff'] . rw [krullDim_le_iff']
intro I intro I
apply WithBot.coe_mono exact WithBot.coe_mono <| height_le_of_le <| le_maximalIdeal I.2.1
apply height_le_of_le
apply le_maximalIdeal
exact I.2.1
. simp only [height_le_krullDim] . simp only [height_le_krullDim]
/-- The height of a prime `𝔭` is greater than `n` if and only if there is a chain of primes less than `𝔭` /-- The height of a prime `𝔭` is greater than `n` if and only if there is a chain of primes less than `𝔭`
with length `n + 1`. -/ with length `n + 1`. -/
lemma lt_height_iff' {𝔭 : PrimeSpectrum R} {n : ℕ∞} : lemma lt_height_iff' {𝔭 : PrimeSpectrum R} {n : ℕ∞} :
n < height 𝔭 ↔ ∃ c : List (PrimeSpectrum R), c.Chain' (· < ·) ∧ (∀ 𝔮 ∈ c, 𝔮 < 𝔭) ∧ c.length = n + 1 := by n < height 𝔭 ↔ ∃ c : List (PrimeSpectrum R), c.Chain' (· < ·) ∧ (∀ 𝔮 ∈ c, 𝔮 < 𝔭) ∧ c.length = n + 1 := by
match n with induction' n using ENat.recTopCoe with n
| => . simp
constructor <;> intro h <;> exfalso . rw [←(ENat.add_one_le_iff <| ENat.coe_ne_top _)]
. exact (not_le.mpr h) le_top
. tauto
| (n : ) =>
have (m : ℕ∞) : n < m ↔ (n + 1 : ℕ∞) ≤ m := by
symm
show (n + 1 ≤ m ↔ _ )
apply ENat.add_one_le_iff
exact ENat.coe_ne_top _
rw [this]
unfold Ideal.height
show ((↑(n + 1):ℕ∞) ≤ _) ↔ ∃c, _ ∧ _ ∧ ((_ : WithTop ) = (_:ℕ∞)) show ((↑(n + 1):ℕ∞) ≤ _) ↔ ∃c, _ ∧ _ ∧ ((_ : WithTop ) = (_:ℕ∞))
rw [{J | J < 𝔭}.le_chainHeight_iff] rw [Ideal.height, Set.le_chainHeight_iff]
show (∃ c, (List.Chain' _ c ∧ ∀𝔮, 𝔮 ∈ c → 𝔮 < 𝔭) ∧ _) ↔ _ show (∃ c, (List.Chain' _ c ∧ ∀𝔮, 𝔮 ∈ c → 𝔮 < 𝔭) ∧ _) ↔ _
constructor <;> rintro ⟨c, hc⟩ <;> use c norm_cast
. tauto simp_rw [and_assoc]
. change _ ∧ _ ∧ (List.length c : ℕ∞) = n + 1 at hc
norm_cast at hc
tauto
/-- Form of `lt_height_iff''` for rewriting with the height coerced to `WithBot ℕ∞`. -/ /-- Form of `lt_height_iff''` for rewriting with the height coerced to `WithBot ℕ∞`. -/
lemma lt_height_iff'' {𝔭 : PrimeSpectrum R} {n : ℕ∞} : lemma lt_height_iff'' {𝔭 : PrimeSpectrum R} {n : ℕ∞} :
@ -203,30 +194,24 @@ lemma lt_height_iff'' {𝔭 : PrimeSpectrum R} {n : ℕ∞} :
--some propositions that would be nice to be able to eventually --some propositions that would be nice to be able to eventually
/-- The prime spectrum of the zero ring is empty. -/ /-- The prime spectrum of the zero ring is empty. -/
lemma primeSpectrum_empty_of_subsingleton (x : PrimeSpectrum R) [Subsingleton R] : False := lemma primeSpectrum_empty_of_subsingleton [Subsingleton R] : IsEmpty <| PrimeSpectrum R where
x.1.ne_top_iff_one.1 x.2.1 <| Eq.substr (Subsingleton.elim 1 (0 : R)) x.1.zero_mem false x := x.1.ne_top_iff_one.1 x.2.1 <| Eq.substr (Subsingleton.elim 1 (0 : R)) x.1.zero_mem
/-- A CommRing has empty prime spectrum if and only if it is the zero ring. -/ /-- A CommRing has empty prime spectrum if and only if it is the zero ring. -/
lemma primeSpectrum_empty_iff : IsEmpty (PrimeSpectrum R) ↔ Subsingleton R := by lemma primeSpectrum_empty_iff : IsEmpty (PrimeSpectrum R) ↔ Subsingleton R := by
constructor constructor <;> contrapose
. contrapose . rw [not_isEmpty_iff, ←not_nontrivial_iff_subsingleton, not_not]
rw [not_isEmpty_iff, ←not_nontrivial_iff_subsingleton, not_not]
apply PrimeSpectrum.instNonemptyPrimeSpectrum apply PrimeSpectrum.instNonemptyPrimeSpectrum
. intro h . intro hneg h
by_contra hneg exact hneg primeSpectrum_empty_of_subsingleton
rw [not_isEmpty_iff] at hneg
rcases hneg with ⟨a, ha⟩
exact primeSpectrum_empty_of_subsingleton ⟨a, ha⟩
/-- A ring has Krull dimension -∞ if and only if it is the zero ring -/ /-- A ring has Krull dimension -∞ if and only if it is the zero ring -/
lemma dim_eq_bot_iff : krullDim R = ⊥ ↔ Subsingleton R := by lemma dim_eq_bot_iff : krullDim R = ⊥ ↔ Subsingleton R := by
unfold Ideal.krullDim rw [Ideal.krullDim, ←primeSpectrum_empty_iff, iSup_eq_bot]
rw [←primeSpectrum_empty_iff, iSup_eq_bot]
constructor <;> intro h constructor <;> intro h
. rw [←not_nonempty_iff] . rw [←not_nonempty_iff]
rintro ⟨a, ha⟩ rintro ⟨a, ha⟩
specialize h ⟨a, ha⟩ cases h ⟨a, ha⟩
tauto
. rw [h.forall_iff] . rw [h.forall_iff]
trivial trivial
@ -246,7 +231,7 @@ lemma not_maximal_of_lt_prime {p : Ideal R} {q : Ideal R} (hq : IsPrime q) (h :
/-- Krull dimension is ≤ 0 if and only if all primes are maximal. -/ /-- Krull dimension is ≤ 0 if and only if all primes are maximal. -/
lemma dim_le_zero_iff : krullDim R ≤ 0 ↔ ∀ I : PrimeSpectrum R, IsMaximal I.asIdeal := by lemma dim_le_zero_iff : krullDim R ≤ 0 ↔ ∀ I : PrimeSpectrum R, IsMaximal I.asIdeal := by
show ((_ : WithBot ℕ∞) ≤ (0 : )) ↔ _ show ((_ : WithBot ℕ∞) ≤ (0 : )) ↔ _
rw [krullDim_le_iff R 0] rw [krullDim_le_iff]
constructor <;> intro h I constructor <;> intro h I
. contrapose! h . contrapose! h
have ⟨𝔪, h𝔪⟩ := I.asIdeal.exists_le_maximal (IsPrime.ne_top I.IsPrime) have ⟨𝔪, h𝔪⟩ := I.asIdeal.exists_le_maximal (IsPrime.ne_top I.IsPrime)
@ -294,26 +279,23 @@ lemma dim_eq_zero_iff [Nontrivial R] : krullDim R = 0 ↔ ∀ I : PrimeSpectrum
/-- In a field, the unique prime ideal is the zero ideal. -/ /-- In a field, the unique prime ideal is the zero ideal. -/
@[simp] @[simp]
lemma field_prime_bot {K: Type _} [Field K] {P : Ideal K} : IsPrime P ↔ P = ⊥ := by lemma field_prime_bot {K: Type _} [Field K] {P : Ideal K} : IsPrime P ↔ P = ⊥ := by
constructor refine' ⟨fun primeP => Or.elim (eq_bot_or_top P) _ _, fun botP => _⟩
· intro primeP · intro P_top; exact P_top
obtain T := eq_bot_or_top P . intro P_bot; exact False.elim (primeP.ne_top P_bot)
have : ¬P = := IsPrime.ne_top primeP · rw [botP]
tauto
· intro botP
rw [botP]
exact bot_prime exact bot_prime
/-- In a field, all primes have height 0. -/ /-- In a field, all primes have height 0. -/
lemma field_prime_height_bot {K: Type _} [Nontrivial K] [Field K] (P : PrimeSpectrum K) : height P = ⊥ := by lemma field_prime_height_zero {K: Type _} [Nontrivial K] [Field K] (P : PrimeSpectrum K) : height P = 0 := by
have : IsPrime P.asIdeal := P.IsPrime have : IsPrime P.asIdeal := P.IsPrime
rw [field_prime_bot] at this rw [field_prime_bot] at this
have : P = ⊥ := PrimeSpectrum.ext P ⊥ this have : P = ⊥ := PrimeSpectrum.ext P ⊥ this
rwa [height_bot_iff_bot] rwa [height_zero_iff_bot]
/-- The Krull dimension of a field is 0. -/ /-- The Krull dimension of a field is 0. -/
lemma dim_field_eq_zero {K : Type _} [Field K] : krullDim K = 0 := by lemma dim_field_eq_zero {K : Type _} [Field K] : krullDim K = 0 := by
unfold krullDim unfold krullDim
simp only [field_prime_height_bot, ciSup_unique] simp only [field_prime_height_zero, ciSup_unique]
/-- A domain with Krull dimension 0 is a field. -/ /-- A domain with Krull dimension 0 is a field. -/
lemma domain_dim_zero.isField {D: Type _} [CommRing D] [IsDomain D] (h: krullDim D = 0) : IsField D := by lemma domain_dim_zero.isField {D: Type _} [CommRing D] [IsDomain D] (h: krullDim D = 0) : IsField D := by
@ -353,7 +335,7 @@ lemma dim_le_one_iff : krullDim R ≤ 1 ↔ Ring.DimensionLEOne R := sorry
applies only to dimension zero rings and domains of dimension 1. -/ applies only to dimension zero rings and domains of dimension 1. -/
lemma dim_le_one_of_dimLEOne : Ring.DimensionLEOne R → krullDim R ≤ 1 := by lemma dim_le_one_of_dimLEOne : Ring.DimensionLEOne R → krullDim R ≤ 1 := by
show _ → ((_ : WithBot ℕ∞) ≤ (1 : )) show _ → ((_ : WithBot ℕ∞) ≤ (1 : ))
rw [krullDim_le_iff R 1] rw [krullDim_le_iff]
intro H p intro H p
apply le_of_not_gt apply le_of_not_gt
intro h intro h
@ -374,12 +356,67 @@ lemma dim_le_one_of_pid [IsDomain R] [IsPrincipalIdealRing R] : krullDim R ≤ 1
rw [dim_le_one_iff] rw [dim_le_one_iff]
exact Ring.DimensionLEOne.principal_ideal_ring R exact Ring.DimensionLEOne.principal_ideal_ring R
private lemma singleton_chainHeight_le_one {α : Type _} {x : α} [Preorder α] : Set.chainHeight {x} ≤ 1 := by
unfold Set.chainHeight
simp only [iSup_le_iff, Nat.cast_le_one]
intro L h
unfold Set.subchain at h
simp only [Set.mem_singleton_iff, Set.mem_setOf_eq] at h
rcases L with (_ | ⟨a,L⟩)
. simp only [List.length_nil, zero_le]
rcases L with (_ | ⟨b,L⟩)
. simp only [List.length_singleton, le_refl]
simp only [List.chain'_cons, List.find?, List.mem_cons, forall_eq_or_imp] at h
rcases h with ⟨⟨h1, _⟩, ⟨rfl, rfl, _⟩⟩
exact absurd h1 (lt_irrefl _)
/-- The ring of polynomials over a field has dimension one. -/ /-- The ring of polynomials over a field has dimension one. -/
lemma polynomial_over_field_dim_one {K : Type} [Nontrivial K] [Field K] : krullDim (Polynomial K) = 1 := by lemma polynomial_over_field_dim_one {K : Type} [Nontrivial K] [Field K] : krullDim (Polynomial K) = 1 := by
rw [le_antisymm_iff] rw [le_antisymm_iff]
let X := @Polynomial.X K _ let X := @Polynomial.X K _
constructor constructor
· exact dim_le_one_of_pid · unfold krullDim
apply @iSup_le (WithBot ℕ∞) _ _ _ _
intro I
have PIR : IsPrincipalIdealRing (Polynomial K) := by infer_instance
by_cases I = ⊥
· rw [← height_zero_iff_bot] at h
simp only [WithBot.coe_le_one, ge_iff_le]
rw [h]
exact bot_le
· push_neg at h
have : I.asIdeal ≠ ⊥ := by
by_contra a
have : I = ⊥ := PrimeSpectrum.ext I ⊥ a
contradiction
have maxI := IsPrime.to_maximal_ideal this
have sngletn : ∀P, P ∈ {J | J < I} ↔ P = ⊥ := by
intro P
constructor
· intro H
simp only [Set.mem_setOf_eq] at H
by_contra x
push_neg at x
have : P.asIdeal ≠ ⊥ := by
by_contra a
have : P = ⊥ := PrimeSpectrum.ext P ⊥ a
contradiction
have maxP := IsPrime.to_maximal_ideal this
have IneTop := IsMaximal.ne_top maxI
have : P ≤ I := le_of_lt H
rw [←PrimeSpectrum.asIdeal_le_asIdeal] at this
have : P.asIdeal = I.asIdeal := Ideal.IsMaximal.eq_of_le maxP IneTop this
have : P = I := PrimeSpectrum.ext P I this
replace H : P ≠ I := ne_of_lt H
contradiction
· intro pBot
simp only [Set.mem_setOf_eq, pBot]
exact lt_of_le_of_ne bot_le h.symm
replace sngletn : {J | J < I} = {⊥} := Set.ext sngletn
unfold height
rw [sngletn]
simp only [WithBot.coe_le_one, ge_iff_le]
exact singleton_chainHeight_le_one
· suffices : ∃I : PrimeSpectrum (Polynomial K), 1 ≤ (height I : WithBot ℕ∞) · suffices : ∃I : PrimeSpectrum (Polynomial K), 1 ≤ (height I : WithBot ℕ∞)
· obtain ⟨I, h⟩ := this · obtain ⟨I, h⟩ := this
have : (height I : WithBot ℕ∞) ≤ ⨆ (I : PrimeSpectrum (Polynomial K)), ↑(height I) := by have : (height I : WithBot ℕ∞) ≤ ⨆ (I : PrimeSpectrum (Polynomial K)), ↑(height I) := by

View file

@ -1,13 +1,3 @@
import Mathlib.RingTheory.Ideal.Operations
import Mathlib.RingTheory.FiniteType
import Mathlib.Order.Height
import Mathlib.RingTheory.Polynomial.Quotient
import Mathlib.RingTheory.PrincipalIdealDomain
import Mathlib.RingTheory.DedekindDomain.Basic
import Mathlib.RingTheory.Ideal.Quotient
import Mathlib.RingTheory.Localization.AtPrime
import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic
import Mathlib.Order.ConditionallyCompleteLattice.Basic
import CommAlg.krull import CommAlg.krull
section ChainLemma section ChainLemma
@ -132,7 +122,6 @@ lemma ht_adjoin_x_eq_ht_add_one [Nontrivial R] (I : PrimeSpectrum R) : height I
apply hl.2 apply hl.2
exact hb exact hb
#check ( : ℕ∞)
/- /-
dim R + 1 ≤ dim R[X] dim R + 1 ≤ dim R[X]
-/ -/

View file

@ -22,7 +22,8 @@ private lemma singleton_bot_chainHeight_one {α : Type} [Preorder α] [Bot α] :
exact absurd h1 (lt_irrefl _) exact absurd h1 (lt_irrefl _)
/-- The ring of polynomials over a field has dimension one. -/ /-- The ring of polynomials over a field has dimension one. -/
lemma polynomial_over_field_dim_one {K : Type} [Nontrivial K] [Field K] : krullDim (Polynomial K) = 1 := by -- It's the exact same lemma as in krull.lean, added ' to avoid conflict
lemma polynomial_over_field_dim_one' {K : Type} [Nontrivial K] [Field K] : krullDim (Polynomial K) = 1 := by
rw [le_antisymm_iff] rw [le_antisymm_iff]
let X := @Polynomial.X K _ let X := @Polynomial.X K _
constructor constructor