diff --git a/HilbertFunction.lean b/HilbertFunction.lean index 21715f0..4ed1a2b 100644 --- a/HilbertFunction.lean +++ b/HilbertFunction.lean @@ -5,6 +5,7 @@ import Mathlib.RingTheory.Ideal.AssociatedPrime import Mathlib.RingTheory.Artinian import Mathlib.Order.Height + -- Setting for "library_search" set_option maxHeartbeats 0 macro "ls" : tactic => `(tactic|library_search) @@ -74,9 +75,8 @@ noncomputable def dimensionring { A: Type _} noncomputable def dimensionmodule ( A : Type _) (M : Type _) [CommRing A] [AddCommGroup M] [Module A M] := krullDim (PrimeSpectrum (A โงธ ((โŠค : Submodule A M).annihilator)) ) end --- 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 + + -- Definition of homogeneous ideal @@ -91,8 +91,6 @@ def HomogeneousPrime (๐’œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [Dir -- Definition of homogeneous maximal ideal def HomogeneousMax (๐’œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [DirectSum.GCommRing ๐’œ] (I : Ideal (โจ i, ๐’œ i)):= (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] @@ -101,10 +99,16 @@ def HomogeneousMax (๐’œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [Direc end +-- If A_0 is Artinian and local, then A is graded local +lemma Graded_local_if_zero_component_Artinian_and_local (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) +[โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +[DirectSum.GCommRing ๐’œ] [DirectSum.Gmodule ๐’œ ๐“œ] (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) : โˆƒ ( I : Ideal ((โจ i, ๐’œ i))),(HomogeneousMax ๐’œ I) := sorry + + -- @[BH, 4.1.3] when d โ‰ฅ 1 -- If M is a finite graed R-Mod of dimension d โ‰ฅ 1, then the Hilbert function H(M, n) is of polynomial type (d - 1) -theorem hilbert_polynomial_ge1 (d : โ„•) (d1 : 1 โ‰ค d) (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +theorem Hilbert_polynomial_ge1 (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)) @@ -118,7 +122,7 @@ theorem hilbert_polynomial_ge1 (d : โ„•) (d1 : 1 โ‰ค d) (๐’œ : โ„ค โ†’ Type _) -- @[BH, 4.1.3] when d = 0 -- If M is a finite graed R-Mod of dimension zero, then the Hilbert function H(M, n) = 0 for n >> 0 -theorem hilbert_polynomial_0 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] +theorem Hilbert_polynomial_0 (๐’œ : โ„ค โ†’ Type _) (๐“œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [โˆ€ i, AddCommGroup (๐“œ i)] [DirectSum.GCommRing ๐’œ] [DirectSum.Gmodule ๐’œ ๐“œ] (art: IsArtinianRing (๐’œ 0)) (loc : LocalRing (๐’œ 0)) (fingen : IsNoetherian (โจ i, ๐’œ i) (โจ i, ๐“œ i)) @@ -169,7 +173,7 @@ def Component_of_graded_as_addsubgroup (๐’œ : โ„ค โ†’ Type _) sorry --- @ Quotient of a graded ring R by a graded ideal p is a graded R-Mod, preserving each component +-- @Quotient of a graded ring R by a graded ideal p is a graded R-Mod, preserving each component instance Quotient_of_graded_is_graded (๐’œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [DirectSum.GCommRing ๐’œ] (p : Ideal (โจ i, ๐’œ i)) (hp : Ideal.IsHomogeneous' ๐’œ p) @@ -177,11 +181,3 @@ instance Quotient_of_graded_is_graded sorry - --- -- @Graded submodule --- instance Graded_submodule --- (๐’œ : โ„ค โ†’ Type _) [โˆ€ i, AddCommGroup (๐’œ i)] [DirectSum.GCommRing ๐’œ] --- (p : Ideal (โจ i, ๐’œ i)) (hp : Ideal.IsHomogeneous' ๐’œ p) --- : DirectSum.Gmodule ๐’œ (fun i => (๐’œ i)โงธ(Component_of_graded_as_addsubgroup ๐’œ p hp i)) := by --- sorry - diff --git a/PolyType.lean b/PolyType.lean index 07c1be7..e69eec2 100644 --- a/PolyType.lean +++ b/PolyType.lean @@ -122,6 +122,11 @@ lemma Poly_constant (F : Polynomial โ„š) (c : โ„š) : simp ยท sorry + +-- Shifting doesn't change the polynomial type +lemma Poly_shifting (f : โ„ค โ†’ โ„ค) (g : โ„ค โ†’ โ„ค) (hf : PolyType f d) (s : โ„ค) (hfg : โˆ€ (n : โ„ค), f (n + s) = g (n)) : PolyType g d := by + sorry + -- PolyType 0 = constant function lemma PolyType_0 (f : โ„ค โ†’ โ„ค) : (PolyType f 0) โ†” (โˆƒ (c : โ„ค), โˆƒ (N : โ„ค), โˆ€ (n : โ„ค), (N โ‰ค n โ†’ f n = c) โˆง c โ‰  0) := by constructor @@ -228,6 +233,8 @@ lemma b_to_a (f : โ„ค โ†’ โ„ค) (d : โ„•) : PolyType f d โ†’ (โˆƒ (c : โ„ค), โˆƒ exact this1 end + + -- @Additive lemma of length for a SES -- Given a SES 0 โ†’ A โ†’ B โ†’ C โ†’ 0, then length (A) - length (B) + length (C) = 0 section