aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/3217.v
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2014-04-05 16:00:31 -0400
committerGravatar Pierre Boutillier <pierre.boutillier@ens-lyon.org>2014-04-10 14:42:17 +0200
commitf2b60f5911d7e881ff369663a7ed1f1bddfddb1d (patch)
tree944841ad03b5191d9bead4a9b5225a783eee578a /test-suite/bugs/closed/3217.v
parent3c87a1cb2ae89284c70c6b9333a88a06fc4d060f (diff)
Test case for bug #3217
It was fixed in c3feef4ed5dec126f1144dec91eee9c0f0522a94. The test case uses [Timeout 2] to test for "Coq runs instantaneously rather than running out of memory". Hopefully this is robust enough.
Diffstat (limited to 'test-suite/bugs/closed/3217.v')
-rw-r--r--test-suite/bugs/closed/3217.v36
1 files changed, 36 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/3217.v b/test-suite/bugs/closed/3217.v
new file mode 100644
index 000000000..ec846bf95
--- /dev/null
+++ b/test-suite/bugs/closed/3217.v
@@ -0,0 +1,36 @@
+(** [Set Implicit Arguments] causes Coq to run out of memory on [Qed] before c3feef4ed5dec126f1144dec91eee9c0f0522a94 *)
+Set Implicit Arguments.
+
+Variable LEM: forall P : Prop, sumbool P (P -> False).
+
+Definition pmap := option (nat -> option nat).
+
+Definition pmplus (oha ohb: pmap) : pmap :=
+ match oha, ohb with
+ | Some ha, Some hb =>
+ if LEM (oha = ohb) then None else None
+ | _, _ => None
+ end.
+
+Definition pmemp: pmap := Some (fun _ => None).
+
+Lemma foo:
+ True ->
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ (pmplus pmemp
+ pmemp))))))))))))
+ =
+ None -> True.
+Proof.
+ auto.
+Timeout 2 Qed.