Skip to content

Commit 7bec33e

Browse files
committed
WIP
1 parent a09a371 commit 7bec33e

File tree

1 file changed

+47
-47
lines changed

1 file changed

+47
-47
lines changed

src/monalg.v

Lines changed: 47 additions & 47 deletions
Original file line numberDiff line numberDiff line change
@@ -449,7 +449,10 @@ Qed.
449449
Lemma fgscaleA c1 c2 g : c1 *:g (c2 *:g g) = (c1 * c2) *:g g.
450450
Proof. by apply/malgP=> x; rewrite !fgscaleE mulrA. Qed.
451451

452-
Lemma fgscale1r D: 1 *:g D = D.
452+
Lemma fgscale0r g : 0 *:g g = 0.
453+
Proof. by apply/malgP=> k; rewrite fgscaleE mul0r mcoeff0. Qed.
454+
455+
Lemma fgscale1r g : 1 *:g g = g.
453456
Proof. by apply/malgP=> k; rewrite !fgscaleE mul1r. Qed.
454457

455458
Lemma fgscaleDr c g1 g2 : c *:g (g1 + g2) = c *:g g1 + c *:g g2.
@@ -458,37 +461,28 @@ Proof. by apply/malgP=> k; rewrite !(mcoeffD, fgscaleE) mulrDr. Qed.
458461
Lemma fgscaleDl g c1 c2: (c1 + c2) *:g g = c1 *:g g + c2 *:g g.
459462
Proof. by apply/malgP=> x; rewrite !(mcoeffD, fgscaleE) mulrDl. Qed.
460463

461-
End MalgSemiRingTheory.
462-
463-
(* -------------------------------------------------------------------- *)
464-
Section MalgRingTheory.
465-
466-
Context {K : choiceType} {R : ringType}.
467-
468-
Implicit Types (g : {malg R[K]}).
469-
470-
(* TODO: Add a semi-module structure and generalize this section *)
471-
HB.instance Definition _ := GRing.Zmodule_isLmodule.Build R {malg R[K]}
472-
fgscaleA fgscale1r fgscaleDr fgscaleDl.
464+
HB.instance Definition _ := GRing.Nmodule_isLSemiModule.Build R {malg R [K]}
465+
fgscaleA fgscale0r fgscale1r fgscaleDr fgscaleDl.
473466

474467
Lemma malgZ_def c g : c *: g = fgscale c g.
475468
Proof. by []. Qed.
476469

477470
Lemma mcoeffZ c g k : (c *: g)@_k = c * g@_k.
478471
Proof. exact/fgscaleE. Qed.
479472

480-
(* FIXME: make the production of a LRMorphism fail below *)
481-
(* HB.instance Definition _ m := *)
482-
(* GRing.isLinear.Build R [lmodType R of {malg R[K]}] R *%R (mcoeff m) *)
483-
(* (fun c g => mcoeffZ c g m). *)
473+
(* FIXME: this instance has to be declared after the RMorphism instance of *)
474+
(* [mcoeff 1%M] to produce the LRMorphism instance. *)
475+
(* HB.instance Definition _ k := *)
476+
(* GRing.isSemilinear.Build R {malg R[K]} R *%R (mcoeff k) *)
477+
(* (fun c g => mcoeffZ c g k, mcoeffD k). *)
484478

485479
Lemma msuppZ_le c g : msupp (c *: g) `<=` msupp g.
486480
Proof.
487481
apply/fsubsetP=> k; rewrite -!mcoeff_neq0 mcoeffZ.
488482
by apply/contraTneq=> ->; rewrite mulr0 negbK.
489483
Qed.
490484

491-
End MalgRingTheory.
485+
End MalgSemiRingTheory.
492486

493487
(* -------------------------------------------------------------------- *)
494488
Section MalgLmodTheoryIntegralDomain.
@@ -586,8 +580,8 @@ move=> k _ /mcoeff_outdom g1k.
586580
by rewrite big1 => // k' _; rewrite g1k mul0r monalgU0.
587581
Qed.
588582

589-
Lemma fgmulrw (d1 d2 : {fset K}) g1 g2 : msupp g1 `<=` d1 -> msupp g2 `<=` d2
590-
-> fgmul g1 g2 = \sum_(k2 <- d2) \sum_(k1 <- d1) g1 *M_[k1, k2] g2.
583+
Lemma fgmulrw (d1 d2 : {fset K}) g1 g2 : msupp g1 `<=` d1 -> msupp g2 `<=` d2 ->
584+
fgmul g1 g2 = \sum_(k2 <- d2) \sum_(k1 <- d1) g1 *M_[k1, k2] g2.
591585
Proof. by move=> le_d1 le_d2; rewrite (fgmullw le_d1 le_d2) exchange_big. Qed.
592586

593587
Definition fgmullwl (d1 : {fset K}) {g1 g2} (le : msupp g1 `<=` d1) :=
@@ -678,6 +672,10 @@ HB.instance Definition _ := GRing.Nmodule_isSemiRing.Build {malg R[K]}
678672

679673
End MalgSemiRingType.
680674

675+
(* TODO: HB.saturate *)
676+
HB.instance Definition _ (K : monomType) (R : ringType) :=
677+
GRing.SemiRing.on {malg R[K]}.
678+
681679
(* -------------------------------------------------------------------- *)
682680
Section MalgSemiRingTheory.
683681

@@ -799,45 +797,39 @@ rewrite !raddf_sum !big1 ?addr0 //= => k; rewrite in_fsetD1 => /andP [ne1_k _].
799797
by rewrite mcoeffU mul1m (negbTE ne1_k).
800798
Qed.
801799

800+
(* FIXME: this instance declaration fails if the [Linear] instance is *)
801+
(* declared first. *)
802802
HB.instance Definition _ :=
803-
GRing.isMultiplicative.Build {malg R[K]} R (@mcoeff K R 1%M)
803+
GRing.isMultiplicative.Build {malg R[K]} R (mcoeff 1%M)
804804
mcoeff1g_is_multiplicative.
805805

806-
End MalgSemiRingTheory.
807-
808-
(* -------------------------------------------------------------------- *)
809-
Section MalgRingTheory.
810-
811-
Context {K : monomType} {R : ringType}.
812-
813-
Implicit Types (g : {malg R[K]}) (k l : K).
814-
815-
HB.instance Definition _ := GRing.SemiRing.on {malg R[K]}.
816-
817806
Lemma mul_malgC c g : c%:MP * g = c *: g.
818807
Proof.
819808
rewrite malgM_def malgZ_def fgmulUg.
820809
by apply/eq_bigr=> /= k _; rewrite mul1m.
821810
Qed.
822811

823-
(* FIXME: building Linear instance here so as to not trigger the creation
824-
of a LRMorphism that fails on above command (but is built just below anyway) *)
825-
HB.instance Definition _ m :=
826-
GRing.isScalable.Build R {malg R[K]} R *%R (mcoeff m)
827-
(fun c => (mcoeffZ c)^~ m).
828-
829812
Lemma fgscaleAl c g1 g2 : c *: (g1 * g2) = (c *: g1) * g2.
830813
Proof. by rewrite -!mul_malgC mulrA. Qed.
831814

832-
HB.instance Definition _ := GRing.Lmodule_isLalgebra.Build R {malg R[K]}
833-
fgscaleAl.
815+
HB.instance Definition _ :=
816+
GRing.LSemiModule_isLSemiAlgebra.Build R {malg R[K]} fgscaleAl.
834817

835-
End MalgRingTheory.
818+
End MalgSemiRingTheory.
819+
820+
(* FIXME: the [Linear] instance moved from above *)
821+
HB.instance Definition _ (K : choiceType) (R : semiRingType) k :=
822+
GRing.isScalable.Build R {malg R[K]} R *%R (mcoeff k)
823+
(fun c g => mcoeffZ c g k).
824+
825+
(* FIXME: HB.saturate? *)
826+
HB.instance Definition _ (K : monomType) (R : semiRingType) :=
827+
GRing.Linear.on (mcoeff 1%M : {malg R[K]} -> R).
836828

837829
(* -------------------------------------------------------------------- *)
838-
Section MalgComRingType.
830+
Section MalgComSemiRingType.
839831

840-
Context {K : conomType} {R : comRingType}.
832+
Context {K : conomType} {R : comSemiRingType}.
841833

842834
Lemma fgmulC : @commutative {malg R[K]} _ *%R.
843835
Proof.
@@ -846,12 +838,20 @@ apply/eq_bigr=> /= k1 _; apply/eq_bigr=> /= k2 _.
846838
by rewrite mulrC [X in X==k]mulmC.
847839
Qed.
848840

849-
HB.instance Definition _ := GRing.Ring_hasCommutativeMul.Build (malg K R)
850-
fgmulC.
841+
HB.instance Definition _ :=
842+
GRing.SemiRing_hasCommutativeMul.Build {malg R[K]} fgmulC.
843+
844+
HB.instance Definition _ :=
845+
GRing.LSemiAlgebra_isComSemiAlgebra.Build R {malg R[K]}.
851846

852-
HB.instance Definition _ := GRing.Lalgebra_isComAlgebra.Build R {malg R[K]}.
847+
End MalgComSemiRingType.
853848

854-
End MalgComRingType.
849+
(* FIXME: HB.saturate *)
850+
HB.instance Definition _ (K : monomType) (R : ringType) :=
851+
GRing.Lmodule.on {malg R[K]}.
852+
HB.instance Definition _ (K : conomType) (R : comRingType) :=
853+
GRing.Lmodule.on {malg R[K]}.
854+
(* /FIXME *)
855855

856856
(* -------------------------------------------------------------------- *)
857857
Section MalgMorphism.

0 commit comments

Comments
 (0)