@@ -567,13 +567,51 @@ section monoid_algebra
567
567
568
568
namespace add_monoid_algebra
569
569
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
571
609
572
610
variables {R : Type *} {M : Type *} [add_comm_monoid M]
573
611
574
612
lemma mv_polynomial_aeval_of_surjective_of_closure [comm_semiring R] {S : set M}
575
613
(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) :=
577
615
begin
578
616
refine λ f, induction_on f (λ m, _) _ _,
579
617
{ 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
593
631
begin
594
632
obtain ⟨S, hS⟩ := h.out,
595
633
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)
597
635
end
598
636
599
637
end add_monoid_algebra
600
638
601
639
namespace monoid_algebra
602
640
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
604
680
605
681
variables {R : Type *} {M : Type *} [comm_monoid M]
606
682
@@ -623,7 +699,7 @@ begin
623
699
end
624
700
625
701
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
627
703
628
704
end monoid_algebra
629
705
0 commit comments