diff --git a/.DS_Store b/.DS_Store index 6d98085..4d3a1f2 100644 Binary files a/.DS_Store and b/.DS_Store differ diff --git a/CommAlg/grant.lean b/CommAlg/grant.lean index ce5041d..c12b189 100644 --- a/CommAlg/grant.lean +++ b/CommAlg/grant.lean @@ -83,6 +83,16 @@ height ๐”ญ > n โ†” โˆƒ c : List (PrimeSpectrum R), c.Chain' (ยท < ยท) โˆง (โˆ€ 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' _ + +lemma height_le_iff {๐”ญ : PrimeSpectrum R} {n : โ„•โˆž} : + height ๐”ญ โ‰ค n โ†” โˆ€ c : List (PrimeSpectrum R), c โˆˆ {I : PrimeSpectrum R | I < ๐”ญ}.subchain โˆง c.length = n + 1 := by + sorry + lemma krullDim_nonneg_of_nontrivial [Nontrivial R] : โˆƒ n : โ„•โˆž, Ideal.krullDim R = n := by have h := dim_eq_bot_iff.not.mpr (not_subsingleton R) lift (Ideal.krullDim R) to โ„•โˆž using h with k @@ -116,19 +126,103 @@ lemma dim_eq_bot_iff : krullDim R = โŠฅ โ†” Subsingleton R := by constructor <;> intro h . rw [โ†not_nonempty_iff] rintro โŸจa, haโŸฉ - -- specialize h โŸจa, haโŸฉ + specialize h โŸจa, haโŸฉ tauto . rw [h.forall_iff] trivial - #check (sorry : False) #check (sorryAx) #check (4 : WithBot โ„•โˆž) #check List.sum +#check (_ โˆˆ (_ : List _)) +variable (ฮฑ : Type ) +#synth Membership ฮฑ (List ฮฑ) +#check bot_lt_iff_ne_bot -- #check ((4 : โ„•โˆž) : WithBot (WithTop โ„•)) -- #check ( (Set.chainHeight s) : WithBot (โ„•โˆž)) -variable (P : PrimeSpectrum R) +/-- 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 + rw [krullDim_le_iff R 1] + -- unfold Ring.DimensionLEOne + intro H p + -- have Hp := H p.asIdeal + -- have Hp := fun h => (Hp h) p.IsPrime + apply le_of_not_gt + intro h + rcases ((lt_height_iff'' R).mp h) with โŸจc, โŸจhc1, hc2, hc3โŸฉโŸฉ + norm_cast at hc3 + rw [List.chain'_iff_get] at hc1 + specialize hc1 0 (by rw [hc3]; simp) + -- generalize hq0 : List.get _ _ = q0 at hc1 + set q0 : PrimeSpectrum R := List.get c { val := 0, isLt := _ } + set q1 : PrimeSpectrum R := List.get c { val := 1, isLt := _ } + -- have hq0 : q0 โˆˆ c := List.get_mem _ _ _ + -- have hq1 : q1 โˆˆ c := List.get_mem _ _ _ + specialize hc2 q1 (List.get_mem _ _ _) + -- have aa := (bot_le : (โŠฅ : Ideal R) โ‰ค q0.asIdeal) + change q0.asIdeal < q1.asIdeal at hc1 + have q1nbot := Trans.trans (bot_le : โŠฅ โ‰ค q0.asIdeal) hc1 + specialize H q1.asIdeal (bot_lt_iff_ne_bot.mp q1nbot) q1.IsPrime + -- change q1.asIdeal < p.asIdeal at hc2 + apply IsPrime.ne_top p.IsPrime + apply (IsCoatom.lt_iff H.out).mp + exact hc2 + --refine Iff.mp radical_eq_top (?_ (id (Eq.symm hc3))) +end Krull -#check {J | J < P}.le_chainHeight_iff (n := 4) +section iSupWithBot + +variable {ฮฑ : Type _} [CompleteSemilatticeSup ฮฑ] {I : Type _} (f : I โ†’ ฮฑ) + +#synth SupSet (WithBot โ„•โˆž) + +protected lemma WithBot.iSup_ge_coe_iff {a : ฮฑ} : + (a โ‰ค โจ† i : I, (f i : WithBot ฮฑ) ) โ†” โˆƒ i : I, f i โ‰ฅ a := by + rw [WithBot.coe_le_iff] + sorry + +end iSupWithBot + +section myGreatElab +open Lean Meta Elab + +syntax (name := lhsStx) "lhs% " term:max : term +syntax (name := rhsStx) "rhs% " term:max : term + +@[term_elab lhsStx, term_elab rhsStx] +def elabLhsStx : Term.TermElab := fun stx expectedType? => + match stx with + | `(lhs% $t) => do + let (lhs, _, eq) โ† mkExpected expectedType? + discard <| Term.elabTermEnsuringType t eq + return lhs + | `(rhs% $t) => do + let (_, rhs, eq) โ† mkExpected expectedType? + discard <| Term.elabTermEnsuringType t eq + return rhs + | _ => throwUnsupportedSyntax +where + mkExpected expectedType? := do + let ฮฑ โ† + if let some expectedType := expectedType? then + pure expectedType + else + mkFreshTypeMVar + let lhs โ† mkFreshExprMVar ฮฑ + let rhs โ† mkFreshExprMVar ฮฑ + let u โ† getLevel ฮฑ + let eq := mkAppN (.const ``Eq [u]) #[ฮฑ, lhs, rhs] + return (lhs, rhs, eq) + +#check lhs% (add_comm 1 2) +#check rhs% (add_comm 1 2) +#check lhs% (add_comm _ _ : _ = 1 + 2) + +example (x y : ฮฑ) (h : x = y) : lhs% h = rhs% h := h + +def lhsOf {ฮฑ : Sort _} {x y : ฮฑ} (h : x = y) : ฮฑ := x + +#check lhsOf (add_comm 1 2) \ No newline at end of file diff --git a/CommAlg/hilbertpolynomial.lean b/CommAlg/hilbertpolynomial.lean new file mode 100644 index 0000000..64650dc --- /dev/null +++ b/CommAlg/hilbertpolynomial.lean @@ -0,0 +1,137 @@ +import Mathlib.Order.KrullDimension +import Mathlib.Order.JordanHolder +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.Order.Height +import Mathlib.RingTheory.Ideal.Basic +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.LinearAlgebra.Finsupp +import Mathlib.RingTheory.GradedAlgebra.Basic +import Mathlib.RingTheory.GradedAlgebra.HomogeneousIdeal +import Mathlib.Algebra.Module.GradedModule +import Mathlib.RingTheory.Ideal.AssociatedPrime +import Mathlib.RingTheory.Noetherian +import Mathlib.RingTheory.Artinian +import Mathlib.Algebra.Module.GradedModule +import Mathlib.RingTheory.Noetherian +import Mathlib.RingTheory.Finiteness +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.RingTheory.FiniteType +import Mathlib.Order.Height +import Mathlib.RingTheory.PrincipalIdealDomain +import Mathlib.RingTheory.DedekindDomain.Basic +import Mathlib.RingTheory.Ideal.Quotient +import Mathlib.RingTheory.Localization.AtPrime +import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Algebra.DirectSum.Ring +import Mathlib.RingTheory.Ideal.LocalRing + +-- Setting for "library_search" +set_option maxHeartbeats 0 +macro "ls" : tactic => `(tactic|library_search) + +-- New tactic "obviously" +macro "obviously" : tactic => + `(tactic| ( + first + | dsimp; simp; done; dbg_trace "it was dsimp simp" + | simp; done; dbg_trace "it was simp" + | tauto; done; dbg_trace "it was tauto" + | simp; tauto; done; dbg_trace "it was simp tauto" + | rfl; done; dbg_trace "it was rfl" + | norm_num; done; dbg_trace "it was norm_num" + | /-change (@Eq โ„ _ _);-/ linarith; done; dbg_trace "it was linarith" + -- | gcongr; done + | ring; done; dbg_trace "it was ring" + | trivial; done; dbg_trace "it was trivial" + -- | nlinarith; done + | fail "No, this is not obvious.")) + +-- @[BH, 1.5.6 (b)(ii)] +lemma ss (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] + [DirectSum.GCommRing ๐’œ] [DirectSum.Gmodule ๐’œ ๐“œ] (p : associatedPrimes (โจ i, ๐’œ i) (โจ i, ๐“œ i)) : true := by + sorry +-- Ideal.IsHomogeneous ๐’œ p + + + + +noncomputable def length ( A : Type _) (M : Type _) + [CommRing A] [AddCommGroup M] [Module A M] := Set.chainHeight {M' : Submodule A M | M' < โŠค} + + +def HomogeneousPrime { A ฯƒ : Type _} [CommRing A] [SetLike ฯƒ A] [AddSubmonoidClass ฯƒ A] (๐’œ : โ„ค โ†’ ฯƒ) [GradedRing ๐’œ] (I : Ideal A):= (Ideal.IsPrime I) โˆง (Ideal.IsHomogeneous ๐’œ I) +def HomogeneousMax { A ฯƒ : Type _} [CommRing A] [SetLike ฯƒ A] [AddSubmonoidClass ฯƒ A] (๐’œ : โ„ค โ†’ ฯƒ) [GradedRing ๐’œ] (I : Ideal A):= (Ideal.IsMaximal I) โˆง (Ideal.IsHomogeneous ๐’œ I) + +--theorem monotone_stabilizes_iff_noetherian : +-- (โˆ€ f : โ„• โ†’o Submodule R M, โˆƒ n, โˆ€ m, n โ‰ค m โ†’ f n = f m) โ†” IsNoetherian R M := by +-- rw [isNoetherian_iff_wellFounded, WellFounded.monotone_chain_condition] + +open GradedMonoid.GSmul +open DirectSum + +instance tada1 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค ) : SMul (๐’œ 0) (๐“œ i) + where smul x y := @Eq.rec โ„ค (0+i) (fun a _ => ๐“œ a) (GradedMonoid.GSmul.smul x y) i (zero_add i) + +lemma mylem (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [h : DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค) (a : ๐’œ 0) (m : ๐“œ i) : + of _ _ (a โ€ข m) = of _ _ a โ€ข of _ _ m := by + refine' Eq.trans _ (Gmodule.of_smul_of ๐’œ ๐“œ a m).symm + refine' of_eq_of_gradedMonoid_eq _ + exact Sigma.ext (zero_add _).symm <| eq_rec_heq _ _ + +instance tada2 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [h : DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค ) : SMulWithZero (๐’œ 0) (๐“œ i) := by + letI := SMulWithZero.compHom (โจ i, ๐“œ i) (of ๐’œ 0).toZeroHom + exact Function.Injective.smulWithZero (of ๐“œ i).toZeroHom Dfinsupp.single_injective (mylem ๐’œ ๐“œ i) + +instance tada3 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [h : DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค ): Module (๐’œ 0) (๐“œ i) := by + letI := Module.compHom (โจ j, ๐“œ j) (ofZeroRingHom ๐’œ) + exact Dfinsupp.single_injective.module (๐’œ 0) (of ๐“œ i) (mylem ๐’œ ๐“œ i) + +noncomputable def hilbert_function (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] + [DirectSum.GCommRing ๐’œ] + [DirectSum.Gmodule ๐’œ ๐“œ] (hilb : โ„ค โ†’ โ„ค) := โˆ€ i, hilb i = (ENat.toNat (length (๐’œ 0) (๐“œ i))) + +noncomputable def dimensionring { A: Type _} + [CommRing A] := krullDim (PrimeSpectrum A) + + +noncomputable def dimensionmodule ( A : Type _) (M : Type _) + [CommRing A] [AddCommGroup M] [Module A M] := krullDim (PrimeSpectrum (A โงธ ((โŠค : Submodule A M).annihilator)) ) + +-- lemma graded_local (๐’œ : โ„ค โ†’ Type _) [SetLike (โจ i, ๐’œ i)] (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +-- [DirectSum.GCommRing ๐’œ] +-- [DirectSum.Gmodule ๐’œ ๐“œ] (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) : โˆƒ ( I : Ideal ((โจ i, ๐’œ i))),(HomogeneousMax ๐’œ I) := sorry + + + + +@[simp] +def PolyType (f : โ„ค โ†’ โ„ค) (d : โ„•) := โˆƒ Poly : Polynomial โ„š, โˆƒ (N : โ„ค), โˆ€ (n : โ„ค), N โ‰ค n โ†’ f n = Polynomial.eval (n : โ„š) Poly โˆง d = Polynomial.degree Poly + +-- @[BH, 4.1.3] +theorem hilbert_polynomial (d : โ„•) (d1 : 1 โ‰ค d) (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +[DirectSum.GCommRing ๐’œ] +[DirectSum.Gmodule ๐’œ ๐“œ] (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) +(fingen : IsNoetherian (โจ i, ๐’œ i) (โจ i, ๐“œ i)) +(findim : dimensionmodule (โจ i, ๐’œ i) (โจ i, ๐“œ i) = d) (hilb : โ„ค โ†’ โ„ค) + (Hhilb: hilbert_function ๐’œ ๐“œ hilb) +: PolyType hilb (d - 1) := by + sorry + +-- @ +lemma Graded_quotient (๐’œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)][DirectSum.GCommRing ๐’œ] + : true := by + sorry + + +-- @Existence of a chain of submodules of graded submoduels of f.g graded R-mod M +lemma Exist_chain_of_graded_submodules (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] + [DirectSum.GCommRing ๐’œ] [DirectSum.Gmodule ๐’œ ๐“œ] (fingen : IsNoetherian (โจ i, ๐’œ i) (โจ i, ๐“œ i)) + : true := by + sorry +#print Exist_chain_of_graded_submodules +#print Finset \ No newline at end of file diff --git a/CommAlg/jayden(krull-dim-zero).lean b/CommAlg/jayden(krull-dim-zero).lean index 3651102..523eb90 100644 --- a/CommAlg/jayden(krull-dim-zero).lean +++ b/CommAlg/jayden(krull-dim-zero).lean @@ -1,97 +1,183 @@ import Mathlib.RingTheory.Ideal.Basic +import Mathlib.RingTheory.Ideal.Operations import Mathlib.RingTheory.JacobsonIdeal import Mathlib.RingTheory.Noetherian import Mathlib.Order.KrullDimension import Mathlib.RingTheory.Artinian import Mathlib.RingTheory.Ideal.Quotient import Mathlib.RingTheory.Nilpotent -import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic import Mathlib.AlgebraicGeometry.PrimeSpectrum.Maximal +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian import Mathlib.Data.Finite.Defs import Mathlib.Order.Height import Mathlib.RingTheory.DedekindDomain.Basic import Mathlib.RingTheory.Localization.AtPrime import Mathlib.Order.ConditionallyCompleteLattice.Basic import Mathlib.Algebra.Ring.Pi -import Mathlib.Topology.NoetherianSpace +import Mathlib.RingTheory.Finiteness + --- copy from krull.lean; the name of Krull dimension for rings is changed to krullDim' since krullDim already exists in the librrary namespace Ideal -variable (R : Type _) [CommRing R] (I : PrimeSpectrum R) +variable (R : Type _) [CommRing R] (P : PrimeSpectrum R) -noncomputable def height : โ„•โˆž := Set.chainHeight {J : PrimeSpectrum R | J < I} - -noncomputable def krullDim' (R : Type) [CommRing R] : WithBot โ„•โˆž := โจ† (I : PrimeSpectrum R), height R I --- copy ends - --- 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] : - IsNoetherianRing R โˆง krullDim' R = 0 โ†” IsArtinianRing R := by sorry - - -#check IsNoetherianRing - -#check krullDim - --- Repeats the definition of the length of a module by Monalisa -variable (M : Type _) [AddCommMonoid M] [Module R M] - --- change the definition of length -noncomputable def length := Set.chainHeight {M' : Submodule R M | M' < โŠค} - -#check length --- Stacks Lemma 10.53.6: R is Artinian iff R has finite length as an R-mod -lemma IsArtinian_iff_finite_length : IsArtinianRing R โ†” โˆƒ n : โ„•, length R R โ‰ค n := by sorry - --- Stacks Lemma 10.53.3: R is Artinian iff R has finitely many maximal ideals -lemma IsArtinian_iff_finite_max_ideal : IsArtinianRing R โ†” Finite (MaximalSpectrum R) := by sorry - --- Stacks Lemma 10.53.4: R Artinian => Jacobson ideal of R is nilpotent -lemma Jacobson_of_Artinian_is_nilpotent : IsArtinianRing R โ†’ IsNilpotent (Ideal.jacobson (โŠค : Ideal R)) := by sorry - - --- Stacks Definition 10.32.1: An ideal is locally nilpotent --- if every element is nilpotent -namespace Ideal -class IsLocallyNilpotent (I : Ideal R) : Prop := - h : โˆ€ x โˆˆ I, IsNilpotent x - -end Ideal - -#check Ideal.IsLocallyNilpotent - --- 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 --- its maximal ideals. Also, all primes are maximal - -lemma product_of_localization_at_maximal_ideal : Finite (MaximalSpectrum R) - โˆง Ideal.IsLocallyNilpotent (Ideal.jacobson (โŠค : Ideal R)) โ†’ Pi.commRing (MaximalSpectrum R) Localization.AtPrime R I - := by sorry --- Haven't finished this. - --- Stacks Lemma 10.31.5: R is Noetherian iff Spec(R) is a Noetherian space -lemma ring_Noetherian_iff_spec_Noetherian : IsNoetherianRing R - โ†” TopologicalSpace.NoetherianSpace (PrimeSpectrum R) := by sorry --- Use TopologicalSpace.NoetherianSpace.exists_finset_irreducible : --- Every closed subset of a noetherian space is a finite union --- of irreducible closed subsets. +noncomputable def height : โ„•โˆž := Set.chainHeight {J : PrimeSpectrum R | J < P} +noncomputable def krullDim (R : Type) [CommRing R] : + WithBot โ„•โˆž := โจ† (I : PrimeSpectrum R), height R I -- Stacks Lemma 10.26.1 (Should already exists) -- (1) The closure of a prime P is V(P) -- (2) the irreducible closed subsets are V(P) for P prime -- (3) the irreducible components are V(P) for P minimal prime --- Stacks Lemma 10.32.5: R Noetherian. I,J ideals. If J โŠ‚ โˆšI, then J ^ n โŠ‚ I for some n +-- Stacks Definition 10.32.1: An ideal is locally nilpotent +-- if every element is nilpotent +class IsLocallyNilpotent (I : Ideal R) : Prop := + h : โˆ€ x โˆˆ I, IsNilpotent x +#check Ideal.IsLocallyNilpotent +end Ideal + + +-- Repeats the definition of the length of a module by Monalisa +variable (R : Type _) [CommRing R] (I J : Ideal R) +variable (M : Type _) [AddCommMonoid M] [Module R M] + +-- change the definition of length of a module +namespace Module +noncomputable def length := Set.chainHeight {M' : Submodule R M | M' < โŠค} +end Module + +-- Stacks Lemma 10.31.5: R is Noetherian iff Spec(R) is a Noetherian space +example [IsNoetherianRing R] : + TopologicalSpace.NoetherianSpace (PrimeSpectrum R) := + inferInstance + +instance ring_Noetherian_of_spec_Noetherian + [TopologicalSpace.NoetherianSpace (PrimeSpectrum R)] : + IsNoetherianRing R where + noetherian := by sorry + +lemma ring_Noetherian_iff_spec_Noetherian : IsNoetherianRing R + โ†” TopologicalSpace.NoetherianSpace (PrimeSpectrum R) := by + constructor + intro RisNoetherian + -- how do I apply an instance to prove one direction? + + +-- Use TopologicalSpace.NoetherianSpace.exists_finset_irreducible : +-- Every closed subset of a noetherian space is a finite union +-- of irreducible closed subsets. + +-- Stacks Lemma 10.32.5: R Noetherian. I,J ideals. +-- If J โŠ‚ โˆšI, then J ^ n โŠ‚ I for some n. In particular, locally nilpotent +-- and nilpotent are the same for Noetherian rings +lemma containment_radical_power_containment : + IsNoetherianRing R โˆง J โ‰ค Ideal.radical I โ†’ โˆƒ n : โ„•, J ^ n โ‰ค I := by + rintro โŸจRisNoetherian, containmentโŸฉ + rw [isNoetherianRing_iff_ideal_fg] at RisNoetherian + specialize RisNoetherian (Ideal.radical I) + -- rcases RisNoetherian with โŸจS, SgeneratesโŸฉ + have containment2 : โˆƒ n : โ„•, (Ideal.radical I) ^ n โ‰ค I := by + apply Ideal.exists_radical_pow_le_of_fg I RisNoetherian + cases' containment2 with n containment2' + have containment3 : J ^ n โ‰ค (Ideal.radical I) ^ n := by + apply Ideal.pow_mono containment + use n + apply le_trans containment3 containment2' +-- The above can be proven using the following quicker theorem that is in the wrong place. +-- Ideal.exists_pow_le_of_le_radical_of_fG + + +-- 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, +lemma length_eq_dim_if_maximal_annihilates {I : Ideal R} [Ideal.IsMaximal I] + : I โ€ข (โŠค : Submodule R M) = 0 + โ†’ Module.length R M = Module.rank RโงธI Mโงธ(I โ€ข (โŠค : Submodule R M)) := by sorry + +-- Does lean know that M/IM is a R/I module? + +-- Stacks Lemma 10.52.8: I is a finitely generated maximal ideal of R. +-- M is a finite R-mod and I^nM=0. Then length of M is finite. +lemma power_zero_finite_length : Ideal.FG I โ†’ Ideal.IsMaximal I โ†’ Module.Finite R M + โ†’ (โˆƒ n : โ„•, (I ^ n) โ€ข (โŠค : Submodule R M) = 0) + โ†’ (โˆƒ m : โ„•, Module.length R M โ‰ค m) := by + intro IisFG IisMaximal MisFinite power + rcases power with โŸจn, npowerโŸฉ + -- how do I get a generating set? + + +-- Stacks Lemma 10.53.3: R is Artinian iff R has finitely many maximal ideals +lemma IsArtinian_iff_finite_max_ideal : + IsArtinianRing R โ†” Finite (MaximalSpectrum R) := by sorry + +-- Stacks Lemma 10.53.4: R Artinian => Jacobson ideal of R is nilpotent +lemma Jacobson_of_Artinian_is_nilpotent : + IsArtinianRing R โ†’ IsNilpotent (Ideal.jacobson (โŠค : Ideal R)) := by sorry + +-- 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 +-- its maximal ideals. Also, all primes are maximal + +-- lemma product_of_localization_at_maximal_ideal : Finite (MaximalSpectrum R) +-- โˆง + +def jaydensRing : Type _ := sorry + -- โˆ€ I : MaximalSpectrum R, Localization.AtPrime R I + +instance : CommRing jaydensRing := sorry -- this should come for free, don't even need to state it + +def foo : jaydensRing โ‰ƒ+* R where + toFun := _ + invFun := _ + left_inv := _ + right_inv := _ + map_mul' := _ + map_add' := _ + -- Ideal.IsLocallyNilpotent (Ideal.jacobson (โŠค : Ideal R)) โ†’ + -- Pi.commRing (MaximalSpectrum R) Localization.AtPrime R I + -- := by sorry +-- Haven't finished this. + +-- Stacks Lemma 10.53.6: R is Artinian iff R has finite length as an R-mod +lemma IsArtinian_iff_finite_length : + IsArtinianRing R โ†” (โˆƒ n : โ„•, Module.length R R โ‰ค n) := by sorry + +-- Lemma: if R has finite length as R-mod, then R is Noetherian +lemma finite_length_is_Noetherian : + (โˆƒ n : โ„•, Module.length R R โ‰ค n) โ†’ IsNoetherianRing R := by sorry + +-- Lemma: if R is Artinian then all the prime ideals are maximal +lemma primes_of_Artinian_are_maximal : + IsArtinianRing R โ†’ Ideal.IsPrime I โ†’ Ideal.IsMaximal I := by sorry + +-- 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 dim_zero_Noetherian_iff_Artinian (R : Type _) [CommRing R] : + IsNoetherianRing R โˆง Ideal.krullDim R = 0 โ†” IsArtinianRing R := by + constructor + sorry + intro RisArtinian + constructor + apply finite_length_is_Noetherian + rwa [IsArtinian_iff_finite_length] at RisArtinian + sorry + + + + + + + + + + --- how to use namespace -namespace something -end something -open something diff --git a/CommAlg/krull.lean b/CommAlg/krull.lean index 78e2a36..06e6d0d 100644 --- a/CommAlg/krull.lean +++ b/CommAlg/krull.lean @@ -25,11 +25,11 @@ variable {R : Type _} [CommRing R] (I : PrimeSpectrum R) noncomputable def height : โ„•โˆž := Set.chainHeight {J : PrimeSpectrum R | J < I} -noncomputable def krullDim (R : Type) [CommRing R] : WithBot โ„•โˆž := โจ† (I : PrimeSpectrum R), height I +noncomputable def krullDim (R : Type _) [CommRing R] : WithBot โ„•โˆž := โจ† (I : PrimeSpectrum R), height I lemma height_def : height I = Set.chainHeight {J : PrimeSpectrum R | J < I} := rfl -lemma krullDim_def (R : Type) [CommRing R] : krullDim R = (โจ† (I : PrimeSpectrum R), height I : WithBot โ„•โˆž) := rfl -lemma krullDim_def' (R : Type) [CommRing R] : krullDim R = iSup (ฮป I : PrimeSpectrum R => (height I : WithBot โ„•โˆž)) := rfl +lemma krullDim_def (R : Type _) [CommRing R] : krullDim R = (โจ† (I : PrimeSpectrum R), height I : WithBot โ„•โˆž) := rfl +lemma krullDim_def' (R : Type _) [CommRing R] : krullDim R = iSup (ฮป I : PrimeSpectrum R => (height I : WithBot โ„•โˆž)) := rfl noncomputable instance : CompleteLattice (WithBot (โ„•โˆž)) := WithBot.WithTop.completeLattice @@ -42,13 +42,13 @@ lemma height_le_of_le {I J : PrimeSpectrum R} (I_le_J : I โ‰ค J) : height I โ‰ค lemma krullDim_le_iff (R : Type _) [CommRing R] (n : โ„•) : krullDim R โ‰ค n โ†” โˆ€ I : PrimeSpectrum R, (height I : WithBot โ„•โˆž) โ‰ค โ†‘n := iSup_le_iff (ฮฑ := WithBot โ„•โˆž) -lemma krullDim_le_iff' (R : Type) [CommRing R] (n : โ„•โˆž) : +lemma krullDim_le_iff' (R : Type _) [CommRing R] (n : โ„•โˆž) : krullDim R โ‰ค n โ†” โˆ€ I : PrimeSpectrum R, (height I : WithBot โ„•โˆž) โ‰ค โ†‘n := iSup_le_iff (ฮฑ := WithBot โ„•โˆž) -lemma le_krullDim_iff (R : Type) [CommRing R] (n : โ„•) : +lemma le_krullDim_iff (R : Type _) [CommRing R] (n : โ„•) : n โ‰ค krullDim R โ†” โˆƒ I : PrimeSpectrum R, n โ‰ค (height I : WithBot โ„•โˆž) := by sorry -lemma le_krullDim_iff' (R : Type) [CommRing R] (n : โ„•โˆž) : +lemma le_krullDim_iff' (R : Type _) [CommRing R] (n : โ„•โˆž) : n โ‰ค krullDim R โ†” โˆƒ I : PrimeSpectrum R, n โ‰ค (height I : WithBot โ„•โˆž) := by sorry @[simp] @@ -94,10 +94,16 @@ lemma dim_eq_bot_iff : krullDim R = โŠฅ โ†” Subsingleton R := by . rw [h.forall_iff] trivial -lemma krullDim_nonneg_of_nontrivial [Nontrivial R] : โˆƒ n : โ„•โˆž, Ideal.krullDim R = n := by +lemma krullDim_nonneg_of_nontrivial (R : Type _) [CommRing R] [Nontrivial R] : โˆƒ n : โ„•โˆž, Ideal.krullDim R = n := by have h := dim_eq_bot_iff.not.mpr (not_subsingleton R) lift (Ideal.krullDim R) to โ„•โˆž using h with k use k + +lemma dim_eq_zero_iff [Nontrivial R] : krullDim R = 0 โ†” โˆ€ I : PrimeSpectrum R, IsMaximal I.asIdeal := by + constructor <;> intro h + . intro I + sorry + . sorry @[simp] lemma field_prime_bot {K: Type _} [Field K] (P : Ideal K) : IsPrime P โ†” P = โŠฅ := by @@ -130,8 +136,6 @@ lemma dim_field_eq_zero {K : Type _} [Field K] : krullDim K = 0 := by simp [field_prime_height_zero] lemma isField.dim_zero {D: Type _} [CommRing D] [IsDomain D] (h: krullDim D = 0) : IsField D := by - unfold krullDim at h - simp [height] at h by_contra x rw [Ring.not_isField_iff_exists_prime] at x obtain โŸจP, โŸจh1, primePโŸฉโŸฉ := x @@ -140,16 +144,16 @@ lemma isField.dim_zero {D: Type _} [CommRing D] [IsDomain D] (h: krullDim D = 0) by_contra a have : P = โŠฅ := by rwa [PrimeSpectrum.ext_iff] at a contradiction - have PgtBot : P' > โŠฅ := Ne.bot_lt h2 - have pos_height : ยฌ โ†‘(Set.chainHeight {J | J < P'}) โ‰ค 0 := by - have : โŠฅ โˆˆ {J | J < P'} := PgtBot + have pos_height : ยฌ (height P') โ‰ค 0 := by + have : โŠฅ โˆˆ {J | J < P'} := Ne.bot_lt h2 have : {J | J < P'}.Nonempty := Set.nonempty_of_mem this + unfold height rw [โ†Set.one_le_chainHeight_iff] at this exact not_le_of_gt (Iff.mp ENat.one_le_iff_pos this) - have nonpos_height : (Set.chainHeight {J | J < P'}) โ‰ค 0 := by - have : (โจ† (I : PrimeSpectrum D), (Set.chainHeight {J | J < I} : WithBot โ„•โˆž)) โ‰ค 0 := h.le - rw [iSup_le_iff] at this - exact Iff.mp WithBot.coe_le_zero (this P') + have nonpos_height : height P' โ‰ค 0 := by + have := height_le_krullDim P' + rw [h] at this + aesop contradiction lemma dim_eq_zero_iff_field {D: Type _} [CommRing D] [IsDomain D] : krullDim D = 0 โ†” IsField D := by @@ -160,17 +164,67 @@ lemma dim_eq_zero_iff_field {D: Type _} [CommRing D] [IsDomain D] : krullDim D = exact dim_field_eq_zero #check Ring.DimensionLEOne +-- 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 + rw [krullDim_le_iff R 1] + intro H p + apply le_of_not_gt + intro h + rcases (lt_height_iff''.mp h) with โŸจc, โŸจhc1, hc2, hc3โŸฉโŸฉ + norm_cast at hc3 + rw [List.chain'_iff_get] at hc1 + specialize hc1 0 (by rw [hc3]; simp) + set q0 : PrimeSpectrum R := List.get c { val := 0, isLt := _ } + set q1 : PrimeSpectrum R := List.get c { val := 1, isLt := _ } + specialize hc2 q1 (List.get_mem _ _ _) + change q0.asIdeal < q1.asIdeal at hc1 + have q1nbot := Trans.trans (bot_le : โŠฅ โ‰ค q0.asIdeal) hc1 + specialize H q1.asIdeal (bot_lt_iff_ne_bot.mp q1nbot) q1.IsPrime + apply IsPrime.ne_top p.IsPrime + apply (IsCoatom.lt_iff H.out).mp + exact hc2 + lemma dim_le_one_of_pid [IsDomain R] [IsPrincipalIdealRing R] : krullDim R โ‰ค 1 := by rw [dim_le_one_iff] exact Ring.DimensionLEOne.principal_ideal_ring R lemma dim_le_dim_polynomial_add_one [Nontrivial R] : - krullDim R โ‰ค krullDim (Polynomial R) + 1 := sorry + krullDim R + 1 โ‰ค krullDim (Polynomial R) := sorry lemma dim_eq_dim_polynomial_add_one [Nontrivial R] [IsNoetherianRing R] : - krullDim R = krullDim (Polynomial R) + 1 := sorry + krullDim R + 1 = krullDim (Polynomial R) := sorry lemma height_eq_dim_localization : height I = krullDim (Localization.AtPrime I.asIdeal) := sorry diff --git a/CommAlg/monalisa.lean b/CommAlg/monalisa.lean new file mode 100644 index 0000000..d3a272e --- /dev/null +++ b/CommAlg/monalisa.lean @@ -0,0 +1,140 @@ +import Mathlib.Order.KrullDimension +import Mathlib.Order.JordanHolder +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.Order.Height +import Mathlib.RingTheory.Ideal.Basic +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.LinearAlgebra.Finsupp +import Mathlib.RingTheory.GradedAlgebra.Basic +import Mathlib.RingTheory.GradedAlgebra.HomogeneousIdeal +import Mathlib.Algebra.Module.GradedModule +import Mathlib.RingTheory.Ideal.AssociatedPrime +import Mathlib.RingTheory.Noetherian +import Mathlib.RingTheory.Artinian +import Mathlib.Algebra.Module.GradedModule +import Mathlib.RingTheory.Noetherian +import Mathlib.RingTheory.Finiteness +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.RingTheory.FiniteType +import Mathlib.Order.Height +import Mathlib.RingTheory.PrincipalIdealDomain +import Mathlib.RingTheory.DedekindDomain.Basic +import Mathlib.RingTheory.Ideal.Quotient +import Mathlib.RingTheory.Localization.AtPrime +import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Algebra.DirectSum.Ring +import Mathlib.RingTheory.Ideal.LocalRing +import Mathlib +import Mathlib.Algebra.MonoidAlgebra.Basic +import Mathlib.Data.Finset.Sort +import Mathlib.Order.Height +import Mathlib.Order.KrullDimension +import Mathlib.Order.JordanHolder +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.Order.Height +import Mathlib.RingTheory.Ideal.Basic +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.LinearAlgebra.Finsupp +import Mathlib.RingTheory.GradedAlgebra.Basic +import Mathlib.RingTheory.GradedAlgebra.HomogeneousIdeal +import Mathlib.Algebra.Module.GradedModule +import Mathlib.RingTheory.Ideal.AssociatedPrime +import Mathlib.RingTheory.Noetherian +import Mathlib.RingTheory.Artinian +import Mathlib.Algebra.Module.GradedModule +import Mathlib.RingTheory.Noetherian +import Mathlib.RingTheory.Finiteness +import Mathlib.RingTheory.Ideal.Operations + + + + +noncomputable def length ( A : Type _) (M : Type _) + [CommRing A] [AddCommGroup M] [Module A M] := Set.chainHeight {M' : Submodule A M | M' < โŠค} + + +def HomogeneousPrime { A ฯƒ : Type _} [CommRing A] [SetLike ฯƒ A] [AddSubmonoidClass ฯƒ A] (๐’œ : โ„ค โ†’ ฯƒ) [GradedRing ๐’œ] (I : Ideal A):= (Ideal.IsPrime I) โˆง (Ideal.IsHomogeneous ๐’œ I) + + +def HomogeneousMax { A ฯƒ : Type _} [CommRing A] [SetLike ฯƒ A] [AddSubmonoidClass ฯƒ A] (๐’œ : โ„ค โ†’ ฯƒ) [GradedRing ๐’œ] (I : Ideal A):= (Ideal.IsMaximal I) โˆง (Ideal.IsHomogeneous ๐’œ I) + +--theorem monotone_stabilizes_iff_noetherian : +-- (โˆ€ f : โ„• โ†’o Submodule R M, โˆƒ n, โˆ€ m, n โ‰ค m โ†’ f n = f m) โ†” IsNoetherian R M := by +-- rw [isNoetherian_iff_wellFounded, WellFounded.monotone_chain_condition] + +open GradedMonoid.GSmul + +open DirectSum + +instance tada1 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค ) : SMul (๐’œ 0) (๐“œ i) + where smul x y := @Eq.rec โ„ค (0+i) (fun a _ => ๐“œ a) (GradedMonoid.GSmul.smul x y) i (zero_add i) + +lemma mylem (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [h : DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค) (a : ๐’œ 0) (m : ๐“œ i) : + of _ _ (a โ€ข m) = of _ _ a โ€ข of _ _ m := by + refine' Eq.trans _ (Gmodule.of_smul_of ๐’œ ๐“œ a m).symm + refine' of_eq_of_gradedMonoid_eq _ + exact Sigma.ext (zero_add _).symm <| eq_rec_heq _ _ + +instance tada2 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [h : DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค ) : SMulWithZero (๐’œ 0) (๐“œ i) := by + letI := SMulWithZero.compHom (โจ i, ๐“œ i) (of ๐’œ 0).toZeroHom + exact Function.Injective.smulWithZero (of ๐“œ i).toZeroHom Dfinsupp.single_injective (mylem ๐’œ ๐“œ i) + +instance tada3 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] + [h : DirectSum.Gmodule ๐’œ ๐“œ] (i : โ„ค ): Module (๐’œ 0) (๐“œ i) := by + letI := Module.compHom (โจ j, ๐“œ j) (ofZeroRingHom ๐’œ) + exact Dfinsupp.single_injective.module (๐’œ 0) (of ๐“œ i) (mylem ๐’œ ๐“œ i) + + -- (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) + +noncomputable def dummyhil_function (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] + [DirectSum.GCommRing ๐’œ] + [DirectSum.Gmodule ๐’œ ๐“œ] (hilb : โ„ค โ†’ โ„•โˆž) := โˆ€ i, hilb i = (length (๐’œ 0) (๐“œ i)) + + +lemma hilbertz (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] + [DirectSum.GCommRing ๐’œ] + [DirectSum.Gmodule ๐’œ ๐“œ] + (finlen : โˆ€ i, (length (๐’œ 0) (๐“œ i)) < โŠค ) : โ„ค โ†’ โ„ค := by + intro i + let h := dummyhil_function ๐’œ ๐“œ + simp at h + let n : โ„ค โ†’ โ„• := fun i โ†ฆ WithTop.untop _ (finlen i).ne + have hn : โˆ€ i, (n i : โ„•โˆž) = length (๐’œ 0) (๐“œ i) := fun i โ†ฆ WithTop.coe_untop _ _ + have' := hn i + exact ((n i) : โ„ค ) + + +noncomputable def hilbert_function (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] + [DirectSum.GCommRing ๐’œ] + [DirectSum.Gmodule ๐’œ ๐“œ] (hilb : โ„ค โ†’ โ„ค) := โˆ€ i, hilb i = (ENat.toNat (length (๐’œ 0) (๐“œ i))) + + + +noncomputable def dimensionring { A: Type _} + [CommRing A] := krullDim (PrimeSpectrum A) + + +noncomputable def dimensionmodule ( A : Type _) (M : Type _) + [CommRing A] [AddCommGroup M] [Module A M] := krullDim (PrimeSpectrum (A โงธ ((โŠค : Submodule A M).annihilator)) ) + +-- lemma graded_local (๐’œ : โ„ค โ†’ Type _) [SetLike (โจ i, ๐’œ i)] (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +-- [DirectSum.GCommRing ๐’œ] +-- [DirectSum.Gmodule ๐’œ ๐“œ] (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) : โˆƒ ( I : Ideal ((โจ i, ๐’œ i))),(HomogeneousMax ๐’œ I) := sorry + + +def PolyType (f : โ„ค โ†’ โ„ค) (d : โ„•) := โˆƒ Poly : Polynomial โ„š, โˆƒ (N : โ„ค), โˆ€ (n : โ„ค), N โ‰ค n โ†’ f n = Polynomial.eval (n : โ„š) Poly โˆง d = Polynomial.degree Poly + + + +theorem hilbert_polynomial (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +[DirectSum.GCommRing ๐’œ] +[DirectSum.Gmodule ๐’œ ๐“œ] (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) (fingen : IsNoetherian (โจ i, ๐’œ i) (โจ i, ๐“œ i)) +(findim : โˆƒ d : โ„• , dimensionmodule (โจ i, ๐’œ i) (โจ i, ๐“œ i) = d):True := sorry + +-- Semiring A] + +-- variable [SetLike ฯƒ A] \ No newline at end of file diff --git a/CommAlg/sayantan(dim_eq_dim_polynomial_add_one).lean b/CommAlg/sayantan(dim_eq_dim_polynomial_add_one).lean new file mode 100644 index 0000000..ea6f7cd --- /dev/null +++ b/CommAlg/sayantan(dim_eq_dim_polynomial_add_one).lean @@ -0,0 +1,52 @@ +import Mathlib.RingTheory.Ideal.Basic +import Mathlib.Order.Height +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 + +namespace Ideal + +variable {R : Type _} [CommRing R] (I : PrimeSpectrum R) +noncomputable def height : โ„•โˆž := Set.chainHeight {J : PrimeSpectrum R | J < I} +noncomputable def krullDim (R : Type) [CommRing R] : WithBot โ„•โˆž := โจ† (I : PrimeSpectrum R), height I + +lemma height_def : height I = Set.chainHeight {J : PrimeSpectrum R | J < I} := rfl +lemma krullDim_def (R : Type) [CommRing R] : krullDim R = (โจ† (I : PrimeSpectrum R), height I : WithBot โ„•โˆž) := rfl +lemma krullDim_def' (R : Type) [CommRing R] : krullDim R = iSup (ฮป I : PrimeSpectrum R => (height I : WithBot โ„•โˆž)) := rfl + +noncomputable instance : CompleteLattice (WithBot (โ„•โˆž)) := WithBot.WithTop.completeLattice + +lemma dim_le_dim_polynomial_add_one [Nontrivial R] : + krullDim R + 1 โ‰ค krullDim (Polynomial R) := sorry -- Others are working on it + +-- private lemma sum_succ_of_succ_sum {ฮน : Type} (a : โ„•โˆž) [inst : Nonempty ฮน] : +-- (โจ† (x : ฮน), a + 1) = (โจ† (x : ฮน), a) + 1 := by +-- have : a + 1 = (โจ† (x : ฮน), a) + 1 := by rw [ciSup_const] +-- have : a + 1 = (โจ† (x : ฮน), a + 1) := Eq.symm ciSup_const +-- simp + +lemma dim_eq_dim_polynomial_add_one [Nontrivial R] [IsNoetherianRing R] : + krullDim R + 1 = krullDim (Polynomial R) := by + rw [le_antisymm_iff] + constructor + ยท exact dim_le_dim_polynomial_add_one + ยท unfold krullDim + have htPBdd : โˆ€ (P : PrimeSpectrum (Polynomial R)), (height P : WithBot โ„•โˆž) โ‰ค (โจ† (I : PrimeSpectrum R), โ†‘(height I + 1)) := by + intro P + have : โˆƒ (I : PrimeSpectrum R), (height P : WithBot โ„•โˆž) โ‰ค โ†‘(height I + 1) := by + sorry + obtain โŸจI, IPโŸฉ := this + have : (โ†‘(height I + 1) : WithBot โ„•โˆž) โ‰ค โจ† (I : PrimeSpectrum R), โ†‘(height I + 1) := by + apply @le_iSup (WithBot โ„•โˆž) _ _ _ I + apply ge_trans this IP + have oneOut : (โจ† (I : PrimeSpectrum R), (height I : WithBot โ„•โˆž) + 1) โ‰ค (โจ† (I : PrimeSpectrum R), โ†‘(height I)) + 1 := by + have : โˆ€ P : PrimeSpectrum R, (height P : WithBot โ„•โˆž) + 1 โ‰ค (โจ† (I : PrimeSpectrum R), โ†‘(height I)) + 1 := + fun P โ†ฆ (by apply add_le_add_right (@le_iSup (WithBot โ„•โˆž) _ _ _ P) 1) + apply iSup_le + apply this + simp + intro P + exact ge_trans oneOut (htPBdd P) diff --git a/CommAlg/sayantan(dim_eq_zero_iff_field).lean b/CommAlg/sayantan(dim_eq_zero_iff_field).lean deleted file mode 100644 index fd712f0..0000000 --- a/CommAlg/sayantan(dim_eq_zero_iff_field).lean +++ /dev/null @@ -1,82 +0,0 @@ -import Mathlib.RingTheory.Ideal.Basic -import Mathlib.Order.Height -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 - -namespace Ideal - -variable {R : Type _} [CommRing R] (I : PrimeSpectrum R) -noncomputable def height : โ„•โˆž := Set.chainHeight {J : PrimeSpectrum R | J < I} -noncomputable def krullDim (R : Type) [CommRing R] : WithBot โ„•โˆž := โจ† (I : PrimeSpectrum R), height I - -lemma height_def : height I = Set.chainHeight {J : PrimeSpectrum R | J < I} := rfl -lemma krullDim_def (R : Type) [CommRing R] : krullDim R = (โจ† (I : PrimeSpectrum R), height I : WithBot โ„•โˆž) := rfl -lemma krullDim_def' (R : Type) [CommRing R] : krullDim R = iSup (ฮป I : PrimeSpectrum R => (height I : WithBot โ„•โˆž)) := rfl - -@[simp] -lemma field_prime_bot {K: Type _} [Field K] (P : Ideal K) : IsPrime P โ†” P = โŠฅ := by - constructor - ยท intro primeP - obtain T := eq_bot_or_top P - have : ยฌP = โŠค := IsPrime.ne_top primeP - tauto - ยท intro botP - rw [botP] - exact bot_prime - -lemma field_prime_height_zero {K: Type _} [Field K] (P : PrimeSpectrum K) : height P = 0 := by - unfold height - simp - by_contra spec - change _ โ‰  _ at spec - rw [โ† Set.nonempty_iff_ne_empty] at spec - obtain โŸจJ, JlP : J < PโŸฉ := spec - have P0 : IsPrime P.asIdeal := P.IsPrime - have J0 : IsPrime J.asIdeal := J.IsPrime - rw [field_prime_bot] at P0 J0 - have : J.asIdeal = P.asIdeal := Eq.trans J0 (Eq.symm P0) - have : J = P := PrimeSpectrum.ext J P this - have : J โ‰  P := ne_of_lt JlP - contradiction - -lemma dim_field_eq_zero {K : Type _} [Field K] : krullDim K = 0 := by - unfold krullDim - simp [field_prime_height_zero] - -noncomputable -instance : CompleteLattice (WithBot โ„•โˆž) := - inferInstanceAs <| CompleteLattice (WithBot (WithTop โ„•)) - -lemma isField.dim_zero {D: Type _} [CommRing D] [IsDomain D] (h: krullDim D = 0) : IsField D := by - unfold krullDim at h - simp [height] at h - by_contra x - rw [Ring.not_isField_iff_exists_prime] at x - obtain โŸจP, โŸจh1, primePโŸฉโŸฉ := x - let P' : PrimeSpectrum D := PrimeSpectrum.mk P primeP - have h2 : P' โ‰  โŠฅ := by - by_contra a - have : P = โŠฅ := by rwa [PrimeSpectrum.ext_iff] at a - contradiction - have PgtBot : P' > โŠฅ := Ne.bot_lt h2 - have pos_height : ยฌ โ†‘(Set.chainHeight {J | J < P'}) โ‰ค 0 := by - have : โŠฅ โˆˆ {J | J < P'} := PgtBot - have : {J | J < P'}.Nonempty := Set.nonempty_of_mem this - rw [โ†Set.one_le_chainHeight_iff] at this - exact not_le_of_gt (Iff.mp ENat.one_le_iff_pos this) - have nonpos_height : (Set.chainHeight {J | J < P'}) โ‰ค 0 := by - have : (โจ† (I : PrimeSpectrum D), (Set.chainHeight {J | J < I} : WithBot โ„•โˆž)) โ‰ค 0 := h.le - rw [iSup_le_iff] at this - exact Iff.mp WithBot.coe_le_zero (this P') - contradiction - -lemma dim_eq_zero_iff_field {D: Type _} [CommRing D] [IsDomain D] : krullDim D = 0 โ†” IsField D := by - constructor - ยท exact isField.dim_zero - ยท intro fieldD - let h : Field D := IsField.toField fieldD - exact dim_field_eq_zero