@@ -1082,65 +1082,41 @@ Lemma mmapMNn n : {morph mmap f h: x / x *- n} . Proof. exact: raddfMNn. Qed.
10821082
10831083End Additive.
10841084
1085- Section CommrMultiplicative .
1085+ Section Multiplicative .
10861086
1087- Context {K : monomType} {R S : pzSemiRingType }.
1088- Context (f : {rmorphism R -> S}) ( h : {mmorphism K -> S}).
1087+ Context {K : monomType} {R : pzSemiRingType} { S : pzSemiAlgType R }.
1088+ Context (h : {mmorphism K -> S}).
10891089
10901090Implicit Types (c : R) (g : {malg R[K]}).
10911091
1092- Lemma mmapZ c g : (c *: g)^[f,h] = f c * g^[f,h].
1093- Proof .
1094- rewrite (mmapEw (msuppZ_le _ _)) mmapE big_distrr /=.
1095- by apply/eq_bigr=> k _; rewrite linearZ rmorphM /= mulrA.
1096- Qed .
1092+ Local Notation "g ^[ h ]" := (g ^[GRing.in_alg S, h]).
10971093
1098- Lemma mmapC c : c%:MP^[f, h] = f c .
1094+ Lemma mmapC c : c%:MP^[h] = c%:A .
10991095Proof . by rewrite mmapU mmorph1 mulr1. Qed .
11001096
1101- Lemma mmap1 : 1^[f,h] = 1.
1102- Proof . by rewrite mmapC rmorph1. Qed .
1103-
1104- Hypothesis commr_f: forall g m m', GRing.comm (f g@_m) (h m').
1097+ Lemma mmap1 : 1^[h] = 1.
1098+ Proof . by rewrite mmapC scale1r. Qed .
11051099
1106- Lemma commr_mmap_is_multiplicative : multiplicative (mmap f h).
1100+ Lemma mmap_is_multiplicative : multiplicative (mmap (GRing.in_alg S) h).
11071101Proof .
11081102split => [g1 g2|]; last by rewrite mmap1.
11091103rewrite malgME raddf_sum mulr_suml /=; apply: eq_bigr=> i _.
11101104rewrite raddf_sum mulr_sumr /=; apply: eq_bigr=> j _.
1111- by rewrite mmapU /= rmorphM mmorphM -mulrA [X in _*X=_]mulrA commr_f !mulrA .
1105+ by rewrite mmapU/= !mulr_algl -!(scalerAl, scalerCA) scalerA -mmorphM .
11121106Qed .
11131107
1114- End CommrMultiplicative.
1115-
1116- (* -------------------------------------------------------------------- *)
1117- Section Multiplicative.
1118-
1119- Context {K : monomType} {R : pzSemiRingType} {S : comPzSemiRingType}.
1120- Context (f : {rmorphism R -> S}) (h : {mmorphism K -> S}).
1121-
1122- Lemma mmap_is_multiplicative : multiplicative (mmap f h).
1123- Proof . by apply/commr_mmap_is_multiplicative=> g m m'; apply/mulrC. Qed .
1124-
11251108HB.instance Definition _ :=
1126- GRing.isMultiplicative.Build {malg R[K]} S (mmap f h)
1109+ GRing.isMultiplicative.Build {malg R[K]} S (mmap (GRing.in_alg S) h)
11271110 mmap_is_multiplicative.
11281111
1129- End Multiplicative.
1130-
1131- (* -------------------------------------------------------------------- *)
1132- Section Linear.
1133-
1134- Context {K : monomType} {R : comPzSemiRingType} (h : {mmorphism K -> R}).
1135-
1136- Lemma mmap_is_linear : scalable_for *%R (mmap idfun h).
1137- Proof . by move=> /= c g; rewrite -mul_malgC rmorphM /= mmapC. Qed .
1112+ Lemma mmapZ : scalable (mmap (@GRing.in_alg _ _) h).
1113+ Proof . by move=> /= c g; rewrite -mul_malgC rmorphM/= mmapC mulr_algl. Qed .
11381114
11391115HB.instance Definition _ :=
1140- GRing.isScalable.Build R {malg R[K]} R * %R (mmap idfun h)
1141- mmap_is_linear .
1116+ GRing.isScalable.Build R {malg R[K]} S *: %R (mmap (GRing.in_alg S) h)
1117+ mmapZ .
11421118
1143- End Linear .
1119+ End Multiplicative .
11441120End MalgMorphism.
11451121
11461122(* -------------------------------------------------------------------- *)
0 commit comments