summaryrefslogtreecommitdiff
path: root/test-suite/bugs/closed/2083.v
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
commit7cfc4e5146be5666419451bdd516f1f3f264d24a (patch)
treee4197645da03dc3c7cc84e434cc31d0a0cca7056 /test-suite/bugs/closed/2083.v
parent420f78b2caeaaddc6fe484565b2d0e49c66888e5 (diff)
Imported Upstream version 8.5~beta1+dfsg
Diffstat (limited to 'test-suite/bugs/closed/2083.v')
-rw-r--r--test-suite/bugs/closed/2083.v27
1 files changed, 27 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/2083.v b/test-suite/bugs/closed/2083.v
new file mode 100644
index 00000000..5f17f7af
--- /dev/null
+++ b/test-suite/bugs/closed/2083.v
@@ -0,0 +1,27 @@
+Require Import Program Arith.
+
+Program Fixpoint check_n (n : nat) (P : { i | i < n } -> bool) (p : nat)
+ (H : forall (i : { i | i < n }), i < p -> P i = true)
+ {measure (n - p)} :
+ Exc (forall (p : { i | i < n}), P p = true) :=
+ match le_lt_dec n p with
+ | left _ => value _
+ | right cmp =>
+ if dec (P p) then
+ check_n n P (S p) _
+ else
+ error
+ end.
+
+Require Import Omega.
+
+Solve Obligations with program_simpl ; auto with *; try omega.
+
+Next Obligation.
+ apply H. simpl. omega.
+Defined.
+
+Next Obligation.
+ case (le_lt_dec p i) ; intros. assert(i = p) by omega. subst.
+ revert H0. clear_subset_proofs. auto.
+ apply H. simpl. assumption. Defined.