summaryrefslogtreecommitdiff
path: root/plugins/setoid_ring/InitialRing.v
diff options
context:
space:
mode:
Diffstat (limited to 'plugins/setoid_ring/InitialRing.v')
-rw-r--r--plugins/setoid_ring/InitialRing.v84
1 files changed, 56 insertions, 28 deletions
diff --git a/plugins/setoid_ring/InitialRing.v b/plugins/setoid_ring/InitialRing.v
index 9c690e2b..f5db2754 100644
--- a/plugins/setoid_ring/InitialRing.v
+++ b/plugins/setoid_ring/InitialRing.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Import Zbool.
@@ -48,12 +50,19 @@ Section ZMORPHISM.
Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
Notation "x == y" := (req x y).
Variable Rsth : Setoid_Theory R req.
- Add Setoid R req Rsth as R_setoid3.
+ Add Parametric Relation : R req
+ reflexivity proved by Rsth.(@Equivalence_Reflexive _ _)
+ symmetry proved by Rsth.(@Equivalence_Symmetric _ _)
+ transitivity proved by Rsth.(@Equivalence_Transitive _ _)
+ as R_setoid3.
Ltac rrefl := gen_reflexivity Rsth.
Variable Reqe : ring_eq_ext radd rmul ropp req.
- Add Morphism radd : radd_ext3. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext3. exact (Rmul_ext Reqe). Qed.
- Add Morphism ropp : ropp_ext3. exact (Ropp_ext Reqe). Qed.
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext3.
+ Proof. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext3.
+ Proof. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp with signature (req ==> req) as ropp_ext3.
+ Proof. exact (Ropp_ext Reqe). Qed.
Fixpoint gen_phiPOS1 (p:positive) : R :=
match p with
@@ -103,7 +112,8 @@ Section ZMORPHISM.
Section ALMOST_RING.
Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
- Add Morphism rsub : rsub_ext3. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext3.
+ Proof. exact (ARsub_ext Rsth Reqe ARth). Qed.
Ltac norm := gen_srewrite Rsth Reqe ARth.
Ltac add_push := gen_add_push radd Rsth Reqe ARth.
@@ -151,7 +161,8 @@ Section ZMORPHISM.
Variable Rth : ring_theory 0 1 radd rmul rsub ropp req.
Let ARth := Rth_ARth Rsth Reqe Rth.
- Add Morphism rsub : rsub_ext4. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext4.
+ Proof. exact (ARsub_ext Rsth Reqe ARth). Qed.
Ltac norm := gen_srewrite Rsth Reqe ARth.
Ltac add_push := gen_add_push radd Rsth Reqe ARth.
@@ -255,7 +266,11 @@ Section NMORPHISM.
Notation "0" := rO. Notation "1" := rI.
Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y).
Variable Rsth : Setoid_Theory R req.
- Add Setoid R req Rsth as R_setoid4.
+ Add Parametric Relation : R req
+ reflexivity proved by Rsth.(@Equivalence_Reflexive _ _)
+ symmetry proved by Rsth.(@Equivalence_Symmetric _ _)
+ transitivity proved by Rsth.(@Equivalence_Transitive _ _)
+ as R_setoid4.
Ltac rrefl := gen_reflexivity Rsth.
Variable SReqe : sring_eq_ext radd rmul req.
Variable SRth : semi_ring_theory 0 1 radd rmul req.
@@ -265,8 +280,10 @@ Section NMORPHISM.
Let rsub := (@SRsub R radd).
Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
Notation "x == y" := (req x y).
- Add Morphism radd : radd_ext4. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext4. exact (Rmul_ext Reqe). Qed.
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext4.
+ Proof. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext4.
+ Proof. exact (Rmul_ext Reqe). Qed.
Ltac norm := gen_srewrite_sr Rsth Reqe ARth.
Definition gen_phiN1 x :=
@@ -374,15 +391,23 @@ Section NWORDMORPHISM.
Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
Notation "x == y" := (req x y).
Variable Rsth : Setoid_Theory R req.
- Add Setoid R req Rsth as R_setoid5.
+ Add Parametric Relation : R req
+ reflexivity proved by Rsth.(@Equivalence_Reflexive _ _)
+ symmetry proved by Rsth.(@Equivalence_Symmetric _ _)
+ transitivity proved by Rsth.(@Equivalence_Transitive _ _)
+ as R_setoid5.
Ltac rrefl := gen_reflexivity Rsth.
Variable Reqe : ring_eq_ext radd rmul ropp req.
- Add Morphism radd : radd_ext5. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext5. exact (Rmul_ext Reqe). Qed.
- Add Morphism ropp : ropp_ext5. exact (Ropp_ext Reqe). Qed.
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext5.
+ Proof. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext5.
+ Proof. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp with signature (req ==> req) as ropp_ext5.
+ Proof. exact (Ropp_ext Reqe). Qed.
Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
- Add Morphism rsub : rsub_ext7. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext7.
+ Proof. exact (ARsub_ext Rsth Reqe ARth). Qed.
Ltac norm := gen_srewrite Rsth Reqe ARth.
Ltac add_push := gen_add_push radd Rsth Reqe ARth.
@@ -555,12 +580,20 @@ Section GEN_DIV.
Variable morph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi.
(* Useful tactics *)
- Add Setoid R req Rsth as R_set1.
+ Add Parametric Relation : R req
+ reflexivity proved by Rsth.(@Equivalence_Reflexive _ _)
+ symmetry proved by Rsth.(@Equivalence_Symmetric _ _)
+ transitivity proved by Rsth.(@Equivalence_Transitive _ _)
+ as R_set1.
Ltac rrefl := gen_reflexivity Rsth.
- Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed.
- Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed.
- Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext.
+ Proof. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext.
+ Proof. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp with signature (req ==> req) as ropp_ext.
+ Proof. exact (Ropp_ext Reqe). Qed.
+ Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext.
+ Proof. exact (ARsub_ext Rsth Reqe ARth). Qed.
Ltac rsimpl := gen_srewrite Rsth Reqe ARth.
Definition triv_div x y :=
@@ -859,8 +892,3 @@ Ltac isZcst t :=
(* *)
| _ => constr:(false)
end.
-
-
-
-
-