aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2019-12-24 13:23:14 +0100
committerEmilio Jesus Gallego Arias2019-12-26 12:27:17 +0100
commite17ee7629006dd119393727a3203bc6e602d5473 (patch)
tree4d52aabee98537c35ac70d9d926d742a80d28332
parentfeea1d0377e2fb083efe74cd241e7867d008d5be (diff)
Deprecate the "omega with *" syntax.
Changes to the test-suite were backported from PR #11288.
-rw-r--r--doc/changelog/04-tactics/11337-omega-with-depr.rst6
-rw-r--r--plugins/omega/g_omega.mlg5
-rw-r--r--test-suite/success/Omega.v2
-rw-r--r--test-suite/success/OmegaPre.v40
4 files changed, 31 insertions, 22 deletions
diff --git a/doc/changelog/04-tactics/11337-omega-with-depr.rst b/doc/changelog/04-tactics/11337-omega-with-depr.rst
new file mode 100644
index 0000000000..25e929e030
--- /dev/null
+++ b/doc/changelog/04-tactics/11337-omega-with-depr.rst
@@ -0,0 +1,6 @@
+- **Deprecated:**
+ The undocumented ``omega with`` tactic variant has been deprecated,
+ using ``lia`` is the recommended replacement, tho the old semantics
+ of ``omega with *`` can be recovered with ``zify; omega``
+ (`#11337 <https://github.com/coq/coq/pull/11337>`_,
+ by Emilio Jesus Gallego Arias).
diff --git a/plugins/omega/g_omega.mlg b/plugins/omega/g_omega.mlg
index 84964a7bd2..a464aa7287 100644
--- a/plugins/omega/g_omega.mlg
+++ b/plugins/omega/g_omega.mlg
@@ -45,13 +45,16 @@ let omega_tactic l =
(Tacticals.New.tclREPEAT (Tacticals.New.tclTHENLIST tacs))
(omega_solver)
+let omega_with_deprecation =
+ Deprecation.make ~since:"8.11.0" ~note:"Use lia instead." ()
+
}
TACTIC EXTEND omega
| [ "omega" ] -> { omega_tactic [] }
END
-TACTIC EXTEND omega'
+TACTIC EXTEND omega' DEPRECATED { omega_with_deprecation }
| [ "omega" "with" ne_ident_list(l) ] ->
{ omega_tactic (List.map Names.Id.to_string l) }
| [ "omega" "with" "*" ] ->
diff --git a/test-suite/success/Omega.v b/test-suite/success/Omega.v
index 470e4f0580..5e0f90d59b 100644
--- a/test-suite/success/Omega.v
+++ b/test-suite/success/Omega.v
@@ -90,5 +90,5 @@ Qed.
(* Postponed... problem with goals of the form "(n*m=0)%nat -> (n*m=0)%Z" *)
Lemma lem10 : forall n m:nat, le n (plus n (mult n m)).
Proof.
-intros; omega with *.
+intros; zify; omega.
Qed.
diff --git a/test-suite/success/OmegaPre.v b/test-suite/success/OmegaPre.v
index 17531064cc..0223255067 100644
--- a/test-suite/success/OmegaPre.v
+++ b/test-suite/success/OmegaPre.v
@@ -16,112 +16,112 @@ Open Scope Z_scope.
Goal forall a:Z, Z.max a a = a.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall a b:Z, Z.max a b = Z.max b a.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall a b c:Z, Z.max a (Z.max b c) = Z.max (Z.max a b) c.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall a b:Z, Z.max a b + Z.min a b = a + b.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall a:Z, (Z.abs a)*(Z.sgn a) = a.
intros.
zify.
-intuition; subst; omega. (* pure multiplication: omega alone can't do it *)
+intuition; subst; zify; omega. (* pure multiplication: zify; omega alone can't do it *)
Qed.
Goal forall a:Z, Z.abs a = a -> a >= 0.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall a:Z, Z.sgn a = a -> a = 1 \/ a = 0 \/ a = -1.
intros.
-omega with *.
+zify; omega.
Qed.
(* zify_nat *)
Goal forall m: nat, (m<2)%nat -> (0<= m+m <=2)%nat.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m:nat, (m<1)%nat -> (m=0)%nat.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m: nat, (m<=100)%nat -> (0<= m+m <=200)%nat.
intros.
-omega with *.
+zify; omega.
Qed.
(* 2000 instead of 200: works, but quite slow *)
Goal forall m: nat, (m*m>=0)%nat.
intros.
-omega with *.
+zify; omega.
Qed.
(* zify_positive *)
Goal forall m: positive, (m<2)%positive -> (2 <= m+m /\ m+m <= 2)%positive.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m:positive, (m<2)%positive -> (m=1)%positive.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m: positive, (m<=1000)%positive -> (2<=m+m/\m+m <=2000)%positive.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m: positive, (m*m>=1)%positive.
intros.
-omega with *.
+zify; omega.
Qed.
(* zify_N *)
Goal forall m:N, (m<2)%N -> (0 <= m+m /\ m+m <= 2)%N.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m:N, (m<1)%N -> (m=0)%N.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m:N, (m<=1000)%N -> (0<=m+m/\m+m <=2000)%N.
intros.
-omega with *.
+zify; omega.
Qed.
Goal forall m:N, (m*m>=0)%N.
intros.
-omega with *.
+zify; omega.
Qed.
(* mix of datatypes *)
Goal forall p, Z.of_N (N.of_nat (N.to_nat (Npos p))) = Zpos p.
intros.
-omega with *.
+zify; omega.
Qed.