aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins/setoid_ring
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-10-05 14:20:28 +0200
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-10-05 23:47:32 +0200
commit6526933e3d6a6392aa6bd5235ea0180bb68b6f94 (patch)
tree83fdc73e5ef1c0e7bdb8b083b7887d09855ce38d /plugins/setoid_ring
parent2aac4ae818fec0d409da31ef9da83796d871d687 (diff)
[ltac] Warning for deprecated `Add Setoid` and `Add Morphism` forms.
The manual has long stated that these forms are deprecated. We add a warning for them, as indeed `Add Morphism` is an "proof evil" [*] command, and we may want to remove it in the future. We've also fixed the stdlib not to emit the warning. [*] https://ncatlab.org/nlab/show/principle+of+equivalence
Diffstat (limited to 'plugins/setoid_ring')
-rw-r--r--plugins/setoid_ring/Field_theory.v24
-rw-r--r--plugins/setoid_ring/InitialRing.v50
-rw-r--r--plugins/setoid_ring/Ring_polynom.v16
-rw-r--r--plugins/setoid_ring/Ring_theory.v39
4 files changed, 88 insertions, 41 deletions
diff --git a/plugins/setoid_ring/Field_theory.v b/plugins/setoid_ring/Field_theory.v
index 56b985aa3..88e2cb1da 100644
--- a/plugins/setoid_ring/Field_theory.v
+++ b/plugins/setoid_ring/Field_theory.v
@@ -56,11 +56,16 @@ Let rI_neq_rO := AFth.(AF_1_neq_0).
Let rdiv_def := AFth.(AFdiv_def).
Let rinv_l := AFth.(AFinv_l).
-Add Morphism radd : radd_ext. Proof. exact (Radd_ext Reqe). Qed.
-Add Morphism rmul : rmul_ext. Proof. exact (Rmul_ext Reqe). Qed.
-Add Morphism ropp : ropp_ext. Proof. exact (Ropp_ext Reqe). Qed.
-Add Morphism rsub : rsub_ext. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed.
-Add Morphism rinv : rinv_ext. Proof. exact SRinv_ext. 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.
+Add Morphism rinv with signature (req ==> req) as rinv_ext.
+Proof. exact SRinv_ext. Qed.
Let eq_trans := Setoid.Seq_trans _ _ Rsth.
Let eq_sym := Setoid.Seq_sym _ _ Rsth.
@@ -1609,9 +1614,12 @@ Section Complete.
Variable Rsth : Setoid_Theory R req.
Add Setoid R req Rsth as R_setoid3.
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.
Section AlmostField.
diff --git a/plugins/setoid_ring/InitialRing.v b/plugins/setoid_ring/InitialRing.v
index 98ffff432..bd4e94687 100644
--- a/plugins/setoid_ring/InitialRing.v
+++ b/plugins/setoid_ring/InitialRing.v
@@ -51,9 +51,12 @@ Section ZMORPHISM.
Add Setoid R req Rsth 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 +106,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 +155,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.
@@ -265,8 +270,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 :=
@@ -377,12 +384,16 @@ Section NWORDMORPHISM.
Add Setoid R req Rsth 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.
@@ -557,10 +568,14 @@ Section GEN_DIV.
(* Useful tactics *)
Add Setoid R req Rsth 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 +874,3 @@ Ltac isZcst t :=
(* *)
| _ => constr:(false)
end.
-
-
-
-
-
diff --git a/plugins/setoid_ring/Ring_polynom.v b/plugins/setoid_ring/Ring_polynom.v
index ac54d862c..a94f8d8df 100644
--- a/plugins/setoid_ring/Ring_polynom.v
+++ b/plugins/setoid_ring/Ring_polynom.v
@@ -59,10 +59,18 @@ Section MakeRingPol.
Infix "?=!" := ceqb. Notation "[ x ]" := (phi x).
(* Useful tactics *)
- 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.
Ltac add_push := gen_add_push radd Rsth Reqe ARth.
diff --git a/plugins/setoid_ring/Ring_theory.v b/plugins/setoid_ring/Ring_theory.v
index 8dda5ecd3..335a68d70 100644
--- a/plugins/setoid_ring/Ring_theory.v
+++ b/plugins/setoid_ring/Ring_theory.v
@@ -254,8 +254,12 @@ Section ALMOST_RING.
Section SEMI_RING.
Variable SReqe : sring_eq_ext radd rmul req.
- Add Morphism radd : radd_ext1. exact (SRadd_ext SReqe). Qed.
- Add Morphism rmul : rmul_ext1. exact (SRmul_ext SReqe). Qed.
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext1.
+ Proof. exact (SRadd_ext SReqe). Qed.
+
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext1.
+ Proof. exact (SRmul_ext SReqe). Qed.
+
Variable SRth : semi_ring_theory 0 1 radd rmul req.
(** Every semi ring can be seen as an almost ring, by taking :
@@ -323,9 +327,15 @@ Section ALMOST_RING.
Notation "- x" := (ropp x).
Variable Reqe : ring_eq_ext radd rmul ropp req.
- Add Morphism radd : radd_ext2. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext2. exact (Rmul_ext Reqe). Qed.
- Add Morphism ropp : ropp_ext2. exact (Ropp_ext Reqe). Qed.
+
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext2.
+ Proof. exact (Radd_ext Reqe). Qed.
+
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext2.
+ Proof. exact (Rmul_ext Reqe). Qed.
+
+ Add Morphism ropp with signature (req ==> req) as ropp_ext2.
+ Proof. exact (Ropp_ext Reqe). Qed.
Section RING.
Variable Rth : ring_theory 0 1 radd rmul rsub ropp req.
@@ -393,14 +403,25 @@ Section ALMOST_RING.
Notation "?=!" := ceqb. Notation "[ x ]" := (phi x).
Variable Csth : Equivalence ceq.
Variable Ceqe : ring_eq_ext cadd cmul copp ceq.
+
Add Setoid C ceq Csth as C_setoid.
- Add Morphism cadd : cadd_ext. exact (Radd_ext Ceqe). Qed.
- Add Morphism cmul : cmul_ext. exact (Rmul_ext Ceqe). Qed.
- Add Morphism copp : copp_ext. exact (Ropp_ext Ceqe). Qed.
+
+ Add Morphism cadd with signature (ceq ==> ceq ==> ceq) as cadd_ext.
+ Proof. exact (Radd_ext Ceqe). Qed.
+
+ Add Morphism cmul with signature (ceq ==> ceq ==> ceq) as cmul_ext.
+ Proof. exact (Rmul_ext Ceqe). Qed.
+
+ Add Morphism copp with signature (ceq ==> ceq) as copp_ext.
+ Proof. exact (Ropp_ext Ceqe). Qed.
+
Variable Cth : ring_theory cO cI cadd cmul csub copp ceq.
Variable Smorph : semi_morph 0 1 radd rmul req cO cI cadd cmul ceqb phi.
Variable phi_ext : forall x y, ceq x y -> [x] == [y].
- Add Morphism phi : phi_ext1. exact phi_ext. Qed.
+
+ Add Morphism phi with signature (ceq ==> req) as phi_ext1.
+ Proof. exact phi_ext. Qed.
+
Lemma Smorph_opp x : [-!x] == -[x].
Proof.
rewrite <- (Rth.(Radd_0_l) [-!x]).