mirror of
https://github.com/GTBarkley/comm_alg.git
synced 2024-12-26 07:38:36 -06:00
Added standard graded assumption
This commit is contained in:
parent
5fa7d286cd
commit
fe64034b11
2 changed files with 5 additions and 4 deletions
|
@ -6,6 +6,7 @@ import Mathlib.RingTheory.Artinian
|
||||||
import Mathlib.Order.Height
|
import Mathlib.Order.Height
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
-- Setting for "library_search"
|
-- Setting for "library_search"
|
||||||
set_option maxHeartbeats 0
|
set_option maxHeartbeats 0
|
||||||
macro "ls" : tactic => `(tactic|library_search)
|
macro "ls" : tactic => `(tactic|library_search)
|
||||||
|
@ -43,7 +44,7 @@ noncomputable def length ( A : Type _) (M : Type _)
|
||||||
[CommRing A] [AddCommGroup M] [Module A M] := Set.chainHeight {M' : Submodule A M | M' < ⊤}
|
[CommRing A] [AddCommGroup M] [Module A M] := Set.chainHeight {M' : Submodule A M | M' < ⊤}
|
||||||
|
|
||||||
-- Make instance of M_i being an R_0-module
|
-- Make instance of M_i being an R_0-module
|
||||||
instance tada1 (𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)] [DirectSum.GCommRing 𝒜]
|
instance tada1 (𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)] [DirectSum.GCommRing 𝒜]
|
||||||
[DirectSum.Gmodule 𝒜 𝓜] (i : ℤ ) : SMul (𝒜 0) (𝓜 i)
|
[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)
|
where smul x y := @Eq.rec ℤ (0+i) (fun a _ => 𝓜 a) (GradedMonoid.GSmul.smul x y) i (zero_add i)
|
||||||
|
|
||||||
|
|
|
@ -203,7 +203,7 @@ theorem Hilbert_polynomial_d_ge_1_reduced
|
||||||
(d : ℕ) (d1 : 1 ≤ d)
|
(d : ℕ) (d1 : 1 ≤ d)
|
||||||
(𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)]
|
(𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)]
|
||||||
[DirectSum.GCommRing 𝒜]
|
[DirectSum.GCommRing 𝒜]
|
||||||
[DirectSum.Gmodule 𝒜 𝓜] (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0))
|
[DirectSum.Gmodule 𝒜 𝓜] (st: StandardGraded 𝒜) (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0))
|
||||||
(fingen : IsNoetherian (⨁ i, 𝒜 i) (⨁ i, 𝓜 i))
|
(fingen : IsNoetherian (⨁ i, 𝒜 i) (⨁ i, 𝓜 i))
|
||||||
(findim : dimensionmodule (⨁ i, 𝒜 i) (⨁ i, 𝓜 i) = d)
|
(findim : dimensionmodule (⨁ i, 𝒜 i) (⨁ i, 𝓜 i) = d)
|
||||||
(hilb : ℤ → ℤ) (Hhilb: hilbert_function 𝒜 𝓜 hilb)
|
(hilb : ℤ → ℤ) (Hhilb: hilbert_function 𝒜 𝓜 hilb)
|
||||||
|
@ -217,7 +217,7 @@ theorem Hilbert_polynomial_d_ge_1_reduced
|
||||||
-- If M is a finite graed R-Mod of dimension zero, then the Hilbert function H(M, n) = 0 for n >> 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_d_0 (𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)]
|
theorem Hilbert_polynomial_d_0 (𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)]
|
||||||
[DirectSum.GCommRing 𝒜]
|
[DirectSum.GCommRing 𝒜]
|
||||||
[DirectSum.Gmodule 𝒜 𝓜] (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0))
|
[DirectSum.Gmodule 𝒜 𝓜] (st: StandardGraded 𝒜) (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0))
|
||||||
(fingen : IsNoetherian (⨁ i, 𝒜 i) (⨁ i, 𝓜 i))
|
(fingen : IsNoetherian (⨁ i, 𝒜 i) (⨁ i, 𝓜 i))
|
||||||
(findim : dimensionmodule (⨁ i, 𝒜 i) (⨁ i, 𝓜 i) = 0)
|
(findim : dimensionmodule (⨁ i, 𝒜 i) (⨁ i, 𝓜 i) = 0)
|
||||||
(hilb : ℤ → ℤ) (Hhilb : hilbert_function 𝒜 𝓜 hilb)
|
(hilb : ℤ → ℤ) (Hhilb : hilbert_function 𝒜 𝓜 hilb)
|
||||||
|
@ -230,7 +230,7 @@ theorem Hilbert_polynomial_d_0 (𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [
|
||||||
theorem Hilbert_polynomial_d_0_reduced
|
theorem Hilbert_polynomial_d_0_reduced
|
||||||
(𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)]
|
(𝒜 : ℤ → Type _) (𝓜 : ℤ → Type _) [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (𝓜 i)]
|
||||||
[DirectSum.GCommRing 𝒜]
|
[DirectSum.GCommRing 𝒜]
|
||||||
[DirectSum.Gmodule 𝒜 𝓜] (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0))
|
[DirectSum.Gmodule 𝒜 𝓜] (st: StandardGraded 𝒜) (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0))
|
||||||
(fingen : IsNoetherian (⨁ i, 𝒜 i) (⨁ i, 𝓜 i))
|
(fingen : IsNoetherian (⨁ i, 𝒜 i) (⨁ i, 𝓜 i))
|
||||||
(findim : dimensionmodule (⨁ i, 𝒜 i) (⨁ i, 𝓜 i) = 0)
|
(findim : dimensionmodule (⨁ i, 𝒜 i) (⨁ i, 𝓜 i) = 0)
|
||||||
(hilb : ℤ → ℤ) (Hhilb : hilbert_function 𝒜 𝓜 hilb)
|
(hilb : ℤ → ℤ) (Hhilb : hilbert_function 𝒜 𝓜 hilb)
|
||||||
|
|
Loading…
Reference in a new issue