aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-06-14 15:30:31 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-06-14 15:30:31 +0000
commit6bbce8f236e96fde855e3fa90c6334e89201f077 (patch)
tree6ec5359c82496bc14b8cdb2c2cf076c37f49f0a6
parenta70eb94dc94a41c9e931b7c13ac3ca254f0bcda9 (diff)
Regression files for bugs #2304 and #2490.
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14204 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2304.v4
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2490.v28
2 files changed, 32 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/shouldsucceed/2304.v b/test-suite/bugs/closed/shouldsucceed/2304.v
new file mode 100644
index 000000000..1ac2702b0
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2304.v
@@ -0,0 +1,4 @@
+(* This used to fail with an anomaly NotASort at some time *)
+Class A (O: Type): Type := a: O -> Type.
+Fail Goal forall (x: a tt), @a x = @a x.
+
diff --git a/test-suite/bugs/closed/shouldsucceed/2490.v b/test-suite/bugs/closed/shouldsucceed/2490.v
new file mode 100644
index 000000000..3efd742bb
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2490.v
@@ -0,0 +1,28 @@
+Class Rel A := rel : A -> Prop.
+Class Rel2 A := rel2 : A -> Prop.
+Class Property A (o : A -> Prop) := rel_property : True.
+
+Class Op A := op : A -> A.
+Instance my_op {A} {r : Rel A} {p : Property A r} : Op A := fun x => x.
+
+Section test.
+Context A (r2 : Rel2 A) (p2 : Property A r2).
+
+(* 8.3 yields an error: "Could not find an instance for "Op A" in environment".
+ Trunk, on the other hand, uses [r2] as an instance of [Rel] and hence finds
+the previously defined instance of [Op]. Trunk's behavior is incorrect. *)
+Lemma test (x : A) : op x = op x.
+Admitted.
+
+(* If we don't refer to the canonical name, Coq will actually complain, as it
+should. *)
+Lemma test (x : A) : my_op x = my_op x.
+
+End test.
+
+(* Now we make an instance of [Op] without dependent condition. *)
+Instance my_op_nondep {A} {r : Rel A} : Op A := fun x => x.
+
+Section test2.
+Context A (r2 : Rel2 A).
+