mirror of
https://github.com/GTBarkley/comm_alg.git
synced 2024-12-26 07:38:36 -06:00
Completed polynomial_over_field_dim_one
This commit is contained in:
parent
bcb867258a
commit
e735a5254f
1 changed files with 19 additions and 3 deletions
|
@ -7,6 +7,20 @@ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic
|
||||||
|
|
||||||
namespace Ideal
|
namespace Ideal
|
||||||
|
|
||||||
|
private lemma singleton_chainHeight_one {α : Type} [Preorder α] [Bot α] : Set.chainHeight {(⊥ : α)} ≤ 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]
|
||||||
|
@ -27,7 +41,7 @@ lemma polynomial_over_field_dim_one {K : Type} [Nontrivial K] [Field K] : krullD
|
||||||
have : I = ⊥ := PrimeSpectrum.ext I ⊥ a
|
have : I = ⊥ := PrimeSpectrum.ext I ⊥ a
|
||||||
contradiction
|
contradiction
|
||||||
have maxI := IsPrime.to_maximal_ideal this
|
have maxI := IsPrime.to_maximal_ideal this
|
||||||
have singleton : ∀P, P ∈ {J | J < I} ↔ P = ⊥ := by
|
have sngletn : ∀P, P ∈ {J | J < I} ↔ P = ⊥ := by
|
||||||
intro P
|
intro P
|
||||||
constructor
|
constructor
|
||||||
· intro H
|
· intro H
|
||||||
|
@ -49,9 +63,11 @@ lemma polynomial_over_field_dim_one {K : Type} [Nontrivial K] [Field K] : krullD
|
||||||
· intro pBot
|
· intro pBot
|
||||||
simp only [Set.mem_setOf_eq, pBot]
|
simp only [Set.mem_setOf_eq, pBot]
|
||||||
exact lt_of_le_of_ne bot_le h.symm
|
exact lt_of_le_of_ne bot_le h.symm
|
||||||
replace singleton : {J | J < I} = {⊥} := Set.ext singleton
|
replace sngletn : {J | J < I} = {⊥} := Set.ext sngletn
|
||||||
unfold height
|
unfold height
|
||||||
sorry
|
rw [sngletn]
|
||||||
|
simp only [WithBot.coe_le_one, ge_iff_le]
|
||||||
|
exact singleton_chainHeight_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
|
||||||
|
|
Loading…
Reference in a new issue