aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2017-11-23 14:28:21 +0100
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2018-02-14 16:14:10 +0100
commit8cd6ddb98c12b6aba002781158180ffb68aba02f (patch)
tree82a3b98a604149a8ae984cb25986e915d9d155a5 /kernel
parent41893cb647fbdce87b40acd5763e837370d61ece (diff)
Abstract further the type of VM bytecode compilation.
This reduces the possibility to wreak havoc while making the API nicer.
Diffstat (limited to 'kernel')
-rw-r--r--kernel/cemitcodes.ml7
-rw-r--r--kernel/cemitcodes.mli10
-rw-r--r--kernel/csymtable.ml11
3 files changed, 15 insertions, 13 deletions
diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml
index eeea19c12..bc6140ef4 100644
--- a/kernel/cemitcodes.ml
+++ b/kernel/cemitcodes.ml
@@ -22,6 +22,7 @@ type reloc_info =
| Reloc_getglobal of Names.Constant.t
type patch = reloc_info * int
+type patches = patch list
let patch_char4 buff pos c1 c2 c3 c4 =
Bytes.unsafe_set buff pos c1;
@@ -54,6 +55,12 @@ let patch_int buff patches =
*)
Bytes.unsafe_to_string buff
+let patch buff pl f =
+ let map (r, pos) = (pos, f r) in
+ let patches = CList.map_left map pl in
+ let buff = patch_int buff patches in
+ buff
+
(* Buffering of bytecode *)
let out_buffer = ref(Bytes.create 1024)
diff --git a/kernel/cemitcodes.mli b/kernel/cemitcodes.mli
index fee45aafd..a2ff2d301 100644
--- a/kernel/cemitcodes.mli
+++ b/kernel/cemitcodes.mli
@@ -6,18 +6,14 @@ type reloc_info =
| Reloc_const of structured_constant
| Reloc_getglobal of Constant.t
-type patch = reloc_info * int
-
-(* A virer *)
-val subst_patch : Mod_subst.substitution -> patch -> patch
-
+type patches
type emitcodes
val length : emitcodes -> int
-val patch_int : emitcodes -> ((*pos*)int * int) list -> emitcodes
+val patch : emitcodes -> patches -> (reloc_info -> int) -> emitcodes
-type to_patch = emitcodes * (patch list) * fv
+type to_patch = emitcodes * patches * fv
val subst_to_patch : Mod_subst.substitution -> to_patch -> to_patch
diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml
index bbd284bc1..38044071f 100644
--- a/kernel/csymtable.ml
+++ b/kernel/csymtable.ml
@@ -205,13 +205,12 @@ and slot_for_fv env fv =
assert false
and eval_to_patch env (buff,pl,fv) =
- let patch = function
- | Reloc_annot a, pos -> (pos, slot_for_annot a)
- | Reloc_const sc, pos -> (pos, slot_for_str_cst sc)
- | Reloc_getglobal kn, pos -> (pos, slot_for_getglobal env kn)
+ let slots = function
+ | Reloc_annot a -> slot_for_annot a
+ | Reloc_const sc -> slot_for_str_cst sc
+ | Reloc_getglobal kn -> slot_for_getglobal env kn
in
- let patches = List.map_left patch pl in
- let buff = patch_int buff patches in
+ let buff = patch buff pl slots in
let vm_env = Array.map (slot_for_fv env) fv in
let tc = tcode_of_code buff (length buff) in
eval_tcode tc vm_env