aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4717.v
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-07-08 11:19:13 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-11-23 12:57:48 +0100
commit1b09c7b0802c85ea72931720a7cb4fbf9ab5e211 (patch)
treef8b7fdc5ae6216a8890f2e5a88d98816441b735a /test-suite/bugs/closed/4717.v
parent22ae762fa8940028f6a3d8a5fd4147d5ca3b53b9 (diff)
Using is_conv rather than eq_constr to find `nat` or `Z` in omega.
Moving at the same to a passing "env sigma" style rather than passing "gl". Not that it is strictly necessary, but since we had to move functions taking only a "sigma" to functions taking also a "env", we eventually adopted the "env sigma" style. (The "gl" style would have been as good.) This answers wish #4717.
Diffstat (limited to 'test-suite/bugs/closed/4717.v')
-rw-r--r--test-suite/bugs/closed/4717.v35
1 files changed, 35 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4717.v b/test-suite/bugs/closed/4717.v
new file mode 100644
index 000000000..4562ed1f1
--- /dev/null
+++ b/test-suite/bugs/closed/4717.v
@@ -0,0 +1,35 @@
+(* Omega being smarter on recognizing nat and Z *)
+
+Require Import Omega.
+
+Definition nat' := nat.
+
+Theorem le_not_eq_lt : forall (n m:nat),
+ n <= m ->
+ n <> m :> nat' ->
+ n < m.
+Proof.
+ intros.
+ omega.
+Qed.
+
+Goal forall (x n : nat'), x = x + n - n.
+Proof.
+ intros.
+ omega.
+Qed.
+
+Require Import ZArith.
+
+Open Scope Z_scope.
+
+Definition Z' := Z.
+
+Theorem Zle_not_eq_lt : forall n m,
+ n <= m ->
+ n <> m :> Z' ->
+ n < m.
+Proof.
+ intros.
+ omega.
+Qed.