aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4121.v
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2015-04-09 17:49:42 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2015-04-09 17:50:32 +0200
commit6158ec51adc31814fde0293f54151c19a5f3b1e4 (patch)
tree284a751d8ded260dce6dba330e42ee20be9b0d91 /test-suite/bugs/closed/4121.v
parent968c8af13deaa3871aca7fc7086f1a5dc5769fde (diff)
Fix declarations of instances to perform restriction of universe
instances as definitions and lemmas do (fixes bug# 4121).
Diffstat (limited to 'test-suite/bugs/closed/4121.v')
-rw-r--r--test-suite/bugs/closed/4121.v15
1 files changed, 15 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4121.v b/test-suite/bugs/closed/4121.v
new file mode 100644
index 000000000..5f8c411ca
--- /dev/null
+++ b/test-suite/bugs/closed/4121.v
@@ -0,0 +1,15 @@
+(* -*- coq-prog-args: ("-emacs" "-nois") -*- *)
+(* File reduced by coq-bug-finder from original input, then from 830 lines to 47 lines, then from 25 lines to 11 lines *)
+(* coqc version 8.5beta1 (March 2015) compiled on Mar 11 2015 18:51:36 with OCaml 4.01.0
+ coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (8dbfee5c5f897af8186cb1bdfb04fd4f88eca677) *)
+
+Set Universe Polymorphism.
+Class Contr_internal (A : Type) := BuildContr { center : A }.
+Arguments center A {_}.
+Class Contr (A : Type) : Type := Contr_is_trunc : Contr_internal A.
+Hint Extern 0 => progress change Contr_internal with Contr in * : typeclass_instances.
+Definition contr_paths_contr0 {A} `{Contr A} : Contr A := {| center := center A |}.
+Instance contr_paths_contr1 {A} `{Contr A} : Contr A := {| center := center A |}.
+Check @contr_paths_contr0@{i}.
+Check @contr_paths_contr1@{i}. (* Error: Universe instance should have length 2 *)
+(** It should have length 1, just like contr_paths_contr0 *) \ No newline at end of file