aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4429.v
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2015-11-19 18:40:32 +0100
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2015-11-19 20:16:45 +0100
commit574e510ba069f1747ecb1e5a17cf86c902d79d44 (patch)
tree694fa60567202af7e66f71b5a9fb727d2753dc24 /test-suite/bugs/closed/4429.v
parentcfc0fc0075784e75783c9b4482fd3f4b858a44bf (diff)
Fix bug #4429: eauto with arith: 70x performance regression in Coq 8.5.
The issue was due to the fact that unfold hints are given a priority of 4 by default. As eauto was now using hint priority rather than the number of goals produced to order the application of hints, unfold were almost always used too late. We fixed this by manually giving them a priority of 1 in the eauto tactic. Also fixed the relative order of proof depth w.r.t. hint priority. It should not be observable except for breadth-first search, which is seldom used.
Diffstat (limited to 'test-suite/bugs/closed/4429.v')
-rw-r--r--test-suite/bugs/closed/4429.v31
1 files changed, 31 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4429.v b/test-suite/bugs/closed/4429.v
new file mode 100644
index 000000000..bf0e570ab
--- /dev/null
+++ b/test-suite/bugs/closed/4429.v
@@ -0,0 +1,31 @@
+Require Import Arith.Compare_dec.
+Require Import Unicode.Utf8.
+
+Fixpoint my_nat_iter (n : nat) {A} (f : A → A) (x : A) : A :=
+ match n with
+ | O => x
+ | S n' => f (my_nat_iter n' f x)
+ end.
+
+Definition gcd_IT_F (f : nat * nat → nat) (mn : nat * nat) : nat :=
+ match mn with
+ | (0, 0) => 0
+ | (0, S n') => S n'
+ | (S m', 0) => S m'
+ | (S m', S n') =>
+ match le_gt_dec (S m') (S n') with
+ | left _ => f (S m', S n' - S m')
+ | right _ => f (S m' - S n', S n')
+ end
+ end.
+
+Axiom max_correct_l : ∀ m n : nat, m <= max m n.
+Axiom max_correct_r : ∀ m n : nat, n <= max m n.
+
+Hint Resolve max_correct_l max_correct_r : arith.
+
+Theorem foo : ∀ p p' p'' : nat, p'' < S (max p (max p' p'')).
+Proof.
+ intros.
+ Timeout 3 eauto with arith.
+Qed.