aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES1
-rw-r--r--plugins/omega/PreOmega.v6
-rw-r--r--test-suite/bugs/opened/6602.v17
3 files changed, 23 insertions, 1 deletions
diff --git a/CHANGES b/CHANGES
index 32a533d14..7ea23eeb7 100644
--- a/CHANGES
+++ b/CHANGES
@@ -41,6 +41,7 @@ Tactics
such as "x := 5 : Z" (see BZ#148). This could be disabled via
Unset Omega UseLocalDefs.
- The tactic "romega" is also aware now of the bodies of context variables.
+- The tactic "zify" resp. "omega with N" is now aware of N.pred.
- Tactic "decide equality" now able to manage constructors which
contain proofs.
- Added tactics reset ltac profile, show ltac profile (and variants)
diff --git a/plugins/omega/PreOmega.v b/plugins/omega/PreOmega.v
index 8da45e0ad..93103e026 100644
--- a/plugins/omega/PreOmega.v
+++ b/plugins/omega/PreOmega.v
@@ -26,7 +26,7 @@ Local Open Scope Z_scope.
- on Z: Z.min, Z.max, Z.abs, Z.sgn are translated in term of <= < =
- on nat: + * - S O pred min max Pos.to_nat N.to_nat Z.abs_nat
- on positive: Zneg Zpos xI xO xH + * - Pos.succ Pos.pred Pos.min Pos.max Pos.of_succ_nat
- - on N: N0 Npos + * - N.succ N.min N.max N.of_nat Z.abs_N
+ - on N: N0 Npos + * - N.pred N.succ N.min N.max N.of_nat Z.abs_N
*)
@@ -391,6 +391,10 @@ Ltac zify_N_op :=
| H : context [ Z.of_N (N.sub ?a ?b) ] |- _ => rewrite (N2Z.inj_sub_max a b) in H
| |- context [ Z.of_N (N.sub ?a ?b) ] => rewrite (N2Z.inj_sub_max a b)
+ (* pred -> minus ... -1 -> Z.max (Z.sub ... -1) 0 *)
+ | H : context [ Z.of_N (N.pred ?a) ] |- _ => rewrite (N.pred_sub a) in H
+ | |- context [ Z.of_N (N.pred ?a) ] => rewrite (N.pred_sub a)
+
(* N.succ -> Z.succ *)
| H : context [ Z.of_N (N.succ ?a) ] |- _ => rewrite (N2Z.inj_succ a) in H
| |- context [ Z.of_N (N.succ ?a) ] => rewrite (N2Z.inj_succ a)
diff --git a/test-suite/bugs/opened/6602.v b/test-suite/bugs/opened/6602.v
new file mode 100644
index 000000000..3690adf90
--- /dev/null
+++ b/test-suite/bugs/opened/6602.v
@@ -0,0 +1,17 @@
+Require Import Omega.
+
+Lemma test_nat:
+ forall n, (5 + pred n <= 5 + n).
+Proof.
+ intros.
+ zify.
+ omega.
+Qed.
+
+Lemma test_N:
+ forall n, (5 + N.pred n <= 5 + n)%N.
+Proof.
+ intros.
+ zify.
+ omega.
+Qed.