summaryrefslogtreecommitdiff
path: root/test-suite/bugs/opened/3948.v
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2016-12-27 16:53:30 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2016-12-27 16:53:30 +0100
commita4c7f8bd98be2a200489325ff7c5061cf80ab4f3 (patch)
tree26dd9c4aa142597ee09c887ef161d5f0fa5077b6 /test-suite/bugs/opened/3948.v
parent164c6861860e6b52818c031f901ffeff91fca16a (diff)
Imported Upstream version 8.6upstream/8.6
Diffstat (limited to 'test-suite/bugs/opened/3948.v')
-rw-r--r--test-suite/bugs/opened/3948.v25
1 files changed, 25 insertions, 0 deletions
diff --git a/test-suite/bugs/opened/3948.v b/test-suite/bugs/opened/3948.v
new file mode 100644
index 00000000..16581308
--- /dev/null
+++ b/test-suite/bugs/opened/3948.v
@@ -0,0 +1,25 @@
+Module Type S.
+Parameter t : Type.
+End S.
+
+Module Bar(X : S).
+Proof.
+ Definition elt := X.t.
+ Axiom fold : elt.
+End Bar.
+
+Module Make (X: S) := Bar(X).
+
+Declare Module X : S.
+
+Module Type Interface.
+ Parameter constant : unit.
+End Interface.
+
+Module DepMap : Interface.
+ Module Dom := Make(X).
+ Definition constant : unit :=
+ let _ := @Dom.fold in tt.
+End DepMap.
+
+Print Assumptions DepMap.constant. \ No newline at end of file