@@ -903,65 +903,42 @@ Lemma mmapMNn n : {morph mmap f h: x / x *- n} . Proof. exact: raddfMNn. Qed.
903
903
904
904
End Additive.
905
905
906
- Section CommrMultiplicative .
906
+ Section Multiplicative .
907
907
908
- Context {K : monomType} {R S : pzSemiRingType }.
909
- Context (f : {rmorphism R -> S}) ( h : {mmorphism K -> S}).
908
+ Context {K : monomType} {R : pzSemiRingType} { S : semiAlgType R }.
909
+ Context (h : {mmorphism K -> S}).
910
910
911
911
Implicit Types (c : R) (g : {malg R[K]}).
912
912
913
- Lemma mmapZ c g : (c *: g)^[f,h] = f c * g^[f,h].
914
- Proof .
915
- rewrite (mmapEw (msuppZ_le _ _)) mmapE big_distrr /=.
916
- by apply/eq_bigr=> k _; rewrite linearZ rmorphM /= mulrA.
917
- Qed .
913
+ (* FIXME: the implicit status of in_alg *)
914
+ Local Notation "g ^[ h ]" := (g ^[@GRing.in_alg _ _, h]).
918
915
919
- Lemma mmapC c : c%:MP^[f, h] = f c .
916
+ Lemma mmapC c : c%:MP^[h] = c%:A .
920
917
Proof . by rewrite mmapU mmorph1 mulr1. Qed .
921
918
922
- Lemma mmap1 : 1^[f,h] = 1.
923
- Proof . by rewrite mmapC rmorph1. Qed .
924
-
925
- Hypothesis commr_f: forall g m m', GRing.comm (f g@_m) (h m').
919
+ Lemma mmap1 : 1^[h] = 1.
920
+ Proof . by rewrite mmapC scale1r. Qed .
926
921
927
- Lemma commr_mmap_is_multiplicative : multiplicative (mmap f h).
922
+ Lemma mmap_is_multiplicative : multiplicative (mmap (@GRing.in_alg _ _) h).
928
923
Proof .
929
924
split => [g1 g2|]; last by rewrite mmap1.
930
925
rewrite malgME raddf_sum mulr_suml /=; apply: eq_bigr=> i _.
931
926
rewrite raddf_sum mulr_sumr /=; apply: eq_bigr=> j _.
932
- by rewrite mmapU /= rmorphM mmorphM -mulrA [X in _*X=_]mulrA commr_f !mulrA .
927
+ by rewrite mmapU/= !mulr_algl -!(scalerAl, scalerCA) scalerA -mmorphM .
933
928
Qed .
934
929
935
- End CommrMultiplicative.
936
-
937
- (* -------------------------------------------------------------------- *)
938
- Section Multiplicative.
939
-
940
- Context {K : monomType} {R : pzSemiRingType} {S : comPzSemiRingType}.
941
- Context (f : {rmorphism R -> S}) (h : {mmorphism K -> S}).
942
-
943
- Lemma mmap_is_multiplicative : multiplicative (mmap f h).
944
- Proof . by apply/commr_mmap_is_multiplicative=> g m m'; apply/mulrC. Qed .
945
-
946
930
HB.instance Definition _ :=
947
- GRing.isMultiplicative.Build {malg R[K]} S (mmap f h)
931
+ GRing.isMultiplicative.Build {malg R[K]} S (mmap (@GRing.in_alg _ _) h)
948
932
mmap_is_multiplicative.
949
933
950
- End Multiplicative.
951
-
952
- (* -------------------------------------------------------------------- *)
953
- Section Linear.
954
-
955
- Context {K : monomType} {R : comPzSemiRingType} (h : {mmorphism K -> R}).
956
-
957
- Lemma mmap_is_linear : scalable_for *%R (mmap idfun h).
958
- Proof . by move=> /= c g; rewrite -mul_malgC rmorphM /= mmapC. Qed .
934
+ Lemma mmapZ : scalable (mmap (@GRing.in_alg _ _) h).
935
+ Proof . by move=> /= c g; rewrite -mul_malgC rmorphM/= mmapC mulr_algl. Qed .
959
936
960
937
HB.instance Definition _ :=
961
- GRing.isScalable.Build R {malg R[K]} R * %R (mmap idfun h)
962
- mmap_is_linear .
938
+ GRing.isScalable.Build R {malg R[K]} S *: %R (mmap (@GRing.in_alg _ _) h)
939
+ mmapZ .
963
940
964
- End Linear .
941
+ End Multiplicative .
965
942
End MalgMorphism.
966
943
967
944
(* -------------------------------------------------------------------- *)
0 commit comments