aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/csymtable.ml
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <Enrico.Tassi@inria.fr>2014-01-06 15:13:32 +0100
committerGravatar Enrico Tassi <Enrico.Tassi@inria.fr>2014-01-06 16:54:15 +0100
commitc44a12ab9136b8d13cec0d9c1f3837b6f92eb283 (patch)
tree75c4334f93dacae8e7eaf93691f448718d5997b3 /kernel/csymtable.ml
parent2f6e3a8a453c3fa29bbc660a929c5140916c76b3 (diff)
STM: fix worker crash when doing vm_compute
Kudos to Maximes for finding the culprit in no time! Values of type 'Pre_env.key' store in the OCaml state the 'address' of an already evaluated constant in the VM's C state. Such values are not sent to work processes. The worker is going to re-evaluate the constant, but just once, since the cache is cleared only when the env is marshalled (via ephemerons).
Diffstat (limited to 'kernel/csymtable.ml')
-rw-r--r--kernel/csymtable.ml8
1 files changed, 5 insertions, 3 deletions
diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml
index 0111cf74d..b606384ad 100644
--- a/kernel/csymtable.ml
+++ b/kernel/csymtable.ml
@@ -68,8 +68,10 @@ exception NotEvaluated
let key rk =
match !rk with
- | Some k -> (*Pp.msgnl (str"found at: "++int k);*) k
- | _ -> raise NotEvaluated
+ | None -> raise NotEvaluated
+ | Some k -> (*Pp.msgnl (str"found at: "++int k);*)
+ try Ephemeron.get k
+ with Ephemeron.InvalidKey -> raise NotEvaluated
(************************)
(* traduction des patch *)
@@ -104,7 +106,7 @@ let rec slot_for_getglobal env kn =
| BCallias kn' -> slot_for_getglobal env kn'
| BCconstant -> set_global (val_of_constant kn) in
(*Pp.msgnl(str"value stored at: "++int pos);*)
- rk := Some pos;
+ rk := Some (Ephemeron.create pos);
pos
and slot_for_fv env fv =