From f87b9c9564b21b2f531629311787f329c5851af7 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 17 Jan 2018 08:01:18 -0500 Subject: Extend `zify_N` with knowledge about `N.pred` by doing the same thing as `zify_nat` does for `nat.pred`. This fixes #6602. --- CHANGES | 1 + plugins/omega/PreOmega.v | 6 +++++- test-suite/bugs/opened/6602.v | 17 +++++++++++++++++ 3 files changed, 23 insertions(+), 1 deletion(-) create mode 100644 test-suite/bugs/opened/6602.v diff --git a/CHANGES b/CHANGES index 545925026..8d75d543d 100644 --- a/CHANGES +++ b/CHANGES @@ -32,6 +32,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. -- cgit v1.2.3