diff options
author | Emilio Jesus Gallego Arias <e+git@x80.org> | 2017-10-05 14:20:28 +0200 |
---|---|---|
committer | Emilio Jesus Gallego Arias <e+git@x80.org> | 2017-10-05 23:47:32 +0200 |
commit | 6526933e3d6a6392aa6bd5235ea0180bb68b6f94 (patch) | |
tree | 83fdc73e5ef1c0e7bdb8b083b7887d09855ce38d | |
parent | 2aac4ae818fec0d409da31ef9da83796d871d687 (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
-rw-r--r-- | plugins/ltac/rewrite.ml | 10 | ||||
-rw-r--r-- | plugins/micromega/EnvRing.v | 16 | ||||
-rw-r--r-- | plugins/setoid_ring/Field_theory.v | 24 | ||||
-rw-r--r-- | plugins/setoid_ring/InitialRing.v | 50 | ||||
-rw-r--r-- | plugins/setoid_ring/Ring_polynom.v | 16 | ||||
-rw-r--r-- | plugins/setoid_ring/Ring_theory.v | 39 | ||||
-rw-r--r-- | test-suite/success/setoid_test.v | 10 | ||||
-rw-r--r-- | test-suite/success/setoid_test2.v | 16 | ||||
-rw-r--r-- | theories/FSets/FSetProperties.v | 2 | ||||
-rw-r--r-- | theories/QArith/Qreduction.v | 8 |
10 files changed, 128 insertions, 63 deletions
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index fd791a910..1809f0fcd 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1935,7 +1935,12 @@ let default_morphism sign m = let evars, mor = resolve_one_typeclass env (goalevars evars) morph in mor, proper_projection sigma mor morph +let warn_add_setoid_deprecated = + CWarnings.create ~name:"add-setoid" ~category:"deprecated" (fun () -> + Pp.(str "Add Setoid is deprecated, please use Add Parametric Relation.")) + let add_setoid global binders a aeq t n = + warn_add_setoid_deprecated ?loc:a.CAst.loc (); init_setoid (); let _lemma_refl = declare_instance_refl global binders a aeq n (mkappc "Seq_refl" [a;aeq;t]) in let _lemma_sym = declare_instance_sym global binders a aeq n (mkappc "Seq_sym" [a;aeq;t]) in @@ -1954,7 +1959,12 @@ let make_tactic name = let tacname = Qualid (Loc.tag tacpath) in TacArg (Loc.tag @@ TacCall (Loc.tag (tacname, []))) +let warn_add_morphism_deprecated = + CWarnings.create ~name:"add-morphism" ~category:"deprecated" (fun () -> + Pp.(str "Add Morphism f : id is deprecated, please use Add Morphism f with signature (...) as id")) + let add_morphism_infer glob m n = + warn_add_morphism_deprecated ?loc:m.CAst.loc (); init_setoid (); let poly = Flags.is_universe_polymorphism () in let instance_id = add_suffix n "_Proper" in diff --git a/plugins/micromega/EnvRing.v b/plugins/micromega/EnvRing.v index 56b3d480e..ae4857a77 100644 --- a/plugins/micromega/EnvRing.v +++ b/plugins/micromega/EnvRing.v @@ -56,10 +56,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/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]). diff --git a/test-suite/success/setoid_test.v b/test-suite/success/setoid_test.v index 1f24ef2a6..c8dfcd2cb 100644 --- a/test-suite/success/setoid_test.v +++ b/test-suite/success/setoid_test.v @@ -33,7 +33,8 @@ Qed. Add Setoid set same setoid_set as setsetoid. -Add Morphism In : In_ext. +Add Morphism In with signature (eq ==> same ==> iff) as In_ext. +Proof. unfold same; intros a s t H; elim (H a); auto. Qed. @@ -50,10 +51,9 @@ simpl; right. apply (H2 H1). Qed. -Add Morphism Add : Add_ext. +Add Morphism Add with signature (eq ==> same ==> same) as Add_ext. split; apply add_aux. assumption. - rewrite H. reflexivity. Qed. @@ -90,7 +90,7 @@ Qed. Parameter P : set -> Prop. Parameter P_ext : forall s t : set, same s t -> P s -> P t. -Add Morphism P : P_extt. +Add Morphism P with signature (same ==> iff) as P_extt. intros; split; apply P_ext; (assumption || apply (Seq_sym _ _ setoid_set); assumption). Qed. @@ -113,7 +113,7 @@ Definition f: forall A : Set, A -> A := fun A x => x. Add Relation (id A) (rel A) as eq_rel. -Add Morphism (@f A) : f_morph. +Add Morphism (@f A) with signature (eq ==> eq) as f_morph. Proof. unfold rel, f. trivial. Qed. diff --git a/test-suite/success/setoid_test2.v b/test-suite/success/setoid_test2.v index 6baf79701..79467e549 100644 --- a/test-suite/success/setoid_test2.v +++ b/test-suite/success/setoid_test2.v @@ -134,8 +134,8 @@ Axiom SetoidS2 : Setoid_Theory S2 eqS2. Add Setoid S2 eqS2 SetoidS2 as S2setoid. Axiom f : S1 -> nat -> S2. -Add Morphism f : f_compat. Admitted. -Add Morphism f : f_compat2. Admitted. +Add Morphism f with signature (eqS1 ==> eq ==> eqS2) as f_compat. Admitted. +Add Morphism f with signature (eqS1 ==> eq ==> eqS2) as f_compat2. Admitted. Theorem test1: forall x y, (eqS1 x y) -> (eqS2 (f x 0) (f y 0)). intros. @@ -151,7 +151,7 @@ Theorem test1': forall x y, (eqS1 x y) -> (eqS2 (f x 0) (f y 0)). Qed. Axiom g : S1 -> S2 -> nat. -Add Morphism g : g_compat. Admitted. +Add Morphism g with signature (eqS1 ==> eqS2 ==> eq) as g_compat. Admitted. Axiom P : nat -> Prop. Theorem test2: @@ -190,13 +190,13 @@ Theorem test5: Qed. Axiom f_test6 : S2 -> Prop. -Add Morphism f_test6 : f_test6_compat. Admitted. +Add Morphism f_test6 with signature (eqS2 ==> iff) as f_test6_compat. Admitted. Axiom g_test6 : bool -> S2. -Add Morphism g_test6 : g_test6_compat. Admitted. +Add Morphism g_test6 with signature (eq ==> eqS2) as g_test6_compat. Admitted. Axiom h_test6 : S1 -> bool. -Add Morphism h_test6 : h_test6_compat. Admitted. +Add Morphism h_test6 with signature (eqS1 ==> eq) as h_test6_compat. Admitted. Theorem test6: forall E1 E2, (eqS1 E1 E2) -> (f_test6 (g_test6 (h_test6 E2))) -> @@ -223,7 +223,7 @@ Add Setoid S1_test8 eqS1_test8 SetoidS1_test8 as S1_test8setoid. Instance eqS1_test8_default : DefaultRelation eqS1_test8. Axiom f_test8 : S2 -> S1_test8. -Add Morphism f_test8 : f_compat_test8. Admitted. +Add Morphism f_test8 with signature (eqS2 ==> eqS1_test8) as f_compat_test8. Admitted. Axiom eqS1_test8': S1_test8 -> S1_test8 -> Prop. Axiom SetoidS1_test8' : Setoid_Theory S1_test8 eqS1_test8'. @@ -233,7 +233,7 @@ Add Setoid S1_test8 eqS1_test8' SetoidS1_test8' as S1_test8setoid'. (S1_test8, eqS1_test8'). However this does not happen and there is still no syntax for it ;-( *) Axiom g_test8 : S1_test8 -> S2. -Add Morphism g_test8 : g_compat_test8. Admitted. +Add Morphism g_test8 with signature (eqS1_test8 ==> eqS2) as g_compat_test8. Admitted. Theorem test8: forall x x': S2, (eqS2 x x') -> diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v index 25b042ca9..0041bfa1c 100644 --- a/theories/FSets/FSetProperties.v +++ b/theories/FSets/FSetProperties.v @@ -762,7 +762,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E). rewrite (cardinal_2 (s:=remove x s') (s':=s') (x:=x)); eauto with set. Qed. - Add Morphism cardinal : cardinal_m. + Add Morphism cardinal with signature (Equal ==> Logic.eq) as cardinal_m. Proof. exact Equal_cardinal. Qed. diff --git a/theories/QArith/Qreduction.v b/theories/QArith/Qreduction.v index 88e1298fb..5d055b547 100644 --- a/theories/QArith/Qreduction.v +++ b/theories/QArith/Qreduction.v @@ -101,7 +101,7 @@ Proof. - apply Qred_complete. Qed. -Add Morphism Qred : Qred_comp. +Add Morphism Qred with signature (Qeq ==> Qeq) as Qred_comp. Proof. intros. now rewrite !Qred_correct. Qed. @@ -125,19 +125,19 @@ Proof. intros; unfold Qminus'; apply Qred_correct; auto. Qed. -Add Morphism Qplus' : Qplus'_comp. +Add Morphism Qplus' with signature (Qeq ==> Qeq ==> Qeq) as Qplus'_comp. Proof. intros; unfold Qplus'. rewrite H, H0; auto with qarith. Qed. -Add Morphism Qmult' : Qmult'_comp. +Add Morphism Qmult' with signature (Qeq ==> Qeq ==> Qeq) as Qmult'_comp. Proof. intros; unfold Qmult'. rewrite H, H0; auto with qarith. Qed. -Add Morphism Qminus' : Qminus'_comp. +Add Morphism Qminus' with signature (Qeq ==> Qeq ==> Qeq) as Qminus'_comp. Proof. intros; unfold Qminus'. rewrite H, H0; auto with qarith. |