1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71 72 73 74
|
(* Accompanying material to https://hal.inria.fr/hal-02478907 *)
From Coq Require Import ssreflect ssrfun ZArith.
From HB Require Import structures.
Declare Scope hb_scope.
Delimit Scope hb_scope with G.
Open Scope hb_scope.
(* Bottom mixin in Fig. 1. *)
HB.mixin Record Monoid_of_Type M := {
zero : M;
add : M -> M -> M;
addrA : associative add;
add0r : left_id zero add;
addr0 : right_id zero add;
}.
HB.structure Definition Monoid := { M of Monoid_of_Type M }.
Notation "0" := zero : hb_scope.
Infix "+" := (@add _) : hb_scope.
(* Bottom right mixin in Fig. 1. *)
HB.mixin Record AbelianGroup_of_Monoid A of Monoid A := {
opp : A -> A;
addrC : commutative (add : A -> A -> A);
addNr : left_inverse zero opp add;
}.
HB.structure Definition AbelianGroup :=
{ A of Monoid A & AbelianGroup_of_Monoid A }.
Notation "- x" := (@opp _ x) : hb_scope.
Notation "x - y" := (x + - y) : hb_scope.
(* Top right mixin in Fig. 1. *)
HB.mixin Record Ring_of_AbelianGroup R of AbelianGroup R := {
one : R;
mul : R -> R -> R;
mulrA : associative mul;
mul1r : left_id one mul;
mulr1 : right_id one mul;
mulrDl : left_distributive mul add;
mulrDr : right_distributive mul add;
}.
HB.structure Definition Ring :=
{ R of AbelianGroup R & Ring_of_AbelianGroup R }.
Notation "1" := one : hb_scope.
Infix "*" := (@mul _) : hb_scope.
Lemma addrN {R : AbelianGroup.type} : right_inverse (zero : R) opp add.
Proof. by move=> x; rewrite addrC addNr. Qed.
(********)
(* TEST *)
(********)
Print Monoid.type.
(* Monoid.type := { sort : Type; ... } *)
Check @add.
(* add : forall M : Monoid.type, M -> M -> M *)
Check @addNr.
(* addNr : forall R : Ring.type, left_inverse zero opp add *)
Check @addrC. (* is now an axiom of abelian groups *)
(* addrC : forall R : AbelianGroup.type, commutative add *)
HB.instance
Definition Z_Monoid_axioms : Monoid_of_Type Z :=
Monoid_of_Type.Build Z 0%Z Z.add Z.add_assoc Z.add_0_l Z.add_0_r.
(********************************************************)
(* This test from V1 FAILS in V2, and is repaired in V3 *)
(********************************************************)
Fail
Definition Z_Ring_axioms : Ring_of_Monoid Z :=
Ring_of_Monoid.Build Z 1%Z Z.opp Z.mul
Z.add_opp_diag_l Z.add_opp_diag_r Z.mul_assoc Z.mul_1_l Z.mul_1_r
Z.mul_add_distr_r Z.mul_add_distr_l.
|