aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/nativecode.ml
diff options
context:
space:
mode:
Diffstat (limited to 'kernel/nativecode.ml')
-rw-r--r--kernel/nativecode.ml41
1 files changed, 22 insertions, 19 deletions
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 096aa7066..ad5b04f3d 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -5,10 +5,10 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Errors
+
+open CErrors
open Names
open Term
-open Context
open Declarations
open Util
open Nativevalues
@@ -1841,31 +1841,32 @@ and apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml =
in
let auxdefs = List.fold_right get_rel_val fv_rel auxdefs in
let auxdefs = List.fold_right get_named_val fv_named auxdefs in
- let lvl = rel_context_length env.env_rel_context in
+ let lvl = Context.Rel.length env.env_rel_context in
let fv_rel = List.map (fun (n,_) -> MLglobal (Grel (lvl-n))) fv_rel in
let fv_named = List.map (fun (id,_) -> MLglobal (Gnamed id)) fv_named in
let aux_name = fresh_lname Anonymous in
auxdefs, MLlet(aux_name, ml, mkMLapp (MLlocal aux_name) (Array.of_list (fv_rel@fv_named)))
and compile_rel env sigma univ auxdefs n =
- let (_,body,_) = lookup_rel n env.env_rel_context in
- let n = rel_context_length env.env_rel_context - n in
- match body with
- | Some t ->
+ let open Context.Rel in
+ let n = length env.env_rel_context - n in
+ let open Declaration in
+ match lookup n env.env_rel_context with
+ | LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in
Glet(Grel n, code)::auxdefs
- | None ->
+ | LocalAssum _ ->
Glet(Grel n, MLprimitive (Mk_rel n))::auxdefs
and compile_named env sigma univ auxdefs id =
- let (_,body,_) = lookup_named id env.env_named_context in
- match body with
- | Some t ->
+ let open Context.Named.Declaration in
+ match Context.Named.lookup id env.env_named_context with
+ | LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in
Glet(Gnamed id, code)::auxdefs
- | None ->
+ | LocalAssum _ ->
Glet(Gnamed id, MLprimitive (Mk_var id))::auxdefs
let compile_constant env sigma prefix ~interactive con cb =
@@ -1879,7 +1880,7 @@ let compile_constant env sigma prefix ~interactive con cb =
| Def t ->
let t = Mod_subst.force_constr t in
let code = lambda_of_constr env sigma t in
- if !Flags.debug then Pp.msg_debug (Pp.str "Generated lambda code");
+ if !Flags.debug then Feedback.msg_debug (Pp.str "Generated lambda code");
let is_lazy = is_lazy prefix t in
let code = if is_lazy then mk_lazy code else code in
let name =
@@ -1894,11 +1895,11 @@ let compile_constant env sigma prefix ~interactive con cb =
let (auxdefs,code) = compile_with_fv env sigma (Some univ) [] (Some l) code in
(auxdefs,mkMLlam [|univ|] code)
in
- if !Flags.debug then Pp.msg_debug (Pp.str "Generated mllambda code");
+ if !Flags.debug then Feedback.msg_debug (Pp.str "Generated mllambda code");
let code =
optimize_stk (Glet(Gconstant ("",(con,u)),code)::auxdefs)
in
- if !Flags.debug then Pp.msg_debug (Pp.str "Optimized mllambda code");
+ if !Flags.debug then Feedback.msg_debug (Pp.str "Optimized mllambda code");
code, name
| _ ->
let i = push_symbol (SymbConst con) in
@@ -1945,13 +1946,15 @@ let compile_constant env sigma prefix ~interactive con cb =
arg|])))::
[Glet(gn, mkMLlam [|c_uid|] code)], Linked prefix
-let loaded_native_files = ref ([] : string list)
+module StringOrd = struct type t = string let compare = String.compare end
+module StringSet = Set.Make(StringOrd)
+
+let loaded_native_files = ref StringSet.empty
-let is_loaded_native_file s = String.List.mem s !loaded_native_files
+let is_loaded_native_file s = StringSet.mem s !loaded_native_files
let register_native_file s =
- if not (is_loaded_native_file s) then
- loaded_native_files := s :: !loaded_native_files
+ loaded_native_files := StringSet.add s !loaded_native_files
let is_code_loaded ~interactive name =
match !name with