aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/success/ROmega2.v
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2007-07-09 13:49:57 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2007-07-09 13:49:57 +0000
commit9c4f65ebe124d972ae1506e388bb206cb7298ce4 (patch)
tree0aaaa3021fb1597c5a9e553dd03c46b563bc105d /test-suite/success/ROmega2.v
parent2ea3dc4db81e6513810da086a65f9c8292d4bebf (diff)
Improvements / Bug fixes for ROmega
----------------------------------- All romega tests in the test-suite are now bug-free. The only known remaining limitation of romega with respect to omega is that it cannot handle stuff on nat. * Equivalences A<->B are now understood by romega (as well as omega), and seen as (A->B)/\(B->A). There might be a smarter way to procede, for instance having a primitive Iff construct and trying to break equivalences as late as possible. * Conclusion-as-Pprop issue: After the resolution by the abstract omega machinery, useless parts are discarded from the reification process by replacing them with Pprop construct (see really_useful_prop). This allow to decrease the size of the proof terms and speed up their normalisation, I guess. But when such Pprop are created in the conclusion, this leads to failure, since concl is negated, and this is donc only if it is decidable. And introducing some Pprop might change the decidability status of the concl: for instance, Pfalse is decidable, whereas Pprop False is considered as _not_ decidable. Quick fix: no more really_useful_prop applied on concl (needs careful computation of useful_var). * NEGATE_CONTRADICT(_INV): This trace instrution comes in fact in two flavors, according to a boolean flag. We now translate to O_NEGATE_CONTRADICT_INV if this flag is false. (fix Besson's bug #1298) * EXACT_DIVIDE: could be used on NeqTerm and not only on EqTerm. * h_step indexes: The abstract omega machinery can introduce new hyps. In the list of hyps, they appears _before_ the regular one (but after the goal seen as an hyp by negating it). But the normalization steps were applied to regular hyps thanks to their indexes counted _before_ the addition extra hyps. * extra hyps (a)normal forms: extra hyps and variables are initially of the shape poly(v1,...,v(n-1)) = vn but O_STATE was expecting them in form 0 = poly(...) + -vn (by the way, SPLIT_INEQ should be checked someday). Since the above is one weekend's worth of debugging, there might well remain some more bugs :-(. For the record, here's the less painful way to debug a failed romega run: - activate debug flag in omega.ml and refl_omega.ml - at the bottom of refl_omega, replace normalise_vm_in_concl with convert_no_check (see comment there): this allow to skip the usually _huge_ error message about "Impossible to unify True with ..." - run the romega - try to run Qed, and enjoy the nice errror message about a (omega_tactic ? ? ? ?) that should be reducible to True. Here starts the real debug work... git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@9962 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'test-suite/success/ROmega2.v')
-rw-r--r--test-suite/success/ROmega2.v19
1 files changed, 16 insertions, 3 deletions
diff --git a/test-suite/success/ROmega2.v b/test-suite/success/ROmega2.v
index 9d47c9f65..a3be2898c 100644
--- a/test-suite/success/ROmega2.v
+++ b/test-suite/success/ROmega2.v
@@ -4,6 +4,20 @@ Require Import ZArith ROmega.
Open Scope Z_scope.
+
+(* First a simplified version used during debug of romega on Test46 *)
+Lemma Test46_simplified :
+forall v1 v2 v5 : Z,
+0 = v2 + v5 ->
+0 < v5 ->
+0 < v2 ->
+4*v2 <> 5*v1.
+intros.
+romega.
+Qed.
+
+
+(* The complete problem *)
Lemma Test46 :
forall v1 v2 v3 v4 v5 : Z,
((2 * v4) + (5)) + (8 * v2) <= ((4 * v4) + (3 * v4)) + (5 * v4) ->
@@ -23,6 +37,5 @@ forall v1 v2 v3 v4 v5 : Z,
((7 * v1) + (1 * v3)) + ((2 * v3) + (1 * v3)) >= ((6 * v5) + (4)) + ((1) + (9))
-> False.
intros.
-(*romega.*)
-Admitted.
-
+romega.
+Qed.