aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorcoqbot-app[bot]2021-01-25 15:31:49 +0000
committerGitHub2021-01-25 15:31:49 +0000
commitcb5e21268ca0da4d26c3ea7e129b1af89d799a8c (patch)
tree39640c2c5408ee11c15da1fdb2a3033754d24d5e
parent0a6444c522c18c634fe1030436ea82f326bada9a (diff)
parenta631b260c54e78faadc5dfc1ef3d319f19e1b615 (diff)
Merge PR #13779: Properly implement local references in Summary.
Reviewed-by: gares
-rw-r--r--library/summary.ml25
-rw-r--r--library/summary.mli2
-rwxr-xr-xtest-suite/misc/non-marshalable-state.sh30
-rw-r--r--test-suite/misc/non-marshalable-state/_CoqProject9
-rw-r--r--test-suite/misc/non-marshalable-state/src/evil.mlg15
-rw-r--r--test-suite/misc/non-marshalable-state/src/evil_plugin.mlpack1
-rw-r--r--test-suite/misc/non-marshalable-state/src/good.mlg16
-rw-r--r--test-suite/misc/non-marshalable-state/src/good_plugin.mlpack1
-rw-r--r--test-suite/misc/non-marshalable-state/theories/evil.v5
-rw-r--r--test-suite/misc/non-marshalable-state/theories/good.v5
10 files changed, 95 insertions, 14 deletions
diff --git a/library/summary.ml b/library/summary.ml
index 221ac868fa..572467ada3 100644
--- a/library/summary.ml
+++ b/library/summary.ml
@@ -131,28 +131,27 @@ let ref ?freeze ~name x = fst @@ ref_tag ?freeze ~name x
module Local = struct
-type 'a local_ref = ('a CEphemeron.key * 'a Dyn.tag) ref
+type 'a local_ref = 'a CEphemeron.key ref * 'a CEphemeron.key Dyn.tag
-let set r v = r := (CEphemeron.create v, snd !r)
+let set (r, tag) v = r := CEphemeron.create v
-let get r =
- let key, name = !r in
- try CEphemeron.get key
+let get (key, name) =
+ try CEphemeron.get !key
with CEphemeron.InvalidKey ->
let { init_function } = DynMap.find name !sum_map in
init_function ();
- CEphemeron.get (fst !r)
+ CEphemeron.get !key
-let ref ?(freeze=fun x -> x) ~name init =
+let ref (type a) ~name (init : a) : a local_ref =
let () = check_name (mangle name) in
- let tag : 'a Dyn.tag = Dyn.create (mangle name) in
- let r = pervasives_ref (CEphemeron.create init, tag) in
+ let tag : a CEphemeron.key Dyn.tag = Dyn.create (mangle name) in
+ let r = pervasives_ref (CEphemeron.create init) in
let () = sum_map := DynMap.add tag
- { freeze_function = (fun ~marshallable -> freeze (get r));
- unfreeze_function = (set r);
- init_function = (fun () -> set r init) } !sum_map
+ { freeze_function = (fun ~marshallable -> !r);
+ unfreeze_function = (fun v -> r := v);
+ init_function = (fun () -> r := CEphemeron.create init) } !sum_map
in
- r
+ (r, tag)
let (!) = get
let (:=) = set
diff --git a/library/summary.mli b/library/summary.mli
index 7c5e1bee6f..a6f94a49ae 100644
--- a/library/summary.mli
+++ b/library/summary.mli
@@ -53,7 +53,7 @@ val ref_tag : ?freeze:(marshallable:bool -> 'a -> 'a) -> name:string -> 'a -> 'a
module Local : sig
type 'a local_ref
- val ref : ?freeze:('a -> 'a) -> name:string -> 'a -> 'a local_ref
+ val ref : name:string -> 'a -> 'a local_ref
val (:=) : 'a local_ref -> 'a -> unit
val (!) : 'a local_ref -> 'a
diff --git a/test-suite/misc/non-marshalable-state.sh b/test-suite/misc/non-marshalable-state.sh
new file mode 100755
index 0000000000..eef7786ebc
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state.sh
@@ -0,0 +1,30 @@
+#!/usr/bin/env bash
+
+set -e
+
+export COQBIN=$BIN
+export PATH=$COQBIN:$PATH
+
+cd misc/non-marshalable-state/
+
+coq_makefile -f _CoqProject -o Makefile
+
+make clean
+
+make src/evil_plugin.cmxs
+make src/good_plugin.cmxs
+
+RC=1
+# must fail
+coqc -async-proofs on -I src -Q theories Marshal theories/evil.v 2> log1 1>&2 || RC=0
+# for this reason
+grep -q 'Marshalling error' log1 || RC=1
+
+# must work
+coqc -async-proofs off -I src -Q theories Marshal theories/evil.v
+
+# must work
+coqc -async-proofs on -I src -Q theories Marshal theories/good.v
+
+
+exit $RC
diff --git a/test-suite/misc/non-marshalable-state/_CoqProject b/test-suite/misc/non-marshalable-state/_CoqProject
new file mode 100644
index 0000000000..09e68d866c
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/_CoqProject
@@ -0,0 +1,9 @@
+-Q theories Marshal
+-I src
+
+src/evil.mlg
+src/good.mlg
+src/evil_plugin.mlpack
+src/good_plugin.mlpack
+theories/evil.v
+theories/good.v
diff --git a/test-suite/misc/non-marshalable-state/src/evil.mlg b/test-suite/misc/non-marshalable-state/src/evil.mlg
new file mode 100644
index 0000000000..59b2b5a8ac
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/src/evil.mlg
@@ -0,0 +1,15 @@
+DECLARE PLUGIN "evil_plugin"
+
+{
+
+let state = Summary.ref
+ ~name:"elpi-compiler-cache"
+ None
+
+}
+
+VERNAC COMMAND EXTEND magic CLASSIFIED AS SIDEFF
+| [ "magic" ] -> {
+ state := Some (fun () -> ())
+}
+END
diff --git a/test-suite/misc/non-marshalable-state/src/evil_plugin.mlpack b/test-suite/misc/non-marshalable-state/src/evil_plugin.mlpack
new file mode 100644
index 0000000000..6382aa69e1
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/src/evil_plugin.mlpack
@@ -0,0 +1 @@
+Evil
diff --git a/test-suite/misc/non-marshalable-state/src/good.mlg b/test-suite/misc/non-marshalable-state/src/good.mlg
new file mode 100644
index 0000000000..c6b9cbefd5
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/src/good.mlg
@@ -0,0 +1,16 @@
+DECLARE PLUGIN "good_plugin"
+
+{
+
+let state = Summary.Local.ref
+ ~name:"elpi-compiler-cache"
+ None
+
+}
+
+VERNAC COMMAND EXTEND magic CLASSIFIED AS SIDEFF
+| [ "magic" ] -> {
+ let open Summary.Local in
+ state := Some (fun () -> ())
+}
+END
diff --git a/test-suite/misc/non-marshalable-state/src/good_plugin.mlpack b/test-suite/misc/non-marshalable-state/src/good_plugin.mlpack
new file mode 100644
index 0000000000..cd9dd73b78
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/src/good_plugin.mlpack
@@ -0,0 +1 @@
+Good
diff --git a/test-suite/misc/non-marshalable-state/theories/evil.v b/test-suite/misc/non-marshalable-state/theories/evil.v
new file mode 100644
index 0000000000..661482a975
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/theories/evil.v
@@ -0,0 +1,5 @@
+Declare ML Module "evil_plugin".
+
+magic.
+
+Lemma x : True. Proof. trivial. Qed.
diff --git a/test-suite/misc/non-marshalable-state/theories/good.v b/test-suite/misc/non-marshalable-state/theories/good.v
new file mode 100644
index 0000000000..eab9a043e1
--- /dev/null
+++ b/test-suite/misc/non-marshalable-state/theories/good.v
@@ -0,0 +1,5 @@
+Declare ML Module "good_plugin".
+
+magic.
+
+Lemma x : True. Proof. trivial. Qed.