Added comments to everything

This commit is contained in:
chelseaandmadrid 2023-06-14 14:44:43 -07:00
parent 08b1fd3e7a
commit 900bf12da2
2 changed files with 17 additions and 43 deletions

View file

@ -29,8 +29,6 @@ macro "obviously" : tactic =>
-- @Definitions (to be classified) -- @Definitions (to be classified)
section section
open GradedMonoid.GSmul open GradedMonoid.GSmul
@ -81,15 +79,20 @@ end
-- [DirectSum.Gmodule 𝒜 𝓜] (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0)) : ∃ ( I : Ideal ((⨁ i, 𝒜 i))),(HomogeneousMax 𝒜 I) := sorry -- [DirectSum.Gmodule 𝒜 𝓜] (art: IsArtinianRing (𝒜 0)) (loc : LocalRing (𝒜 0)) : ∃ ( I : Ideal ((⨁ i, 𝒜 i))),(HomogeneousMax 𝒜 I) := sorry
-- Definition(s) of homogeneous ideals -- Definition of homogeneous ideal
def Ideal.IsHomogeneous' (𝒜 : → Type _) def Ideal.IsHomogeneous' (𝒜 : → Type _)
[∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜]
(I : Ideal (⨁ i, 𝒜 i)) := ∀ (i : ) (I : Ideal (⨁ i, 𝒜 i)) := ∀ (i : )
⦃r : (⨁ i, 𝒜 i)⦄, r ∈ I → DirectSum.of _ i ( r i : 𝒜 i) ∈ I ⦃r : (⨁ i, 𝒜 i)⦄, r ∈ I → DirectSum.of _ i ( r i : 𝒜 i) ∈ I
-- Definition of homogeneous prime ideal
def HomogeneousPrime (𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] (I : Ideal (⨁ i, 𝒜 i)):= (Ideal.IsPrime I) ∧ (Ideal.IsHomogeneous' 𝒜 I) def HomogeneousPrime (𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] (I : Ideal (⨁ i, 𝒜 i)):= (Ideal.IsPrime I) ∧ (Ideal.IsHomogeneous' 𝒜 I)
-- Definition of homogeneous maximal ideal
def HomogeneousMax (𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] (I : Ideal (⨁ i, 𝒜 i)):= (Ideal.IsMaximal I) ∧ (Ideal.IsHomogeneous' 𝒜 I) def HomogeneousMax (𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] (I : Ideal (⨁ i, 𝒜 i)):= (Ideal.IsMaximal I) ∧ (Ideal.IsHomogeneous' 𝒜 I)
--theorem monotone_stabilizes_iff_noetherian : --theorem monotone_stabilizes_iff_noetherian :
-- (∀ f : →o Submodule R M, ∃ n, ∀ m, n ≤ m → f n = f m) ↔ IsNoetherian R M := by -- (∀ 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] -- rw [isNoetherian_iff_wellFounded, WellFounded.monotone_chain_condition]
@ -99,12 +102,6 @@ end
-- @[BH, 4.1.3] when d ≥ 1 -- @[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) -- 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)]
@ -165,14 +162,7 @@ lemma Associated_prime_of_graded_is_graded
-- sorry -- sorry
-- instance sdfasdf -- Each component of a graded ring is an additive subgroup
-- (𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜]
-- (p : Ideal (⨁ i, 𝒜 i)) (hp : Ideal.IsHomogeneous' 𝒜 p)
-- : ∀ i, AddCommGroup (p i) := by
-- sorry
def Component_of_graded_as_addsubgroup (𝒜 : → Type _) def Component_of_graded_as_addsubgroup (𝒜 : → Type _)
[∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜]
(p : Ideal (⨁ i, 𝒜 i)) (hp : Ideal.IsHomogeneous' 𝒜 p) (i : ) : AddSubgroup (𝒜 i) := by (p : Ideal (⨁ i, 𝒜 i)) (hp : Ideal.IsHomogeneous' 𝒜 p) (i : ) : AddSubgroup (𝒜 i) := by
@ -186,10 +176,12 @@ instance Quotient_of_graded_is_graded
: DirectSum.Gmodule 𝒜 (fun i => (𝒜 i)(Component_of_graded_as_addsubgroup 𝒜 p hp i)) := by : DirectSum.Gmodule 𝒜 (fun i => (𝒜 i)(Component_of_graded_as_addsubgroup 𝒜 p hp i)) := by
sorry sorry
-- @Graded submodule
instance Graded_submodule
(𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜] -- -- @Graded submodule
(p : Ideal (⨁ i, 𝒜 i)) (hp : Ideal.IsHomogeneous' 𝒜 p) -- instance Graded_submodule
: DirectSum.Gmodule 𝒜 (fun i => (𝒜 i)(Component_of_graded_as_addsubgroup 𝒜 p hp i)) := by -- (𝒜 : → Type _) [∀ i, AddCommGroup (𝒜 i)] [DirectSum.GCommRing 𝒜]
sorry -- (p : Ideal (⨁ i, 𝒜 i)) (hp : Ideal.IsHomogeneous' 𝒜 p)
-- : DirectSum.Gmodule 𝒜 (fun i => (𝒜 i)(Component_of_graded_as_addsubgroup 𝒜 p hp i)) := by
-- sorry

View file

@ -1,24 +1,5 @@
import Mathlib
import Mathlib.Algebra.MonoidAlgebra.Basic
import Mathlib.Data.Finset.Sort
import Mathlib.Order.Height import Mathlib.Order.Height
import Mathlib.Order.KrullDimension
import Mathlib.Order.JordanHolder
import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic 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
-- Setting for "library_search" -- Setting for "library_search"
set_option maxHeartbeats 0 set_option maxHeartbeats 0
@ -86,6 +67,7 @@ example : Polynomial.eval (100 : ) F = (2 : ) := by
end section end section
-- @[BH, 4.1.2] -- @[BH, 4.1.2]
-- All the polynomials are in [X], all the functions are considered as -- All the polynomials are in [X], all the functions are considered as
noncomputable section noncomputable section
@ -247,6 +229,7 @@ lemma b_to_a (f : ) (d : ) : PolyType f d → (∃ (c : ), ∃
end end
-- @Additive lemma of length for a SES -- @Additive lemma of length for a SES
-- Given a SES 0 → A → B → C → 0, then length (A) - length (B) + length (C) = 0
section section
-- variable {R M N : Type _} [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] -- variable {R M N : Type _} [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N]
-- (f : M →[R] N) -- (f : M →[R] N)
@ -305,4 +288,3 @@ end section