aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs
diff options
context:
space:
mode:
authorGravatar Pierre Letouzey <pierre.letouzey@inria.fr>2015-12-21 18:56:43 +0100
committerGravatar Pierre Letouzey <pierre.letouzey@inria.fr>2015-12-22 15:05:47 +0100
commit5122a39888cfc6afd2383d59465324dd67b69f4a (patch)
treeecfecd5ee4f2d89b6d61ab3638a30dfce6048af2 /test-suite/bugs
parent840cefca77a48ad68641539cd26d8d2e8c9dc031 (diff)
Inclusion of functors with restricted signature is now forbidden (fix #3746)
The previous behavior was to include the interface of such a functor, possibly leading to the creation of unexpected axioms, see bug report #3746. In the case of non-functor module with restricted signature, we could simply refer to the original objects (strengthening), but for a functor, the inner objects have no existence yet. As said in the new error message, a simple workaround is hence to first instantiate the functor, then include the local instance: Module LocalInstance := Funct(Args). Include LocalInstance. By the way, the mod_type_alg field is now filled more systematically, cf new comments in declarations.mli. This way, we could use it to know whether a module had been given a restricted signature (via ":"). Earlier, some mod_type_alg were None in situations not handled by the extraction (MEapply of module type). Some code refactoring on the fly.
Diffstat (limited to 'test-suite/bugs')
-rw-r--r--test-suite/bugs/closed/3746.v92
1 files changed, 92 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/3746.v b/test-suite/bugs/closed/3746.v
new file mode 100644
index 000000000..a9463f94b
--- /dev/null
+++ b/test-suite/bugs/closed/3746.v
@@ -0,0 +1,92 @@
+
+(* Bug report #3746 : Include and restricted signature *)
+
+Module Type MT. Parameter p : nat. End MT.
+Module Type EMPTY. End EMPTY.
+Module Empty. End Empty.
+
+(* Include of an applied functor with restricted sig :
+ Used to create axioms (bug report #3746), now forbidden. *)
+
+Module F (X:EMPTY) : MT.
+ Definition p := 0.
+End F.
+
+Module InclFunctRestr.
+ Fail Include F(Empty).
+End InclFunctRestr.
+
+(* A few variants (indirect restricted signature), also forbidden. *)
+
+Module F1 := F.
+Module F2 (X:EMPTY) := F X.
+
+Module F3a (X:EMPTY). Definition p := 0. End F3a.
+Module F3 (X:EMPTY) : MT := F3a X.
+
+Module InclFunctRestrBis.
+ Fail Include F1(Empty).
+ Fail Include F2(Empty).
+ Fail Include F3(Empty).
+End InclFunctRestrBis.
+
+(* Recommended workaround: manual instance before the include. *)
+
+Module InclWorkaround.
+ Module Temp := F(Empty).
+ Include Temp.
+End InclWorkaround.
+
+Compute InclWorkaround.p.
+Print InclWorkaround.p.
+Print Assumptions InclWorkaround.p. (* Closed under the global context *)
+
+
+
+(* Related situations which are ok, just to check *)
+
+(* A) Include of non-functor with restricted signature :
+ creates a proxy to initial stuff *)
+
+Module M : MT.
+ Definition p := 0.
+End M.
+
+Module InclNonFunct.
+ Include M.
+End InclNonFunct.
+
+Definition check : InclNonFunct.p = M.p := eq_refl.
+Print Assumptions InclNonFunct.p. (* Closed *)
+
+
+(* B) Include of a module type with opaque content:
+ The opaque content is "copy-pasted". *)
+
+Module Type SigOpaque.
+ Definition p : nat. Proof. exact 0. Qed.
+End SigOpaque.
+
+Module InclSigOpaque.
+ Include SigOpaque.
+End InclSigOpaque.
+
+Compute InclSigOpaque.p.
+Print InclSigOpaque.p.
+Print Assumptions InclSigOpaque.p. (* Closed *)
+
+
+(* C) Include of an applied functor with opaque proofs :
+ opaque proof "copy-pasted" (and substituted). *)
+
+Module F' (X:EMPTY).
+ Definition p : nat. Proof. exact 0. Qed.
+End F'.
+
+Module InclFunctOpa.
+ Include F'(Empty).
+End InclFunctOpa.
+
+Compute InclFunctOpa.p.
+Print InclFunctOpa.p.
+Print Assumptions InclFunctOpa.p. (* Closed *)