Almost completed polynomial_over_field_dim_one

This commit is contained in:
Sayantan Santra 2023-06-16 00:22:40 -07:00
parent d4a2a416f5
commit d2836ad8f8
Signed by: SinTan1729
GPG key ID: EB3E68BFBA25C85F
2 changed files with 54 additions and 17 deletions

View file

@ -60,7 +60,8 @@ lemma height_le_krullDim (I : PrimeSpectrum R) : height I ≤ krullDim R :=
le_iSup (λ I : PrimeSpectrum R => (height I : WithBot ℕ∞)) I le_iSup (λ I : PrimeSpectrum R => (height I : WithBot ℕ∞)) I
/-- 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. -/
lemma height_bot_iff_bot {D: Type} [CommRing D] [IsDomain D] (P : PrimeSpectrum D) : height P = ⊥ ↔ P = ⊥ := by @[simp]
lemma height_bot_iff_bot {D: Type} [CommRing D] [IsDomain D] {P : PrimeSpectrum D} : height P = ⊥ ↔ P = ⊥ := by
constructor constructor
· intro h · intro h
unfold height at h unfold height at h
@ -263,7 +264,7 @@ 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 constructor
· intro primeP · intro primeP
obtain T := eq_bot_or_top P obtain T := eq_bot_or_top P
@ -274,9 +275,13 @@ lemma field_prime_bot {K: Type _} [Field K] (P : Ideal K) : IsPrime P ↔ P =
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 _} [Field K] (P : PrimeSpectrum K) : height P = ⊥ := by lemma field_prime_height_bot {K: Type _} [Nontrivial K] [Field K] {P : PrimeSpectrum K} : height P = ⊥ := by
-- This should be doable by using field_prime_height_bot -- This should be doable by
-- and height_bot_iff_bot -- have : IsPrime P.asIdeal := P.IsPrime
-- rw [field_prime_bot] at this
-- have : P = ⊥ := PrimeSpectrum.ext P ⊥ this
-- rw [height_bot_iff_bot]
-- Need to check what's happening
rw [bot_eq_zero] rw [bot_eq_zero]
unfold height unfold height
simp only [Set.chainHeight_eq_zero_iff] simp only [Set.chainHeight_eq_zero_iff]

View file

@ -1,39 +1,71 @@
import CommAlg.krull import CommAlg.krull
import Mathlib.RingTheory.Ideal.Operations import Mathlib.RingTheory.Ideal.Operations
import Mathlib.RingTheory.FiniteType
import Mathlib.Order.Height import Mathlib.Order.Height
import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.RingTheory.PrincipalIdealDomain
import Mathlib.RingTheory.DedekindDomain.Basic import Mathlib.RingTheory.DedekindDomain.Basic
import Mathlib.RingTheory.Ideal.Quotient
import Mathlib.RingTheory.Ideal.MinimalPrime
import Mathlib.RingTheory.Localization.AtPrime
import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic
import Mathlib.Order.ConditionallyCompleteLattice.Basic
namespace Ideal namespace Ideal
/-- 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
-- unfold krullDim
rw [le_antisymm_iff] rw [le_antisymm_iff]
let X := @Polynomial.X K _
constructor constructor
· · unfold krullDim
sorry apply @iSup_le (WithBot ℕ∞) _ _ _ _
intro I
have PIR : IsPrincipalIdealRing (Polynomial K) := by infer_instance
by_cases I = ⊥
· rw [← height_bot_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 singleton : ∀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 singleton : {J | J < I} = {⊥} := Set.ext singleton
unfold height
sorry
· 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
apply @le_iSup (WithBot ℕ∞) _ _ _ I apply @le_iSup (WithBot ℕ∞) _ _ _ I
exact le_trans h this exact le_trans h this
have primeX : Prime Polynomial.X := @Polynomial.prime_X K _ _ have primeX : Prime Polynomial.X := @Polynomial.prime_X K _ _
let X := @Polynomial.X K _
have : IsPrime (span {X}) := by have : IsPrime (span {X}) := by
refine Iff.mpr (span_singleton_prime ?hp) primeX refine (span_singleton_prime ?hp).mpr primeX
exact Polynomial.X_ne_zero exact Polynomial.X_ne_zero
let P := PrimeSpectrum.mk (span {X}) this let P := PrimeSpectrum.mk (span {X}) this
unfold height unfold height
use P use P
have : ⊥ ∈ {J | J < P} := by have : ⊥ ∈ {J | J < P} := by
simp only [Set.mem_setOf_eq] simp only [Set.mem_setOf_eq, bot_lt_iff_ne_bot]
rw [bot_lt_iff_ne_bot]
suffices : P.asIdeal ≠ ⊥ suffices : P.asIdeal ≠ ⊥
· by_contra x · by_contra x
rw [PrimeSpectrum.ext_iff] at x rw [PrimeSpectrum.ext_iff] at x