diff options
| author | Vincent Laporte | 2019-10-29 16:18:00 +0000 |
|---|---|---|
| committer | Vincent Laporte | 2020-03-24 09:49:07 +0100 |
| commit | b99e95486d3f66c29cc831eb57c018c32b7479f5 (patch) | |
| tree | 13d10f1ba3e3f0aa11ffa50e110dcc23a8576b5f | |
| parent | d48757f04b1a23e2d47448da843409a7b44bd091 (diff) | |
“auto with zarith”: use “lia” rather than “omega”
| -rw-r--r-- | doc/changelog/04-tactics/11018-lia-in-auto-with-zarith.rst | 7 | ||||
| -rw-r--r-- | theories/omega/Omega.v | 51 |
2 files changed, 33 insertions, 25 deletions
diff --git a/doc/changelog/04-tactics/11018-lia-in-auto-with-zarith.rst b/doc/changelog/04-tactics/11018-lia-in-auto-with-zarith.rst new file mode 100644 index 0000000000..d510416990 --- /dev/null +++ b/doc/changelog/04-tactics/11018-lia-in-auto-with-zarith.rst @@ -0,0 +1,7 @@ +- **Changed:** The :g:`auto with zarith` tactic and variations (including :tacn:`intuition`) + may now call the :tacn:`lia` tactic instead of :tacn:`omega` + (when the `Omega` module is loaded); + more goals may be automatically solved, + fewer section variables will be captured spuriously + (`#11018 <https://github.com/coq/coq/pull/11018>`_, + by Vincent Laporte). diff --git a/theories/omega/Omega.v b/theories/omega/Omega.v index 9c2e8a9212..10a5aa47b3 100644 --- a/theories/omega/Omega.v +++ b/theories/omega/Omega.v @@ -19,6 +19,7 @@ Require Export ZArith_base. Require Export OmegaLemmas. Require Export PreOmega. +Require Import Lia. Declare ML Module "omega_plugin". @@ -28,28 +29,28 @@ Hint Resolve Z.le_refl Z.add_comm Z.add_assoc Z.mul_comm Z.mul_assoc Z.add_0_l Require Export Zhints. -Hint Extern 10 (_ = _ :>nat) => abstract omega: zarith. -Hint Extern 10 (_ <= _) => abstract omega: zarith. -Hint Extern 10 (_ < _) => abstract omega: zarith. -Hint Extern 10 (_ >= _) => abstract omega: zarith. -Hint Extern 10 (_ > _) => abstract omega: zarith. - -Hint Extern 10 (_ <> _ :>nat) => abstract omega: zarith. -Hint Extern 10 (~ _ <= _) => abstract omega: zarith. -Hint Extern 10 (~ _ < _) => abstract omega: zarith. -Hint Extern 10 (~ _ >= _) => abstract omega: zarith. -Hint Extern 10 (~ _ > _) => abstract omega: zarith. - -Hint Extern 10 (_ = _ :>Z) => abstract omega: zarith. -Hint Extern 10 (_ <= _)%Z => abstract omega: zarith. -Hint Extern 10 (_ < _)%Z => abstract omega: zarith. -Hint Extern 10 (_ >= _)%Z => abstract omega: zarith. -Hint Extern 10 (_ > _)%Z => abstract omega: zarith. - -Hint Extern 10 (_ <> _ :>Z) => abstract omega: zarith. -Hint Extern 10 (~ (_ <= _)%Z) => abstract omega: zarith. -Hint Extern 10 (~ (_ < _)%Z) => abstract omega: zarith. -Hint Extern 10 (~ (_ >= _)%Z) => abstract omega: zarith. -Hint Extern 10 (~ (_ > _)%Z) => abstract omega: zarith. - -Hint Extern 10 False => abstract omega: zarith. +Hint Extern 10 (_ = _ :>nat) => abstract lia: zarith. +Hint Extern 10 (_ <= _) => abstract lia: zarith. +Hint Extern 10 (_ < _) => abstract lia: zarith. +Hint Extern 10 (_ >= _) => abstract lia: zarith. +Hint Extern 10 (_ > _) => abstract lia: zarith. + +Hint Extern 10 (_ <> _ :>nat) => abstract lia: zarith. +Hint Extern 10 (~ _ <= _) => abstract lia: zarith. +Hint Extern 10 (~ _ < _) => abstract lia: zarith. +Hint Extern 10 (~ _ >= _) => abstract lia: zarith. +Hint Extern 10 (~ _ > _) => abstract lia: zarith. + +Hint Extern 10 (_ = _ :>Z) => abstract lia: zarith. +Hint Extern 10 (_ <= _)%Z => abstract lia: zarith. +Hint Extern 10 (_ < _)%Z => abstract lia: zarith. +Hint Extern 10 (_ >= _)%Z => abstract lia: zarith. +Hint Extern 10 (_ > _)%Z => abstract lia: zarith. + +Hint Extern 10 (_ <> _ :>Z) => abstract lia: zarith. +Hint Extern 10 (~ (_ <= _)%Z) => abstract lia: zarith. +Hint Extern 10 (~ (_ < _)%Z) => abstract lia: zarith. +Hint Extern 10 (~ (_ >= _)%Z) => abstract lia: zarith. +Hint Extern 10 (~ (_ > _)%Z) => abstract lia: zarith. + +Hint Extern 10 False => abstract lia: zarith. |
