Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 3773525

Browse files
feat(ring_theory/finiteness): add lemmas (#7409)
I add here some preliminary lemmas to prove that a monoid is finitely generated iff the monoid algebra is.
1 parent 67dad97 commit 3773525

File tree

2 files changed

+96
-13
lines changed

2 files changed

+96
-13
lines changed

src/algebra/monoid_algebra.lean

Lines changed: 15 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -616,9 +616,8 @@ section span
616616
variables [semiring k] [mul_one_class G]
617617

618618
lemma mem_span_support (f : monoid_algebra k G) :
619-
f ∈ submodule.span k (monoid_algebra.of k G '' (f.support : set G)) :=
620-
by rw [monoid_algebra.of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single,
621-
finsupp.mem_supported]
619+
f ∈ submodule.span k (of k G '' (f.support : set G)) :=
620+
by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported]
622621

623622
end span
624623

@@ -879,11 +878,16 @@ def of [add_zero_class G] : multiplicative G →* add_monoid_algebra k G :=
879878
map_one' := rfl,
880879
map_mul' := λ a b, by { rw [single_mul_single, one_mul], refl } }
881880

881+
/-- Embedding of a monoid into its monoid algebra, having `G` as source. -/
882+
def of' : G → add_monoid_algebra k G := λ a, single a 1
883+
882884
end
883885

884886
@[simp] lemma of_apply [add_zero_class G] (a : multiplicative G) : of k G a = single a.to_add 1 :=
885887
rfl
886888

889+
@[simp] lemma of'_apply (a : G) : of' k G a = single a 1 := rfl
890+
887891
lemma of_injective [nontrivial k] [add_zero_class G] : function.injective (of k G) :=
888892
λ a b h, by simpa using (single_eq_single_iff _ _ _ _).mp h
889893

@@ -927,12 +931,15 @@ end misc_theorems
927931

928932
section span
929933

930-
variables [semiring k] [add_zero_class G]
934+
variables [semiring k]
935+
936+
lemma mem_span_support [add_zero_class G] (f : add_monoid_algebra k G) :
937+
f ∈ submodule.span k (of k G '' (f.support : set G)) :=
938+
by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported]
931939

932-
lemma mem_span_support (f : add_monoid_algebra k G) :
933-
f ∈ submodule.span k (add_monoid_algebra.of k G '' (f.support : set G)) :=
934-
by rw [add_monoid_algebra.of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single,
935-
finsupp.mem_supported]
940+
lemma mem_span_support' (f : add_monoid_algebra k G) :
941+
f ∈ submodule.span k (of' k G '' (f.support : set G)) :=
942+
by rw [of', ← finsupp.supported_eq_span_single, finsupp.mem_supported]
936943

937944
end span
938945

src/ring_theory/finiteness.lean

Lines changed: 81 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -567,13 +567,51 @@ section monoid_algebra
567567

568568
namespace add_monoid_algebra
569569

570-
open algebra add_submonoid
570+
open algebra add_submonoid submodule
571+
572+
section span
573+
574+
variables {R : Type*} {M : Type*} [comm_semiring R] [add_monoid M]
575+
576+
lemma mem_adjoint_support (f : add_monoid_algebra R M) :
577+
f ∈ adjoin R (of' R M '' (f.support : set M)) :=
578+
begin
579+
suffices : span R (of R M '' (f.support : set M)) ≤
580+
(adjoin R (of R M '' (f.support : set M))).to_submodule,
581+
{ exact this (mem_span_support f) },
582+
rw submodule.span_le,
583+
exact subset_adjoin
584+
end
585+
586+
lemma support_gen_of_gen {S : set (add_monoid_algebra R M)} (hS : algebra.adjoin R S = ⊤) :
587+
algebra.adjoin R (⋃ f ∈ S, (of' R M '' (f.support : set M))) = ⊤ :=
588+
begin
589+
refine le_antisymm le_top _,
590+
rw [← hS, adjoin_le_iff],
591+
intros f hf,
592+
have hincl : of' R M '' (f.support : set M) ⊆
593+
⋃ (g : add_monoid_algebra R M) (H : g ∈ S), of' R M '' (g.support : set M),
594+
{ intros s hs,
595+
exact set.mem_bUnion_iff.2 ⟨f, ⟨hf, hs⟩⟩ },
596+
exact adjoin_mono hincl (mem_adjoint_support f)
597+
end
598+
599+
lemma support_gen_of_gen' {S : set (add_monoid_algebra R M)} (hS : algebra.adjoin R S = ⊤) :
600+
algebra.adjoin R (of' R M '' (⋃ f ∈ S, (f.support : set M))) = ⊤ :=
601+
begin
602+
suffices : of' R M '' (⋃ f ∈ S, (f.support : set M)) = ⋃ f ∈ S, (of' R M '' (f.support : set M)),
603+
{ rw this,
604+
exact support_gen_of_gen hS },
605+
simp only [set.image_Union]
606+
end
607+
608+
end span
571609

572610
variables {R : Type*} {M : Type*} [add_comm_monoid M]
573611

574612
lemma mv_polynomial_aeval_of_surjective_of_closure [comm_semiring R] {S : set M}
575613
(hS : closure S = ⊤) : function.surjective (mv_polynomial.aeval
576-
(λ (s : S), of R M (multiplicative.of_add ↑s)) : mv_polynomial S R → add_monoid_algebra R M) :=
614+
(λ (s : S), of' R M ↑s) : mv_polynomial S R → add_monoid_algebra R M) :=
577615
begin
578616
refine λ f, induction_on f (λ m, _) _ _,
579617
{ have : m ∈ closure S := hS.symm ▸ mem_top _,
@@ -593,14 +631,52 @@ instance ft_of_fg [comm_ring R] [h : add_monoid.fg M] : finite_type R (add_monoi
593631
begin
594632
obtain ⟨S, hS⟩ := h.out,
595633
exact (finite_type.mv_polynomial R (S : set M)).of_surjective (mv_polynomial.aeval
596-
(λ (s : (S : set M)), of R M s.1)) (mv_polynomial_aeval_of_surjective_of_closure hS)
634+
(λ (s : (S : set M)), of' R M ↑s)) (mv_polynomial_aeval_of_surjective_of_closure hS)
597635
end
598636

599637
end add_monoid_algebra
600638

601639
namespace monoid_algebra
602640

603-
open algebra submonoid
641+
open algebra submonoid submodule
642+
643+
section span
644+
645+
variables {R : Type*} {M : Type*} [comm_semiring R] [monoid M]
646+
647+
lemma mem_adjoint_support (f : monoid_algebra R M) :
648+
f ∈ adjoin R (monoid_algebra.of R M '' (f.support : set M)) :=
649+
begin
650+
suffices : span R (of R M '' (f.support : set M)) ≤
651+
(adjoin R (of R M '' (f.support : set M))).to_submodule,
652+
{ exact this (mem_span_support f) },
653+
rw submodule.span_le,
654+
exact subset_adjoin
655+
end
656+
657+
lemma support_gen_of_gen {S : set (monoid_algebra R M)} (hS : algebra.adjoin R S = ⊤) :
658+
algebra.adjoin R (⋃ f ∈ S, ((of R M) '' (f.support : set M))) = ⊤ :=
659+
begin
660+
refine le_antisymm le_top _,
661+
rw [← hS, adjoin_le_iff],
662+
intros f hf,
663+
have hincl : (of R M) '' (f.support : set M) ⊆
664+
⋃ (g : monoid_algebra R M) (H : g ∈ S), of R M '' (g.support : set M),
665+
{ intros s hs,
666+
exact set.mem_bUnion_iff.2 ⟨f, ⟨hf, hs⟩⟩ },
667+
exact adjoin_mono hincl (mem_adjoint_support f)
668+
end
669+
670+
lemma support_gen_of_gen' {S : set (monoid_algebra R M)} (hS : algebra.adjoin R S = ⊤) :
671+
algebra.adjoin R (of R M '' (⋃ f ∈ S, (f.support : set M))) = ⊤ :=
672+
begin
673+
suffices : of R M '' (⋃ f ∈ S, (f.support : set M)) = ⋃ f ∈ S, (of R M '' (f.support : set M)),
674+
{ rw this,
675+
exact support_gen_of_gen hS },
676+
simp only [set.image_Union]
677+
end
678+
679+
end span
604680

605681
variables {R : Type*} {M : Type*} [comm_monoid M]
606682

@@ -623,7 +699,7 @@ begin
623699
end
624700

625701
instance ft_of_fg [comm_ring R] [monoid.fg M] : finite_type R (monoid_algebra R M) :=
626-
add_monoid_algebra.ft_of_fg.equiv (monoid_algebra.to_additive_alg_equiv R M).symm
702+
add_monoid_algebra.ft_of_fg.equiv (to_additive_alg_equiv R M).symm
627703

628704
end monoid_algebra
629705

0 commit comments

Comments
 (0)