From d8093626b49b539bec283285ea37ba50e79f69d4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 27 Nov 2017 23:02:51 +0100 Subject: Use large arrays in the checker demarshaller. This allows to work around the size limitation of vanilla OCaml arrays on 32-bit platforms, which is rather easy to hit. --- checker/analyze.ml | 57 +++++++++++++++++++++++++++++++++++++++++++++++++---- checker/analyze.mli | 18 ++++++++++++++--- checker/votour.ml | 22 ++++++++++----------- 3 files changed, 79 insertions(+), 18 deletions(-) (limited to 'checker') diff --git a/checker/analyze.ml b/checker/analyze.ml index df75d5b93..2f3ee0356 100644 --- a/checker/analyze.ml +++ b/checker/analyze.ml @@ -55,6 +55,55 @@ let magic_number = "\132\149\166\190" (** Memory reification *) +module LargeArray : +sig + type 'a t + val empty : 'a t + val length : 'a t -> int + val make : int -> 'a -> 'a t + val get : 'a t -> int -> 'a + val set : 'a t -> int -> 'a -> unit +end = +struct + + let max_length = Sys.max_array_length + + type 'a t = 'a array array * 'a array + (** Invariants: + - All subarrays of the left array have length [max_length]. + - The right array has length < [max_length]. + *) + + let empty = [||], [||] + + let length (vl, vr) = + (max_length * Array.length vl) + Array.length vr + + let make n x = + let k = n / max_length in + let r = n mod max_length in + let vl = Array.init k (fun _ -> Array.make max_length x) in + let vr = Array.make r x in + (vl, vr) + + let get (vl, vr) n = + let k = n / max_length in + let r = n mod max_length in + let len = Array.length vl in + if k < len then vl.(k).(r) + else if k == len then vr.(r) + else invalid_arg "index out of bounds" + + let set (vl, vr) n x = + let k = n / max_length in + let r = n mod max_length in + let len = Array.length vl in + if k < len then vl.(k).(r) <- x + else if k == len then vr.(r) <- x + else invalid_arg "index out of bounds" + +end + type repr = | RInt of int | RBlock of (int * int) (* tag × len *) @@ -82,7 +131,7 @@ end module type S = sig type input - val parse : input -> (data * obj array) + val parse : input -> (data * obj LargeArray.t) end module Make(M : Input) = @@ -261,7 +310,7 @@ let parse_object chan = let parse chan = let (magic, len, _, _, size) = parse_header chan in let () = assert (magic = magic_number) in - let memory = Array.make size (Struct ((-1), [||])) in + let memory = LargeArray.make size (Struct ((-1), [||])) in let current_object = ref 0 in let fill_obj = function | RPointer n -> @@ -272,7 +321,7 @@ let parse chan = data, None | RString s -> let data = Ptr !current_object in - let () = memory.(!current_object) <- String s in + let () = LargeArray.set memory !current_object (String s) in let () = incr current_object in data, None | RBlock (tag, 0) -> @@ -282,7 +331,7 @@ let parse chan = | RBlock (tag, len) -> let data = Ptr !current_object in let nblock = Array.make len (Atm (-1)) in - let () = memory.(!current_object) <- Struct (tag, nblock) in + let () = LargeArray.set memory !current_object (Struct (tag, nblock)) in let () = incr current_object in data, Some nblock | RCode addr -> diff --git a/checker/analyze.mli b/checker/analyze.mli index 42efcf01d..1b6c94759 100644 --- a/checker/analyze.mli +++ b/checker/analyze.mli @@ -8,8 +8,20 @@ type obj = | Struct of int * data array (* tag × data *) | String of string -val parse_channel : in_channel -> (data * obj array) -val parse_string : string -> (data * obj array) +module LargeArray : +sig + type 'a t + val empty : 'a t + val length : 'a t -> int + val make : int -> 'a -> 'a t + val get : 'a t -> int -> 'a + val set : 'a t -> int -> 'a -> unit +end +(** A data structure similar to arrays but allowing to overcome the 2^22 length + limitation on 32-bit architecture. *) + +val parse_channel : in_channel -> (data * obj LargeArray.t) +val parse_string : string -> (data * obj LargeArray.t) (** {6 Functorized version} *) @@ -26,7 +38,7 @@ end module type S = sig type input - val parse : input -> (data * obj array) + val parse : input -> (data * obj LargeArray.t) (** Return the entry point and the reification of the memory out of a marshalled structure. *) end diff --git a/checker/votour.ml b/checker/votour.ml index b7c898232..77c9999c4 100644 --- a/checker/votour.ml +++ b/checker/votour.ml @@ -77,8 +77,8 @@ struct type obj = data - let memory = ref [||] - let sizes = ref [||] + let memory = ref LargeArray.empty + let sizes = ref LargeArray.empty (** size, in words *) let ws = Sys.word_size / 8 @@ -86,10 +86,10 @@ struct let rec init_size seen k = function | Int _ | Atm _ | Fun _ -> k 0 | Ptr p -> - if seen.(p) then k 0 + if LargeArray.get seen p then k 0 else - let () = seen.(p) <- true in - match (!memory).(p) with + let () = LargeArray.set seen p true in + match LargeArray.get !memory p with | Struct (tag, os) -> let len = Array.length os in let rec fold i accu k = @@ -97,30 +97,30 @@ struct else init_size seen (fun n -> fold (succ i) (accu + 1 + n) k) os.(i) in - fold 0 1 (fun size -> let () = (!sizes).(p) <- size in k size) + fold 0 1 (fun size -> let () = LargeArray.set !sizes p size in k size) | String s -> let size = 2 + (String.length s / ws) in - let () = (!sizes).(p) <- size in + let () = LargeArray.set !sizes p size in k size let size = function | Int _ | Atm _ | Fun _ -> 0 - | Ptr p -> (!sizes).(p) + | Ptr p -> LargeArray.get !sizes p let repr = function | Int i -> INT i | Atm t -> BLOCK (t, [||]) | Fun _ -> OTHER | Ptr p -> - match (!memory).(p) with + match LargeArray.get !memory p with | Struct (tag, os) -> BLOCK (tag, os) | String s -> STRING s let input ch = let obj, mem = parse_channel ch in let () = memory := mem in - let () = sizes := Array.make (Array.length mem) (-1) in - let seen = Array.make (Array.length mem) false in + let () = sizes := LargeArray.make (LargeArray.length mem) (-1) in + let seen = LargeArray.make (LargeArray.length mem) false in let () = init_size seen ignore obj in obj -- cgit v1.2.3