package coq

  1. Overview
  2. Docs
Legend:
Page
Library
Module
Module type
Parameter
Class
Class type
Source

Source file vmsymtable.ml

1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *         Copyright INRIA, CNRS and contributors             *)
(* <O___,, * (see version control and CREDITS file for authors & dates) *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

(* Created by Bruno Barras for Benjamin Grégoire as part of the
   bytecode-based reduction machine, Oct 2004 *)
(* Bug fix #1419 by Jean-Marc Notin, Mar 2007 *)

(* This file manages the table of global symbols for the bytecode machine *)

open Util
open Names
open Vmvalues
open Vmemitcodes
open Vmbytecodes
open Declarations
open Environ
open Vmbytegen

module NamedDecl = Context.Named.Declaration
module RelDecl = Context.Rel.Declaration

type vm_global = values array

(* interpreter *)
external coq_interprete : tcode -> values -> atom array -> vm_global -> Vmvalues.vm_env -> int -> values =
  "coq_interprete_byte" "coq_interprete_ml"

(* table for structured constants and switch annotations *)

module HashMap (M : Hashtbl.HashedType) :
sig
  type +'a t
  val empty : 'a t
  val add : M.t -> 'a -> 'a t -> 'a t
  val find : M.t -> 'a t -> 'a
end =
struct
  type 'a t = (M.t * 'a) list Int.Map.t
  let empty = Int.Map.empty
  let add k v m =
    let hash = M.hash k in
    try Int.Map.modify hash (fun _ l -> (k, v) :: l) m
    with Not_found -> Int.Map.add hash [k, v] m
  let find k m =
    let hash = M.hash k in
    let l = Int.Map.find hash m in
    List.assoc_f M.equal k l
end

module SConstTable = HashMap (struct
  type t = structured_constant
  let equal = eq_structured_constant
  let hash = hash_structured_constant
end)

module AnnotTable = HashMap (struct
  type t = annot_switch
  let equal = eq_annot_switch
  let hash = hash_annot_switch
end)

module GlobVal :
sig
  type key = int
  type t
  val empty : int -> t
  val push : t -> values -> key * t
  val vm_global : t -> vm_global
  (** Warning: due to sharing, the resulting value can be modified by
      persistent operations. When calling this function one must ensure
      that no other context modification is performed before the value drops
      out of scope. *)
end =
struct
  type key = int

  type view =
  | Root of values array
  | DSet of int * values * view ref

  type t = {
    data : view ref;
    size : int;
    (** number of actual elements in the array *)
  }

  let empty n = {
    data = ref (Root (Array.make n crazy_val));
    size = 0;
  }

  let rec rerootk t k = match !t with
  | Root _ -> k ()
  | DSet (i, v, t') ->
    let next () = match !t' with
    | Root a as n ->
      let v' = Array.unsafe_get a i in
      let () = Array.unsafe_set a i v in
      let () = t := n in
      let () = t' := DSet (i, v', t) in
      k ()
    | DSet _ -> assert false
    in
    rerootk t' next

  let reroot t = rerootk t (fun () -> ())

  let push { data = t; size = i } v =
    let () = reroot t in
    match !t with
    | DSet _ -> assert false
    | Root a as n ->
      let len = Array.length a in
      let data =
        if i < len then
          let old = Array.unsafe_get a i in
          if old == v then t
          else
            let () = Array.unsafe_set a i v in
            let res = ref n in
            let () = t := DSet (i, old, res) in
            res
        else
          let nlen = 2 * len + 1 in
          let nlen = min nlen Sys.max_array_length in
          let () = assert (i < nlen) in
          let a' = Array.make nlen crazy_val in
          let () = Array.blit a 0 a' 0 len in
          let () = Array.unsafe_set a' i v in
          let res = ref (Root a') in
          let () = t := DSet (i, crazy_val, res) in
          res
      in
      (i, { data; size = i + 1 })

  let vm_global { data; _ } =
    let () = reroot data in
    match !data with
    | DSet _ -> assert false
    | Root a -> (a : vm_global)

end


(****************)
(* Code linking *)
(****************)

(* Global Table *)

(** [global_table] contains values of global constants, switch annotations,
    and structured constants. *)

type global_table = {
  glob_val : GlobVal.t;
  glob_sconst : GlobVal.key SConstTable.t;
  glob_annot : GlobVal.key AnnotTable.t;
  glob_prim : GlobVal.key array;
}

let get_global_data table = GlobVal.vm_global table.glob_val

let set_global v rtable =
  let table = !rtable in
  let (n, glob_val) = GlobVal.push table.glob_val v in
  let table = { table with glob_val } in
  let () = rtable := table in
  n

(*************************************************************)
(*** Mise a jour des valeurs des variables et des constantes *)
(*************************************************************)

exception NotEvaluated

let key rk =
  match !rk with
  | None -> raise NotEvaluated
  | Some k ->
      try CEphemeron.get k
      with CEphemeron.InvalidKey -> raise NotEvaluated

(************************)
(* traduction des patch *)

(* slot_for_*, calcul la valeur de l'objet, la place
   dans la table global, rend sa position dans la table *)

let slot_for_str_cst key table =
  try SConstTable.find key !table.glob_sconst
  with Not_found ->
    let n = set_global (val_of_str_const key) table in
    let glob_sconst = SConstTable.add key n !table.glob_sconst in
    let () = table := { !table with glob_sconst } in
    n

let slot_for_annot key table =
  try AnnotTable.find key !table.glob_annot
  with Not_found ->
    let n = set_global (val_of_annot_switch key) table in
    let glob_annot = AnnotTable.add key n !table.glob_annot in
    let () = table := { !table with glob_annot } in
    n

(* Initialization of OCaml primitives *)
let eval_caml_prim = function
| CAML_Arraymake -> 0, Vmvalues.parray_make
| CAML_Arrayget -> 1, Vmvalues.parray_get
| CAML_Arraydefault -> 2, Vmvalues.parray_get_default
| CAML_Arrayset -> 3, Vmvalues.parray_set
| CAML_Arraycopy -> 4, Vmvalues.parray_copy
| CAML_Arraylength -> 5, Vmvalues.parray_length

let make_static_prim table =
  let fold table prim =
    let _, v = eval_caml_prim prim in
    let key, table = GlobVal.push table v in
    table, key
  in
  (* Keep synchronized *)
  let prims = [
    CAML_Arraymake;
    CAML_Arrayget;
    CAML_Arraydefault;
    CAML_Arrayset;
    CAML_Arraycopy;
    CAML_Arraylength;
  ] in
  let table, prims = CList.fold_left_map fold table prims in
  table, Array.of_list prims

let slot_for_caml_prim op table =
  !table.glob_prim.(fst (eval_caml_prim op))

let rec slot_for_getglobal env sigma kn table =
  let (cb,(_,rk)) = lookup_constant_key kn env in
  try key rk
  with NotEvaluated ->
(*    Pp.msgnl(str"not yet evaluated");*)
    let pos =
      match cb.const_body_code with
      | None -> set_global (val_of_constant kn) table
      | Some code ->
        match code with
        | BCdefined (code, fv) ->
           let v = eval_to_patch env sigma (code, fv) table in
           set_global v table
        | BCalias kn' -> slot_for_getglobal env sigma kn' table
        | BCconstant -> set_global (val_of_constant kn) table
    in
(*Pp.msgnl(str"value stored at: "++int pos);*)
    rk := Some (CEphemeron.create pos);
    pos

and slot_for_fv env sigma fv table =
  let fill_fv_cache cache id v_of_id env_of_id b =
    let v, d =
      match b with
      | None -> v_of_id id, Id.Set.empty
      | Some c ->
        let v = val_of_constr (env_of_id id env) sigma c table in
        v, Environ.global_vars_set env c
    in
    build_lazy_val cache (v, d); v in
  let val_of_rel i = val_of_rel (nb_rel env - i) in
  let idfun _ x = x in
  match fv with
  | FVnamed id ->
      let nv = lookup_named_val id env in
      begin match force_lazy_val nv with
      | None ->
         env |> lookup_named id |> NamedDecl.get_value |> fill_fv_cache nv id val_of_named idfun
      | Some (v, _) -> v
      end
  | FVrel i ->
      let rv = lookup_rel_val i env in
      begin match force_lazy_val rv with
      | None ->
        env |> lookup_rel i |> RelDecl.get_value |> fill_fv_cache rv i val_of_rel env_of_rel
      | Some (v, _) -> v
      end
  | FVevar evk -> val_of_evar evk
  | FVuniv_var _idu ->
    assert false

and eval_to_patch env sigma (code, fv) table =
  let slots = function
    | Reloc_annot a -> slot_for_annot a table
    | Reloc_const sc -> slot_for_str_cst sc table
    | Reloc_getglobal kn -> slot_for_getglobal env sigma kn table
    | Reloc_caml_prim op -> slot_for_caml_prim op table
  in
  let tc = patch code slots in
  let vm_env =
    (* Environment should look like a closure, so free variables start at slot 2. *)
    let a = Array.make (Array.length fv + 2) crazy_val in
    a.(1) <- Obj.magic 2;
    let iter i fv =
      let v = slot_for_fv env sigma fv table in
      a.(i + 2) <- v
    in
    let () = Array.iteri iter fv in
    a
  in
  let global = get_global_data !table in
  let v = coq_interprete tc crazy_val (get_atom_rel ()) global (inj_env vm_env) 0 in
  v

and val_of_constr env sigma c table =
  match compile ~fail_on_error:true env sigma c with
  | Some v -> eval_to_patch env sigma v table
  | None -> assert false

let global_table =
  let glob_val = GlobVal.empty 4096 in
  (* Register OCaml primitives statically *)
  let glob_val, glob_prim = make_static_prim glob_val in
  ref {
    glob_val; glob_prim;
    glob_sconst = SConstTable.empty;
    glob_annot = AnnotTable.empty;
  }

let val_of_constr env sigma c =
  let v = val_of_constr env sigma c global_table in
  v

let vm_interp code v env k =
  coq_interprete code v (get_atom_rel ()) (get_global_data !global_table) env k
OCaml

Innovation. Community. Security.