From 7173aefe8d9e171bac69d087b87a9276d30ba70b Mon Sep 17 00:00:00 2001 From: GTBarkley Date: Wed, 14 Jun 2023 17:43:16 +0000 Subject: [PATCH 1/3] changed type of 1 in dim_le_one_of_dimLEOne --- CommAlg/krull.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/CommAlg/krull.lean b/CommAlg/krull.lean index 06e6d0d..27c3d7e 100644 --- a/CommAlg/krull.lean +++ b/CommAlg/krull.lean @@ -197,7 +197,8 @@ height ๐”ญ > (n : WithBot โ„•โˆž) โ†” โˆƒ c : List (PrimeSpectrum R), c.Chain' /-- The converse of this is false, because the definition of "dimension โ‰ค 1" in mathlib 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 : โ„•)) rw [krullDim_le_iff R 1] intro H p apply le_of_not_gt From 3c4cfeab65b8d9d4d1c07282cf30596f3d017231 Mon Sep 17 00:00:00 2001 From: GTBarkley Date: Wed, 14 Jun 2023 20:28:19 +0000 Subject: [PATCH 2/3] proved dim_le_zero_iff --- CommAlg/grant.lean | 42 ++++++++++++++++++++ CommAlg/krull.lean | 97 +++++++++++++++++++++++++++++++++------------- 2 files changed, 111 insertions(+), 28 deletions(-) diff --git a/CommAlg/grant.lean b/CommAlg/grant.lean index c12b189..76a470c 100644 --- a/CommAlg/grant.lean +++ b/CommAlg/grant.lean @@ -171,6 +171,48 @@ lemma dim_le_one_of_dimLEOne : Ring.DimensionLEOne R โ†’ krullDim R โ‰ค (1 : apply (IsCoatom.lt_iff H.out).mp exact hc2 --refine Iff.mp radical_eq_top (?_ (id (Eq.symm hc3))) + +lemma not_maximal_of_lt_prime {p : Ideal R} {q : Ideal R} (hq : IsPrime q) (h : p < q) : ยฌIsMaximal p := by + intro hp + apply IsPrime.ne_top hq + apply (IsCoatom.lt_iff hp.out).mp + exact h + +lemma dim_le_zero_iff : krullDim R โ‰ค 0 โ†” โˆ€ I : PrimeSpectrum R, IsMaximal I.asIdeal := by + show ((_ : WithBot โ„•โˆž) โ‰ค (0 : โ„•)) โ†” _ + rw [krullDim_le_iff R 0] + constructor <;> intro h I + . contrapose! h + have โŸจ๐”ช, h๐”ชโŸฉ := I.asIdeal.exists_le_maximal (IsPrime.ne_top I.IsPrime) + let ๐”ชp := (โŸจ๐”ช, IsMaximal.isPrime h๐”ช.1โŸฉ : PrimeSpectrum R) + have hstrct : I < ๐”ชp := by + apply lt_of_le_of_ne h๐”ช.2 + intro hcontr + rw [hcontr] at h + exact h h๐”ช.1 + use ๐”ชp + show (_ : WithBot โ„•โˆž) > (0 : โ„•โˆž) + rw [_root_.lt_height_iff''] + use [I] + constructor + . exact List.chain'_singleton _ + . constructor + . intro I' hI' + simp at hI' + rwa [hI'] + . simp + . contrapose! h + change (_ : WithBot โ„•โˆž) > (0 : โ„•โˆž) at h + rw [_root_.lt_height_iff''] at h + obtain โŸจc, _, hc2, hc3โŸฉ := h + norm_cast at hc3 + rw [List.length_eq_one] at hc3 + obtain โŸจ๐”ฎ, h๐”ฎโŸฉ := hc3 + use ๐”ฎ + specialize hc2 ๐”ฎ (h๐”ฎ โ–ธ (List.mem_singleton.mpr rfl)) + apply not_maximal_of_lt_prime _ I.IsPrime + exact hc2 + end Krull section iSupWithBot diff --git a/CommAlg/krull.lean b/CommAlg/krull.lean index 27c3d7e..d58a065 100644 --- a/CommAlg/krull.lean +++ b/CommAlg/krull.lean @@ -65,6 +65,34 @@ lemma krullDim_eq_height [LocalRing R] : krullDim R = height (closedPoint R) := exact I.2.1 . simp +lemma lt_height_iff' {๐”ญ : PrimeSpectrum R} {n : โ„•โˆž} : +height ๐”ญ > n โ†” โˆƒ c : List (PrimeSpectrum R), c.Chain' (ยท < ยท) โˆง (โˆ€ ๐”ฎ โˆˆ c, ๐”ฎ < ๐”ญ) โˆง c.length = n + 1 := by + rcases n with _ | n + . constructor <;> intro h <;> exfalso + . exact (not_le.mpr h) le_top + . tauto + have (m : โ„•โˆž) : m > some n โ†” m โ‰ฅ some (n + 1) := 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 โ„•) = (_:โ„•โˆž)) + rw [{J | J < ๐”ญ}.le_chainHeight_iff] + show (โˆƒ c, (List.Chain' _ c โˆง โˆ€๐”ฎ, ๐”ฎ โˆˆ c โ†’ ๐”ฎ < ๐”ญ) โˆง _) โ†” _ + constructor <;> rintro โŸจc, hcโŸฉ <;> use c + . tauto + . change _ โˆง _ โˆง (List.length c : โ„•โˆž) = n + 1 at hc + norm_cast at hc + tauto + +lemma lt_height_iff'' {๐”ญ : PrimeSpectrum R} {n : โ„•โˆž} : +height ๐”ญ > (n : WithBot โ„•โˆž) โ†” โˆƒ c : List (PrimeSpectrum R), c.Chain' (ยท < ยท) โˆง (โˆ€ ๐”ฎ โˆˆ c, ๐”ฎ < ๐”ญ) โˆง c.length = n + 1 := by + show (_ < _) โ†” _ + rw [WithBot.coe_lt_coe] + exact lt_height_iff' + #check height_le_krullDim --some propositions that would be nice to be able to eventually @@ -99,6 +127,47 @@ lemma krullDim_nonneg_of_nontrivial (R : Type _) [CommRing R] [Nontrivial R] : lift (Ideal.krullDim R) to โ„•โˆž using h with k use k +lemma not_maximal_of_lt_prime {p : Ideal R} {q : Ideal R} (hq : IsPrime q) (h : p < q) : ยฌIsMaximal p := by + intro hp + apply IsPrime.ne_top hq + apply (IsCoatom.lt_iff hp.out).mp + exact h + +lemma dim_le_zero_iff : krullDim R โ‰ค 0 โ†” โˆ€ I : PrimeSpectrum R, IsMaximal I.asIdeal := by + show ((_ : WithBot โ„•โˆž) โ‰ค (0 : โ„•)) โ†” _ + rw [krullDim_le_iff R 0] + constructor <;> intro h I + . contrapose! h + have โŸจ๐”ช, h๐”ชโŸฉ := I.asIdeal.exists_le_maximal (IsPrime.ne_top I.IsPrime) + let ๐”ชp := (โŸจ๐”ช, IsMaximal.isPrime h๐”ช.1โŸฉ : PrimeSpectrum R) + have hstrct : I < ๐”ชp := by + apply lt_of_le_of_ne h๐”ช.2 + intro hcontr + rw [hcontr] at h + exact h h๐”ช.1 + use ๐”ชp + show (_ : WithBot โ„•โˆž) > (0 : โ„•โˆž) + rw [lt_height_iff''] + use [I] + constructor + . exact List.chain'_singleton _ + . constructor + . intro I' hI' + simp at hI' + rwa [hI'] + . simp + . contrapose! h + change (_ : WithBot โ„•โˆž) > (0 : โ„•โˆž) at h + rw [lt_height_iff''] at h + obtain โŸจc, _, hc2, hc3โŸฉ := h + norm_cast at hc3 + rw [List.length_eq_one] at hc3 + obtain โŸจ๐”ฎ, h๐”ฎโŸฉ := hc3 + use ๐”ฎ + specialize hc2 ๐”ฎ (h๐”ฎ โ–ธ (List.mem_singleton.mpr rfl)) + apply not_maximal_of_lt_prime I.IsPrime + exact hc2 + lemma dim_eq_zero_iff [Nontrivial R] : krullDim R = 0 โ†” โˆ€ I : PrimeSpectrum R, IsMaximal I.asIdeal := by constructor <;> intro h . intro I @@ -167,34 +236,6 @@ lemma dim_eq_zero_iff_field {D: Type _} [CommRing D] [IsDomain D] : krullDim D = -- This lemma is false! lemma dim_le_one_iff : krullDim R โ‰ค 1 โ†” Ring.DimensionLEOne R := sorry -lemma lt_height_iff' {๐”ญ : PrimeSpectrum R} {n : โ„•โˆž} : -height ๐”ญ > n โ†” โˆƒ c : List (PrimeSpectrum R), c.Chain' (ยท < ยท) โˆง (โˆ€ ๐”ฎ โˆˆ c, ๐”ฎ < ๐”ญ) โˆง c.length = n + 1 := by - rcases n with _ | n - . constructor <;> intro h <;> exfalso - . exact (not_le.mpr h) le_top - . tauto - have (m : โ„•โˆž) : m > some n โ†” m โ‰ฅ some (n + 1) := 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 โ„•) = (_:โ„•โˆž)) - rw [{J | J < ๐”ญ}.le_chainHeight_iff] - show (โˆƒ c, (List.Chain' _ c โˆง โˆ€๐”ฎ, ๐”ฎ โˆˆ c โ†’ ๐”ฎ < ๐”ญ) โˆง _) โ†” _ - constructor <;> rintro โŸจc, hcโŸฉ <;> use c - . tauto - . change _ โˆง _ โˆง (List.length c : โ„•โˆž) = n + 1 at hc - norm_cast at hc - tauto - -lemma lt_height_iff'' {๐”ญ : PrimeSpectrum R} {n : โ„•โˆž} : -height ๐”ญ > (n : WithBot โ„•โˆž) โ†” โˆƒ c : List (PrimeSpectrum R), c.Chain' (ยท < ยท) โˆง (โˆ€ ๐”ฎ โˆˆ c, ๐”ฎ < ๐”ญ) โˆง c.length = n + 1 := by - show (_ < _) โ†” _ - rw [WithBot.coe_lt_coe] - exact lt_height_iff' - /-- The converse of this is false, because the definition of "dimension โ‰ค 1" in mathlib applies only to dimension zero rings and domains of dimension 1. -/ lemma dim_le_one_of_dimLEOne : Ring.DimensionLEOne R โ†’ krullDim R โ‰ค 1 := by From 46908891497cd2510c470a563e127b48843a2b76 Mon Sep 17 00:00:00 2001 From: GTBarkley Date: Wed, 14 Jun 2023 21:35:36 +0000 Subject: [PATCH 3/3] proved dim_eq_zero_iff --- CommAlg/krull.lean | 15 ++++++++++----- 1 file changed, 10 insertions(+), 5 deletions(-) diff --git a/CommAlg/krull.lean b/CommAlg/krull.lean index d58a065..b2d0230 100644 --- a/CommAlg/krull.lean +++ b/CommAlg/krull.lean @@ -169,11 +169,16 @@ lemma dim_le_zero_iff : krullDim R โ‰ค 0 โ†” โˆ€ I : PrimeSpectrum R, IsMaximal exact hc2 lemma dim_eq_zero_iff [Nontrivial R] : krullDim R = 0 โ†” โˆ€ I : PrimeSpectrum R, IsMaximal I.asIdeal := by - constructor <;> intro h - . intro I - sorry - . sorry - + rw [โ†dim_le_zero_iff] + obtain โŸจn, hnโŸฉ := krullDim_nonneg_of_nontrivial R + have : n โ‰ฅ 0 := zero_le n + change _ โ‰ค _ at this + rw [โ†WithBot.coe_le_coe,โ†hn] at this + change (0 : WithBot โ„•โˆž) โ‰ค _ at this + constructor <;> intro h' + rw [h'] + exact le_antisymm h' this + @[simp] lemma field_prime_bot {K: Type _} [Field K] (P : Ideal K) : IsPrime P โ†” P = โŠฅ := by constructor