aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2014-09-17 11:42:05 +0200
committerGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2014-09-17 12:37:02 +0200
commit9de1edd730eeb3cada742a27a36bc70178eda6d8 (patch)
tree341d93576b40d8fc4bb2e6abd6096a69df6b09a6 /test-suite/bugs/closed
parentd34e1eed232c84590ddb80d70db9f7f7cf13584a (diff)
While resolving typeclass evars in clenv, touch only the ones that appear in the
clenv's value and type, ensuring we don't try to solve unrelated goals (fixes bug#3633).
Diffstat (limited to 'test-suite/bugs/closed')
-rw-r--r--test-suite/bugs/closed/3633.v21
1 files changed, 21 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/3633.v b/test-suite/bugs/closed/3633.v
new file mode 100644
index 000000000..6f89ef6db
--- /dev/null
+++ b/test-suite/bugs/closed/3633.v
@@ -0,0 +1,21 @@
+Set Typeclasses Strict Resolution.
+Class Contr (A : Type) := { center : A }.
+Definition foo {A} `{Contr A} : A.
+Proof.
+ apply center.
+ Undo.
+ (* Ensure the constraints are solved indepentently, otherwise a frozen ?A
+ makes a search for Contr ?A fail when finishing to apply (fun x => x) *)
+ apply (fun x => x), center.
+
+(* Toplevel input, characters 7-17:
+Error:
+Unable to satisfy the following constraints:
+In environment:
+A : Type
+H : Contr A
+
+?A : "Type"
+?Contr : "Contr ?X8"
+
+ *) \ No newline at end of file