diff options
| author | Kazuhiko Sakaguchi | 2019-09-20 18:58:51 +0200 |
|---|---|---|
| committer | Kazuhiko Sakaguchi | 2019-09-24 00:25:15 +0200 |
| commit | e7b83e2a07fffaba0b19c62ef1a02c40d6ef603a (patch) | |
| tree | 64daeb1e336ff86934cc5b529d9d6926ed8bfad8 | |
| parent | dc690e7067aa91a05472b5d573cb463223ef4dec (diff) | |
Make `zify` does work for `Z.to_N`
| -rw-r--r-- | doc/changelog/04-tactics/10774-zify-Z_to_N.rst | 3 | ||||
| -rw-r--r-- | plugins/micromega/ZifyInst.v | 4 | ||||
| -rw-r--r-- | test-suite/micromega/bug_9162.v | 8 |
3 files changed, 15 insertions, 0 deletions
diff --git a/doc/changelog/04-tactics/10774-zify-Z_to_N.rst b/doc/changelog/04-tactics/10774-zify-Z_to_N.rst new file mode 100644 index 0000000000..ed46cb101e --- /dev/null +++ b/doc/changelog/04-tactics/10774-zify-Z_to_N.rst @@ -0,0 +1,3 @@ +- The :tacn:`zify` tactic is now aware of `Z.to_N`. + (`#10774 <https://github.com/coq/coq/pull/10774>`_ fixes + `#9162 <https://github.com/coq/coq/issues/9162>`_, by Kazuhiko Sakaguchi). diff --git a/plugins/micromega/ZifyInst.v b/plugins/micromega/ZifyInst.v index 3c44113604..1217e8a5f7 100644 --- a/plugins/micromega/ZifyInst.v +++ b/plugins/micromega/ZifyInst.v @@ -172,6 +172,10 @@ Program Instance Op_Z_of_N : UnOp Z.of_N := { TUOp := (fun x => x) ; TUOpInj := fun x => eq_refl (Z.of_N x) }. Add UnOp Op_Z_of_N. +Instance Op_Z_to_N : UnOp Z.to_N := + { TUOp := fun x => Z.max 0 x ; TUOpInj := ltac:(now intro x; destruct x) }. +Add UnOp Op_Z_to_N. + Instance Op_Z_neg : UnOp Z.neg := { TUOp := Z.opp ; TUOpInj := (fun x => eq_refl (Zneg x))}. Add UnOp Op_Z_neg. diff --git a/test-suite/micromega/bug_9162.v b/test-suite/micromega/bug_9162.v new file mode 100644 index 0000000000..4aedf57faf --- /dev/null +++ b/test-suite/micromega/bug_9162.v @@ -0,0 +1,8 @@ +Require Import ZArith Lia. +Local Open Scope Z_scope. + +Goal Z.of_N (Z.to_N 0) = 0. +Proof. lia. Qed. + +Goal forall q, (Z.of_N (Z.to_N 0) = 0 -> q = 0) -> Z.of_N (Z.to_N 0) = q. +Proof. lia. Qed. |
