aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4628.v
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2016-05-23 18:02:08 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2016-05-23 18:20:32 +0200
commit265be83a546b0bec6d01f6650f7489442293cb0e (patch)
tree86878969836c43d570b2ba4f307b229c60d60935 /test-suite/bugs/closed/4628.v
parentbf1a1070d6cd111385baf59569feea2e0db3eb7c (diff)
Hints/Univs: fix bug #4628 anomalies
Fix handling of non-polymorphic hints built from polymorphic values, or simply producing new universes. We have to record the side effects of global hints built from constrs which are not polymorphic when they declare global universes, which might need to be discharged at the end of sections too. Also issue a warning when a Hint is declared for a polymorphic reference but the Hint isn't polymorphic itself (this used to produce an anomaly). For [using] hints, treat all lemmas as polymorphic, refreshing their universes at each use, as is done for their existentials (also used to produce an anomaly).
Diffstat (limited to 'test-suite/bugs/closed/4628.v')
-rw-r--r--test-suite/bugs/closed/4628.v46
1 files changed, 46 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4628.v b/test-suite/bugs/closed/4628.v
new file mode 100644
index 000000000..7d4a15d68
--- /dev/null
+++ b/test-suite/bugs/closed/4628.v
@@ -0,0 +1,46 @@
+Module first.
+ Polymorphic Record BAR (A:Type) :=
+ { foo: A->Prop; bar: forall (x y: A), foo x -> foo y}.
+
+Section A.
+Context {A:Type}.
+
+Set Printing Universes.
+
+Hint Resolve bar.
+Goal forall (P:BAR A) x y, foo _ P x -> foo _ P y.
+intros.
+eauto.
+Qed.
+End A.
+End first.
+
+Module firstbest.
+ Polymorphic Record BAR (A:Type) :=
+ { foo: A->Prop; bar: forall (x y: A), foo x -> foo y}.
+
+Section A.
+Context {A:Type}.
+
+Set Printing Universes.
+
+Polymorphic Hint Resolve bar.
+Goal forall (P:BAR A) x y, foo _ P x -> foo _ P y.
+intros.
+eauto.
+Qed.
+End A.
+End firstbest.
+
+Module second.
+Axiom foo: Set.
+Axiom foo': Set.
+
+Polymorphic Record BAR (A:Type) :=
+ { bar: foo' -> foo}.
+Set Printing Universes.
+
+Lemma baz@{i}: forall (P:BAR@{Set} nat), foo' -> foo.
+ eauto using bar.
+Qed.
+End second.